diff --git a/thys/Native_Word/Native_Word_Test.thy b/thys/Native_Word/Native_Word_Test.thy --- a/thys/Native_Word/Native_Word_Test.thy +++ b/thys/Native_Word/Native_Word_Test.thy @@ -1,514 +1,522 @@ (* Title: Native_Word_Test.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Test cases\ theory Native_Word_Test imports Uint64 Uint32 Uint16 Uint8 Uint Native_Cast_Uint "HOL-Library.Code_Test" "Word_Lib.Bit_Shifts_Infix_Syntax" begin +export_code + nat_of_uint8 uint8_of_nat + nat_of_uint16 uint16_of_nat + nat_of_uint32 uint32_of_nat + nat_of_uint64 uint64_of_nat + nat_of_uint uint_of_nat + in SML file "/tmp/foo.ML" + section \Tests for @{typ uint32}\ context includes bit_operations_syntax begin abbreviation (input) sshiftr_uint32 (infixl ">>>" 55) where \w >>> n \ signed_drop_bit_uint32 n w\ definition test_uint32 where "test_uint32 \ (([ 0x100000001, -1, -4294967291, 0xFFFFFFFF, 0x12345678 , 0x5A AND 0x36 , 0x5A OR 0x36 , 0x5A XOR 0x36 , NOT 0x5A , 5 + 6, -5 + 6, -6 + 5, -5 + (- 6), 0xFFFFFFFFF + 1 , 5 - 3, 3 - 5 , 5 * 3, -5 * 3, -5 * -4, 0x12345678 * 0x87654321 , 5 div 3, -5 div 3, -5 div -3, 5 div -3 , 5 mod 3, -5 mod 3, -5 mod -3, 5 mod -3 , set_bit 5 4 True, set_bit (- 5) 2 True, set_bit 5 0 False, set_bit (- 5) 1 False , set_bit 5 32 True, set_bit 5 32 False, set_bit (- 5) 32 True, set_bit (- 5) 32 False , 1 << 2, -1 << 3, 1 << 32, 1 << 0 , 100 >> 3, -100 >> 3, 100 >> 32, -100 >> 32 , 100 >>> 3, -100 >>> 3, 100 >>> 32, -100 >>> 32] :: uint32 list) = [ 1, 4294967295, 5, 4294967295, 305419896 , 18 , 126 , 108 , 4294967205 , 11, 1, 4294967295, 4294967285, 0 , 2, 4294967294 , 15, 4294967281, 20, 1891143032 , 1, 1431655763, 0, 0 , 2, 2, 4294967291, 5 , 21, 4294967295, 4, 4294967289 , 5, 5, 4294967291, 4294967291 , 4, 4294967288, 0, 1 , 12, 536870899, 0, 0 , 12, 4294967283, 0, 4294967295]) \ ([ (0x5 :: uint32) = 0x5, (0x5 :: uint32) = 0x6 , (0x5 :: uint32) < 0x5, (0x5 :: uint32) < 0x6, (-5 :: uint32) < 6, (6 :: uint32) < -5 , (0x5 :: uint32) \ 0x5, (0x5 :: uint32) \ 0x4, (-5 :: uint32) \ 6, (6 :: uint32) \ -5 , (0x7FFFFFFF :: uint32) < 0x80000000, (0xFFFFFFFF :: uint32) < 0, (0x80000000 :: uint32) < 0x7FFFFFFF , bit (0x7FFFFFFF :: uint32) 0, bit (0x7FFFFFFF :: uint32) 31, bit (0x80000000 :: uint32) 31, bit (0x80000000 :: uint32) 32 ] = [ True, False , False, True, False, True , True, False, False, True , True, False, False , True, False, True, False ]) \ ([integer_of_uint32 0, integer_of_uint32 0x7FFFFFFF, integer_of_uint32 0x80000000, integer_of_uint32 0xAAAAAAAA] = [0, 0x7FFFFFFF, 0x80000000, 0xAAAAAAAA])" export_code test_uint32 checking SML Haskell? OCaml? Scala notepad begin have test_uint32 by eval have test_uint32 by code_simp have test_uint32 by normalization end definition test_uint32' :: uint32 where "test_uint32' = 0 + 10 - 14 * 3 div 6 mod 3 << 3 >> 2" ML \val 0wx12 = @{code test_uint32'}\ lemma "x AND y = x OR (y :: uint32)" quickcheck[random, expect=counterexample] quickcheck[exhaustive, expect=counterexample] oops lemma "(x :: uint32) AND x = x OR x" quickcheck[narrowing, expect=no_counterexample] by transfer simp lemma "(f :: uint32 \ unit) = g" quickcheck[narrowing, size=3, expect=no_counterexample] by(simp add: fun_eq_iff) end section \Tests for @{typ uint16}\ context includes bit_operations_syntax begin abbreviation (input) sshiftr_uint16 (infixl ">>>" 55) where \w >>> n \ signed_drop_bit_uint16 n w\ definition test_uint16 where "test_uint16 \ (([ 0x10001, -1, -65535, 0xFFFF, 0x1234 , 0x5A AND 0x36 , 0x5A OR 0x36 , 0x5A XOR 0x36 , NOT 0x5A , 5 + 6, -5 + 6, -6 + 5, -5 + -6, 0xFFFF + 1 , 5 - 3, 3 - 5 , 5 * 3, -5 * 3, -5 * -4, 0x1234 * 0x8765 , 5 div 3, -5 div 3, -5 div -3, 5 div -3 , 5 mod 3, -5 mod 3, -5 mod -3, 5 mod -3 , set_bit 5 4 True, set_bit (- 5) 2 True, set_bit 5 0 False, set_bit (- 5) 1 False , set_bit 5 32 True, set_bit 5 32 False, set_bit (- 5) 32 True, set_bit (- 5) 32 False , 1 << 2, -1 << 3, 1 << 16, 1 << 0 , 100 >> 3, -100 >> 3, 100 >> 16, -100 >> 16 , 100 >>> 3, -100 >>> 3, 100 >>> 16, -100 >>> 16] :: uint16 list) = [ 1, 65535, 1, 65535, 4660 , 18 , 126 , 108 , 65445 , 11, 1, 65535, 65525, 0 , 2, 65534 , 15, 65521, 20, 39556 , 1, 21843, 0, 0 , 2, 2, 65531, 5 , 21, 65535, 4, 65529 , 5, 5, 65531, 65531 , 4, 65528, 0, 1 , 12, 8179, 0, 0 , 12, 65523, 0, 65535]) \ ([ (0x5 :: uint16) = 0x5, (0x5 :: uint16) = 0x6 , (0x5 :: uint16) < 0x5, (0x5 :: uint16) < 0x6, (-5 :: uint16) < 6, (6 :: uint16) < -5 , (0x5 :: uint16) \ 0x5, (0x5 :: uint16) \ 0x4, (-5 :: uint16) \ 6, (6 :: uint16) \ -5 , (0x7FFF :: uint16) < 0x8000, (0xFFFF :: uint16) < 0, (0x8000 :: uint16) < 0x7FFF , bit (0x7FFF :: uint16) 0, bit (0x7FFF :: uint16) 15, bit (0x8000 :: uint16) 15, bit (0x8000 :: uint16) 16 ] = [ True, False , False, True, False, True , True, False, False, True , True, False, False , True, False, True, False ]) \ ([integer_of_uint16 0, integer_of_uint16 0x7FFF, integer_of_uint16 0x8000, integer_of_uint16 0xAAAA] = [0, 0x7FFF, 0x8000, 0xAAAA])" export_code test_uint16 checking Haskell? Scala export_code test_uint16 checking SML_word notepad begin have test_uint16 by code_simp have test_uint16 by normalization end lemma "(x :: uint16) AND x = x OR x" quickcheck[narrowing, expect=no_counterexample] by transfer simp lemma "(f :: uint16 \ unit) = g" quickcheck[narrowing, size=3, expect=no_counterexample] by(simp add: fun_eq_iff) end section \Tests for @{typ uint8}\ context includes bit_operations_syntax begin abbreviation (input) sshiftr_uint8 (infixl ">>>" 55) where \w >>> n \ signed_drop_bit_uint8 n w\ definition test_uint8 where "test_uint8 \ (([ 0x101, -1, -255, 0xFF, 0x12 , 0x5A AND 0x36 , 0x5A OR 0x36 , 0x5A XOR 0x36 , NOT 0x5A , 5 + 6, -5 + 6, -6 + 5, -5 + -6, 0xFF + 1 , 5 - 3, 3 - 5 , 5 * 3, -5 * 3, -5 * -4, 0x12 * 0x87 , 5 div 3, -5 div 3, -5 div -3, 5 div -3 , 5 mod 3, -5 mod 3, -5 mod -3, 5 mod -3 , set_bit 5 4 True, set_bit (- 5) 2 True, set_bit 5 0 False, set_bit (- 5) 1 False , set_bit 5 32 True, set_bit 5 32 False, set_bit (- 5) 32 True, set_bit (- 5) 32 False , 1 << 2, -1 << 3, 1 << 8, 1 << 0 , 100 >> 3, -100 >> 3, 100 >> 8, -100 >> 8 , 100 >>> 3, -100 >>> 3, 100 >>> 8, -100 >>> 8] :: uint8 list) = [ 1, 255, 1, 255, 18 , 18 , 126 , 108 , 165 , 11, 1, 255, 245, 0 , 2, 254 , 15, 241, 20, 126 , 1, 83, 0, 0 , 2, 2, 251, 5 , 21, 255, 4, 249 , 5, 5, 251, 251 , 4, 248, 0, 1 , 12, 19, 0, 0 , 12, 243, 0, 255]) \ ([ (0x5 :: uint8) = 0x5, (0x5 :: uint8) = 0x6 , (0x5 :: uint8) < 0x5, (0x5 :: uint8) < 0x6, (-5 :: uint8) < 6, (6 :: uint8) < -5 , (0x5 :: uint8) \ 0x5, (0x5 :: uint8) \ 0x4, (-5 :: uint8) \ 6, (6 :: uint8) \ -5 , (0x7F :: uint8) < 0x80, (0xFF :: uint8) < 0, (0x80 :: uint8) < 0x7F , bit (0x7F :: uint8) 0, bit (0x7F :: uint8) 7, bit (0x80 :: uint8) 7, bit (0x80 :: uint8) 8 ] = [ True, False , False, True, False, True , True, False, False, True , True, False, False , True, False, True, False ]) \ ([integer_of_uint8 0, integer_of_uint8 0x7F, integer_of_uint8 0x80, integer_of_uint8 0xAA] = [0, 0x7F, 0x80, 0xAA])" export_code test_uint8 checking SML Haskell? Scala notepad begin have test_uint8 by eval have test_uint8 by code_simp have test_uint8 by normalization end ML_val \val true = @{code test_uint8}\ definition test_uint8' :: uint8 where "test_uint8' = 0 + 10 - 14 * 3 div 6 mod 3 << 3 >> 2" ML \val 0wx12 = @{code test_uint8'}\ lemma "x AND y = x OR (y :: uint8)" quickcheck[random, expect=counterexample] quickcheck[exhaustive, expect=counterexample] oops lemma "(x :: uint8) AND x = x OR x" quickcheck[narrowing, expect=no_counterexample] by transfer simp lemma "(f :: uint8 \ unit) = g" quickcheck[narrowing, size=3, expect=no_counterexample] by(simp add: fun_eq_iff) section \Tests for @{typ "uint"}\ context includes bit_operations_syntax begin abbreviation (input) sshiftr_uint (infixl ">>>" 55) where \w >>> n \ signed_drop_bit_uint n w\ definition "test_uint \ let test_list1 = (let HS = uint_of_int (2 ^ (dflt_size - 1)) in ([ HS + HS + 1, -1, -HS - HS + 5, HS + (HS - 1), 0x12 , 0x5A AND 0x36 , 0x5A OR 0x36 , 0x5A XOR 0x36 , NOT 0x5A , 5 + 6, -5 + 6, -6 + 5, -5 + -6, HS + (HS - 1) + 1 , 5 - 3, 3 - 5 , 5 * 3, -5 * 3, -5 * -4, 0x12345678 * 0x87654321] @ (if dflt_size > 4 then [ 5 div 3, -5 div 3, -5 div -3, 5 div -3 , 5 mod 3, -5 mod 3, -5 mod -3, 5 mod -3 , set_bit 5 4 True, set_bit (- 5) 2 True, set_bit 5 0 False, set_bit (- 5) 1 False , set_bit 5 dflt_size True, set_bit 5 dflt_size False, set_bit (- 5) dflt_size True, set_bit (- 5) dflt_size False , 1 << 2, -1 << 3, push_bit dflt_size 1, 1 << 0 , 31 >> 3, -1 >> 3, 31 >> dflt_size, -1 >> dflt_size , 15 >>> 2, -1 >>> 3, 15 >>> dflt_size, -1 >>> dflt_size] else []) :: uint list)); test_list2 = (let S = wivs_shift in ([ 1, -1, -S + 5, S - 1, 0x12 , 0x5A AND 0x36 , 0x5A OR 0x36 , 0x5A XOR 0x36 , NOT 0x5A , 5 + 6, -5 + 6, -6 + 5, -5 + -6, 0 , 5 - 3, 3 - 5 , 5 * 3, -5 * 3, -5 * -4, 0x12345678 * 0x87654321] @ (if dflt_size > 4 then [ 5 div 3, (S - 5) div 3, (S - 5) div (S - 3), 5 div (S - 3) , 5 mod 3, (S - 5) mod 3, (S - 5) mod (S - 3), 5 mod (S - 3) , set_bit 5 4 True, -1, set_bit 5 0 False, -7 , 5, 5, -5, -5 , 4, -8, 0, 1 , 3, (S >> 3) - 1, 0, 0 , 3, (S >> 1) + (S >> 1) - 1, 0, -1] else []) :: int list)); test_list_c1 = (let HS = uint_of_int ((2^(dflt_size - 1))) in [ (0x5 :: uint) = 0x5, (0x5 :: uint) = 0x6 , (0x5 :: uint) < 0x5, (0x5 :: uint) < 0x6, (-5 :: uint) < 6, (6 :: uint) < -5 , (0x5 :: uint) \ 0x5, (0x5 :: uint) \ 0x4, (-5 :: uint) \ 6, (6 :: uint) \ -5 , (HS - 1) < HS, (HS + HS - 1) < 0, HS < HS - 1 , bit (HS - 1) 0, bit (HS - 1 :: uint) (dflt_size - 1), bit (HS :: uint) (dflt_size - 1), bit (HS :: uint) dflt_size ]); test_list_c2 = [ True, False , False, dflt_size\2, dflt_size=3, dflt_size\3 , True, False, dflt_size=3, dflt_size\3 , True, False, False , dflt_size\1, False, True, False ] in test_list1 = map uint_of_int test_list2 \ test_list_c1 = test_list_c2" export_code test_uint checking SML Haskell? OCaml? Scala lemma "test_uint" quickcheck[exhaustive, expect=no_counterexample] oops \ \FIXME: prove correctness of test by reflective means (not yet supported)\ lemma "x AND y = x OR (y :: uint)" quickcheck[random, expect=counterexample] quickcheck[exhaustive, expect=counterexample] oops lemma "(x :: uint) AND x = x OR x" quickcheck[narrowing, expect=no_counterexample] by transfer simp lemma "(f :: uint \ unit) = g" quickcheck[narrowing, size=3, expect=no_counterexample] by(simp add: fun_eq_iff) section \Tests for @{typ uint64} \ context includes bit_operations_syntax begin abbreviation (input) sshiftr_uint64 (infixl ">>>" 55) where \w >>> n \ signed_drop_bit_uint64 n w\ definition test_uint64 where "test_uint64 \ (([ 0x10000000000000001, -1, -9223372036854775808, 0xFFFFFFFFFFFFFFFF, 0x1234567890ABCDEF , 0x5A AND 0x36 , 0x5A OR 0x36 , 0x5A XOR 0x36 , NOT 0x5A , 5 + 6, -5 + 6, -6 + 5, -5 + (- 6), 0xFFFFFFFFFFFFFFFFFF + 1 , 5 - 3, 3 - 5 , 5 * 3, -5 * 3, -5 * -4, 0x1234567890ABCDEF * 0xFEDCBA0987654321 , 5 div 3, -5 div 3, -5 div -3, 5 div -3 , 5 mod 3, -5 mod 3, -5 mod -3, 5 mod -3 , set_bit 5 4 True, set_bit (- 5) 2 True, set_bit 5 0 False, set_bit (- 5) 1 False , set_bit 5 64 True, set_bit 5 64 False, set_bit (- 5) 64 True, set_bit (- 5) 64 False , 1 << 2, -1 << 3, 1 << 64, 1 << 0 , 100 >> 3, -100 >> 3, 100 >> 64, -100 >> 64 , 100 >>> 3, -100 >>> 3, 100 >>> 64, -100 >>> 64] :: uint64 list) = [ 1, 18446744073709551615, 9223372036854775808, 18446744073709551615, 1311768467294899695 , 18 , 126 , 108 , 18446744073709551525 , 11, 1, 18446744073709551615, 18446744073709551605, 0 , 2, 18446744073709551614 , 15, 18446744073709551601, 20, 14000077364136384719 , 1, 6148914691236517203, 0, 0 , 2, 2, 18446744073709551611, 5 , 21, 18446744073709551615, 4, 18446744073709551609 , 5, 5, 18446744073709551611, 18446744073709551611 , 4, 18446744073709551608, 0, 1 , 12, 2305843009213693939, 0, 0 , 12, 18446744073709551603, 0, 18446744073709551615]) \ ([ (0x5 :: uint64) = 0x5, (0x5 :: uint64) = 0x6 , (0x5 :: uint64) < 0x5, (0x5 :: uint64) < 0x6, (-5 :: uint64) < 6, (6 :: uint64) < -5 , (0x5 :: uint64) \ 0x5, (0x5 :: uint64) \ 0x4, (-5 :: uint64) \ 6, (6 :: uint64) \ -5 , (0x7FFFFFFFFFFFFFFF :: uint64) < 0x8000000000000000, (0xFFFFFFFFFFFFFFFF :: uint64) < 0, (0x8000000000000000 :: uint64) < 0x7FFFFFFFFFFFFFFF , bit (0x7FFFFFFFFFFFFFFF :: uint64) 0, bit (0x7FFFFFFFFFFFFFFF :: uint64) 63, bit (0x8000000000000000 :: uint64) 63, bit (0x8000000000000000 :: uint64) 64 ] = [ True, False , False, True, False, True , True, False, False, True , True, False, False , True, False, True, False ]) \ ([integer_of_uint64 0, integer_of_uint64 0x7FFFFFFFFFFFFFFF, integer_of_uint64 0x8000000000000000, integer_of_uint64 0xAAAAAAAAAAAAAAAA] = [0, 0x7FFFFFFFFFFFFFFF, 0x8000000000000000, 0xAAAAAAAAAAAAAAAA])" value [nbe] "[0x10000000000000001, -1, -9223372036854775808, 0xFFFFFFFFFFFFFFFF, 0x1234567890ABCDEF , 0x5A AND 0x36 , 0x5A OR 0x36 , 0x5A XOR 0x36 , NOT 0x5A , 5 + 6, -5 + 6, -6 + 5, -5 + (- 6), 0xFFFFFFFFFFFFFFFFFF + 1 , 5 - 3, 3 - 5 , 5 * 3, -5 * 3, -5 * -4, 0x1234567890ABCDEF * 0xFEDCBA0987654321 , 5 div 3, -5 div 3, -5 div -3, 5 div -3 , 5 mod 3, -5 mod 3, -5 mod -3, 5 mod -3 , set_bit 5 4 True, set_bit (- 5) 2 True, set_bit 5 0 False, set_bit (- 5) 1 False , set_bit 5 64 True, set_bit 5 64 False, set_bit (- 5) 64 True, set_bit (- 5) 64 False , 1 << 2, -1 << 3, 1 << 64, 1 << 0 , 100 >> 3, -100 >> 3, 100 >> 64, -100 >> 64 , 100 >>> 3, -100 >>> 3, 100 >>> 64, -100 >>> 64] :: uint64 list" export_code test_uint64 checking SML Haskell? OCaml? Scala notepad begin have test_uint64 by eval have test_uint64 by code_simp have test_uint64 by normalization end ML_val \val true = @{code test_uint64}\ definition test_uint64' :: uint64 where "test_uint64' = 0 + 10 - 14 * 3 div 6 mod 3 << 3 >> 2" end section \Tests for casts\ definition test_casts :: bool where "test_casts \ map uint8_of_uint32 [10, 0, 0xFE, 0xFFFFFFFF] = [10, 0, 0xFE, 0xFF] \ map uint8_of_uint64 [10, 0, 0xFE, 0xFFFFFFFFFFFFFFFF] = [10, 0, 0xFE, 0xFF] \ map uint32_of_uint8 [10, 0, 0xFF] = [10, 0, 0xFF] \ map uint64_of_uint8 [10, 0, 0xFF] = [10, 0, 0xFF]" definition test_casts' :: bool where "test_casts' \ map uint8_of_uint16 [10, 0, 0xFE, 0xFFFF] = [10, 0, 0xFE, 0xFF] \ map uint16_of_uint8 [10, 0, 0xFF] = [10, 0, 0xFF] \ map uint16_of_uint32 [10, 0, 0xFFFE, 0xFFFFFFFF] = [10, 0, 0xFFFE, 0xFFFF] \ map uint16_of_uint64 [10, 0, 0xFFFE, 0xFFFFFFFFFFFFFFFF] = [10, 0, 0xFFFE, 0xFFFF] \ map uint32_of_uint16 [10, 0, 0xFFFF] = [10, 0, 0xFFFF] \ map uint64_of_uint16 [10, 0, 0xFFFF] = [10, 0, 0xFFFF]" definition test_casts'' :: bool where "test_casts'' \ map uint32_of_uint64 [10, 0, 0xFFFFFFFE, 0xFFFFFFFFFFFFFFFF] = [10, 0, 0xFFFFFFFE, 0xFFFFFFFF] \ map uint64_of_uint32 [10, 0, 0xFFFFFFFF] = [10, 0, 0xFFFFFFFF]" export_code test_casts test_casts'' checking SML Haskell? Scala export_code test_casts'' checking OCaml? export_code test_casts' checking Haskell? Scala notepad begin have test_casts by eval have test_casts by normalization have test_casts by code_simp have test_casts' by normalization have test_casts' by code_simp have test_casts'' by eval have test_casts'' by normalization have test_casts'' by code_simp end ML \ val true = @{code test_casts} val true = @{code test_casts''} \ definition test_casts_uint :: bool where "test_casts_uint \ map uint_of_uint32 ([0, 10] @ (if dflt_size < 32 then [push_bit (dflt_size - 1) 1, 0xFFFFFFFF] else [0xFFFFFFFF])) = [0, 10] @ (if dflt_size < 32 then [push_bit (dflt_size - 1) 1, (push_bit dflt_size 1) - 1] else [0xFFFFFFFF]) \ map uint32_of_uint [0, 10, if dflt_size < 32 then push_bit (dflt_size - 1) 1 else 0xFFFFFFFF] = [0, 10, if dflt_size < 32 then push_bit (dflt_size - 1) 1 else 0xFFFFFFFF] \ map uint_of_uint64 [0, 10, push_bit (dflt_size - 1) 1, 0xFFFFFFFFFFFFFFFF] = [0, 10, push_bit (dflt_size - 1) 1, (push_bit dflt_size 1) - 1] \ map uint64_of_uint [0, 10, push_bit (dflt_size - 1) 1] = [0, 10, push_bit (dflt_size - 1) 1]" definition test_casts_uint' :: bool where "test_casts_uint' \ map uint_of_uint16 [0, 10, 0xFFFF] = [0, 10, 0xFFFF] \ map uint16_of_uint [0, 10, 0xFFFF] = [0, 10, 0xFFFF]" definition test_casts_uint'' :: bool where "test_casts_uint'' \ map uint_of_uint8 [0, 10, 0xFF] = [0, 10, 0xFF] \ map uint8_of_uint [0, 10, 0xFF] = [0, 10, 0xFF]" end end end diff --git a/thys/Native_Word/Uint16.thy b/thys/Native_Word/Uint16.thy --- a/thys/Native_Word/Uint16.thy +++ b/thys/Native_Word/Uint16.thy @@ -1,532 +1,536 @@ (* Title: Uint16.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of 16 bits\ theory Uint16 imports Code_Target_Word_Base Word_Type_Copies begin text \ Restriction for ML code generation: This theory assumes that the ML system provides a Word16 implementation (mlton does, but PolyML 5.5 does not). Therefore, the code setup lives in the target \SML_word\ rather than \SML\. This ensures that code generation still works as long as \uint16\ is not involved. For the target \SML\ itself, no special code generation for this type is set up. Nevertheless, it should work by emulation via \<^typ>\16 word\ if the theory \Code_Target_Bits_Int\ is imported. Restriction for OCaml code generation: OCaml does not provide an int16 type, so no special code generation for this type is set up. \ section \Type definition and primitive operations\ typedef uint16 = \UNIV :: 16 word set\ .. global_interpretation uint16: word_type_copy Abs_uint16 Rep_uint16 using type_definition_uint16 by (rule word_type_copy.intro) setup_lifting type_definition_uint16 declare uint16.of_word_of [code abstype] declare Quotient_uint16 [transfer_rule] instantiation uint16 :: \{comm_ring_1, semiring_modulo, equal, linorder}\ begin lift_definition zero_uint16 :: uint16 is 0 . lift_definition one_uint16 :: uint16 is 1 . lift_definition plus_uint16 :: \uint16 \ uint16 \ uint16\ is \(+)\ . lift_definition uminus_uint16 :: \uint16 \ uint16\ is uminus . lift_definition minus_uint16 :: \uint16 \ uint16 \ uint16\ is \(-)\ . lift_definition times_uint16 :: \uint16 \ uint16 \ uint16\ is \(*)\ . lift_definition divide_uint16 :: \uint16 \ uint16 \ uint16\ is \(div)\ . lift_definition modulo_uint16 :: \uint16 \ uint16 \ uint16\ is \(mod)\ . lift_definition equal_uint16 :: \uint16 \ uint16 \ bool\ is \HOL.equal\ . lift_definition less_eq_uint16 :: \uint16 \ uint16 \ bool\ is \(\)\ . lift_definition less_uint16 :: \uint16 \ uint16 \ bool\ is \(<)\ . global_interpretation uint16: word_type_copy_ring Abs_uint16 Rep_uint16 by standard (fact zero_uint16.rep_eq one_uint16.rep_eq plus_uint16.rep_eq uminus_uint16.rep_eq minus_uint16.rep_eq times_uint16.rep_eq divide_uint16.rep_eq modulo_uint16.rep_eq equal_uint16.rep_eq less_eq_uint16.rep_eq less_uint16.rep_eq)+ instance proof - show \OFCLASS(uint16, comm_ring_1_class)\ by (rule uint16.of_class_comm_ring_1) show \OFCLASS(uint16, semiring_modulo_class)\ by (fact uint16.of_class_semiring_modulo) show \OFCLASS(uint16, equal_class)\ by (fact uint16.of_class_equal) show \OFCLASS(uint16, linorder_class)\ by (fact uint16.of_class_linorder) qed end instantiation uint16 :: ring_bit_operations begin lift_definition bit_uint16 :: \uint16 \ nat \ bool\ is bit . lift_definition not_uint16 :: \uint16 \ uint16\ is NOT . lift_definition and_uint16 :: \uint16 \ uint16 \ uint16\ is \Bit_Operations.and\ . lift_definition or_uint16 :: \uint16 \ uint16 \ uint16\ is \Bit_Operations.or\ . lift_definition xor_uint16 :: \uint16 \ uint16 \ uint16\ is \Bit_Operations.xor\ . lift_definition mask_uint16 :: \nat \ uint16\ is mask . lift_definition push_bit_uint16 :: \nat \ uint16 \ uint16\ is push_bit . lift_definition drop_bit_uint16 :: \nat \ uint16 \ uint16\ is drop_bit . lift_definition signed_drop_bit_uint16 :: \nat \ uint16 \ uint16\ is signed_drop_bit . lift_definition take_bit_uint16 :: \nat \ uint16 \ uint16\ is take_bit . lift_definition set_bit_uint16 :: \nat \ uint16 \ uint16\ is Bit_Operations.set_bit . lift_definition unset_bit_uint16 :: \nat \ uint16 \ uint16\ is unset_bit . lift_definition flip_bit_uint16 :: \nat \ uint16 \ uint16\ is flip_bit . global_interpretation uint16: word_type_copy_bits Abs_uint16 Rep_uint16 signed_drop_bit_uint16 by standard (fact bit_uint16.rep_eq not_uint16.rep_eq and_uint16.rep_eq or_uint16.rep_eq xor_uint16.rep_eq mask_uint16.rep_eq push_bit_uint16.rep_eq drop_bit_uint16.rep_eq signed_drop_bit_uint16.rep_eq take_bit_uint16.rep_eq set_bit_uint16.rep_eq unset_bit_uint16.rep_eq flip_bit_uint16.rep_eq)+ instance by (fact uint16.of_class_ring_bit_operations) end lift_definition uint16_of_nat :: \nat \ uint16\ is word_of_nat . lift_definition nat_of_uint16 :: \uint16 \ nat\ is unat . lift_definition uint16_of_int :: \int \ uint16\ is word_of_int . lift_definition int_of_uint16 :: \uint16 \ int\ is uint . context includes integer.lifting begin lift_definition Uint16 :: \integer \ uint16\ is word_of_int . lift_definition integer_of_uint16 :: \uint16 \ integer\ is uint . end global_interpretation uint16: word_type_copy_more Abs_uint16 Rep_uint16 signed_drop_bit_uint16 uint16_of_nat nat_of_uint16 uint16_of_int int_of_uint16 Uint16 integer_of_uint16 apply standard apply (simp_all add: uint16_of_nat.rep_eq nat_of_uint16.rep_eq uint16_of_int.rep_eq int_of_uint16.rep_eq Uint16.rep_eq integer_of_uint16.rep_eq integer_eq_iff) done instantiation uint16 :: "{size, msb, lsb, set_bit, bit_comprehension}" begin lift_definition size_uint16 :: \uint16 \ nat\ is size . lift_definition msb_uint16 :: \uint16 \ bool\ is msb . lift_definition lsb_uint16 :: \uint16 \ bool\ is lsb . text \Workaround: avoid name space clash by spelling out \<^text>\lift_definition\ explicitly.\ definition set_bit_uint16 :: \uint16 \ nat \ bool \ uint16\ where set_bit_uint16_eq: \set_bit_uint16 a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ context includes lifting_syntax begin lemma set_bit_uint16_transfer [transfer_rule]: \(cr_uint16 ===> (=) ===> (\) ===> cr_uint16) Generic_set_bit.set_bit Generic_set_bit.set_bit\ by (simp only: set_bit_eq [abs_def] set_bit_uint16_eq [abs_def]) transfer_prover end lift_definition set_bits_uint16 :: \(nat \ bool) \ uint16\ is set_bits . lift_definition set_bits_aux_uint16 :: \(nat \ bool) \ nat \ uint16 \ uint16\ is set_bits_aux . global_interpretation uint16: word_type_copy_misc Abs_uint16 Rep_uint16 signed_drop_bit_uint16 uint16_of_nat nat_of_uint16 uint16_of_int int_of_uint16 Uint16 integer_of_uint16 16 set_bits_aux_uint16 by (standard; transfer) simp_all instance using uint16.of_class_bit_comprehension uint16.of_class_set_bit uint16.of_class_lsb by simp_all standard end section \Code setup\ code_printing code_module Uint16 \ (SML_word) \(* Test that words can handle numbers between 0 and 15 *) val _ = if 4 <= Word.wordSize then () else raise (Fail ("wordSize less than 4")); structure Uint16 : sig val set_bit : Word16.word -> IntInf.int -> bool -> Word16.word val shiftl : Word16.word -> IntInf.int -> Word16.word val shiftr : Word16.word -> IntInf.int -> Word16.word val shiftr_signed : Word16.word -> IntInf.int -> Word16.word val test_bit : Word16.word -> IntInf.int -> bool end = struct fun set_bit x n b = let val mask = Word16.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n)) in if b then Word16.orb (x, mask) else Word16.andb (x, Word16.notb mask) end fun shiftl x n = Word16.<< (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr x n = Word16.>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr_signed x n = Word16.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun test_bit x n = Word16.andb (x, Word16.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))) <> Word16.fromInt 0 end; (* struct Uint16 *)\ code_reserved SML_word Uint16 code_printing code_module Uint16 \ (Haskell) \module Uint16(Int16, Word16) where import Data.Int(Int16) import Data.Word(Word16)\ code_reserved Haskell Uint16 text \Scala provides unsigned 16-bit numbers as Char.\ code_printing code_module Uint16 \ (Scala) \object Uint16 { def set_bit(x: scala.Char, n: BigInt, b: Boolean) : scala.Char = if (b) (x | (1.toChar << n.intValue)).toChar else (x & (1.toChar << n.intValue).unary_~).toChar def shiftl(x: scala.Char, n: BigInt) : scala.Char = (x << n.intValue).toChar def shiftr(x: scala.Char, n: BigInt) : scala.Char = (x >>> n.intValue).toChar def shiftr_signed(x: scala.Char, n: BigInt) : scala.Char = (x.toShort >> n.intValue).toChar def test_bit(x: scala.Char, n: BigInt) : Boolean = (x & (1.toChar << n.intValue)) != 0 } /* object Uint16 */\ code_reserved Scala Uint16 text \ Avoid @{term Abs_uint16} in generated code, use @{term Rep_uint16'} instead. The symbolic implementations for code\_simp use @{term Rep_uint16}. The new destructor @{term Rep_uint16'} is executable. As the simplifier is given the [code abstract] equations literally, we cannot implement @{term Rep_uint16} directly, because that makes code\_simp loop. If code generation raises Match, some equation probably contains @{term Rep_uint16} ([code abstract] equations for @{typ uint16} may use @{term Rep_uint16} because these instances will be folded away.) To convert @{typ "16 word"} values into @{typ uint16}, use @{term "Abs_uint16'"}. \ definition Rep_uint16' where [simp]: "Rep_uint16' = Rep_uint16" lemma Rep_uint16'_transfer [transfer_rule]: "rel_fun cr_uint16 (=) (\x. x) Rep_uint16'" unfolding Rep_uint16'_def by(rule uint16.rep_transfer) lemma Rep_uint16'_code [code]: "Rep_uint16' x = (BITS n. bit x n)" by transfer (simp add: set_bits_bit_eq) lift_definition Abs_uint16' :: "16 word \ uint16" is "\x :: 16 word. x" . lemma Abs_uint16'_code [code]: "Abs_uint16' x = Uint16 (integer_of_int (uint x))" including integer.lifting by transfer simp declare [[code drop: "term_of_class.term_of :: uint16 \ _"]] lemma term_of_uint16_code [code]: defines "TR \ typerep.Typerep" and "bit0 \ STR ''Numeral_Type.bit0''" shows "term_of_class.term_of x = Code_Evaluation.App (Code_Evaluation.Const (STR ''Uint16.uint16.Abs_uint16'') (TR (STR ''fun'') [TR (STR ''Word.word'') [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR (STR ''Numeral_Type.num1'') []]]]]], TR (STR ''Uint16.uint16'') []])) (term_of_class.term_of (Rep_uint16' x))" by(simp add: term_of_anything) lemma Uint16_code [code]: "Rep_uint16 (Uint16 i) = word_of_int (int_of_integer_symbolic i)" unfolding Uint16_def int_of_integer_symbolic_def by(simp add: Abs_uint16_inverse) code_printing type_constructor uint16 \ (SML_word) "Word16.word" and (Haskell) "Uint16.Word16" and (Scala) "scala.Char" | constant Uint16 \ (SML_word) "Word16.fromLargeInt (IntInf.toLarge _)" and (Haskell) "(Prelude.fromInteger _ :: Uint16.Word16)" and (Haskell_Quickcheck) "(Prelude.fromInteger (Prelude.toInteger _) :: Uint16.Word16)" and (Scala) "_.charValue" | constant "0 :: uint16" \ (SML_word) "(Word16.fromInt 0)" and (Haskell) "(0 :: Uint16.Word16)" and (Scala) "0" | constant "1 :: uint16" \ (SML_word) "(Word16.fromInt 1)" and (Haskell) "(1 :: Uint16.Word16)" and (Scala) "1" | constant "plus :: uint16 \ _ \ _" \ (SML_word) "Word16.+ ((_), (_))" and (Haskell) infixl 6 "+" and (Scala) "(_ +/ _).toChar" | constant "uminus :: uint16 \ _" \ (SML_word) "Word16.~" and (Haskell) "negate" and (Scala) "(- _).toChar" | constant "minus :: uint16 \ _" \ (SML_word) "Word16.- ((_), (_))" and (Haskell) infixl 6 "-" and (Scala) "(_ -/ _).toChar" | constant "times :: uint16 \ _ \ _" \ (SML_word) "Word16.* ((_), (_))" and (Haskell) infixl 7 "*" and (Scala) "(_ */ _).toChar" | constant "HOL.equal :: uint16 \ _ \ bool" \ (SML_word) "!((_ : Word16.word) = _)" and (Haskell) infix 4 "==" and (Scala) infixl 5 "==" | class_instance uint16 :: equal \ (Haskell) - | constant "less_eq :: uint16 \ _ \ bool" \ (SML_word) "Word16.<= ((_), (_))" and (Haskell) infix 4 "<=" and (Scala) infixl 4 "<=" | constant "less :: uint16 \ _ \ bool" \ (SML_word) "Word16.< ((_), (_))" and (Haskell) infix 4 "<" and (Scala) infixl 4 "<" | constant "NOT :: uint16 \ _" \ (SML_word) "Word16.notb" and (Haskell) "Data'_Bits.complement" and (Scala) "_.unary'_~.toChar" | constant "Bit_Operations.and :: uint16 \ _" \ (SML_word) "Word16.andb ((_),/ (_))" and (Haskell) infixl 7 "Data_Bits..&." and (Scala) "(_ & _).toChar" | constant "Bit_Operations.or :: uint16 \ _" \ (SML_word) "Word16.orb ((_),/ (_))" and (Haskell) infixl 5 "Data_Bits..|." and (Scala) "(_ | _).toChar" | constant "Bit_Operations.xor :: uint16 \ _" \ (SML_word) "Word16.xorb ((_),/ (_))" and (Haskell) "Data'_Bits.xor" and (Scala) "(_ ^ _).toChar" definition uint16_div :: "uint16 \ uint16 \ uint16" where "uint16_div x y = (if y = 0 then undefined ((div) :: uint16 \ _) x (0 :: uint16) else x div y)" definition uint16_mod :: "uint16 \ uint16 \ uint16" where "uint16_mod x y = (if y = 0 then undefined ((mod) :: uint16 \ _) x (0 :: uint16) else x mod y)" context includes undefined_transfer begin lemma div_uint16_code [code]: "x div y = (if y = 0 then 0 else uint16_div x y)" unfolding uint16_div_def by transfer (simp add: word_div_def) lemma mod_uint16_code [code]: "x mod y = (if y = 0 then x else uint16_mod x y)" unfolding uint16_mod_def by transfer (simp add: word_mod_def) lemma uint16_div_code [code]: "Rep_uint16 (uint16_div x y) = (if y = 0 then Rep_uint16 (undefined ((div) :: uint16 \ _) x (0 :: uint16)) else Rep_uint16 x div Rep_uint16 y)" unfolding uint16_div_def by transfer simp lemma uint16_mod_code [code]: "Rep_uint16 (uint16_mod x y) = (if y = 0 then Rep_uint16 (undefined ((mod) :: uint16 \ _) x (0 :: uint16)) else Rep_uint16 x mod Rep_uint16 y)" unfolding uint16_mod_def by transfer simp end code_printing constant uint16_div \ (SML_word) "Word16.div ((_), (_))" and (Haskell) "Prelude.div" and (Scala) "(_ '/ _).toChar" | constant uint16_mod \ (SML_word) "Word16.mod ((_), (_))" and (Haskell) "Prelude.mod" and (Scala) "(_ % _).toChar" definition uint16_test_bit :: "uint16 \ integer \ bool" where [code del]: "uint16_test_bit x n = (if n < 0 \ 15 < n then undefined (bit :: uint16 \ _) x n else bit x (nat_of_integer n))" lemma test_bit_uint16_code [code]: "bit x n \ n < 16 \ uint16_test_bit x (integer_of_nat n)" including undefined_transfer integer.lifting unfolding uint16_test_bit_def by (transfer, simp, transfer, simp) lemma uint16_test_bit_code [code]: "uint16_test_bit w n = (if n < 0 \ 15 < n then undefined (bit :: uint16 \ _) w n else bit (Rep_uint16 w) (nat_of_integer n))" unfolding uint16_test_bit_def by (simp add: bit_uint16.rep_eq) code_printing constant uint16_test_bit \ (SML_word) "Uint16.test'_bit" and (Haskell) "Data'_Bits.testBitBounded" and (Scala) "Uint16.test'_bit" definition uint16_set_bit :: "uint16 \ integer \ bool \ uint16" where [code del]: "uint16_set_bit x n b = (if n < 0 \ 15 < n then undefined (set_bit :: uint16 \ _) x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_uint16_code [code]: "set_bit x n b = (if n < 16 then uint16_set_bit x (integer_of_nat n) b else x)" including undefined_transfer integer.lifting unfolding uint16_set_bit_def by(transfer)(auto cong: conj_cong simp add: not_less set_bit_beyond word_size) lemma uint16_set_bit_code [code]: "Rep_uint16 (uint16_set_bit w n b) = (if n < 0 \ 15 < n then Rep_uint16 (undefined (set_bit :: uint16 \ _) w n b) else set_bit (Rep_uint16 w) (nat_of_integer n) b)" including undefined_transfer unfolding uint16_set_bit_def by transfer simp code_printing constant uint16_set_bit \ (SML_word) "Uint16.set'_bit" and (Haskell) "Data'_Bits.setBitBounded" and (Scala) "Uint16.set'_bit" definition uint16_shiftl :: "uint16 \ integer \ uint16" where [code del]: "uint16_shiftl x n = (if n < 0 \ 16 \ n then undefined (push_bit :: nat \ uint16 \ _) x n else push_bit (nat_of_integer n) x)" lemma shiftl_uint16_code [code]: "push_bit n x = (if n < 16 then uint16_shiftl x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint16_shiftl_def by transfer simp lemma uint16_shiftl_code [code]: "Rep_uint16 (uint16_shiftl w n) = (if n < 0 \ 16 \ n then Rep_uint16 (undefined (push_bit :: nat \ uint16 \ _) w n) else push_bit (nat_of_integer n) (Rep_uint16 w))" including undefined_transfer unfolding uint16_shiftl_def by transfer simp code_printing constant uint16_shiftl \ (SML_word) "Uint16.shiftl" and (Haskell) "Data'_Bits.shiftlBounded" and (Scala) "Uint16.shiftl" definition uint16_shiftr :: "uint16 \ integer \ uint16" where [code del]: "uint16_shiftr x n = (if n < 0 \ 16 \ n then undefined (drop_bit :: nat \ uint16 \ _) x n else drop_bit (nat_of_integer n) x)" lemma shiftr_uint16_code [code]: "drop_bit n x = (if n < 16 then uint16_shiftr x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint16_shiftr_def by transfer simp lemma uint16_shiftr_code [code]: "Rep_uint16 (uint16_shiftr w n) = (if n < 0 \ 16 \ n then Rep_uint16 (undefined (drop_bit :: nat \ uint16 \ _) w n) else drop_bit (nat_of_integer n) (Rep_uint16 w))" including undefined_transfer unfolding uint16_shiftr_def by transfer simp code_printing constant uint16_shiftr \ (SML_word) "Uint16.shiftr" and (Haskell) "Data'_Bits.shiftrBounded" and (Scala) "Uint16.shiftr" definition uint16_sshiftr :: "uint16 \ integer \ uint16" where [code del]: "uint16_sshiftr x n = (if n < 0 \ 16 \ n then undefined signed_drop_bit_uint16 n x else signed_drop_bit_uint16 (nat_of_integer n) x)" lemma sshiftr_uint16_code [code]: "signed_drop_bit_uint16 n x = (if n < 16 then uint16_sshiftr x (integer_of_nat n) else if bit x 15 then -1 else 0)" including undefined_transfer integer.lifting unfolding uint16_sshiftr_def by transfer (simp add: not_less signed_drop_bit_beyond word_size) lemma uint16_sshiftr_code [code]: "Rep_uint16 (uint16_sshiftr w n) = (if n < 0 \ 16 \ n then Rep_uint16 (undefined signed_drop_bit_uint16 n w) else signed_drop_bit (nat_of_integer n) (Rep_uint16 w))" including undefined_transfer unfolding uint16_sshiftr_def by transfer simp code_printing constant uint16_sshiftr \ (SML_word) "Uint16.shiftr'_signed" and (Haskell) "(Prelude.fromInteger (Prelude.toInteger (Data'_Bits.shiftrBounded (Prelude.fromInteger (Prelude.toInteger _) :: Uint16.Int16) _)) :: Uint16.Word16)" and (Scala) "Uint16.shiftr'_signed" lemma uint16_msb_test_bit: "msb x \ bit (x :: uint16) 15" by transfer (simp add: msb_word_iff_bit) lemma msb_uint16_code [code]: "msb x \ uint16_test_bit x 15" by (simp add: uint16_test_bit_def uint16_msb_test_bit) lemma uint16_of_int_code [code]: "uint16_of_int i = Uint16 (integer_of_int i)" including integer.lifting by transfer simp lemma int_of_uint16_code [code]: "int_of_uint16 x = int_of_integer (integer_of_uint16 x)" by (simp add: int_of_uint16.rep_eq integer_of_uint16_def) +lemma uint16_of_nat_code [code]: + "uint16_of_nat = uint16_of_int \ int" + by transfer (simp add: fun_eq_iff) + lemma nat_of_uint16_code [code]: "nat_of_uint16 x = nat_of_integer (integer_of_uint16 x)" unfolding integer_of_uint16_def including integer.lifting by transfer simp lemma integer_of_uint16_code [code]: "integer_of_uint16 n = integer_of_int (uint (Rep_uint16' n))" unfolding integer_of_uint16_def by transfer auto code_printing constant "integer_of_uint16" \ (SML_word) "Word16.toInt _ : IntInf.int" and (Haskell) "Prelude.toInteger" and (Scala) "BigInt" section \Quickcheck setup\ definition uint16_of_natural :: "natural \ uint16" where "uint16_of_natural x \ Uint16 (integer_of_natural x)" instantiation uint16 :: "{random, exhaustive, full_exhaustive}" begin definition "random_uint16 \ qc_random_cnv uint16_of_natural" definition "exhaustive_uint16 \ qc_exhaustive_cnv uint16_of_natural" definition "full_exhaustive_uint16 \ qc_full_exhaustive_cnv uint16_of_natural" instance .. end instantiation uint16 :: narrowing begin interpretation quickcheck_narrowing_samples "\i. let x = Uint16 i in (x, 0xFFFF - x)" "0" "Typerep.Typerep (STR ''Uint16.uint16'') []" . definition "narrowing_uint16 d = qc_narrowing_drawn_from (narrowing_samples d) d" declare [[code drop: "partial_term_of :: uint16 itself \ _"]] lemmas partial_term_of_uint16 [code] = partial_term_of_code instance .. end end diff --git a/thys/Native_Word/Uint32.thy b/thys/Native_Word/Uint32.thy --- a/thys/Native_Word/Uint32.thy +++ b/thys/Native_Word/Uint32.thy @@ -1,685 +1,691 @@ (* Title: Uint32.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of 32 bits\ theory Uint32 imports Code_Target_Word_Base Word_Type_Copies begin section \Type definition and primitive operations\ typedef uint32 = \UNIV :: 32 word set\ .. global_interpretation uint32: word_type_copy Abs_uint32 Rep_uint32 using type_definition_uint32 by (rule word_type_copy.intro) setup_lifting type_definition_uint32 declare uint32.of_word_of [code abstype] declare Quotient_uint32 [transfer_rule] instantiation uint32 :: \{comm_ring_1, semiring_modulo, equal, linorder}\ begin lift_definition zero_uint32 :: uint32 is 0 . lift_definition one_uint32 :: uint32 is 1 . lift_definition plus_uint32 :: \uint32 \ uint32 \ uint32\ is \(+)\ . lift_definition uminus_uint32 :: \uint32 \ uint32\ is uminus . lift_definition minus_uint32 :: \uint32 \ uint32 \ uint32\ is \(-)\ . lift_definition times_uint32 :: \uint32 \ uint32 \ uint32\ is \(*)\ . lift_definition divide_uint32 :: \uint32 \ uint32 \ uint32\ is \(div)\ . lift_definition modulo_uint32 :: \uint32 \ uint32 \ uint32\ is \(mod)\ . lift_definition equal_uint32 :: \uint32 \ uint32 \ bool\ is \HOL.equal\ . lift_definition less_eq_uint32 :: \uint32 \ uint32 \ bool\ is \(\)\ . lift_definition less_uint32 :: \uint32 \ uint32 \ bool\ is \(<)\ . global_interpretation uint32: word_type_copy_ring Abs_uint32 Rep_uint32 by standard (fact zero_uint32.rep_eq one_uint32.rep_eq plus_uint32.rep_eq uminus_uint32.rep_eq minus_uint32.rep_eq times_uint32.rep_eq divide_uint32.rep_eq modulo_uint32.rep_eq equal_uint32.rep_eq less_eq_uint32.rep_eq less_uint32.rep_eq)+ instance proof - show \OFCLASS(uint32, comm_ring_1_class)\ by (rule uint32.of_class_comm_ring_1) show \OFCLASS(uint32, semiring_modulo_class)\ by (fact uint32.of_class_semiring_modulo) show \OFCLASS(uint32, equal_class)\ by (fact uint32.of_class_equal) show \OFCLASS(uint32, linorder_class)\ by (fact uint32.of_class_linorder) qed end instantiation uint32 :: ring_bit_operations begin lift_definition bit_uint32 :: \uint32 \ nat \ bool\ is bit . lift_definition not_uint32 :: \uint32 \ uint32\ is NOT . lift_definition and_uint32 :: \uint32 \ uint32 \ uint32\ is \Bit_Operations.and\ . lift_definition or_uint32 :: \uint32 \ uint32 \ uint32\ is \Bit_Operations.or\ . lift_definition xor_uint32 :: \uint32 \ uint32 \ uint32\ is \Bit_Operations.xor\ . lift_definition mask_uint32 :: \nat \ uint32\ is mask . lift_definition push_bit_uint32 :: \nat \ uint32 \ uint32\ is push_bit . lift_definition drop_bit_uint32 :: \nat \ uint32 \ uint32\ is drop_bit . lift_definition signed_drop_bit_uint32 :: \nat \ uint32 \ uint32\ is signed_drop_bit . lift_definition take_bit_uint32 :: \nat \ uint32 \ uint32\ is take_bit . lift_definition set_bit_uint32 :: \nat \ uint32 \ uint32\ is Bit_Operations.set_bit . lift_definition unset_bit_uint32 :: \nat \ uint32 \ uint32\ is unset_bit . lift_definition flip_bit_uint32 :: \nat \ uint32 \ uint32\ is flip_bit . global_interpretation uint32: word_type_copy_bits Abs_uint32 Rep_uint32 signed_drop_bit_uint32 by standard (fact bit_uint32.rep_eq not_uint32.rep_eq and_uint32.rep_eq or_uint32.rep_eq xor_uint32.rep_eq mask_uint32.rep_eq push_bit_uint32.rep_eq drop_bit_uint32.rep_eq signed_drop_bit_uint32.rep_eq take_bit_uint32.rep_eq set_bit_uint32.rep_eq unset_bit_uint32.rep_eq flip_bit_uint32.rep_eq)+ instance by (fact uint32.of_class_ring_bit_operations) end lift_definition uint32_of_nat :: \nat \ uint32\ is word_of_nat . lift_definition nat_of_uint32 :: \uint32 \ nat\ is unat . lift_definition uint32_of_int :: \int \ uint32\ is word_of_int . lift_definition int_of_uint32 :: \uint32 \ int\ is uint . context includes integer.lifting begin lift_definition Uint32 :: \integer \ uint32\ is word_of_int . lift_definition integer_of_uint32 :: \uint32 \ integer\ is uint . end global_interpretation uint32: word_type_copy_more Abs_uint32 Rep_uint32 signed_drop_bit_uint32 uint32_of_nat nat_of_uint32 uint32_of_int int_of_uint32 Uint32 integer_of_uint32 apply standard apply (simp_all add: uint32_of_nat.rep_eq nat_of_uint32.rep_eq uint32_of_int.rep_eq int_of_uint32.rep_eq Uint32.rep_eq integer_of_uint32.rep_eq integer_eq_iff) done instantiation uint32 :: "{size, msb, lsb, set_bit, bit_comprehension}" begin lift_definition size_uint32 :: \uint32 \ nat\ is size . lift_definition msb_uint32 :: \uint32 \ bool\ is msb . lift_definition lsb_uint32 :: \uint32 \ bool\ is lsb . text \Workaround: avoid name space clash by spelling out \<^text>\lift_definition\ explicitly.\ definition set_bit_uint32 :: \uint32 \ nat \ bool \ uint32\ where set_bit_uint32_eq: \set_bit_uint32 a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ context includes lifting_syntax begin lemma set_bit_uint32_transfer [transfer_rule]: \(cr_uint32 ===> (=) ===> (\) ===> cr_uint32) Generic_set_bit.set_bit Generic_set_bit.set_bit\ by (simp only: set_bit_eq [abs_def] set_bit_uint32_eq [abs_def]) transfer_prover end lift_definition set_bits_uint32 :: \(nat \ bool) \ uint32\ is set_bits . lift_definition set_bits_aux_uint32 :: \(nat \ bool) \ nat \ uint32 \ uint32\ is set_bits_aux . global_interpretation uint32: word_type_copy_misc Abs_uint32 Rep_uint32 signed_drop_bit_uint32 uint32_of_nat nat_of_uint32 uint32_of_int int_of_uint32 Uint32 integer_of_uint32 32 set_bits_aux_uint32 by (standard; transfer) simp_all instance using uint32.of_class_bit_comprehension uint32.of_class_set_bit uint32.of_class_lsb by simp_all standard end section \Code setup\ code_printing code_module Uint32 \ (SML) \(* Test that words can handle numbers between 0 and 31 *) val _ = if 5 <= Word.wordSize then () else raise (Fail ("wordSize less than 5")); structure Uint32 : sig val set_bit : Word32.word -> IntInf.int -> bool -> Word32.word val shiftl : Word32.word -> IntInf.int -> Word32.word val shiftr : Word32.word -> IntInf.int -> Word32.word val shiftr_signed : Word32.word -> IntInf.int -> Word32.word val test_bit : Word32.word -> IntInf.int -> bool end = struct fun set_bit x n b = let val mask = Word32.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n)) in if b then Word32.orb (x, mask) else Word32.andb (x, Word32.notb mask) end fun shiftl x n = Word32.<< (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr x n = Word32.>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr_signed x n = Word32.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun test_bit x n = Word32.andb (x, Word32.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))) <> Word32.fromInt 0 end; (* struct Uint32 *)\ code_reserved SML Uint32 code_printing code_module Uint32 \ (Haskell) \module Uint32(Int32, Word32) where import Data.Int(Int32) import Data.Word(Word32)\ code_reserved Haskell Uint32 text \ OCaml and Scala provide only signed 32bit numbers, so we use these and implement sign-sensitive operations like comparisons manually. \ code_printing code_module "Uint32" \ (OCaml) \module Uint32 : sig val less : int32 -> int32 -> bool val less_eq : int32 -> int32 -> bool val set_bit : int32 -> Z.t -> bool -> int32 val shiftl : int32 -> Z.t -> int32 val shiftr : int32 -> Z.t -> int32 val shiftr_signed : int32 -> Z.t -> int32 val test_bit : int32 -> Z.t -> bool end = struct (* negative numbers have their highest bit set, so they are greater than positive ones *) let less x y = if Int32.compare x Int32.zero < 0 then Int32.compare y Int32.zero < 0 && Int32.compare x y < 0 else Int32.compare y Int32.zero < 0 || Int32.compare x y < 0;; let less_eq x y = if Int32.compare x Int32.zero < 0 then Int32.compare y Int32.zero < 0 && Int32.compare x y <= 0 else Int32.compare y Int32.zero < 0 || Int32.compare x y <= 0;; let set_bit x n b = let mask = Int32.shift_left Int32.one (Z.to_int n) in if b then Int32.logor x mask else Int32.logand x (Int32.lognot mask);; let shiftl x n = Int32.shift_left x (Z.to_int n);; let shiftr x n = Int32.shift_right_logical x (Z.to_int n);; let shiftr_signed x n = Int32.shift_right x (Z.to_int n);; let test_bit x n = Int32.compare (Int32.logand x (Int32.shift_left Int32.one (Z.to_int n))) Int32.zero <> 0;; end;; (*struct Uint32*)\ code_reserved OCaml Uint32 code_printing code_module Uint32 \ (Scala) \object Uint32 { def less(x: Int, y: Int) : Boolean = if (x < 0) y < 0 && x < y else y < 0 || x < y def less_eq(x: Int, y: Int) : Boolean = if (x < 0) y < 0 && x <= y else y < 0 || x <= y def set_bit(x: Int, n: BigInt, b: Boolean) : Int = if (b) x | (1 << n.intValue) else x & (1 << n.intValue).unary_~ def shiftl(x: Int, n: BigInt) : Int = x << n.intValue def shiftr(x: Int, n: BigInt) : Int = x >>> n.intValue def shiftr_signed(x: Int, n: BigInt) : Int = x >> n.intValue def test_bit(x: Int, n: BigInt) : Boolean = (x & (1 << n.intValue)) != 0 } /* object Uint32 */\ code_reserved Scala Uint32 text \ OCaml's conversion from Big\_int to int32 demands that the value fits int a signed 32-bit integer. The following justifies the implementation. \ context includes bit_operations_syntax begin definition Uint32_signed :: "integer \ uint32" where "Uint32_signed i = (if i < -(0x80000000) \ i \ 0x80000000 then undefined Uint32 i else Uint32 i)" lemma Uint32_code [code]: "Uint32 i = (let i' = i AND 0xFFFFFFFF in if bit i' 31 then Uint32_signed (i' - 0x100000000) else Uint32_signed i')" including undefined_transfer integer.lifting unfolding Uint32_signed_def apply transfer apply (subst word_of_int_via_signed) apply (auto simp add: push_bit_of_1 mask_eq_exp_minus_1 word_of_int_via_signed cong del: if_cong) done lemma Uint32_signed_code [code]: "Rep_uint32 (Uint32_signed i) = (if i < -(0x80000000) \ i \ 0x80000000 then Rep_uint32 (undefined Uint32 i) else word_of_int (int_of_integer_symbolic i))" unfolding Uint32_signed_def Uint32_def int_of_integer_symbolic_def word_of_integer_def by(simp add: Abs_uint32_inverse) end text \ Avoid @{term Abs_uint32} in generated code, use @{term Rep_uint32'} instead. The symbolic implementations for code\_simp use @{term Rep_uint32}. The new destructor @{term Rep_uint32'} is executable. As the simplifier is given the [code abstract] equations literally, we cannot implement @{term Rep_uint32} directly, because that makes code\_simp loop. If code generation raises Match, some equation probably contains @{term Rep_uint32} ([code abstract] equations for @{typ uint32} may use @{term Rep_uint32} because these instances will be folded away.) To convert @{typ "32 word"} values into @{typ uint32}, use @{term "Abs_uint32'"}. \ definition Rep_uint32' where [simp]: "Rep_uint32' = Rep_uint32" lemma Rep_uint32'_transfer [transfer_rule]: "rel_fun cr_uint32 (=) (\x. x) Rep_uint32'" unfolding Rep_uint32'_def by(rule uint32.rep_transfer) lemma Rep_uint32'_code [code]: "Rep_uint32' x = (BITS n. bit x n)" by transfer (simp add: set_bits_bit_eq) lift_definition Abs_uint32' :: "32 word \ uint32" is "\x :: 32 word. x" . lemma Abs_uint32'_code [code]: "Abs_uint32' x = Uint32 (integer_of_int (uint x))" including integer.lifting by transfer simp declare [[code drop: "term_of_class.term_of :: uint32 \ _"]] lemma term_of_uint32_code [code]: defines "TR \ typerep.Typerep" and "bit0 \ STR ''Numeral_Type.bit0''" shows "term_of_class.term_of x = Code_Evaluation.App (Code_Evaluation.Const (STR ''Uint32.uint32.Abs_uint32'') (TR (STR ''fun'') [TR (STR ''Word.word'') [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR (STR ''Numeral_Type.num1'') []]]]]]], TR (STR ''Uint32.uint32'') []])) (term_of_class.term_of (Rep_uint32' x))" by(simp add: term_of_anything) code_printing type_constructor uint32 \ (SML) "Word32.word" and (Haskell) "Uint32.Word32" and (OCaml) "int32" and (Scala) "Int" and (Eval) "Word32.word" | constant Uint32 \ (SML) "Word32.fromLargeInt (IntInf.toLarge _)" and (Haskell) "(Prelude.fromInteger _ :: Uint32.Word32)" and (Haskell_Quickcheck) "(Prelude.fromInteger (Prelude.toInteger _) :: Uint32.Word32)" and (Scala) "_.intValue" | constant Uint32_signed \ (OCaml) "Z.to'_int32" | constant "0 :: uint32" \ (SML) "(Word32.fromInt 0)" and (Haskell) "(0 :: Uint32.Word32)" and (OCaml) "Int32.zero" and (Scala) "0" | constant "1 :: uint32" \ (SML) "(Word32.fromInt 1)" and (Haskell) "(1 :: Uint32.Word32)" and (OCaml) "Int32.one" and (Scala) "1" | constant "plus :: uint32 \ _ " \ (SML) "Word32.+ ((_), (_))" and (Haskell) infixl 6 "+" and (OCaml) "Int32.add" and (Scala) infixl 7 "+" | constant "uminus :: uint32 \ _" \ (SML) "Word32.~" and (Haskell) "negate" and (OCaml) "Int32.neg" and (Scala) "!(- _)" | constant "minus :: uint32 \ _" \ (SML) "Word32.- ((_), (_))" and (Haskell) infixl 6 "-" and (OCaml) "Int32.sub" and (Scala) infixl 7 "-" | constant "times :: uint32 \ _ \ _" \ (SML) "Word32.* ((_), (_))" and (Haskell) infixl 7 "*" and (OCaml) "Int32.mul" and (Scala) infixl 8 "*" | constant "HOL.equal :: uint32 \ _ \ bool" \ (SML) "!((_ : Word32.word) = _)" and (Haskell) infix 4 "==" and (OCaml) "(Int32.compare _ _ = 0)" and (Scala) infixl 5 "==" | class_instance uint32 :: equal \ (Haskell) - | constant "less_eq :: uint32 \ _ \ bool" \ (SML) "Word32.<= ((_), (_))" and (Haskell) infix 4 "<=" and (OCaml) "Uint32.less'_eq" and (Scala) "Uint32.less'_eq" | constant "less :: uint32 \ _ \ bool" \ (SML) "Word32.< ((_), (_))" and (Haskell) infix 4 "<" and (OCaml) "Uint32.less" and (Scala) "Uint32.less" | constant "NOT :: uint32 \ _" \ (SML) "Word32.notb" and (Haskell) "Data'_Bits.complement" and (OCaml) "Int32.lognot" and (Scala) "_.unary'_~" | constant "Bit_Operations.and :: uint32 \ _" \ (SML) "Word32.andb ((_),/ (_))" and (Haskell) infixl 7 "Data_Bits..&." and (OCaml) "Int32.logand" and (Scala) infixl 3 "&" | constant "Bit_Operations.or :: uint32 \ _" \ (SML) "Word32.orb ((_),/ (_))" and (Haskell) infixl 5 "Data_Bits..|." and (OCaml) "Int32.logor" and (Scala) infixl 1 "|" | constant "Bit_Operations.xor :: uint32 \ _" \ (SML) "Word32.xorb ((_),/ (_))" and (Haskell) "Data'_Bits.xor" and (OCaml) "Int32.logxor" and (Scala) infixl 2 "^" definition uint32_divmod :: "uint32 \ uint32 \ uint32 \ uint32" where "uint32_divmod x y = (if y = 0 then (undefined ((div) :: uint32 \ _) x (0 :: uint32), undefined ((mod) :: uint32 \ _) x (0 :: uint32)) else (x div y, x mod y))" definition uint32_div :: "uint32 \ uint32 \ uint32" where "uint32_div x y = fst (uint32_divmod x y)" definition uint32_mod :: "uint32 \ uint32 \ uint32" where "uint32_mod x y = snd (uint32_divmod x y)" lemma div_uint32_code [code]: "x div y = (if y = 0 then 0 else uint32_div x y)" including undefined_transfer unfolding uint32_divmod_def uint32_div_def by transfer (simp add: word_div_def) lemma mod_uint32_code [code]: "x mod y = (if y = 0 then x else uint32_mod x y)" including undefined_transfer unfolding uint32_mod_def uint32_divmod_def by transfer (simp add: word_mod_def) definition uint32_sdiv :: "uint32 \ uint32 \ uint32" where [code del]: "uint32_sdiv x y = (if y = 0 then undefined ((div) :: uint32 \ _) x (0 :: uint32) else Abs_uint32 (Rep_uint32 x sdiv Rep_uint32 y))" definition div0_uint32 :: "uint32 \ uint32" where [code del]: "div0_uint32 x = undefined ((div) :: uint32 \ _) x (0 :: uint32)" declare [[code abort: div0_uint32]] definition mod0_uint32 :: "uint32 \ uint32" where [code del]: "mod0_uint32 x = undefined ((mod) :: uint32 \ _) x (0 :: uint32)" declare [[code abort: mod0_uint32]] lemma uint32_divmod_code [code]: "uint32_divmod x y = (if 0x80000000 \ y then if x < y then (0, x) else (1, x - y) else if y = 0 then (div0_uint32 x, mod0_uint32 x) else let q = push_bit 1 (uint32_sdiv (drop_bit 1 x) y); r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" including undefined_transfer unfolding uint32_divmod_def uint32_sdiv_def div0_uint32_def mod0_uint32_def less_eq_uint32.rep_eq apply transfer apply (simp add: divmod_via_sdivmod push_bit_eq_mult) done lemma uint32_sdiv_code [code]: "Rep_uint32 (uint32_sdiv x y) = (if y = 0 then Rep_uint32 (undefined ((div) :: uint32 \ _) x (0 :: uint32)) else Rep_uint32 x sdiv Rep_uint32 y)" unfolding uint32_sdiv_def by(simp add: Abs_uint32_inverse) text \ Note that we only need a translation for signed division, but not for the remainder because @{thm uint32_divmod_code} computes both with division only. \ code_printing constant uint32_div \ (SML) "Word32.div ((_), (_))" and (Haskell) "Prelude.div" | constant uint32_mod \ (SML) "Word32.mod ((_), (_))" and (Haskell) "Prelude.mod" | constant uint32_divmod \ (Haskell) "divmod" | constant uint32_sdiv \ (OCaml) "Int32.div" and (Scala) "_ '/ _" definition uint32_test_bit :: "uint32 \ integer \ bool" where [code del]: "uint32_test_bit x n = (if n < 0 \ 31 < n then undefined (bit :: uint32 \ _) x n else bit x (nat_of_integer n))" lemma test_bit_uint32_code [code]: "bit x n \ n < 32 \ uint32_test_bit x (integer_of_nat n)" including undefined_transfer integer.lifting unfolding uint32_test_bit_def by (transfer, simp, transfer, simp) lemma uint32_test_bit_code [code]: "uint32_test_bit w n = (if n < 0 \ 31 < n then undefined (bit :: uint32 \ _) w n else bit (Rep_uint32 w) (nat_of_integer n))" unfolding uint32_test_bit_def by(simp add: bit_uint32.rep_eq) code_printing constant uint32_test_bit \ (SML) "Uint32.test'_bit" and (Haskell) "Data'_Bits.testBitBounded" and (OCaml) "Uint32.test'_bit" and (Scala) "Uint32.test'_bit" and (Eval) "(fn w => fn n => if n < 0 orelse 32 <= n then raise (Fail \"argument to uint32'_test'_bit out of bounds\") else Uint32.test'_bit w n)" definition uint32_set_bit :: "uint32 \ integer \ bool \ uint32" where [code del]: "uint32_set_bit x n b = (if n < 0 \ 31 < n then undefined (set_bit :: uint32 \ _) x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_uint32_code [code]: "set_bit x n b = (if n < 32 then uint32_set_bit x (integer_of_nat n) b else x)" including undefined_transfer integer.lifting unfolding uint32_set_bit_def by(transfer)(auto cong: conj_cong simp add: not_less set_bit_beyond word_size) lemma uint32_set_bit_code [code]: "Rep_uint32 (uint32_set_bit w n b) = (if n < 0 \ 31 < n then Rep_uint32 (undefined (set_bit :: uint32 \ _) w n b) else set_bit (Rep_uint32 w) (nat_of_integer n) b)" including undefined_transfer unfolding uint32_set_bit_def by transfer simp code_printing constant uint32_set_bit \ (SML) "Uint32.set'_bit" and (Haskell) "Data'_Bits.setBitBounded" and (OCaml) "Uint32.set'_bit" and (Scala) "Uint32.set'_bit" and (Eval) "(fn w => fn n => fn b => if n < 0 orelse 32 <= n then raise (Fail \"argument to uint32'_set'_bit out of bounds\") else Uint32.set'_bit x n b)" definition uint32_shiftl :: "uint32 \ integer \ uint32" where [code del]: "uint32_shiftl x n = (if n < 0 \ 32 \ n then undefined (push_bit :: nat \ uint32 \ _) x n else push_bit (nat_of_integer n) x)" lemma shiftl_uint32_code [code]: "push_bit n x = (if n < 32 then uint32_shiftl x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint32_shiftl_def by transfer simp lemma uint32_shiftl_code [code]: "Rep_uint32 (uint32_shiftl w n) = (if n < 0 \ 32 \ n then Rep_uint32 (undefined (push_bit :: nat \ uint32 \ _) w n) else push_bit (nat_of_integer n) (Rep_uint32 w))" including undefined_transfer unfolding uint32_shiftl_def by transfer simp code_printing constant uint32_shiftl \ (SML) "Uint32.shiftl" and (Haskell) "Data'_Bits.shiftlBounded" and (OCaml) "Uint32.shiftl" and (Scala) "Uint32.shiftl" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 32 then raise Fail \"argument to uint32'_shiftl out of bounds\" else Uint32.shiftl x i)" definition uint32_shiftr :: "uint32 \ integer \ uint32" where [code del]: "uint32_shiftr x n = (if n < 0 \ 32 \ n then undefined (drop_bit :: nat \ uint32 \ _) x n else drop_bit (nat_of_integer n) x)" lemma shiftr_uint32_code [code]: "drop_bit n x = (if n < 32 then uint32_shiftr x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint32_shiftr_def by transfer simp lemma uint32_shiftr_code [code]: "Rep_uint32 (uint32_shiftr w n) = (if n < 0 \ 32 \ n then Rep_uint32 (undefined (drop_bit :: nat \ uint32 \ _) w n) else drop_bit (nat_of_integer n) (Rep_uint32 w))" including undefined_transfer unfolding uint32_shiftr_def by transfer simp code_printing constant uint32_shiftr \ (SML) "Uint32.shiftr" and (Haskell) "Data'_Bits.shiftrBounded" and (OCaml) "Uint32.shiftr" and (Scala) "Uint32.shiftr" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 32 then raise Fail \"argument to uint32'_shiftr out of bounds\" else Uint32.shiftr x i)" definition uint32_sshiftr :: "uint32 \ integer \ uint32" where [code del]: "uint32_sshiftr x n = (if n < 0 \ 32 \ n then undefined signed_drop_bit_uint32 n x else signed_drop_bit_uint32 (nat_of_integer n) x)" lemma sshiftr_uint32_code [code]: "signed_drop_bit_uint32 n x = (if n < 32 then uint32_sshiftr x (integer_of_nat n) else if bit x 31 then - 1 else 0)" including undefined_transfer integer.lifting unfolding uint32_sshiftr_def by transfer (simp add: not_less signed_drop_bit_beyond) lemma uint32_sshiftr_code [code]: "Rep_uint32 (uint32_sshiftr w n) = (if n < 0 \ 32 \ n then Rep_uint32 (undefined signed_drop_bit_uint32 n w) else signed_drop_bit (nat_of_integer n) (Rep_uint32 w))" including undefined_transfer unfolding uint32_sshiftr_def by transfer simp code_printing constant uint32_sshiftr \ (SML) "Uint32.shiftr'_signed" and (Haskell) "(Prelude.fromInteger (Prelude.toInteger (Data'_Bits.shiftrBounded (Prelude.fromInteger (Prelude.toInteger _) :: Uint32.Int32) _)) :: Uint32.Word32)" and (OCaml) "Uint32.shiftr'_signed" and (Scala) "Uint32.shiftr'_signed" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 32 then raise Fail \"argument to uint32'_shiftr'_signed out of bounds\" else Uint32.shiftr'_signed x i)" context includes bit_operations_syntax begin lemma uint32_msb_test_bit: "msb x \ bit (x :: uint32) 31" by transfer (simp add: msb_word_iff_bit) lemma msb_uint32_code [code]: "msb x \ uint32_test_bit x 31" by (simp add: uint32_test_bit_def uint32_msb_test_bit) -lemma uint32_of_int_code [code]: "uint32_of_int i = Uint32 (integer_of_int i)" -including integer.lifting by transfer simp +lemma uint32_of_int_code [code]: + "uint32_of_int i = Uint32 (integer_of_int i)" + including integer.lifting by transfer simp -lemma int_of_uint32_code [code]: "int_of_uint32 x = int_of_integer (integer_of_uint32 x)" -including integer.lifting by transfer simp +lemma int_of_uint32_code [code]: + "int_of_uint32 x = int_of_integer (integer_of_uint32 x)" + including integer.lifting by transfer simp + +lemma uint32_of_nat_code [code]: + "uint32_of_nat = uint32_of_int \ int" + by transfer (simp add: fun_eq_iff) lemma nat_of_uint32_code [code]: "nat_of_uint32 x = nat_of_integer (integer_of_uint32 x)" -unfolding integer_of_uint32_def including integer.lifting by transfer simp + unfolding integer_of_uint32_def including integer.lifting by transfer simp definition integer_of_uint32_signed :: "uint32 \ integer" where "integer_of_uint32_signed n = (if bit n 31 then undefined integer_of_uint32 n else integer_of_uint32 n)" lemma integer_of_uint32_signed_code [code]: "integer_of_uint32_signed n = (if bit n 31 then undefined integer_of_uint32 n else integer_of_int (uint (Rep_uint32' n)))" by (simp add: integer_of_uint32_signed_def integer_of_uint32_def) lemma integer_of_uint32_code [code]: "integer_of_uint32 n = (if bit n 31 then integer_of_uint32_signed (n AND 0x7FFFFFFF) OR 0x80000000 else integer_of_uint32_signed n)" proof - have \integer_of_uint32_signed (n AND 0x7FFFFFFF) OR 0x80000000 = Bit_Operations.set_bit 31 (integer_of_uint32_signed (take_bit 31 n))\ by (simp add: take_bit_eq_mask set_bit_eq_or push_bit_eq_mult mask_eq_exp_minus_1) moreover have \integer_of_uint32 n = Bit_Operations.set_bit 31 (integer_of_uint32 (take_bit 31 n))\ if \bit n 31\ proof (rule bit_eqI) fix m from that show \bit (integer_of_uint32 n) m = bit (Bit_Operations.set_bit 31 (integer_of_uint32 (take_bit 31 n))) m\ for m including integer.lifting by transfer (auto simp add: bit_simps dest: bit_imp_le_length) qed ultimately show ?thesis by simp (simp add: integer_of_uint32_signed_def bit_simps) qed end code_printing constant "integer_of_uint32" \ (SML) "IntInf.fromLarge (Word32.toLargeInt _) : IntInf.int" and (Haskell) "Prelude.toInteger" | constant "integer_of_uint32_signed" \ (OCaml) "Z.of'_int32" and (Scala) "BigInt" section \Quickcheck setup\ definition uint32_of_natural :: "natural \ uint32" where "uint32_of_natural x \ Uint32 (integer_of_natural x)" instantiation uint32 :: "{random, exhaustive, full_exhaustive}" begin definition "random_uint32 \ qc_random_cnv uint32_of_natural" definition "exhaustive_uint32 \ qc_exhaustive_cnv uint32_of_natural" definition "full_exhaustive_uint32 \ qc_full_exhaustive_cnv uint32_of_natural" instance .. end instantiation uint32 :: narrowing begin interpretation quickcheck_narrowing_samples "\i. let x = Uint32 i in (x, 0xFFFFFFFF - x)" "0" "Typerep.Typerep (STR ''Uint32.uint32'') []" . definition "narrowing_uint32 d = qc_narrowing_drawn_from (narrowing_samples d) d" declare [[code drop: "partial_term_of :: uint32 itself \ _"]] lemmas partial_term_of_uint32 [code] = partial_term_of_code instance .. end end diff --git a/thys/Native_Word/Uint64.thy b/thys/Native_Word/Uint64.thy --- a/thys/Native_Word/Uint64.thy +++ b/thys/Native_Word/Uint64.thy @@ -1,883 +1,888 @@ (* Title: Uint64.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of 64 bits\ theory Uint64 imports Code_Target_Word_Base Word_Type_Copies begin text \ PolyML (in version 5.7) provides a Word64 structure only when run in 64-bit mode. Therefore, we by default provide an implementation of 64-bit words using \verb$IntInf.int$ and masking. The code target \texttt{SML\_word} replaces this implementation and maps the operations directly to the \verb$Word64$ structure provided by the Standard ML implementations. The \verb$Eval$ target used by @{command value} and @{method eval} dynamically tests at runtime for the version of PolyML and uses PolyML's Word64 structure if it detects a 64-bit version which does not suffer from a division bug found in PolyML 5.6. \ section \Type definition and primitive operations\ typedef uint64 = \UNIV :: 64 word set\ .. global_interpretation uint64: word_type_copy Abs_uint64 Rep_uint64 using type_definition_uint64 by (rule word_type_copy.intro) setup_lifting type_definition_uint64 declare uint64.of_word_of [code abstype] declare Quotient_uint64 [transfer_rule] instantiation uint64 :: \{comm_ring_1, semiring_modulo, equal, linorder}\ begin lift_definition zero_uint64 :: uint64 is 0 . lift_definition one_uint64 :: uint64 is 1 . lift_definition plus_uint64 :: \uint64 \ uint64 \ uint64\ is \(+)\ . lift_definition uminus_uint64 :: \uint64 \ uint64\ is uminus . lift_definition minus_uint64 :: \uint64 \ uint64 \ uint64\ is \(-)\ . lift_definition times_uint64 :: \uint64 \ uint64 \ uint64\ is \(*)\ . lift_definition divide_uint64 :: \uint64 \ uint64 \ uint64\ is \(div)\ . lift_definition modulo_uint64 :: \uint64 \ uint64 \ uint64\ is \(mod)\ . lift_definition equal_uint64 :: \uint64 \ uint64 \ bool\ is \HOL.equal\ . lift_definition less_eq_uint64 :: \uint64 \ uint64 \ bool\ is \(\)\ . lift_definition less_uint64 :: \uint64 \ uint64 \ bool\ is \(<)\ . global_interpretation uint64: word_type_copy_ring Abs_uint64 Rep_uint64 by standard (fact zero_uint64.rep_eq one_uint64.rep_eq plus_uint64.rep_eq uminus_uint64.rep_eq minus_uint64.rep_eq times_uint64.rep_eq divide_uint64.rep_eq modulo_uint64.rep_eq equal_uint64.rep_eq less_eq_uint64.rep_eq less_uint64.rep_eq)+ instance proof - show \OFCLASS(uint64, comm_ring_1_class)\ by (rule uint64.of_class_comm_ring_1) show \OFCLASS(uint64, semiring_modulo_class)\ by (fact uint64.of_class_semiring_modulo) show \OFCLASS(uint64, equal_class)\ by (fact uint64.of_class_equal) show \OFCLASS(uint64, linorder_class)\ by (fact uint64.of_class_linorder) qed end instantiation uint64 :: ring_bit_operations begin lift_definition bit_uint64 :: \uint64 \ nat \ bool\ is bit . lift_definition not_uint64 :: \uint64 \ uint64\ is NOT . lift_definition and_uint64 :: \uint64 \ uint64 \ uint64\ is \Bit_Operations.and\ . lift_definition or_uint64 :: \uint64 \ uint64 \ uint64\ is \Bit_Operations.or\ . lift_definition xor_uint64 :: \uint64 \ uint64 \ uint64\ is \Bit_Operations.xor\ . lift_definition mask_uint64 :: \nat \ uint64\ is mask . lift_definition push_bit_uint64 :: \nat \ uint64 \ uint64\ is push_bit . lift_definition drop_bit_uint64 :: \nat \ uint64 \ uint64\ is drop_bit . lift_definition signed_drop_bit_uint64 :: \nat \ uint64 \ uint64\ is signed_drop_bit . lift_definition take_bit_uint64 :: \nat \ uint64 \ uint64\ is take_bit . lift_definition set_bit_uint64 :: \nat \ uint64 \ uint64\ is Bit_Operations.set_bit . lift_definition unset_bit_uint64 :: \nat \ uint64 \ uint64\ is unset_bit . lift_definition flip_bit_uint64 :: \nat \ uint64 \ uint64\ is flip_bit . global_interpretation uint64: word_type_copy_bits Abs_uint64 Rep_uint64 signed_drop_bit_uint64 by standard (fact bit_uint64.rep_eq not_uint64.rep_eq and_uint64.rep_eq or_uint64.rep_eq xor_uint64.rep_eq mask_uint64.rep_eq push_bit_uint64.rep_eq drop_bit_uint64.rep_eq signed_drop_bit_uint64.rep_eq take_bit_uint64.rep_eq set_bit_uint64.rep_eq unset_bit_uint64.rep_eq flip_bit_uint64.rep_eq)+ instance by (fact uint64.of_class_ring_bit_operations) end lift_definition uint64_of_nat :: \nat \ uint64\ is word_of_nat . lift_definition nat_of_uint64 :: \uint64 \ nat\ is unat . lift_definition uint64_of_int :: \int \ uint64\ is word_of_int . lift_definition int_of_uint64 :: \uint64 \ int\ is uint . context includes integer.lifting begin lift_definition Uint64 :: \integer \ uint64\ is word_of_int . lift_definition integer_of_uint64 :: \uint64 \ integer\ is uint . end global_interpretation uint64: word_type_copy_more Abs_uint64 Rep_uint64 signed_drop_bit_uint64 uint64_of_nat nat_of_uint64 uint64_of_int int_of_uint64 Uint64 integer_of_uint64 apply standard apply (simp_all add: uint64_of_nat.rep_eq nat_of_uint64.rep_eq uint64_of_int.rep_eq int_of_uint64.rep_eq Uint64.rep_eq integer_of_uint64.rep_eq integer_eq_iff) done instantiation uint64 :: "{size, msb, lsb, set_bit, bit_comprehension}" begin lift_definition size_uint64 :: \uint64 \ nat\ is size . lift_definition msb_uint64 :: \uint64 \ bool\ is msb . lift_definition lsb_uint64 :: \uint64 \ bool\ is lsb . text \Workaround: avoid name space clash by spelling out \<^text>\lift_definition\ explicitly.\ definition set_bit_uint64 :: \uint64 \ nat \ bool \ uint64\ where set_bit_uint64_eq: \set_bit_uint64 a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ context includes lifting_syntax begin lemma set_bit_uint64_transfer [transfer_rule]: \(cr_uint64 ===> (=) ===> (\) ===> cr_uint64) Generic_set_bit.set_bit Generic_set_bit.set_bit\ by (simp only: set_bit_eq [abs_def] set_bit_uint64_eq [abs_def]) transfer_prover end lift_definition set_bits_uint64 :: \(nat \ bool) \ uint64\ is set_bits . lift_definition set_bits_aux_uint64 :: \(nat \ bool) \ nat \ uint64 \ uint64\ is set_bits_aux . global_interpretation uint64: word_type_copy_misc Abs_uint64 Rep_uint64 signed_drop_bit_uint64 uint64_of_nat nat_of_uint64 uint64_of_int int_of_uint64 Uint64 integer_of_uint64 64 set_bits_aux_uint64 by (standard; transfer) simp_all instance using uint64.of_class_bit_comprehension uint64.of_class_set_bit uint64.of_class_lsb by simp_all standard end section \Code setup\ text \ For SML, we generate an implementation of unsigned 64-bit words using \verb$IntInf.int$. If @{ML "LargeWord.wordSize > 63"} of the Isabelle/ML runtime environment holds, then we assume that there is also a \Word64\ structure available and accordingly replace the implementation for the target \verb$Eval$. \ code_printing code_module "Uint64" \ (SML) \(* Test that words can handle numbers between 0 and 63 *) val _ = if 6 <= Word.wordSize then () else raise (Fail ("wordSize less than 6")); structure Uint64 : sig eqtype uint64; val zero : uint64; val one : uint64; val fromInt : IntInf.int -> uint64; val toInt : uint64 -> IntInf.int; val toLarge : uint64 -> LargeWord.word; val fromLarge : LargeWord.word -> uint64 val plus : uint64 -> uint64 -> uint64; val minus : uint64 -> uint64 -> uint64; val times : uint64 -> uint64 -> uint64; val divide : uint64 -> uint64 -> uint64; val modulus : uint64 -> uint64 -> uint64; val negate : uint64 -> uint64; val less_eq : uint64 -> uint64 -> bool; val less : uint64 -> uint64 -> bool; val notb : uint64 -> uint64; val andb : uint64 -> uint64 -> uint64; val orb : uint64 -> uint64 -> uint64; val xorb : uint64 -> uint64 -> uint64; val shiftl : uint64 -> IntInf.int -> uint64; val shiftr : uint64 -> IntInf.int -> uint64; val shiftr_signed : uint64 -> IntInf.int -> uint64; val set_bit : uint64 -> IntInf.int -> bool -> uint64; val test_bit : uint64 -> IntInf.int -> bool; end = struct type uint64 = IntInf.int; val mask = 0xFFFFFFFFFFFFFFFF : IntInf.int; val zero = 0 : IntInf.int; val one = 1 : IntInf.int; fun fromInt x = IntInf.andb(x, mask); fun toInt x = x fun toLarge x = LargeWord.fromLargeInt (IntInf.toLarge x); fun fromLarge x = IntInf.fromLarge (LargeWord.toLargeInt x); fun plus x y = IntInf.andb(IntInf.+(x, y), mask); fun minus x y = IntInf.andb(IntInf.-(x, y), mask); fun negate x = IntInf.andb(IntInf.~(x), mask); fun times x y = IntInf.andb(IntInf.*(x, y), mask); fun divide x y = IntInf.div(x, y); fun modulus x y = IntInf.mod(x, y); fun less_eq x y = IntInf.<=(x, y); fun less x y = IntInf.<(x, y); fun notb x = IntInf.andb(IntInf.notb(x), mask); fun orb x y = IntInf.orb(x, y); fun andb x y = IntInf.andb(x, y); fun xorb x y = IntInf.xorb(x, y); val maxWord = IntInf.pow (2, Word.wordSize); fun shiftl x n = if n < maxWord then IntInf.andb(IntInf.<< (x, Word.fromLargeInt (IntInf.toLarge n)), mask) else 0; fun shiftr x n = if n < maxWord then IntInf.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) else 0; val msb_mask = 0x8000000000000000 : IntInf.int; fun shiftr_signed x i = if IntInf.andb(x, msb_mask) = 0 then shiftr x i else if i >= 64 then 0xFFFFFFFFFFFFFFFF else let val x' = shiftr x i val m' = IntInf.andb(IntInf.<<(mask, Word.max(0w64 - Word.fromLargeInt (IntInf.toLarge i), 0w0)), mask) in IntInf.orb(x', m') end; fun test_bit x n = if n < maxWord then IntInf.andb (x, IntInf.<< (1, Word.fromLargeInt (IntInf.toLarge n))) <> 0 else false; fun set_bit x n b = if n < 64 then if b then IntInf.orb (x, IntInf.<< (1, Word.fromLargeInt (IntInf.toLarge n))) else IntInf.andb (x, IntInf.notb (IntInf.<< (1, Word.fromLargeInt (IntInf.toLarge n)))) else x; end \ code_reserved SML Uint64 setup \ let val polyml64 = LargeWord.wordSize > 63; (* PolyML 5.6 has bugs in its Word64 implementation. We test for one such bug and refrain from using Word64 in that case. Testing is done with dynamic code evaluation such that the compiler does not choke on the Word64 structure, which need not be present in a 32bit environment. *) val error_msg = "Buggy Word64 structure"; val test_code = "val _ = if Word64.div (0w18446744073709551611 : Word64.word, 0w3) = 0w6148914691236517203 then ()\n" ^ "else raise (Fail \"" ^ error_msg ^ "\");"; val f = Exn.interruptible_capture (fn () => ML_Compiler.eval ML_Compiler.flags Position.none (ML_Lex.tokenize test_code)) val use_Word64 = polyml64 andalso (case f () of Exn.Res _ => true | Exn.Exn (e as ERROR m) => if String.isSuffix error_msg m then false else Exn.reraise e | Exn.Exn e => Exn.reraise e) ; val newline = "\n"; val content = "structure Uint64 : sig" ^ newline ^ " eqtype uint64;" ^ newline ^ " val zero : uint64;" ^ newline ^ " val one : uint64;" ^ newline ^ " val fromInt : IntInf.int -> uint64;" ^ newline ^ " val toInt : uint64 -> IntInf.int;" ^ newline ^ " val toLarge : uint64 -> LargeWord.word;" ^ newline ^ " val fromLarge : LargeWord.word -> uint64" ^ newline ^ " val plus : uint64 -> uint64 -> uint64;" ^ newline ^ " val minus : uint64 -> uint64 -> uint64;" ^ newline ^ " val times : uint64 -> uint64 -> uint64;" ^ newline ^ " val divide : uint64 -> uint64 -> uint64;" ^ newline ^ " val modulus : uint64 -> uint64 -> uint64;" ^ newline ^ " val negate : uint64 -> uint64;" ^ newline ^ " val less_eq : uint64 -> uint64 -> bool;" ^ newline ^ " val less : uint64 -> uint64 -> bool;" ^ newline ^ " val notb : uint64 -> uint64;" ^ newline ^ " val andb : uint64 -> uint64 -> uint64;" ^ newline ^ " val orb : uint64 -> uint64 -> uint64;" ^ newline ^ " val xorb : uint64 -> uint64 -> uint64;" ^ newline ^ " val shiftl : uint64 -> IntInf.int -> uint64;" ^ newline ^ " val shiftr : uint64 -> IntInf.int -> uint64;" ^ newline ^ " val shiftr_signed : uint64 -> IntInf.int -> uint64;" ^ newline ^ " val set_bit : uint64 -> IntInf.int -> bool -> uint64;" ^ newline ^ " val test_bit : uint64 -> IntInf.int -> bool;" ^ newline ^ "end = struct" ^ newline ^ "" ^ newline ^ "type uint64 = Word64.word;" ^ newline ^ "" ^ newline ^ "val zero = (0wx0 : uint64);" ^ newline ^ "" ^ newline ^ "val one = (0wx1 : uint64);" ^ newline ^ "" ^ newline ^ "fun fromInt x = Word64.fromLargeInt (IntInf.toLarge x);" ^ newline ^ "" ^ newline ^ "fun toInt x = IntInf.fromLarge (Word64.toLargeInt x);" ^ newline ^ "" ^ newline ^ "fun fromLarge x = Word64.fromLarge x;" ^ newline ^ "" ^ newline ^ "fun toLarge x = Word64.toLarge x;" ^ newline ^ "" ^ newline ^ "fun plus x y = Word64.+(x, y);" ^ newline ^ "" ^ newline ^ "fun minus x y = Word64.-(x, y);" ^ newline ^ "" ^ newline ^ "fun negate x = Word64.~(x);" ^ newline ^ "" ^ newline ^ "fun times x y = Word64.*(x, y);" ^ newline ^ "" ^ newline ^ "fun divide x y = Word64.div(x, y);" ^ newline ^ "" ^ newline ^ "fun modulus x y = Word64.mod(x, y);" ^ newline ^ "" ^ newline ^ "fun less_eq x y = Word64.<=(x, y);" ^ newline ^ "" ^ newline ^ "fun less x y = Word64.<(x, y);" ^ newline ^ "" ^ newline ^ "fun set_bit x n b =" ^ newline ^ " let val mask = Word64.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))" ^ newline ^ " in if b then Word64.orb (x, mask)" ^ newline ^ " else Word64.andb (x, Word64.notb mask)" ^ newline ^ " end" ^ newline ^ "" ^ newline ^ "fun shiftl x n =" ^ newline ^ " Word64.<< (x, Word.fromLargeInt (IntInf.toLarge n))" ^ newline ^ "" ^ newline ^ "fun shiftr x n =" ^ newline ^ " Word64.>> (x, Word.fromLargeInt (IntInf.toLarge n))" ^ newline ^ "" ^ newline ^ "fun shiftr_signed x n =" ^ newline ^ " Word64.~>> (x, Word.fromLargeInt (IntInf.toLarge n))" ^ newline ^ "" ^ newline ^ "fun test_bit x n =" ^ newline ^ " Word64.andb (x, Word64.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))) <> Word64.fromInt 0" ^ newline ^ "" ^ newline ^ "val notb = Word64.notb" ^ newline ^ "" ^ newline ^ "fun andb x y = Word64.andb(x, y);" ^ newline ^ "" ^ newline ^ "fun orb x y = Word64.orb(x, y);" ^ newline ^ "" ^ newline ^ "fun xorb x y = Word64.xorb(x, y);" ^ newline ^ "" ^ newline ^ "end (*struct Uint64*)" val target_SML64 = "SML_word"; in (if use_Word64 then Code_Target.set_printings (Code_Symbol.Module ("Uint64", [(Code_Runtime.target, SOME (content, []))])) else I) #> Code_Target.set_printings (Code_Symbol.Module ("Uint64", [(target_SML64, SOME (content, []))])) end \ code_printing code_module Uint64 \ (Haskell) \module Uint64(Int64, Word64) where import Data.Int(Int64) import Data.Word(Word64)\ code_reserved Haskell Uint64 text \ OCaml and Scala provide only signed 64bit numbers, so we use these and implement sign-sensitive operations like comparisons manually. \ code_printing code_module "Uint64" \ (OCaml) \module Uint64 : sig val less : int64 -> int64 -> bool val less_eq : int64 -> int64 -> bool val set_bit : int64 -> Z.t -> bool -> int64 val shiftl : int64 -> Z.t -> int64 val shiftr : int64 -> Z.t -> int64 val shiftr_signed : int64 -> Z.t -> int64 val test_bit : int64 -> Z.t -> bool end = struct (* negative numbers have their highest bit set, so they are greater than positive ones *) let less x y = if Int64.compare x Int64.zero < 0 then Int64.compare y Int64.zero < 0 && Int64.compare x y < 0 else Int64.compare y Int64.zero < 0 || Int64.compare x y < 0;; let less_eq x y = if Int64.compare x Int64.zero < 0 then Int64.compare y Int64.zero < 0 && Int64.compare x y <= 0 else Int64.compare y Int64.zero < 0 || Int64.compare x y <= 0;; let set_bit x n b = let mask = Int64.shift_left Int64.one (Z.to_int n) in if b then Int64.logor x mask else Int64.logand x (Int64.lognot mask);; let shiftl x n = Int64.shift_left x (Z.to_int n);; let shiftr x n = Int64.shift_right_logical x (Z.to_int n);; let shiftr_signed x n = Int64.shift_right x (Z.to_int n);; let test_bit x n = Int64.compare (Int64.logand x (Int64.shift_left Int64.one (Z.to_int n))) Int64.zero <> 0;; end;; (*struct Uint64*)\ code_reserved OCaml Uint64 code_printing code_module Uint64 \ (Scala) \object Uint64 { def less(x: Long, y: Long) : Boolean = if (x < 0) y < 0 && x < y else y < 0 || x < y def less_eq(x: Long, y: Long) : Boolean = if (x < 0) y < 0 && x <= y else y < 0 || x <= y def set_bit(x: Long, n: BigInt, b: Boolean) : Long = if (b) x | (1L << n.intValue) else x & (1L << n.intValue).unary_~ def shiftl(x: Long, n: BigInt) : Long = x << n.intValue def shiftr(x: Long, n: BigInt) : Long = x >>> n.intValue def shiftr_signed(x: Long, n: BigInt) : Long = x >> n.intValue def test_bit(x: Long, n: BigInt) : Boolean = (x & (1L << n.intValue)) != 0 } /* object Uint64 */\ code_reserved Scala Uint64 text \ OCaml's conversion from Big\_int to int64 demands that the value fits int a signed 64-bit integer. The following justifies the implementation. \ context includes bit_operations_syntax begin definition Uint64_signed :: "integer \ uint64" where "Uint64_signed i = (if i < -(0x8000000000000000) \ i \ 0x8000000000000000 then undefined Uint64 i else Uint64 i)" lemma Uint64_code [code]: "Uint64 i = (let i' = i AND 0xFFFFFFFFFFFFFFFF in if bit i' 63 then Uint64_signed (i' - 0x10000000000000000) else Uint64_signed i')" including undefined_transfer integer.lifting unfolding Uint64_signed_def apply transfer apply (subst word_of_int_via_signed) apply (auto simp add: push_bit_of_1 mask_eq_exp_minus_1 word_of_int_via_signed cong del: if_cong) done lemma Uint64_signed_code [code]: "Rep_uint64 (Uint64_signed i) = (if i < -(0x8000000000000000) \ i \ 0x8000000000000000 then Rep_uint64 (undefined Uint64 i) else word_of_int (int_of_integer_symbolic i))" unfolding Uint64_signed_def Uint64_def int_of_integer_symbolic_def word_of_integer_def by(simp add: Abs_uint64_inverse) end text \ Avoid @{term Abs_uint64} in generated code, use @{term Rep_uint64'} instead. The symbolic implementations for code\_simp use @{term Rep_uint64}. The new destructor @{term Rep_uint64'} is executable. As the simplifier is given the [code abstract] equations literally, we cannot implement @{term Rep_uint64} directly, because that makes code\_simp loop. If code generation raises Match, some equation probably contains @{term Rep_uint64} ([code abstract] equations for @{typ uint64} may use @{term Rep_uint64} because these instances will be folded away.) To convert @{typ "64 word"} values into @{typ uint64}, use @{term "Abs_uint64'"}. \ definition Rep_uint64' where [simp]: "Rep_uint64' = Rep_uint64" lemma Rep_uint64'_transfer [transfer_rule]: "rel_fun cr_uint64 (=) (\x. x) Rep_uint64'" unfolding Rep_uint64'_def by(rule uint64.rep_transfer) lemma Rep_uint64'_code [code]: "Rep_uint64' x = (BITS n. bit x n)" by transfer (simp add: set_bits_bit_eq) lift_definition Abs_uint64' :: "64 word \ uint64" is "\x :: 64 word. x" . lemma Abs_uint64'_code [code]: "Abs_uint64' x = Uint64 (integer_of_int (uint x))" including integer.lifting by transfer simp declare [[code drop: "term_of_class.term_of :: uint64 \ _"]] lemma term_of_uint64_code [code]: defines "TR \ typerep.Typerep" and "bit0 \ STR ''Numeral_Type.bit0''" shows "term_of_class.term_of x = Code_Evaluation.App (Code_Evaluation.Const (STR ''Uint64.uint64.Abs_uint64'') (TR (STR ''fun'') [TR (STR ''Word.word'') [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR bit0 [TR (STR ''Numeral_Type.num1'') []]]]]]]], TR (STR ''Uint64.uint64'') []])) (term_of_class.term_of (Rep_uint64' x))" by(simp add: term_of_anything) code_printing type_constructor uint64 \ (SML) "Uint64.uint64" and (Haskell) "Uint64.Word64" and (OCaml) "int64" and (Scala) "Long" | constant Uint64 \ (SML) "Uint64.fromInt" and (Haskell) "(Prelude.fromInteger _ :: Uint64.Word64)" and (Haskell_Quickcheck) "(Prelude.fromInteger (Prelude.toInteger _) :: Uint64.Word64)" and (Scala) "_.longValue" | constant Uint64_signed \ (OCaml) "Z.to'_int64" | constant "0 :: uint64" \ (SML) "Uint64.zero" and (Haskell) "(0 :: Uint64.Word64)" and (OCaml) "Int64.zero" and (Scala) "0" | constant "1 :: uint64" \ (SML) "Uint64.one" and (Haskell) "(1 :: Uint64.Word64)" and (OCaml) "Int64.one" and (Scala) "1" | constant "plus :: uint64 \ _ " \ (SML) "Uint64.plus" and (Haskell) infixl 6 "+" and (OCaml) "Int64.add" and (Scala) infixl 7 "+" | constant "uminus :: uint64 \ _" \ (SML) "Uint64.negate" and (Haskell) "negate" and (OCaml) "Int64.neg" and (Scala) "!(- _)" | constant "minus :: uint64 \ _" \ (SML) "Uint64.minus" and (Haskell) infixl 6 "-" and (OCaml) "Int64.sub" and (Scala) infixl 7 "-" | constant "times :: uint64 \ _ \ _" \ (SML) "Uint64.times" and (Haskell) infixl 7 "*" and (OCaml) "Int64.mul" and (Scala) infixl 8 "*" | constant "HOL.equal :: uint64 \ _ \ bool" \ (SML) "!((_ : Uint64.uint64) = _)" and (Haskell) infix 4 "==" and (OCaml) "(Int64.compare _ _ = 0)" and (Scala) infixl 5 "==" | class_instance uint64 :: equal \ (Haskell) - | constant "less_eq :: uint64 \ _ \ bool" \ (SML) "Uint64.less'_eq" and (Haskell) infix 4 "<=" and (OCaml) "Uint64.less'_eq" and (Scala) "Uint64.less'_eq" | constant "less :: uint64 \ _ \ bool" \ (SML) "Uint64.less" and (Haskell) infix 4 "<" and (OCaml) "Uint64.less" and (Scala) "Uint64.less" | constant "NOT :: uint64 \ _" \ (SML) "Uint64.notb" and (Haskell) "Data'_Bits.complement" and (OCaml) "Int64.lognot" and (Scala) "_.unary'_~" | constant "Bit_Operations.and :: uint64 \ _" \ (SML) "Uint64.andb" and (Haskell) infixl 7 "Data_Bits..&." and (OCaml) "Int64.logand" and (Scala) infixl 3 "&" | constant "Bit_Operations.or :: uint64 \ _" \ (SML) "Uint64.orb" and (Haskell) infixl 5 "Data_Bits..|." and (OCaml) "Int64.logor" and (Scala) infixl 1 "|" | constant "Bit_Operations.xor :: uint64 \ _" \ (SML) "Uint64.xorb" and (Haskell) "Data'_Bits.xor" and (OCaml) "Int64.logxor" and (Scala) infixl 2 "^" definition uint64_divmod :: "uint64 \ uint64 \ uint64 \ uint64" where "uint64_divmod x y = (if y = 0 then (undefined ((div) :: uint64 \ _) x (0 :: uint64), undefined ((mod) :: uint64 \ _) x (0 :: uint64)) else (x div y, x mod y))" definition uint64_div :: "uint64 \ uint64 \ uint64" where "uint64_div x y = fst (uint64_divmod x y)" definition uint64_mod :: "uint64 \ uint64 \ uint64" where "uint64_mod x y = snd (uint64_divmod x y)" lemma div_uint64_code [code]: "x div y = (if y = 0 then 0 else uint64_div x y)" including undefined_transfer unfolding uint64_divmod_def uint64_div_def by transfer (simp add: word_div_def) lemma mod_uint64_code [code]: "x mod y = (if y = 0 then x else uint64_mod x y)" including undefined_transfer unfolding uint64_mod_def uint64_divmod_def by transfer (simp add: word_mod_def) definition uint64_sdiv :: "uint64 \ uint64 \ uint64" where [code del]: "uint64_sdiv x y = (if y = 0 then undefined ((div) :: uint64 \ _) x (0 :: uint64) else Abs_uint64 (Rep_uint64 x sdiv Rep_uint64 y))" definition div0_uint64 :: "uint64 \ uint64" where [code del]: "div0_uint64 x = undefined ((div) :: uint64 \ _) x (0 :: uint64)" declare [[code abort: div0_uint64]] definition mod0_uint64 :: "uint64 \ uint64" where [code del]: "mod0_uint64 x = undefined ((mod) :: uint64 \ _) x (0 :: uint64)" declare [[code abort: mod0_uint64]] lemma uint64_divmod_code [code]: "uint64_divmod x y = (if 0x8000000000000000 \ y then if x < y then (0, x) else (1, x - y) else if y = 0 then (div0_uint64 x, mod0_uint64 x) else let q = push_bit 1 (uint64_sdiv (drop_bit 1 x) y); r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" including undefined_transfer unfolding uint64_divmod_def uint64_sdiv_def div0_uint64_def mod0_uint64_def less_eq_uint64.rep_eq apply transfer apply (simp add: divmod_via_sdivmod push_bit_eq_mult) done lemma uint64_sdiv_code [code]: "Rep_uint64 (uint64_sdiv x y) = (if y = 0 then Rep_uint64 (undefined ((div) :: uint64 \ _) x (0 :: uint64)) else Rep_uint64 x sdiv Rep_uint64 y)" unfolding uint64_sdiv_def by(simp add: Abs_uint64_inverse) text \ Note that we only need a translation for signed division, but not for the remainder because @{thm uint64_divmod_code} computes both with division only. \ code_printing constant uint64_div \ (SML) "Uint64.divide" and (Haskell) "Prelude.div" | constant uint64_mod \ (SML) "Uint64.modulus" and (Haskell) "Prelude.mod" | constant uint64_divmod \ (Haskell) "divmod" | constant uint64_sdiv \ (OCaml) "Int64.div" and (Scala) "_ '/ _" definition uint64_test_bit :: "uint64 \ integer \ bool" where [code del]: "uint64_test_bit x n = (if n < 0 \ 63 < n then undefined (bit :: uint64 \ _) x n else bit x (nat_of_integer n))" lemma bit_uint64_code [code]: "bit x n \ n < 64 \ uint64_test_bit x (integer_of_nat n)" including undefined_transfer integer.lifting unfolding uint64_test_bit_def by transfer (auto dest: bit_imp_le_length) lemma uint64_test_bit_code [code]: "uint64_test_bit w n = (if n < 0 \ 63 < n then undefined (bit :: uint64 \ _) w n else bit (Rep_uint64 w) (nat_of_integer n))" unfolding uint64_test_bit_def by(simp add: bit_uint64.rep_eq) code_printing constant uint64_test_bit \ (SML) "Uint64.test'_bit" and (Haskell) "Data'_Bits.testBitBounded" and (OCaml) "Uint64.test'_bit" and (Scala) "Uint64.test'_bit" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 64 then raise (Fail \"argument to uint64'_test'_bit out of bounds\") else Uint64.test'_bit x i)" definition uint64_set_bit :: "uint64 \ integer \ bool \ uint64" where [code del]: "uint64_set_bit x n b = (if n < 0 \ 63 < n then undefined (set_bit :: uint64 \ _) x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_uint64_code [code]: "set_bit x n b = (if n < 64 then uint64_set_bit x (integer_of_nat n) b else x)" including undefined_transfer integer.lifting unfolding uint64_set_bit_def by(transfer)(auto cong: conj_cong simp add: not_less set_bit_beyond word_size) lemma uint64_set_bit_code [code]: "Rep_uint64 (uint64_set_bit w n b) = (if n < 0 \ 63 < n then Rep_uint64 (undefined (set_bit :: uint64 \ _) w n b) else set_bit (Rep_uint64 w) (nat_of_integer n) b)" including undefined_transfer unfolding uint64_set_bit_def by transfer simp code_printing constant uint64_set_bit \ (SML) "Uint64.set'_bit" and (Haskell) "Data'_Bits.setBitBounded" and (OCaml) "Uint64.set'_bit" and (Scala) "Uint64.set'_bit" and (Eval) "(fn x => fn i => fn b => if i < 0 orelse i >= 64 then raise (Fail \"argument to uint64'_set'_bit out of bounds\") else Uint64.set'_bit x i b)" definition uint64_shiftl :: "uint64 \ integer \ uint64" where [code del]: "uint64_shiftl x n = (if n < 0 \ 64 \ n then undefined (push_bit :: nat \ uint64 \ _) x n else push_bit (nat_of_integer n) x)" lemma shiftl_uint64_code [code]: "push_bit n x = (if n < 64 then uint64_shiftl x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint64_shiftl_def by transfer simp lemma uint64_shiftl_code [code]: "Rep_uint64 (uint64_shiftl w n) = (if n < 0 \ 64 \ n then Rep_uint64 (undefined (push_bit :: nat \ uint64 \ _) w n) else push_bit (nat_of_integer n) (Rep_uint64 w))" including undefined_transfer unfolding uint64_shiftl_def by transfer simp code_printing constant uint64_shiftl \ (SML) "Uint64.shiftl" and (Haskell) "Data'_Bits.shiftlBounded" and (OCaml) "Uint64.shiftl" and (Scala) "Uint64.shiftl" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 64 then raise (Fail \"argument to uint64'_shiftl out of bounds\") else Uint64.shiftl x i)" definition uint64_shiftr :: "uint64 \ integer \ uint64" where [code del]: "uint64_shiftr x n = (if n < 0 \ 64 \ n then undefined (drop_bit :: nat \ uint64 \ _) x n else drop_bit (nat_of_integer n) x)" lemma shiftr_uint64_code [code]: "drop_bit n x = (if n < 64 then uint64_shiftr x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint64_shiftr_def by transfer simp lemma uint64_shiftr_code [code]: "Rep_uint64 (uint64_shiftr w n) = (if n < 0 \ 64 \ n then Rep_uint64 (undefined (drop_bit :: nat \ uint64 \ _) w n) else drop_bit (nat_of_integer n) (Rep_uint64 w))" including undefined_transfer unfolding uint64_shiftr_def by transfer simp code_printing constant uint64_shiftr \ (SML) "Uint64.shiftr" and (Haskell) "Data'_Bits.shiftrBounded" and (OCaml) "Uint64.shiftr" and (Scala) "Uint64.shiftr" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 64 then raise (Fail \"argument to uint64'_shiftr out of bounds\") else Uint64.shiftr x i)" definition uint64_sshiftr :: "uint64 \ integer \ uint64" where [code del]: "uint64_sshiftr x n = (if n < 0 \ 64 \ n then undefined signed_drop_bit_uint64 n x else signed_drop_bit_uint64 (nat_of_integer n) x)" lemma sshiftr_uint64_code [code]: "signed_drop_bit_uint64 n x = (if n < 64 then uint64_sshiftr x (integer_of_nat n) else if bit x 63 then - 1 else 0)" including undefined_transfer integer.lifting unfolding uint64_sshiftr_def by transfer (simp add: not_less signed_drop_bit_beyond) lemma uint64_sshiftr_code [code]: "Rep_uint64 (uint64_sshiftr w n) = (if n < 0 \ 64 \ n then Rep_uint64 (undefined signed_drop_bit_uint64 n w) else signed_drop_bit (nat_of_integer n) (Rep_uint64 w))" including undefined_transfer unfolding uint64_sshiftr_def by transfer simp code_printing constant uint64_sshiftr \ (SML) "Uint64.shiftr'_signed" and (Haskell) "(Prelude.fromInteger (Prelude.toInteger (Data'_Bits.shiftrBounded (Prelude.fromInteger (Prelude.toInteger _) :: Uint64.Int64) _)) :: Uint64.Word64)" and (OCaml) "Uint64.shiftr'_signed" and (Scala) "Uint64.shiftr'_signed" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 64 then raise (Fail \"argument to uint64'_shiftr'_signed out of bounds\") else Uint64.shiftr'_signed x i)" context includes bit_operations_syntax begin lemma uint64_msb_test_bit: "msb x \ bit (x :: uint64) 63" by transfer (simp add: msb_word_iff_bit) lemma msb_uint64_code [code]: "msb x \ uint64_test_bit x 63" by (simp add: uint64_test_bit_def uint64_msb_test_bit) -lemma uint64_of_int_code [code]: "uint64_of_int i = Uint64 (integer_of_int i)" -including integer.lifting by transfer simp +lemma uint64_of_int_code [code]: + "uint64_of_int i = Uint64 (integer_of_int i)" + including integer.lifting by transfer simp lemma int_of_uint64_code [code]: "int_of_uint64 x = int_of_integer (integer_of_uint64 x)" -including integer.lifting by transfer simp + including integer.lifting by transfer simp + +lemma uint64_of_nat_code [code]: + "uint64_of_nat = uint64_of_int \ int" + by transfer (simp add: fun_eq_iff) lemma nat_of_uint64_code [code]: "nat_of_uint64 x = nat_of_integer (integer_of_uint64 x)" -unfolding integer_of_uint64_def including integer.lifting by transfer simp + unfolding integer_of_uint64_def including integer.lifting by transfer simp definition integer_of_uint64_signed :: "uint64 \ integer" where "integer_of_uint64_signed n = (if bit n 63 then undefined integer_of_uint64 n else integer_of_uint64 n)" lemma integer_of_uint64_signed_code [code]: "integer_of_uint64_signed n = (if bit n 63 then undefined integer_of_uint64 n else integer_of_int (uint (Rep_uint64' n)))" by (simp add: integer_of_uint64_signed_def integer_of_uint64_def) lemma integer_of_uint64_code [code]: "integer_of_uint64 n = (if bit n 63 then integer_of_uint64_signed (n AND 0x7FFFFFFFFFFFFFFF) OR 0x8000000000000000 else integer_of_uint64_signed n)" proof - have \integer_of_uint64_signed (n AND 0x7FFFFFFFFFFFFFFF) OR 0x8000000000000000 = Bit_Operations.set_bit 63 (integer_of_uint64_signed (take_bit 63 n))\ by (simp add: take_bit_eq_mask set_bit_eq_or push_bit_eq_mult mask_eq_exp_minus_1) moreover have \integer_of_uint64 n = Bit_Operations.set_bit 63 (integer_of_uint64 (take_bit 63 n))\ if \bit n 63\ proof (rule bit_eqI) fix m from that show \bit (integer_of_uint64 n) m = bit (Bit_Operations.set_bit 63 (integer_of_uint64 (take_bit 63 n))) m\ for m including integer.lifting by transfer (auto simp add: bit_simps dest: bit_imp_le_length) qed ultimately show ?thesis by simp (simp add: integer_of_uint64_signed_def bit_simps) qed end code_printing constant "integer_of_uint64" \ (SML) "Uint64.toInt" and (Haskell) "Prelude.toInteger" | constant "integer_of_uint64_signed" \ (OCaml) "Z.of'_int64" and (Scala) "BigInt" section \Quickcheck setup\ definition uint64_of_natural :: "natural \ uint64" where "uint64_of_natural x \ Uint64 (integer_of_natural x)" instantiation uint64 :: "{random, exhaustive, full_exhaustive}" begin definition "random_uint64 \ qc_random_cnv uint64_of_natural" definition "exhaustive_uint64 \ qc_exhaustive_cnv uint64_of_natural" definition "full_exhaustive_uint64 \ qc_full_exhaustive_cnv uint64_of_natural" instance .. end instantiation uint64 :: narrowing begin interpretation quickcheck_narrowing_samples "\i. let x = Uint64 i in (x, 0xFFFFFFFFFFFFFFFF - x)" "0" "Typerep.Typerep (STR ''Uint64.uint64'') []" . definition "narrowing_uint64 d = qc_narrowing_drawn_from (narrowing_samples d) d" declare [[code drop: "partial_term_of :: uint64 itself \ _"]] lemmas partial_term_of_uint64 [code] = partial_term_of_code instance .. end end diff --git a/thys/Native_Word/Uint8.thy b/thys/Native_Word/Uint8.thy --- a/thys/Native_Word/Uint8.thy +++ b/thys/Native_Word/Uint8.thy @@ -1,605 +1,610 @@ (* Title: Uint8.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Unsigned words of 8 bits\ theory Uint8 imports Code_Target_Word_Base Word_Type_Copies begin text \ Restriction for OCaml code generation: OCaml does not provide an int8 type, so no special code generation for this type is set up. If the theory \Code_Target_Bits_Int\ is imported, the type \uint8\ is emulated via \<^typ>\8 word\. \ section \Type definition and primitive operations\ typedef uint8 = \UNIV :: 8 word set\ .. global_interpretation uint8: word_type_copy Abs_uint8 Rep_uint8 using type_definition_uint8 by (rule word_type_copy.intro) setup_lifting type_definition_uint8 declare uint8.of_word_of [code abstype] declare Quotient_uint8 [transfer_rule] instantiation uint8 :: \{comm_ring_1, semiring_modulo, equal, linorder}\ begin lift_definition zero_uint8 :: uint8 is 0 . lift_definition one_uint8 :: uint8 is 1 . lift_definition plus_uint8 :: \uint8 \ uint8 \ uint8\ is \(+)\ . lift_definition uminus_uint8 :: \uint8 \ uint8\ is uminus . lift_definition minus_uint8 :: \uint8 \ uint8 \ uint8\ is \(-)\ . lift_definition times_uint8 :: \uint8 \ uint8 \ uint8\ is \(*)\ . lift_definition divide_uint8 :: \uint8 \ uint8 \ uint8\ is \(div)\ . lift_definition modulo_uint8 :: \uint8 \ uint8 \ uint8\ is \(mod)\ . lift_definition equal_uint8 :: \uint8 \ uint8 \ bool\ is \HOL.equal\ . lift_definition less_eq_uint8 :: \uint8 \ uint8 \ bool\ is \(\)\ . lift_definition less_uint8 :: \uint8 \ uint8 \ bool\ is \(<)\ . global_interpretation uint8: word_type_copy_ring Abs_uint8 Rep_uint8 by standard (fact zero_uint8.rep_eq one_uint8.rep_eq plus_uint8.rep_eq uminus_uint8.rep_eq minus_uint8.rep_eq times_uint8.rep_eq divide_uint8.rep_eq modulo_uint8.rep_eq equal_uint8.rep_eq less_eq_uint8.rep_eq less_uint8.rep_eq)+ instance proof - show \OFCLASS(uint8, comm_ring_1_class)\ by (rule uint8.of_class_comm_ring_1) show \OFCLASS(uint8, semiring_modulo_class)\ by (fact uint8.of_class_semiring_modulo) show \OFCLASS(uint8, equal_class)\ by (fact uint8.of_class_equal) show \OFCLASS(uint8, linorder_class)\ by (fact uint8.of_class_linorder) qed end instantiation uint8 :: ring_bit_operations begin lift_definition bit_uint8 :: \uint8 \ nat \ bool\ is bit . lift_definition not_uint8 :: \uint8 \ uint8\ is NOT . lift_definition and_uint8 :: \uint8 \ uint8 \ uint8\ is \Bit_Operations.and\ . lift_definition or_uint8 :: \uint8 \ uint8 \ uint8\ is \Bit_Operations.or\ . lift_definition xor_uint8 :: \uint8 \ uint8 \ uint8\ is \Bit_Operations.xor\ . lift_definition mask_uint8 :: \nat \ uint8\ is mask . lift_definition push_bit_uint8 :: \nat \ uint8 \ uint8\ is push_bit . lift_definition drop_bit_uint8 :: \nat \ uint8 \ uint8\ is drop_bit . lift_definition signed_drop_bit_uint8 :: \nat \ uint8 \ uint8\ is signed_drop_bit . lift_definition take_bit_uint8 :: \nat \ uint8 \ uint8\ is take_bit . lift_definition set_bit_uint8 :: \nat \ uint8 \ uint8\ is Bit_Operations.set_bit . lift_definition unset_bit_uint8 :: \nat \ uint8 \ uint8\ is unset_bit . lift_definition flip_bit_uint8 :: \nat \ uint8 \ uint8\ is flip_bit . global_interpretation uint8: word_type_copy_bits Abs_uint8 Rep_uint8 signed_drop_bit_uint8 by standard (fact bit_uint8.rep_eq not_uint8.rep_eq and_uint8.rep_eq or_uint8.rep_eq xor_uint8.rep_eq mask_uint8.rep_eq push_bit_uint8.rep_eq drop_bit_uint8.rep_eq signed_drop_bit_uint8.rep_eq take_bit_uint8.rep_eq set_bit_uint8.rep_eq unset_bit_uint8.rep_eq flip_bit_uint8.rep_eq)+ instance by (fact uint8.of_class_ring_bit_operations) end lift_definition uint8_of_nat :: \nat \ uint8\ is word_of_nat . lift_definition nat_of_uint8 :: \uint8 \ nat\ is unat . lift_definition uint8_of_int :: \int \ uint8\ is word_of_int . lift_definition int_of_uint8 :: \uint8 \ int\ is uint . context includes integer.lifting begin lift_definition Uint8 :: \integer \ uint8\ is word_of_int . lift_definition integer_of_uint8 :: \uint8 \ integer\ is uint . end global_interpretation uint8: word_type_copy_more Abs_uint8 Rep_uint8 signed_drop_bit_uint8 uint8_of_nat nat_of_uint8 uint8_of_int int_of_uint8 Uint8 integer_of_uint8 apply standard apply (simp_all add: uint8_of_nat.rep_eq nat_of_uint8.rep_eq uint8_of_int.rep_eq int_of_uint8.rep_eq Uint8.rep_eq integer_of_uint8.rep_eq integer_eq_iff) done instantiation uint8 :: "{size, msb, lsb, set_bit, bit_comprehension}" begin lift_definition size_uint8 :: \uint8 \ nat\ is size . lift_definition msb_uint8 :: \uint8 \ bool\ is msb . lift_definition lsb_uint8 :: \uint8 \ bool\ is lsb . text \Workaround: avoid name space clash by spelling out \<^text>\lift_definition\ explicitly.\ definition set_bit_uint8 :: \uint8 \ nat \ bool \ uint8\ where set_bit_uint8_eq: \set_bit_uint8 a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ context includes lifting_syntax begin lemma set_bit_uint8_transfer [transfer_rule]: \(cr_uint8 ===> (=) ===> (\) ===> cr_uint8) Generic_set_bit.set_bit Generic_set_bit.set_bit\ by (simp only: set_bit_eq [abs_def] set_bit_uint8_eq [abs_def]) transfer_prover end lift_definition set_bits_uint8 :: \(nat \ bool) \ uint8\ is set_bits . lift_definition set_bits_aux_uint8 :: \(nat \ bool) \ nat \ uint8 \ uint8\ is set_bits_aux . global_interpretation uint8: word_type_copy_misc Abs_uint8 Rep_uint8 signed_drop_bit_uint8 uint8_of_nat nat_of_uint8 uint8_of_int int_of_uint8 Uint8 integer_of_uint8 8 set_bits_aux_uint8 by (standard; transfer) simp_all instance using uint8.of_class_bit_comprehension uint8.of_class_set_bit uint8.of_class_lsb by simp_all standard end section \Code setup\ code_printing code_module Uint8 \ (SML) \(* Test that words can handle numbers between 0 and 3 *) val _ = if 3 <= Word.wordSize then () else raise (Fail ("wordSize less than 3")); structure Uint8 : sig val set_bit : Word8.word -> IntInf.int -> bool -> Word8.word val shiftl : Word8.word -> IntInf.int -> Word8.word val shiftr : Word8.word -> IntInf.int -> Word8.word val shiftr_signed : Word8.word -> IntInf.int -> Word8.word val test_bit : Word8.word -> IntInf.int -> bool end = struct fun set_bit x n b = let val mask = Word8.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n)) in if b then Word8.orb (x, mask) else Word8.andb (x, Word8.notb mask) end fun shiftl x n = Word8.<< (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr x n = Word8.>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun shiftr_signed x n = Word8.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) fun test_bit x n = Word8.andb (x, Word8.<< (0wx1, Word.fromLargeInt (IntInf.toLarge n))) <> Word8.fromInt 0 end; (* struct Uint8 *)\ code_reserved SML Uint8 code_printing code_module Uint8 \ (Haskell) \module Uint8(Int8, Word8) where import Data.Int(Int8) import Data.Word(Word8)\ code_reserved Haskell Uint8 text \ Scala provides only signed 8bit numbers, so we use these and implement sign-sensitive operations like comparisons manually. \ code_printing code_module Uint8 \ (Scala) \object Uint8 { def less(x: Byte, y: Byte) : Boolean = if (x < 0) y < 0 && x < y else y < 0 || x < y def less_eq(x: Byte, y: Byte) : Boolean = if (x < 0) y < 0 && x <= y else y < 0 || x <= y def set_bit(x: Byte, n: BigInt, b: Boolean) : Byte = if (b) (x | (1 << n.intValue)).toByte else (x & (1 << n.intValue).unary_~).toByte def shiftl(x: Byte, n: BigInt) : Byte = (x << n.intValue).toByte def shiftr(x: Byte, n: BigInt) : Byte = ((x & 255) >>> n.intValue).toByte def shiftr_signed(x: Byte, n: BigInt) : Byte = (x >> n.intValue).toByte def test_bit(x: Byte, n: BigInt) : Boolean = (x & (1 << n.intValue)) != 0 } /* object Uint8 */\ code_reserved Scala Uint8 text \ Avoid @{term Abs_uint8} in generated code, use @{term Rep_uint8'} instead. The symbolic implementations for code\_simp use @{term Rep_uint8}. The new destructor @{term Rep_uint8'} is executable. As the simplifier is given the [code abstract] equations literally, we cannot implement @{term Rep_uint8} directly, because that makes code\_simp loop. If code generation raises Match, some equation probably contains @{term Rep_uint8} ([code abstract] equations for @{typ uint8} may use @{term Rep_uint8} because these instances will be folded away.) To convert @{typ "8 word"} values into @{typ uint8}, use @{term "Abs_uint8'"}. \ definition Rep_uint8' where [simp]: "Rep_uint8' = Rep_uint8" lemma Rep_uint8'_transfer [transfer_rule]: "rel_fun cr_uint8 (=) (\x. x) Rep_uint8'" unfolding Rep_uint8'_def by(rule uint8.rep_transfer) lemma Rep_uint8'_code [code]: "Rep_uint8' x = (BITS n. bit x n)" by transfer (simp add: set_bits_bit_eq) lift_definition Abs_uint8' :: "8 word \ uint8" is "\x :: 8 word. x" . lemma Abs_uint8'_code [code]: "Abs_uint8' x = Uint8 (integer_of_int (uint x))" including integer.lifting by transfer simp declare [[code drop: "term_of_class.term_of :: uint8 \ _"]] lemma term_of_uint8_code [code]: defines "TR \ typerep.Typerep" and "bit0 \ STR ''Numeral_Type.bit0''" shows "term_of_class.term_of x = Code_Evaluation.App (Code_Evaluation.Const (STR ''Uint8.uint8.Abs_uint8'') (TR (STR ''fun'') [TR (STR ''Word.word'') [TR bit0 [TR bit0 [TR bit0 [TR (STR ''Numeral_Type.num1'') []]]]], TR (STR ''Uint8.uint8'') []])) (term_of_class.term_of (Rep_uint8' x))" by(simp add: term_of_anything) lemma Uin8_code [code]: "Rep_uint8 (Uint8 i) = word_of_int (int_of_integer_symbolic i)" unfolding Uint8_def int_of_integer_symbolic_def by(simp add: Abs_uint8_inverse) code_printing type_constructor uint8 \ (SML) "Word8.word" and (Haskell) "Uint8.Word8" and (Scala) "Byte" | constant Uint8 \ (SML) "Word8.fromLargeInt (IntInf.toLarge _)" and (Haskell) "(Prelude.fromInteger _ :: Uint8.Word8)" and (Haskell_Quickcheck) "(Prelude.fromInteger (Prelude.toInteger _) :: Uint8.Word8)" and (Scala) "_.byteValue" | constant "0 :: uint8" \ (SML) "(Word8.fromInt 0)" and (Haskell) "(0 :: Uint8.Word8)" and (Scala) "0.toByte" | constant "1 :: uint8" \ (SML) "(Word8.fromInt 1)" and (Haskell) "(1 :: Uint8.Word8)" and (Scala) "1.toByte" | constant "plus :: uint8 \ _ \ _" \ (SML) "Word8.+ ((_), (_))" and (Haskell) infixl 6 "+" and (Scala) "(_ +/ _).toByte" | constant "uminus :: uint8 \ _" \ (SML) "Word8.~" and (Haskell) "negate" and (Scala) "(- _).toByte" | constant "minus :: uint8 \ _" \ (SML) "Word8.- ((_), (_))" and (Haskell) infixl 6 "-" and (Scala) "(_ -/ _).toByte" | constant "times :: uint8 \ _ \ _" \ (SML) "Word8.* ((_), (_))" and (Haskell) infixl 7 "*" and (Scala) "(_ */ _).toByte" | constant "HOL.equal :: uint8 \ _ \ bool" \ (SML) "!((_ : Word8.word) = _)" and (Haskell) infix 4 "==" and (Scala) infixl 5 "==" | class_instance uint8 :: equal \ (Haskell) - | constant "less_eq :: uint8 \ _ \ bool" \ (SML) "Word8.<= ((_), (_))" and (Haskell) infix 4 "<=" and (Scala) "Uint8.less'_eq" | constant "less :: uint8 \ _ \ bool" \ (SML) "Word8.< ((_), (_))" and (Haskell) infix 4 "<" and (Scala) "Uint8.less" | constant "NOT :: uint8 \ _" \ (SML) "Word8.notb" and (Haskell) "Data'_Bits.complement" and (Scala) "_.unary'_~.toByte" | constant "Bit_Operations.and :: uint8 \ _" \ (SML) "Word8.andb ((_),/ (_))" and (Haskell) infixl 7 "Data_Bits..&." and (Scala) "(_ & _).toByte" | constant "Bit_Operations.or :: uint8 \ _" \ (SML) "Word8.orb ((_),/ (_))" and (Haskell) infixl 5 "Data_Bits..|." and (Scala) "(_ | _).toByte" | constant "Bit_Operations.xor :: uint8 \ _" \ (SML) "Word8.xorb ((_),/ (_))" and (Haskell) "Data'_Bits.xor" and (Scala) "(_ ^ _).toByte" definition uint8_divmod :: "uint8 \ uint8 \ uint8 \ uint8" where "uint8_divmod x y = (if y = 0 then (undefined ((div) :: uint8 \ _) x (0 :: uint8), undefined ((mod) :: uint8 \ _) x (0 :: uint8)) else (x div y, x mod y))" definition uint8_div :: "uint8 \ uint8 \ uint8" where "uint8_div x y = fst (uint8_divmod x y)" definition uint8_mod :: "uint8 \ uint8 \ uint8" where "uint8_mod x y = snd (uint8_divmod x y)" lemma div_uint8_code [code]: "x div y = (if y = 0 then 0 else uint8_div x y)" including undefined_transfer unfolding uint8_divmod_def uint8_div_def by transfer (simp add: word_div_def) lemma mod_uint8_code [code]: "x mod y = (if y = 0 then x else uint8_mod x y)" including undefined_transfer unfolding uint8_mod_def uint8_divmod_def by transfer (simp add: word_mod_def) definition uint8_sdiv :: "uint8 \ uint8 \ uint8" where "uint8_sdiv x y = (if y = 0 then undefined ((div) :: uint8 \ _) x (0 :: uint8) else Abs_uint8 (Rep_uint8 x sdiv Rep_uint8 y))" definition div0_uint8 :: "uint8 \ uint8" where [code del]: "div0_uint8 x = undefined ((div) :: uint8 \ _) x (0 :: uint8)" declare [[code abort: div0_uint8]] definition mod0_uint8 :: "uint8 \ uint8" where [code del]: "mod0_uint8 x = undefined ((mod) :: uint8 \ _) x (0 :: uint8)" declare [[code abort: mod0_uint8]] lemma uint8_divmod_code [code]: "uint8_divmod x y = (if 0x80 \ y then if x < y then (0, x) else (1, x - y) else if y = 0 then (div0_uint8 x, mod0_uint8 x) else let q = push_bit 1 (uint8_sdiv (drop_bit 1 x) y); r = x - q * y in if r \ y then (q + 1, r - y) else (q, r))" including undefined_transfer unfolding uint8_divmod_def uint8_sdiv_def div0_uint8_def mod0_uint8_def less_eq_uint8.rep_eq apply transfer apply (simp add: divmod_via_sdivmod push_bit_eq_mult) done lemma uint8_sdiv_code [code]: "Rep_uint8 (uint8_sdiv x y) = (if y = 0 then Rep_uint8 (undefined ((div) :: uint8 \ _) x (0 :: uint8)) else Rep_uint8 x sdiv Rep_uint8 y)" unfolding uint8_sdiv_def by(simp add: Abs_uint8_inverse) text \ Note that we only need a translation for signed division, but not for the remainder because @{thm uint8_divmod_code} computes both with division only. \ code_printing constant uint8_div \ (SML) "Word8.div ((_), (_))" and (Haskell) "Prelude.div" | constant uint8_mod \ (SML) "Word8.mod ((_), (_))" and (Haskell) "Prelude.mod" | constant uint8_divmod \ (Haskell) "divmod" | constant uint8_sdiv \ (Scala) "(_ '/ _).toByte" definition uint8_test_bit :: "uint8 \ integer \ bool" where [code del]: "uint8_test_bit x n = (if n < 0 \ 7 < n then undefined (bit :: uint8 \ _) x n else bit x (nat_of_integer n))" lemma bit_uint8_code [code]: "bit x n \ n < 8 \ uint8_test_bit x (integer_of_nat n)" including undefined_transfer integer.lifting unfolding uint8_test_bit_def by (transfer, simp, transfer, simp) lemma uint8_test_bit_code [code]: "uint8_test_bit w n = (if n < 0 \ 7 < n then undefined (bit :: uint8 \ _) w n else bit (Rep_uint8 w) (nat_of_integer n))" unfolding uint8_test_bit_def by (simp add: bit_uint8.rep_eq) code_printing constant uint8_test_bit \ (SML) "Uint8.test'_bit" and (Haskell) "Data'_Bits.testBitBounded" and (Scala) "Uint8.test'_bit" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 8 then raise (Fail \"argument to uint8'_test'_bit out of bounds\") else Uint8.test'_bit x i)" definition uint8_set_bit :: "uint8 \ integer \ bool \ uint8" where [code del]: "uint8_set_bit x n b = (if n < 0 \ 7 < n then undefined (set_bit :: uint8 \ _) x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_uint8_code [code]: "set_bit x n b = (if n < 8 then uint8_set_bit x (integer_of_nat n) b else x)" including undefined_transfer integer.lifting unfolding uint8_set_bit_def by(transfer)(auto cong: conj_cong simp add: not_less set_bit_beyond word_size) lemma uint8_set_bit_code [code]: "Rep_uint8 (uint8_set_bit w n b) = (if n < 0 \ 7 < n then Rep_uint8 (undefined (set_bit :: uint8 \ _) w n b) else set_bit (Rep_uint8 w) (nat_of_integer n) b)" including undefined_transfer unfolding uint8_set_bit_def by transfer simp code_printing constant uint8_set_bit \ (SML) "Uint8.set'_bit" and (Haskell) "Data'_Bits.setBitBounded" and (Scala) "Uint8.set'_bit" and (Eval) "(fn x => fn i => fn b => if i < 0 orelse i >= 8 then raise (Fail \"argument to uint8'_set'_bit out of bounds\") else Uint8.set'_bit x i b)" definition uint8_shiftl :: "uint8 \ integer \ uint8" where [code del]: "uint8_shiftl x n = (if n < 0 \ 8 \ n then undefined (push_bit :: nat \ uint8 \ _) x n else push_bit (nat_of_integer n) x)" lemma shiftl_uint8_code [code]: "push_bit n x = (if n < 8 then uint8_shiftl x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint8_shiftl_def by transfer simp lemma uint8_shiftl_code [code]: "Rep_uint8 (uint8_shiftl w n) = (if n < 0 \ 8 \ n then Rep_uint8 (undefined (push_bit :: nat \ uint8 \ _) w n) else push_bit (nat_of_integer n) (Rep_uint8 w))" including undefined_transfer unfolding uint8_shiftl_def by transfer simp code_printing constant uint8_shiftl \ (SML) "Uint8.shiftl" and (Haskell) "Data'_Bits.shiftlBounded" and (Scala) "Uint8.shiftl" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 8 then raise (Fail \"argument to uint8'_shiftl out of bounds\") else Uint8.shiftl x i)" definition uint8_shiftr :: "uint8 \ integer \ uint8" where [code del]: "uint8_shiftr x n = (if n < 0 \ 8 \ n then undefined (drop_bit :: _ \ _ \ uint8) x n else drop_bit (nat_of_integer n) x)" lemma shiftr_uint8_code [code]: "drop_bit n x = (if n < 8 then uint8_shiftr x (integer_of_nat n) else 0)" including undefined_transfer integer.lifting unfolding uint8_shiftr_def by transfer simp lemma uint8_shiftr_code [code]: "Rep_uint8 (uint8_shiftr w n) = (if n < 0 \ 8 \ n then Rep_uint8 (undefined (drop_bit :: _ \ _ \ uint8) w n) else drop_bit (nat_of_integer n) (Rep_uint8 w))" including undefined_transfer unfolding uint8_shiftr_def by transfer simp code_printing constant uint8_shiftr \ (SML) "Uint8.shiftr" and (Haskell) "Data'_Bits.shiftrBounded" and (Scala) "Uint8.shiftr" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 8 then raise (Fail \"argument to uint8'_shiftr out of bounds\") else Uint8.shiftr x i)" definition uint8_sshiftr :: "uint8 \ integer \ uint8" where [code del]: "uint8_sshiftr x n = (if n < 0 \ 8 \ n then undefined signed_drop_bit_uint8 n x else signed_drop_bit_uint8 (nat_of_integer n) x)" lemma sshiftr_uint8_code [code]: "signed_drop_bit_uint8 n x = (if n < 8 then uint8_sshiftr x (integer_of_nat n) else if bit x 7 then -1 else 0)" including undefined_transfer integer.lifting unfolding uint8_sshiftr_def by transfer (simp add: not_less signed_drop_bit_beyond word_size) lemma uint8_sshiftr_code [code]: "Rep_uint8 (uint8_sshiftr w n) = (if n < 0 \ 8 \ n then Rep_uint8 (undefined signed_drop_bit_uint8 n w) else signed_drop_bit (nat_of_integer n) (Rep_uint8 w))" including undefined_transfer unfolding uint8_sshiftr_def by transfer simp code_printing constant uint8_sshiftr \ (SML) "Uint8.shiftr'_signed" and (Haskell) "(Prelude.fromInteger (Prelude.toInteger (Data'_Bits.shiftrBounded (Prelude.fromInteger (Prelude.toInteger _) :: Uint8.Int8) _)) :: Uint8.Word8)" and (Scala) "Uint8.shiftr'_signed" and (Eval) "(fn x => fn i => if i < 0 orelse i >= 8 then raise (Fail \"argument to uint8'_sshiftr out of bounds\") else Uint8.shiftr'_signed x i)" context includes bit_operations_syntax begin lemma uint8_msb_test_bit: "msb x \ bit (x :: uint8) 7" by transfer (simp add: msb_word_iff_bit) lemma msb_uint16_code [code]: "msb x \ uint8_test_bit x 7" by (simp add: uint8_test_bit_def uint8_msb_test_bit) -lemma uint8_of_int_code [code]: "uint8_of_int i = Uint8 (integer_of_int i)" -including integer.lifting by transfer simp +lemma uint8_of_int_code [code]: + "uint8_of_int i = Uint8 (integer_of_int i)" + including integer.lifting by transfer simp lemma int_of_uint8_code [code]: "int_of_uint8 x = int_of_integer (integer_of_uint8 x)" by (simp add: int_of_uint8.rep_eq integer_of_uint8_def) +lemma uint8_of_nat_code [code]: + "uint8_of_nat = uint8_of_int \ int" + by transfer (simp add: fun_eq_iff) + lemma nat_of_uint8_code [code]: "nat_of_uint8 x = nat_of_integer (integer_of_uint8 x)" unfolding integer_of_uint8_def including integer.lifting by transfer simp definition integer_of_uint8_signed :: "uint8 \ integer" where "integer_of_uint8_signed n = (if bit n 7 then undefined integer_of_uint8 n else integer_of_uint8 n)" lemma integer_of_uint8_signed_code [code]: "integer_of_uint8_signed n = (if bit n 7 then undefined integer_of_uint8 n else integer_of_int (uint (Rep_uint8' n)))" by (simp add: integer_of_uint8_signed_def integer_of_uint8_def) lemma integer_of_uint8_code [code]: "integer_of_uint8 n = (if bit n 7 then integer_of_uint8_signed (n AND 0x7F) OR 0x80 else integer_of_uint8_signed n)" proof - have \integer_of_uint8_signed (n AND 0x7F) OR 0x80 = Bit_Operations.set_bit 7 (integer_of_uint8_signed (take_bit 7 n))\ by (simp add: take_bit_eq_mask set_bit_eq_or push_bit_eq_mult mask_eq_exp_minus_1) moreover have \integer_of_uint8 n = Bit_Operations.set_bit 7 (integer_of_uint8 (take_bit 7 n))\ if \bit n 7\ proof (rule bit_eqI) fix m from that show \bit (integer_of_uint8 n) m = bit (Bit_Operations.set_bit 7 (integer_of_uint8 (take_bit 7 n))) m\ for m including integer.lifting by transfer (auto simp add: bit_simps dest: bit_imp_le_length) qed ultimately show ?thesis by simp (simp add: integer_of_uint8_signed_def bit_simps) qed end code_printing constant "integer_of_uint8" \ (SML) "IntInf.fromLarge (Word8.toLargeInt _)" and (Haskell) "Prelude.toInteger" | constant "integer_of_uint8_signed" \ (Scala) "BigInt" section \Quickcheck setup\ definition uint8_of_natural :: "natural \ uint8" where "uint8_of_natural x \ Uint8 (integer_of_natural x)" instantiation uint8 :: "{random, exhaustive, full_exhaustive}" begin definition "random_uint8 \ qc_random_cnv uint8_of_natural" definition "exhaustive_uint8 \ qc_exhaustive_cnv uint8_of_natural" definition "full_exhaustive_uint8 \ qc_full_exhaustive_cnv uint8_of_natural" instance .. end instantiation uint8 :: narrowing begin interpretation quickcheck_narrowing_samples "\i. let x = Uint8 i in (x, 0xFF - x)" "0" "Typerep.Typerep (STR ''Uint8.uint8'') []" . definition "narrowing_uint8 d = qc_narrowing_drawn_from (narrowing_samples d) d" declare [[code drop: "partial_term_of :: uint8 itself \ _"]] lemmas partial_term_of_uint8 [code] = partial_term_of_code instance .. end end