Theory Native_Word.Code_Target_Word
chapter ‹Common base for target language implementations of word types›
theory Code_Target_Word
  imports "HOL-Library.Word"
begin
section ‹SML›
text ‹
  The separate code target ‹SML_word› collects setups for the
  code generator that PolyML does not provide.
›
setup ‹Code_Target.add_derived_target ("SML_word", [(Code_ML.target_SML, I)])›
section ‹Haskell›
text ‹
  In the \<^text>‹Data.Bits.Bits› type class, shifts and bit indices are given as
  \<^text>‹Int› rather than \<^text>‹Integer›.
  Additional constants take only parameters of type @{typ integer} rather than @{typ nat}
  and check the preconditions as far as possible (e.g., being non-negative) in a portable way.
›
code_printing code_module Data_Bits ⇀ (Haskell)
‹
module Data_Bits where {
import qualified Data.Bits;
{-
  The ...Bounded functions assume that the Integer argument for the shift
  or bit index fits into an Int, is non-negative and (for types of fixed bit width)
  less than bitSize
-}
infixl 7 .&.;
infixl 6 `xor`;
infixl 5 .|.;
(.&.) :: Data.Bits.Bits a => a -> a -> a;
(.&.) = (Data.Bits..&.);
xor :: Data.Bits.Bits a => a -> a -> a;
xor = Data.Bits.xor;
(.|.) :: Data.Bits.Bits a => a -> a -> a;
(.|.) = (Data.Bits..|.);
complement :: Data.Bits.Bits a => a -> a;
complement = Data.Bits.complement;
testBitUnbounded :: Data.Bits.Bits a => a -> Integer -> Bool;
testBitUnbounded x b
  | b <= toInteger (Prelude.maxBound :: Int) = Data.Bits.testBit x (fromInteger b)
  | otherwise = error ("Bit index too large: " ++ show b)
;
testBitBounded :: Data.Bits.Bits a => a -> Integer -> Bool;
testBitBounded x b = Data.Bits.testBit x (fromInteger b);
shiftlUnbounded :: Data.Bits.Bits a => a -> Integer -> a;
shiftlUnbounded x n
  | n <= toInteger (Prelude.maxBound :: Int) = Data.Bits.shiftL x (fromInteger n)
  | otherwise = error ("Shift operand too large: " ++ show n)
;
shiftlBounded :: Data.Bits.Bits a => a -> Integer -> a;
shiftlBounded x n = Data.Bits.shiftL x (fromInteger n);
shiftrUnbounded :: Data.Bits.Bits a => a -> Integer -> a;
shiftrUnbounded x n
  | n <= toInteger (Prelude.maxBound :: Int) = Data.Bits.shiftR x (fromInteger n)
  | otherwise = error ("Shift operand too large: " ++ show n)
;
shiftrBounded :: (Ord a, Data.Bits.Bits a) => a -> Integer -> a;
shiftrBounded x n = Data.Bits.shiftR x (fromInteger n);
}›
  and 
  (Haskell_Quickcheck)
‹
module Data_Bits where {
import qualified Data.Bits;
{-
  The functions assume that the Int argument for the shift or bit index is
  non-negative and (for types of fixed bit width) less than bitSize
-}
infixl 7 .&.;
infixl 6 `xor`;
infixl 5 .|.;
(.&.) :: Data.Bits.Bits a => a -> a -> a;
(.&.) = (Data.Bits..&.);
xor :: Data.Bits.Bits a => a -> a -> a;
xor = Data.Bits.xor;
(.|.) :: Data.Bits.Bits a => a -> a -> a;
(.|.) = (Data.Bits..|.);
complement :: Data.Bits.Bits a => a -> a;
complement = Data.Bits.complement;
testBitUnbounded :: Data.Bits.Bits a => a -> Prelude.Int -> Bool;
testBitUnbounded = Data.Bits.testBit;
testBitBounded :: Data.Bits.Bits a => a -> Prelude.Int -> Bool;
testBitBounded = Data.Bits.testBit;
shiftlUnbounded :: Data.Bits.Bits a => a -> Prelude.Int -> a;
shiftlUnbounded = Data.Bits.shiftL;
shiftlBounded :: Data.Bits.Bits a => a -> Prelude.Int -> a;
shiftlBounded = Data.Bits.shiftL;
shiftrUnbounded :: Data.Bits.Bits a => a -> Prelude.Int -> a;
shiftrUnbounded = Data.Bits.shiftR;
shiftrBounded :: (Ord a, Data.Bits.Bits a) => a -> Prelude.Int -> a;
shiftrBounded = Data.Bits.shiftR;
}›
code_reserved (Haskell) Data_Bits
end