diff --git a/thys/IP_Addresses/Lib_Word_toString.thy b/thys/IP_Addresses/Lib_Word_toString.thy --- a/thys/IP_Addresses/Lib_Word_toString.thy +++ b/thys/IP_Addresses/Lib_Word_toString.thy @@ -1,180 +1,220 @@ theory Lib_Word_toString imports Lib_Numbers_toString Word_Lib.Word_Lemmas begin +context unique_euclidean_semiring_numeral +begin + +lemma exp_estimate [simp]: + \numeral Num.One \ 2 ^ n\ (is \?P1\) + \numeral Num.One < 2 ^ n \ 1 \ n\ (is \?P2\) + \numeral (Num.Bit0 k) \ 2 ^ n \ 1 \ n \ numeral k \ 2 ^ (n - 1)\ (is \?P3\) + \numeral (Num.Bit0 k) < 2 ^ n \ 1 \ n \ numeral k < 2 ^ (n - 1)\ (is \?P4\) + \numeral (Num.Bit1 k) \ 2 ^ n \ 1 \ n \ numeral k < 2 ^ (n - 1)\ (is \?P5\) + \numeral (Num.Bit1 k) < 2 ^ n \ 1 \ n \ numeral k < 2 ^ (n - 1)\ (is \?P6\) +proof - + show ?P1 + by simp + show ?P2 + using one_less_power power_eq_if by auto + + let ?K = \numeral k :: nat\ + + define m where \m = n - 1\ + then consider \n = 0\ | \n = Suc m\ + by (cases n) simp_all + note Suc = this + + have \2 * ?K \ 2 * 2 ^ m \ ?K \ 2 ^ m\ + by linarith + then have \of_nat (2 * ?K) \ of_nat (2 * 2 ^ m) \ of_nat ?K \ of_nat (2 ^ m)\ + by (simp only: of_nat_le_iff) + then show ?P3 + by (auto intro: Suc) + + have \2 * ?K < 2 * 2 ^ m \ ?K < 2 ^ m\ + by linarith + then have \of_nat (2 * ?K) < of_nat (2 * 2 ^ m) \ of_nat ?K < of_nat (2 ^ m)\ + by (simp only: of_nat_less_iff) + then show ?P4 + by (auto intro: Suc) + + have \Suc (2 * ?K) \ 2 * 2 ^ m \ ?K < 2 ^ m\ + by linarith + then have \of_nat (Suc (2 * ?K)) \ of_nat (2 * 2 ^ m) \ of_nat ?K < of_nat (2 ^ m)\ + by (simp only: of_nat_le_iff of_nat_less_iff) + then show ?P5 + by (auto intro: Suc) + + have \Suc (2 * ?K) < 2 * 2 ^ m \ ?K < 2 ^ m\ + by linarith + then have \of_nat (Suc (2 * ?K)) < of_nat (2 * 2 ^ m) \ of_nat ?K < of_nat (2 ^ m)\ + by (simp only: of_nat_less_iff) + then show ?P6 + by (auto intro: Suc) + +qed + +end + + section\Printing Machine Words\ (*imitation of http://stackoverflow.com/questions/23864965/string-of-nat-in-isabelle*) (*parameters: lc = lower-case w = word to print*) definition string_of_word_single :: "bool \ 'a::len word \ string" where "string_of_word_single lc w \ (if w < 10 then [char_of (48 + unat w)] else if w < 36 then [char_of ((if lc then 87 else 55) + unat w)] else undefined)" text\Example:\ lemma "let word_upto = ((\ i j. map (of_nat \ nat) [i .. j]) :: int \ int \ 32 word list) in map (string_of_word_single False) (word_upto 1 35) = [''1'', ''2'', ''3'', ''4'', ''5'', ''6'', ''7'', ''8'', ''9'', ''A'', ''B'', ''C'', ''D'', ''E'', ''F'', ''G'', ''H'', ''I'', ''J'', ''K'', ''L'', ''M'', ''N'', ''O'', ''P'', ''Q'', ''R'', ''S'', ''T'', ''U'', ''V'', ''W'', ''X'', ''Y'', ''Z'']" by eval (* parameters: lowercase, base, minimum length - 1, to-be-serialized word *) function string_of_word :: "bool \ ('a :: len) word \ nat \ ('a :: len) word \ string" where "string_of_word lc base ml w = (if base < 2 \ LENGTH('a) < 2 then undefined else if w < base \ ml = 0 then string_of_word_single lc w else string_of_word lc base (ml - 1) (w div base) @ string_of_word_single lc (w mod base) )" by pat_completeness auto definition "hex_string_of_word l \ string_of_word True (16 :: ('a::len) word) l" definition "hex_string_of_word0 \ hex_string_of_word 0" (* be careful though, these functions only make sense with words > length 4. With 4 bits, base 16 is not representable. *) definition "dec_string_of_word0 \ string_of_word True 10 0" termination string_of_word apply(relation "measure (\(a,b,c,d). unat d + c)") apply(rule wf_measure) apply(subst in_measure) apply(clarsimp) subgoal for base ml n apply(case_tac "ml \ 0") apply(simp add: less_eq_Suc_le unat_div) apply(simp) apply(subgoal_tac "(n div base) < n") apply(blast intro: unat_mono) apply(rule div_less_dividend_word) apply (auto simp add: not_less word_le_nat_alt) done done declare string_of_word.simps[simp del] lemma "hex_string_of_word0 (0xdeadbeef42 :: 42 word) = ''deadbeef42''" by eval lemma "hex_string_of_word 1 (0x1 :: 5 word) = ''01''" by eval lemma "hex_string_of_word 8 (0xff::32 word) = ''0000000ff''" by eval lemma "dec_string_of_word0 (8::32 word) = ''8''" by eval lemma "dec_string_of_word0 (3::2 word) = ''11''" by eval lemma "dec_string_of_word0 (-1::8 word) = ''255''" by eval lemma string_of_word_single_atoi: "n < 10 \ string_of_word_single True n = [char_of (48 + unat n)]" by(simp add: string_of_word_single_def) (*TODO: move!*) -lemma bintrunc_pos_eq: "x \ 0 \ bintrunc n x = x \ x < 2^n" -apply(rule iffI) - subgoal using bintr_lt2p by metis -by (simp add: mod_pos_pos_trivial no_bintr_alt1; fail) - +lemma bintrunc_pos_eq: "x \ 0 \ take_bit n x = x \ x < 2^n" for x :: int + by (simp add: take_bit_int_eq_self_iff) (*The following lemma [symmetric] as [code_unfold] may give some cool speedup*) lemma string_of_word_base_ten_zeropad: fixes w ::"'a::len word" assumes lena: "LENGTH('a) \ 5" (*word must be long enough to encode 10 = 0xA*) shows "base = 10 \ zero = 0 \ string_of_word True base zero w = string_of_nat (unat w)" - proof(induction True base zero w rule: string_of_word.induct) +proof(induction True base zero w rule: string_of_word.induct) case (1 base ml n) note Word.word_less_no[simp del] note Word.uint_bintrunc[simp del] - have "5 \ n \ bintrunc n 10 = 10" for n - apply(subst bintrunc_pos_eq) - apply(simp; fail) - apply(induction rule: Nat.dec_induct) - by simp+ - with lena have unat_ten: "unat (0xA::'a::len word) = 10" by(simp) + define l where \l = LENGTH('a) - 5\ + with lena have l: \LENGTH('a) = l + 5\ + by simp - have "5 \ n \ bintrunc n 2 = 2" for n - apply(subst bintrunc_pos_eq) - apply(simp; fail) - apply(induction rule: Nat.dec_induct) - by simp+ - with lena have unat_two: "unat (2::'a::len word) = 2" by(simp) + have [simp]: \take_bit LENGTH('a) (10 :: nat) = 10\ + using lena by (auto simp add: take_bit_nat_eq_self_iff l Suc_lessI) + + have [simp]: \take_bit LENGTH('a) (10 :: int) = 10\ + using lena by (auto simp add: take_bit_int_eq_self_iff l) + (smt (verit) zero_less_power) have unat_mod_ten: "unat (n mod 0xA) = unat n mod 10" - apply(subst Word.unat_mod) - apply(subst unat_ten) - by(simp) + by (simp add: nat_take_bit_eq unat_mod) have unat_div_ten: "(unat (n div 0xA)) = unat n div 10" - apply(subst Word.unat_div) - apply(subst unat_ten) - by simp + by (simp add: nat_take_bit_eq unat_div) + have n_less_ten_unat: "n < 0xA \ (unat n < 10)" - apply(rule unat_less_helper) - by(simp) + by (simp add: unat_less_helper) + have "0xA \ n \ 10 \ unat n" - apply(subst(asm) Word.word_le_nat_alt) - apply(subst(asm) unat_ten) - by(simp) + by (simp add: nat_take_bit_eq word_le_nat_alt) + hence n_less_ten_unat_not: "\ n < 0xA \ \ unat n < 10" by fastforce have not_wordlength_too_small: "\ LENGTH('a) < 2" using lena by fastforce have "2 \ (0xA::'a word)" - apply(subst word_le_nat_alt) - apply(subst unat_ten unat_two) - apply(subst unat_two) - by(simp) + by simp hence ten_not_less_two: "\ (0xA::'a word) < 2" by (simp add: Word.word_less_no Word.uint_bintrunc) with 1(2,3) have " \ (base < 2 \ LENGTH(32) < 2)" by(simp) with 1 not_wordlength_too_small have IH: "\ n < 0xA \ string_of_word True 0xA 0 (n div 0xA) = string_of_nat (unat (n div 0xA))" by(simp) show ?case apply(simp add: 1) apply(cases "n < 0xA") subgoal apply(subst(1) string_of_word.simps) apply(subst(1) string_of_nat.simps) apply(simp add: n_less_ten_unat) by(simp add: not_wordlength_too_small ten_not_less_two string_of_word_single_atoi) using sym[OF IH] apply(simp) apply(subst(1) string_of_word.simps) apply(simp) apply(subst(1) string_of_nat.simps) apply(simp) apply(simp add: not_wordlength_too_small ten_not_less_two) apply(subst string_of_word_single_atoi) apply(rule Word.word_mod_less_divisor) - using unat_ten word_gt_0_no apply fastforce - apply(simp add: unat_mod_ten) - apply(rule sym) - apply(drule n_less_ten_unat_not) - apply(simp add: unat_div_ten) - by (simp add: string_of_nat.simps) + apply (auto simp add: not_less) + apply (metis \10 \ n \ 10 \ unat n\ not_le string_of_nat.simps unat_div_ten unat_mod_ten) + done qed (*TODO: one for all words?*) lemma dec_string_of_word0: "dec_string_of_word0 (w8:: 8 word) = string_of_nat (unat w8)" "dec_string_of_word0 (w16:: 16 word) = string_of_nat (unat w16)" "dec_string_of_word0 (w32:: 32 word) = string_of_nat (unat w32)" "dec_string_of_word0 (w64:: 64 word) = string_of_nat (unat w64)" "dec_string_of_word0 (w128:: 128 word) = string_of_nat (unat w128)" unfolding dec_string_of_word0_def using string_of_word_base_ten_zeropad by force+ end diff --git a/thys/Native_Word/Bits_Integer.thy b/thys/Native_Word/Bits_Integer.thy --- a/thys/Native_Word/Bits_Integer.thy +++ b/thys/Native_Word/Bits_Integer.thy @@ -1,688 +1,688 @@ (* Title: Bits_Integer.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Bit operations for target language integers\ theory Bits_Integer imports More_Bits_Int Code_Symbolic_Bits_Int begin lemmas [transfer_rule] = identity_quotient fun_quotient Quotient_integer[folded integer.pcr_cr_eq] lemma undefined_transfer: assumes "Quotient R Abs Rep T" shows "T (Rep undefined) undefined" using assms unfolding Quotient_alt_def by blast bundle undefined_transfer = undefined_transfer[transfer_rule] section \More lemmas about @{typ integer}s\ context includes integer.lifting begin lemma bitval_integer_transfer [transfer_rule]: "(rel_fun (=) pcr_integer) of_bool of_bool" by(auto simp add: of_bool_def integer.pcr_cr_eq cr_integer_def) lemma integer_of_nat_less_0_conv [simp]: "\ integer_of_nat n < 0" by(transfer) simp lemma int_of_integer_pow: "int_of_integer (x ^ n) = int_of_integer x ^ n" by(induct n) simp_all lemma pow_integer_transfer [transfer_rule]: "(rel_fun pcr_integer (rel_fun (=) pcr_integer)) (^) (^)" by(auto 4 3 simp add: integer.pcr_cr_eq cr_integer_def int_of_integer_pow) lemma sub1_lt_0_iff [simp]: "Code_Numeral.sub n num.One < 0 \ False" by(cases n)(simp_all add: Code_Numeral.sub_code) lemma nat_of_integer_numeral [simp]: "nat_of_integer (numeral n) = numeral n" by transfer simp lemma nat_of_integer_sub1_conv_pred_numeral [simp]: "nat_of_integer (Code_Numeral.sub n num.One) = pred_numeral n" by(cases n)(simp_all add: Code_Numeral.sub_code) lemma nat_of_integer_1 [simp]: "nat_of_integer 1 = 1" by transfer simp lemma dup_1 [simp]: "Code_Numeral.dup 1 = 2" by transfer simp section \Bit operations on @{typ integer}\ text \Bit operations on @{typ integer} are the same as on @{typ int}\ -lift_definition bin_rest_integer :: "integer \ integer" is bin_rest . -lift_definition bin_last_integer :: "integer \ bool" is bin_last . +lift_definition bin_rest_integer :: "integer \ integer" is \\k . k div 2\ . +lift_definition bin_last_integer :: "integer \ bool" is odd . lift_definition Bit_integer :: "integer \ bool \ integer" is \\k b. of_bool b + 2 * k\ . end instance integer :: semiring_bit_syntax .. context includes lifting_syntax integer.lifting begin lemma test_bit_integer_transfer [transfer_rule]: \(pcr_integer ===> (=)) bit (!!)\ unfolding test_bit_eq_bit by transfer_prover lemma shiftl_integer_transfer [transfer_rule]: \(pcr_integer ===> (=) ===> pcr_integer) (\k n. push_bit n k) (<<)\ unfolding shiftl_eq_push_bit by transfer_prover lemma shiftr_integer_transfer [transfer_rule]: \(pcr_integer ===> (=) ===> pcr_integer) (\k n. drop_bit n k) (>>)\ unfolding shiftr_eq_drop_bit by transfer_prover end instantiation integer :: lsb begin context includes integer.lifting begin lift_definition lsb_integer :: "integer \ bool" is lsb . instance by (standard; transfer) (fact lsb_odd) end end instantiation integer :: msb begin context includes integer.lifting begin lift_definition msb_integer :: "integer \ bool" is msb . instance .. end end instantiation integer :: set_bit begin context includes integer.lifting begin lift_definition set_bit_integer :: "integer \ nat \ bool \ integer" is set_bit . instance apply standard apply transfer apply (simp add: bit_simps) done end end abbreviation (input) wf_set_bits_integer where "wf_set_bits_integer \ wf_set_bits_int" section \Target language implementations\ text \ Unfortunately, this is not straightforward, because these API functions have different signatures and preconditions on the parameters: \begin{description} \item[Standard ML] Shifts in IntInf are given as word, but not IntInf. \item[Haskell] In the Data.Bits.Bits type class, shifts and bit indices are given as Int rather than Integer. \end{description} Additional constants take only parameters of type @{typ integer} rather than @{typ nat} and check the preconditions as far as possible (e.g., being non-negative) in a portable way. Manual implementations inside code\_printing perform the remaining range checks and convert these @{typ integer}s into the right type. For normalisation by evaluation, we derive custom code equations, because NBE does not know these code\_printing serialisations and would otherwise loop. \ code_identifier code_module Bits_Integer \ (SML) Bits_Int and (OCaml) Bits_Int and (Haskell) Bits_Int and (Scala) Bits_Int code_printing code_module Bits_Integer \ (SML) \structure Bits_Integer : sig val set_bit : IntInf.int -> IntInf.int -> bool -> IntInf.int val shiftl : IntInf.int -> IntInf.int -> IntInf.int val shiftr : IntInf.int -> IntInf.int -> IntInf.int val test_bit : IntInf.int -> IntInf.int -> bool end = struct val maxWord = IntInf.pow (2, Word.wordSize); fun set_bit x n b = if n < maxWord 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 raise (Fail ("Bit index too large: " ^ IntInf.toString n)); fun shiftl x n = if n < maxWord then IntInf.<< (x, Word.fromLargeInt (IntInf.toLarge n)) else raise (Fail ("Shift operand too large: " ^ IntInf.toString n)); fun shiftr x n = if n < maxWord then IntInf.~>> (x, Word.fromLargeInt (IntInf.toLarge n)) else raise (Fail ("Shift operand too large: " ^ IntInf.toString n)); fun test_bit x n = if n < maxWord then IntInf.andb (x, IntInf.<< (1, Word.fromLargeInt (IntInf.toLarge n))) <> 0 else raise (Fail ("Bit index too large: " ^ IntInf.toString n)); end; (*struct Bits_Integer*)\ code_reserved SML Bits_Integer code_printing code_module Bits_Integer \ (OCaml) \module Bits_Integer : sig val shiftl : Z.t -> Z.t -> Z.t val shiftr : Z.t -> Z.t -> Z.t val test_bit : Z.t -> Z.t -> bool end = struct (* We do not need an explicit range checks here, because Big_int.int_of_big_int raises Failure if the argument does not fit into an int. *) let shiftl x n = Z.shift_left x (Z.to_int n);; let shiftr x n = Z.shift_right x (Z.to_int n);; let test_bit x n = Z.testbit x (Z.to_int n);; end;; (*struct Bits_Integer*)\ code_reserved OCaml Bits_Integer code_printing code_module Data_Bits \ (Haskell) \ module Data_Bits where { import qualified Data.Bits; {- The ...Bounded functions assume that the Integer argument for the shift or bit index fits into an Int, is non-negative and (for types of fixed bit width) less than bitSize -} infixl 7 .&.; infixl 6 `xor`; infixl 5 .|.; (.&.) :: Data.Bits.Bits a => a -> a -> a; (.&.) = (Data.Bits..&.); xor :: Data.Bits.Bits a => a -> a -> a; xor = Data.Bits.xor; (.|.) :: Data.Bits.Bits a => a -> a -> a; (.|.) = (Data.Bits..|.); complement :: Data.Bits.Bits a => a -> a; complement = Data.Bits.complement; testBitUnbounded :: Data.Bits.Bits a => a -> Integer -> Bool; testBitUnbounded x b | b <= toInteger (Prelude.maxBound :: Int) = Data.Bits.testBit x (fromInteger b) | otherwise = error ("Bit index too large: " ++ show b) ; testBitBounded :: Data.Bits.Bits a => a -> Integer -> Bool; testBitBounded x b = Data.Bits.testBit x (fromInteger b); setBitUnbounded :: Data.Bits.Bits a => a -> Integer -> Bool -> a; setBitUnbounded x n b | n <= toInteger (Prelude.maxBound :: Int) = if b then Data.Bits.setBit x (fromInteger n) else Data.Bits.clearBit x (fromInteger n) | otherwise = error ("Bit index too large: " ++ show n) ; setBitBounded :: Data.Bits.Bits a => a -> Integer -> Bool -> a; setBitBounded x n True = Data.Bits.setBit x (fromInteger n); setBitBounded x n False = Data.Bits.clearBit x (fromInteger n); shiftlUnbounded :: Data.Bits.Bits a => a -> Integer -> a; shiftlUnbounded x n | n <= toInteger (Prelude.maxBound :: Int) = Data.Bits.shiftL x (fromInteger n) | otherwise = error ("Shift operand too large: " ++ show n) ; shiftlBounded :: Data.Bits.Bits a => a -> Integer -> a; shiftlBounded x n = Data.Bits.shiftL x (fromInteger n); shiftrUnbounded :: Data.Bits.Bits a => a -> Integer -> a; shiftrUnbounded x n | n <= toInteger (Prelude.maxBound :: Int) = Data.Bits.shiftR x (fromInteger n) | otherwise = error ("Shift operand too large: " ++ show n) ; shiftrBounded :: (Ord a, Data.Bits.Bits a) => a -> Integer -> a; shiftrBounded x n = Data.Bits.shiftR x (fromInteger n); }\ and \ \@{theory HOL.Quickcheck_Narrowing} maps @{typ integer} to Haskell's Prelude.Int type instead of Integer. For compatibility with the Haskell target, we nevertheless provide bounded and unbounded functions.\ (Haskell_Quickcheck) \ module Data_Bits where { import qualified Data.Bits; {- The functions assume that the Int argument for the shift or bit index is non-negative and (for types of fixed bit width) less than bitSize -} infixl 7 .&.; infixl 6 `xor`; infixl 5 .|.; (.&.) :: Data.Bits.Bits a => a -> a -> a; (.&.) = (Data.Bits..&.); xor :: Data.Bits.Bits a => a -> a -> a; xor = Data.Bits.xor; (.|.) :: Data.Bits.Bits a => a -> a -> a; (.|.) = (Data.Bits..|.); complement :: Data.Bits.Bits a => a -> a; complement = Data.Bits.complement; testBitUnbounded :: Data.Bits.Bits a => a -> Prelude.Int -> Bool; testBitUnbounded = Data.Bits.testBit; testBitBounded :: Data.Bits.Bits a => a -> Prelude.Int -> Bool; testBitBounded = Data.Bits.testBit; setBitUnbounded :: Data.Bits.Bits a => a -> Prelude.Int -> Bool -> a; setBitUnbounded x n True = Data.Bits.setBit x n; setBitUnbounded x n False = Data.Bits.clearBit x n; setBitBounded :: Data.Bits.Bits a => a -> Prelude.Int -> Bool -> a; setBitBounded x n True = Data.Bits.setBit x n; setBitBounded x n False = Data.Bits.clearBit x n; shiftlUnbounded :: Data.Bits.Bits a => a -> Prelude.Int -> a; shiftlUnbounded = Data.Bits.shiftL; shiftlBounded :: Data.Bits.Bits a => a -> Prelude.Int -> a; shiftlBounded = Data.Bits.shiftL; shiftrUnbounded :: Data.Bits.Bits a => a -> Prelude.Int -> a; shiftrUnbounded = Data.Bits.shiftR; shiftrBounded :: (Ord a, Data.Bits.Bits a) => a -> Prelude.Int -> a; shiftrBounded = Data.Bits.shiftR; }\ code_reserved Haskell Data_Bits code_printing code_module Bits_Integer \ (Scala) \object Bits_Integer { def setBit(x: BigInt, n: BigInt, b: Boolean) : BigInt = if (n.isValidInt) if (b) x.setBit(n.toInt) else x.clearBit(n.toInt) else sys.error("Bit index too large: " + n.toString) def shiftl(x: BigInt, n: BigInt) : BigInt = if (n.isValidInt) x << n.toInt else sys.error("Shift index too large: " + n.toString) def shiftr(x: BigInt, n: BigInt) : BigInt = if (n.isValidInt) x << n.toInt else sys.error("Shift index too large: " + n.toString) def testBit(x: BigInt, n: BigInt) : Boolean = if (n.isValidInt) x.testBit(n.toInt) else sys.error("Bit index too large: " + n.toString) } /* object Bits_Integer */\ code_printing constant "(AND) :: integer \ integer \ integer" \ (SML) "IntInf.andb ((_),/ (_))" and (OCaml) "Z.logand" and (Haskell) "((Data'_Bits..&.) :: Integer -> Integer -> Integer)" and (Haskell_Quickcheck) "((Data'_Bits..&.) :: Prelude.Int -> Prelude.Int -> Prelude.Int)" and (Scala) infixl 3 "&" | constant "(OR) :: integer \ integer \ integer" \ (SML) "IntInf.orb ((_),/ (_))" and (OCaml) "Z.logor" and (Haskell) "((Data'_Bits..|.) :: Integer -> Integer -> Integer)" and (Haskell_Quickcheck) "((Data'_Bits..|.) :: Prelude.Int -> Prelude.Int -> Prelude.Int)" and (Scala) infixl 1 "|" | constant "(XOR) :: integer \ integer \ integer" \ (SML) "IntInf.xorb ((_),/ (_))" and (OCaml) "Z.logxor" and (Haskell) "(Data'_Bits.xor :: Integer -> Integer -> Integer)" and (Haskell_Quickcheck) "(Data'_Bits.xor :: Prelude.Int -> Prelude.Int -> Prelude.Int)" and (Scala) infixl 2 "^" | constant "NOT :: integer \ integer" \ (SML) "IntInf.notb" and (OCaml) "Z.lognot" and (Haskell) "(Data'_Bits.complement :: Integer -> Integer)" and (Haskell_Quickcheck) "(Data'_Bits.complement :: Prelude.Int -> Prelude.Int)" and (Scala) "_.unary'_~" code_printing constant bin_rest_integer \ (SML) "IntInf.div ((_), 2)" and (OCaml) "Z.shift'_right/ _/ 1" and (Haskell) "(Data'_Bits.shiftrUnbounded _ 1 :: Integer)" and (Haskell_Quickcheck) "(Data'_Bits.shiftrUnbounded _ 1 :: Prelude.Int)" and (Scala) "_ >> 1" context includes integer.lifting begin lemma bitNOT_integer_code [code]: fixes i :: integer shows "NOT i = - i - 1" by transfer(simp add: int_not_def) lemma bin_rest_integer_code [code nbe]: "bin_rest_integer i = i div 2" by transfer rule lemma bin_last_integer_code [code]: "bin_last_integer i \ i AND 1 \ 0" by transfer (rule bin_last_conv_AND) lemma bin_last_integer_nbe [code nbe]: "bin_last_integer i \ i mod 2 \ 0" by transfer(simp add: bin_last_def) lemma bitval_bin_last_integer [code_unfold]: "of_bool (bin_last_integer i) = i AND 1" by transfer(rule bitval_bin_last) end definition integer_test_bit :: "integer \ integer \ bool" where "integer_test_bit x n = (if n < 0 then undefined x n else bit x (nat_of_integer n))" declare [[code drop: \bit :: integer \ nat \ bool\]] lemma bit_integer_code [code]: "bit x n \ integer_test_bit x (integer_of_nat n)" by (simp add: integer_test_bit_def) lemma integer_test_bit_code [code]: "integer_test_bit x (Code_Numeral.Neg n) = undefined x (Code_Numeral.Neg n)" "integer_test_bit 0 0 = False" "integer_test_bit 0 (Code_Numeral.Pos n) = False" "integer_test_bit (Code_Numeral.Pos num.One) 0 = True" "integer_test_bit (Code_Numeral.Pos (num.Bit0 n)) 0 = False" "integer_test_bit (Code_Numeral.Pos (num.Bit1 n)) 0 = True" "integer_test_bit (Code_Numeral.Pos num.One) (Code_Numeral.Pos n') = False" "integer_test_bit (Code_Numeral.Pos (num.Bit0 n)) (Code_Numeral.Pos n') = integer_test_bit (Code_Numeral.Pos n) (Code_Numeral.sub n' num.One)" "integer_test_bit (Code_Numeral.Pos (num.Bit1 n)) (Code_Numeral.Pos n') = integer_test_bit (Code_Numeral.Pos n) (Code_Numeral.sub n' num.One)" "integer_test_bit (Code_Numeral.Neg num.One) 0 = True" "integer_test_bit (Code_Numeral.Neg (num.Bit0 n)) 0 = False" "integer_test_bit (Code_Numeral.Neg (num.Bit1 n)) 0 = True" "integer_test_bit (Code_Numeral.Neg num.One) (Code_Numeral.Pos n') = True" "integer_test_bit (Code_Numeral.Neg (num.Bit0 n)) (Code_Numeral.Pos n') = integer_test_bit (Code_Numeral.Neg n) (Code_Numeral.sub n' num.One)" "integer_test_bit (Code_Numeral.Neg (num.Bit1 n)) (Code_Numeral.Pos n') = integer_test_bit (Code_Numeral.Neg (n + num.One)) (Code_Numeral.sub n' num.One)" apply (simp_all add: integer_test_bit_def bit_integer_def) using bin_nth_numeral_simps bit_numeral_int_simps(6) by presburger code_printing constant integer_test_bit \ (SML) "Bits'_Integer.test'_bit" and (OCaml) "Bits'_Integer.test'_bit" and (Haskell) "(Data'_Bits.testBitUnbounded :: Integer -> Integer -> Bool)" and (Haskell_Quickcheck) "(Data'_Bits.testBitUnbounded :: Prelude.Int -> Prelude.Int -> Bool)" and (Scala) "Bits'_Integer.testBit" context includes integer.lifting begin lemma lsb_integer_code [code]: fixes x :: integer shows "lsb x = bit x 0" by transfer(simp add: lsb_int_def) definition integer_set_bit :: "integer \ integer \ bool \ integer" where [code del]: "integer_set_bit x n b = (if n < 0 then undefined x n b else set_bit x (nat_of_integer n) b)" lemma set_bit_integer_code [code]: "set_bit x i b = integer_set_bit x (integer_of_nat i) b" by(simp add: integer_set_bit_def) lemma set_bit_integer_conv_masks: fixes x :: integer shows "set_bit x i b = (if b then x OR (1 << i) else x AND NOT (1 << i))" by transfer (simp add: int_set_bit_False_conv_NAND int_set_bit_True_conv_OR shiftl_eq_push_bit) end code_printing constant integer_set_bit \ (SML) "Bits'_Integer.set'_bit" and (Haskell) "(Data'_Bits.setBitUnbounded :: Integer -> Integer -> Bool -> Integer)" and (Haskell_Quickcheck) "(Data'_Bits.setBitUnbounded :: Prelude.Int -> Prelude.Int -> Bool -> Prelude.Int)" and (Scala) "Bits'_Integer.setBit" text \ OCaml.Big\_int does not have a method for changing an individual bit, so we emulate that with masks. We prefer an Isabelle implementation, because this then takes care of the signs for AND and OR. \ lemma integer_set_bit_code [code]: "integer_set_bit x n b = (if n < 0 then undefined x n b else if b then x OR (push_bit (nat_of_integer n) 1) else x AND NOT (push_bit (nat_of_integer n) 1))" by (auto simp add: integer_set_bit_def not_less set_bit_eq set_bit_def unset_bit_def) definition integer_shiftl :: "integer \ integer \ integer" where [code del]: "integer_shiftl x n = (if n < 0 then undefined x n else push_bit (nat_of_integer n) x)" declare [[code drop: \push_bit :: nat \ integer \ integer\]] lemma shiftl_integer_code [code]: fixes x :: integer shows "push_bit n x = integer_shiftl x (integer_of_nat n)" by(auto simp add: integer_shiftl_def) context includes integer.lifting begin lemma shiftl_integer_conv_mult_pow2: fixes x :: integer shows "x << n = x * 2 ^ n" by (simp add: push_bit_eq_mult shiftl_eq_push_bit) lemma integer_shiftl_code [code]: "integer_shiftl x (Code_Numeral.Neg n) = undefined x (Code_Numeral.Neg n)" "integer_shiftl x 0 = x" "integer_shiftl x (Code_Numeral.Pos n) = integer_shiftl (Code_Numeral.dup x) (Code_Numeral.sub n num.One)" "integer_shiftl 0 (Code_Numeral.Pos n) = 0" apply (simp_all add: integer_shiftl_def numeral_eq_Suc) apply transfer apply (simp add: ac_simps) done end code_printing constant integer_shiftl \ (SML) "Bits'_Integer.shiftl" and (OCaml) "Bits'_Integer.shiftl" and (Haskell) "(Data'_Bits.shiftlUnbounded :: Integer -> Integer -> Integer)" and (Haskell_Quickcheck) "(Data'_Bits.shiftlUnbounded :: Prelude.Int -> Prelude.Int -> Prelude.Int)" and (Scala) "Bits'_Integer.shiftl" definition integer_shiftr :: "integer \ integer \ integer" where [code del]: "integer_shiftr x n = (if n < 0 then undefined x n else drop_bit (nat_of_integer n) x)" declare [[code drop: \drop_bit :: nat \ integer \ integer\]] lemma shiftr_integer_conv_div_pow2: includes integer.lifting fixes x :: integer shows "x >> n = x div 2 ^ n" by (simp add: drop_bit_eq_div shiftr_eq_drop_bit) lemma shiftr_integer_code [code]: fixes x :: integer shows "drop_bit n x = integer_shiftr x (integer_of_nat n)" by(auto simp add: integer_shiftr_def) code_printing constant integer_shiftr \ (SML) "Bits'_Integer.shiftr" and (OCaml) "Bits'_Integer.shiftr" and (Haskell) "(Data'_Bits.shiftrUnbounded :: Integer -> Integer -> Integer)" and (Haskell_Quickcheck) "(Data'_Bits.shiftrUnbounded :: Prelude.Int -> Prelude.Int -> Prelude.Int)" and (Scala) "Bits'_Integer.shiftr" lemma integer_shiftr_code [code]: includes integer.lifting shows "integer_shiftr x (Code_Numeral.Neg n) = undefined x (Code_Numeral.Neg n)" "integer_shiftr x 0 = x" "integer_shiftr 0 (Code_Numeral.Pos n) = 0" "integer_shiftr (Code_Numeral.Pos num.One) (Code_Numeral.Pos n) = 0" "integer_shiftr (Code_Numeral.Pos (num.Bit0 n')) (Code_Numeral.Pos n) = integer_shiftr (Code_Numeral.Pos n') (Code_Numeral.sub n num.One)" "integer_shiftr (Code_Numeral.Pos (num.Bit1 n')) (Code_Numeral.Pos n) = integer_shiftr (Code_Numeral.Pos n') (Code_Numeral.sub n num.One)" "integer_shiftr (Code_Numeral.Neg num.One) (Code_Numeral.Pos n) = -1" "integer_shiftr (Code_Numeral.Neg (num.Bit0 n')) (Code_Numeral.Pos n) = integer_shiftr (Code_Numeral.Neg n') (Code_Numeral.sub n num.One)" "integer_shiftr (Code_Numeral.Neg (num.Bit1 n')) (Code_Numeral.Pos n) = integer_shiftr (Code_Numeral.Neg (Num.inc n')) (Code_Numeral.sub n num.One)" apply (simp_all add: integer_shiftr_def numeral_eq_Suc drop_bit_Suc) apply transfer apply simp apply transfer apply simp apply transfer apply (simp add: add_One) done context includes integer.lifting begin lemma Bit_integer_code [code]: "Bit_integer i False = push_bit 1 i" "Bit_integer i True = (push_bit 1 i) + 1" by (transfer; simp add: shiftl_int_def)+ lemma msb_integer_code [code]: "msb (x :: integer) \ x < 0" by transfer(simp add: msb_int_def) end context includes integer.lifting natural.lifting begin lemma bitAND_integer_unfold [code]: "x AND y = (if x = 0 then 0 else if x = - 1 then y else Bit_integer (bin_rest_integer x AND bin_rest_integer y) (bin_last_integer x \ bin_last_integer y))" by transfer (auto simp add: algebra_simps and_int_rec [of _ \_ * 2\] and_int_rec [of \_ * 2\] and_int_rec [of \1 + _ * 2\] elim!: evenE oddE) lemma bitOR_integer_unfold [code]: "x OR y = (if x = 0 then y else if x = - 1 then - 1 else Bit_integer (bin_rest_integer x OR bin_rest_integer y) (bin_last_integer x \ bin_last_integer y))" by transfer (auto simp add: algebra_simps or_int_rec [of _ \_ * 2\] or_int_rec [of _ \1 + _ * 2\] or_int_rec [of \1 + _ * 2\] elim!: evenE oddE) lemma bitXOR_integer_unfold [code]: "x XOR y = (if x = 0 then y else if x = - 1 then NOT y else Bit_integer (bin_rest_integer x XOR bin_rest_integer y) (\ bin_last_integer x \ bin_last_integer y))" by transfer (auto simp add: algebra_simps xor_int_rec [of _ \_ * 2\] xor_int_rec [of \_ * 2\] xor_int_rec [of \1 + _ * 2\] elim!: evenE oddE) end section \Test code generator setup\ definition bit_integer_test :: "bool" where "bit_integer_test = (([ -1 AND 3, 1 AND -3, 3 AND 5, -3 AND (- 5) , -3 OR 1, 1 OR -3, 3 OR 5, -3 OR (- 5) , NOT 1, NOT (- 3) , -1 XOR 3, 1 XOR (- 3), 3 XOR 5, -5 XOR (- 3) , set_bit 5 4 True, set_bit (- 5) 2 True, set_bit 5 0 False, set_bit (- 5) 1 False , 1 << 2, -1 << 3 , 100 >> 3, -100 >> 3] :: integer list) = [ 3, 1, 1, -7 , -3, -3, 7, -1 , -2, 2 , -4, -4, 6, 6 , 21, -1, 4, -7 , 4, -8 , 12, -13] \ [ (5 :: integer) !! 4, (5 :: integer) !! 2, (-5 :: integer) !! 4, (-5 :: integer) !! 2 , lsb (5 :: integer), lsb (4 :: integer), lsb (-1 :: integer), lsb (-2 :: integer), msb (5 :: integer), msb (0 :: integer), msb (-1 :: integer), msb (-2 :: integer)] = [ False, True, True, False, True, False, True, False, False, False, True, True])" export_code bit_integer_test checking SML Haskell? Haskell_Quickcheck? OCaml? Scala notepad begin have bit_integer_test by eval have bit_integer_test by normalization have bit_integer_test by code_simp end ML_val \val true = @{code bit_integer_test}\ lemma "x AND y = x OR (y :: integer)" quickcheck[random, expect=counterexample] quickcheck[exhaustive, expect=counterexample] oops lemma "(x :: integer) AND x = x OR x" quickcheck[narrowing, expect=no_counterexample] oops lemma "(f :: integer \ unit) = g" quickcheck[narrowing, size=3, expect=no_counterexample] by(simp add: fun_eq_iff) hide_const bit_integer_test hide_fact bit_integer_test_def end diff --git a/thys/Native_Word/Code_Symbolic_Bits_Int.thy b/thys/Native_Word/Code_Symbolic_Bits_Int.thy --- a/thys/Native_Word/Code_Symbolic_Bits_Int.thy +++ b/thys/Native_Word/Code_Symbolic_Bits_Int.thy @@ -1,122 +1,122 @@ (* Title: Code_Symbolic_Bits_Int.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Symbolic implementation of bit operations on int\ theory Code_Symbolic_Bits_Int imports "Word_Lib.Generic_set_bit" "Word_Lib.Least_significant_bit" More_Bits_Int begin section \Implementations of bit operations on \<^typ>\int\ operating on symbolic representation\ lemma test_bit_int_code [code]: "bit (0::int) n = False" "bit (Int.Neg num.One) n = True" "bit (Int.Pos num.One) 0 = True" "bit (Int.Pos (num.Bit0 m)) 0 = False" "bit (Int.Pos (num.Bit1 m)) 0 = True" "bit (Int.Neg (num.Bit0 m)) 0 = False" "bit (Int.Neg (num.Bit1 m)) 0 = True" "bit (Int.Pos num.One) (Suc n) = False" "bit (Int.Pos (num.Bit0 m)) (Suc n) = bit (Int.Pos m) n" "bit (Int.Pos (num.Bit1 m)) (Suc n) = bit (Int.Pos m) n" "bit (Int.Neg (num.Bit0 m)) (Suc n) = bit (Int.Neg m) n" "bit (Int.Neg (num.Bit1 m)) (Suc n) = bit (Int.Neg (Num.inc m)) n" by (simp_all add: Num.add_One bit_Suc) lemma int_not_code [code]: "NOT (0 :: int) = -1" "NOT (Int.Pos n) = Int.Neg (Num.inc n)" "NOT (Int.Neg n) = Num.sub n num.One" by(simp_all add: Num.add_One int_not_def) lemma int_and_code [code]: fixes i j :: int shows "0 AND j = 0" "i AND 0 = 0" "Int.Pos n AND Int.Pos m = (case bitAND_num n m of None \ 0 | Some n' \ Int.Pos n')" "Int.Neg n AND Int.Neg m = NOT (Num.sub n num.One OR Num.sub m num.One)" "Int.Pos n AND Int.Neg num.One = Int.Pos n" "Int.Pos n AND Int.Neg (num.Bit0 m) = Num.sub (bitORN_num (Num.BitM m) n) num.One" "Int.Pos n AND Int.Neg (num.Bit1 m) = Num.sub (bitORN_num (num.Bit0 m) n) num.One" "Int.Neg num.One AND Int.Pos m = Int.Pos m" "Int.Neg (num.Bit0 n) AND Int.Pos m = Num.sub (bitORN_num (Num.BitM n) m) num.One" "Int.Neg (num.Bit1 n) AND Int.Pos m = Num.sub (bitORN_num (num.Bit0 n) m) num.One" apply (simp_all add: int_numeral_bitAND_num Num.add_One sub_inc_One_eq inc_BitM_eq not_minus_numeral_inc_eq flip: int_not_neg_numeral int_or_not_bitORN_num split: option.split) apply (simp_all add: ac_simps) done lemma int_or_code [code]: fixes i j :: int shows "0 OR j = j" "i OR 0 = i" "Int.Pos n OR Int.Pos m = Int.Pos (bitOR_num n m)" "Int.Neg n OR Int.Neg m = NOT (Num.sub n num.One AND Num.sub m num.One)" "Int.Pos n OR Int.Neg num.One = Int.Neg num.One" "Int.Pos n OR Int.Neg (num.Bit0 m) = (case bitANDN_num (Num.BitM m) n of None \ -1 | Some n' \ Int.Neg (Num.inc n'))" "Int.Pos n OR Int.Neg (num.Bit1 m) = (case bitANDN_num (num.Bit0 m) n of None \ -1 | Some n' \ Int.Neg (Num.inc n'))" "Int.Neg num.One OR Int.Pos m = Int.Neg num.One" "Int.Neg (num.Bit0 n) OR Int.Pos m = (case bitANDN_num (Num.BitM n) m of None \ -1 | Some n' \ Int.Neg (Num.inc n'))" "Int.Neg (num.Bit1 n) OR Int.Pos m = (case bitANDN_num (num.Bit0 n) m of None \ -1 | Some n' \ Int.Neg (Num.inc n'))" apply (simp_all add: int_numeral_bitOR_num flip: int_not_neg_numeral) apply (simp_all add: or_int_def int_and_comm int_not_and_bitANDN_num del: int_not_simps(4) split: option.split) apply (simp_all add: Num.add_One) done lemma int_xor_code [code]: fixes i j :: int shows "0 XOR j = j" "i XOR 0 = i" "Int.Pos n XOR Int.Pos m = (case bitXOR_num n m of None \ 0 | Some n' \ Int.Pos n')" "Int.Neg n XOR Int.Neg m = Num.sub n num.One XOR Num.sub m num.One" "Int.Neg n XOR Int.Pos m = NOT (Num.sub n num.One XOR Int.Pos m)" "Int.Pos n XOR Int.Neg m = NOT (Int.Pos n XOR Num.sub m num.One)" by(fold int_not_neg_numeral)(simp_all add: int_numeral_bitXOR_num int_xor_not cong: option.case_cong) -lemma bin_rest_code: "bin_rest i = i >> 1" +lemma bin_rest_code: "i div 2 = i >> 1" for i :: int by (simp add: shiftr_int_def) lemma set_bits_code [code]: "set_bits = Code.abort (STR ''set_bits is unsupported on type int'') (\_. set_bits :: _ \ int)" by simp lemma fixes i :: int shows int_set_bit_True_conv_OR [code]: "set_bit i n True = i OR (1 << n)" and int_set_bit_False_conv_NAND [code]: "set_bit i n False = i AND NOT (1 << n)" and int_set_bit_conv_ops: "set_bit i n b = (if b then i OR (1 << n) else i AND NOT (1 << n))" by(simp_all add: set_bit_int_def bin_set_conv_OR bin_clr_conv_NAND) declare [[code drop: \drop_bit :: nat \ int \ int\]] lemma drop_bit_int_code [code]: fixes i :: int shows "drop_bit 0 i = i" "drop_bit (Suc n) 0 = (0 :: int)" "drop_bit (Suc n) (Int.Pos num.One) = 0" "drop_bit (Suc n) (Int.Pos (num.Bit0 m)) = drop_bit n (Int.Pos m)" "drop_bit (Suc n) (Int.Pos (num.Bit1 m)) = drop_bit n (Int.Pos m)" "drop_bit (Suc n) (Int.Neg num.One) = - 1" "drop_bit (Suc n) (Int.Neg (num.Bit0 m)) = drop_bit n (Int.Neg m)" "drop_bit (Suc n) (Int.Neg (num.Bit1 m)) = drop_bit n (Int.Neg (Num.inc m))" by (simp_all add: shiftr_eq_drop_bit drop_bit_Suc add_One) declare [[code drop: \push_bit :: nat \ int \ int\]] lemma push_bit_int_code [code]: "push_bit 0 i = i" "push_bit (Suc n) i = push_bit n (Int.dup i)" by (simp_all add: ac_simps) lemma int_lsb_code [code]: "lsb (0 :: int) = False" "lsb (Int.Pos num.One) = True" "lsb (Int.Pos (num.Bit0 w)) = False" "lsb (Int.Pos (num.Bit1 w)) = True" "lsb (Int.Neg num.One) = True" "lsb (Int.Neg (num.Bit0 w)) = False" "lsb (Int.Neg (num.Bit1 w)) = True" by simp_all end diff --git a/thys/Native_Word/Code_Target_Bits_Int.thy b/thys/Native_Word/Code_Target_Bits_Int.thy --- a/thys/Native_Word/Code_Target_Bits_Int.thy +++ b/thys/Native_Word/Code_Target_Bits_Int.thy @@ -1,92 +1,92 @@ (* Title: Code_Target_Bits_Int.thy Author: Andreas Lochbihler, ETH Zurich *) chapter \Implementation of bit operations on int by target language operations\ theory Code_Target_Bits_Int imports Bits_Integer "HOL-Library.Code_Target_Int" begin declare [[code drop: "(AND) :: int \ _" "(OR) :: int \ _" "(XOR) :: int \ _" "(NOT) :: int \ _" "lsb :: int \ _" "set_bit :: int \ _" "bit :: int \ _" "push_bit :: _ \ int \ _" "drop_bit :: _ \ int \ _" int_of_integer_symbolic ]] declare bitval_bin_last [code_unfold] lemma [code_unfold]: \bit x n \ x AND (push_bit n 1) \ 0\ for x :: int by (fact bit_iff_and_push_bit_not_eq_0) context includes integer.lifting begin lemma bit_int_code [code]: "bit (int_of_integer x) n = bit x n" by transfer simp lemma and_int_code [code]: "int_of_integer i AND int_of_integer j = int_of_integer (i AND j)" by transfer simp lemma or_int_code [code]: "int_of_integer i OR int_of_integer j = int_of_integer (i OR j)" by transfer simp lemma xor_int_code [code]: "int_of_integer i XOR int_of_integer j = int_of_integer (i XOR j)" by transfer simp lemma not_int_code [code]: "NOT (int_of_integer i) = int_of_integer (NOT i)" by transfer simp lemma push_bit_int_code [code]: \push_bit n (int_of_integer x) = int_of_integer (push_bit n x)\ by transfer simp lemma drop_bit_int_code [code]: \drop_bit n (int_of_integer x) = int_of_integer (drop_bit n x)\ by transfer simp lemma take_bit_int_code [code]: \take_bit n (int_of_integer x) = int_of_integer (take_bit n x)\ by transfer simp lemma lsb_int_code [code]: "lsb (int_of_integer x) = lsb x" by transfer simp lemma set_bit_int_code [code]: "set_bit (int_of_integer x) n b = int_of_integer (set_bit x n b)" by transfer simp lemma int_of_integer_symbolic_code [code]: "int_of_integer_symbolic = int_of_integer" by (simp add: int_of_integer_symbolic_def) context begin qualified definition even :: \int \ bool\ where [code_abbrev]: \even = Parity.even\ end lemma [code]: \Code_Target_Bits_Int.even i \ i AND 1 = 0\ by (simp add: Code_Target_Bits_Int.even_def even_iff_mod_2_eq_zero and_one_eq) lemma bin_rest_code: - "bin_rest (int_of_integer i) = int_of_integer (bin_rest_integer i)" + "int_of_integer i div 2 = int_of_integer (bin_rest_integer i)" by transfer simp end end diff --git a/thys/PAC_Checker/PAC_Checker_Relation.thy b/thys/PAC_Checker/PAC_Checker_Relation.thy --- a/thys/PAC_Checker/PAC_Checker_Relation.thy +++ b/thys/PAC_Checker/PAC_Checker_Relation.thy @@ -1,387 +1,387 @@ (* File: PAC_Checker_Relation.thy Author: Mathias Fleury, Daniela Kaufmann, JKU Maintainer: Mathias Fleury, JKU *) theory PAC_Checker_Relation imports PAC_Checker WB_Sort "Native_Word.Uint64" begin section \Various Refinement Relations\ text \When writing this, it was not possible to share the definition with the IsaSAT version.\ definition uint64_nat_rel :: "(uint64 \ nat) set" where \uint64_nat_rel = br nat_of_uint64 (\_. True)\ abbreviation uint64_nat_assn where \uint64_nat_assn \ pure uint64_nat_rel\ instantiation uint32 :: hashable begin definition hashcode_uint32 :: \uint32 \ uint32\ where \hashcode_uint32 n = n\ definition def_hashmap_size_uint32 :: \uint32 itself \ nat\ where \def_hashmap_size_uint32 = (\_. 16)\ \ \same as @{typ nat}\ instance by standard (simp add: def_hashmap_size_uint32_def) end instantiation uint64 :: hashable begin definition hashcode_uint64 :: \uint64 \ uint32\ where \hashcode_uint64 n = (uint32_of_nat (nat_of_uint64 ((n) AND ((2 :: uint64)^32 -1))))\ definition def_hashmap_size_uint64 :: \uint64 itself \ nat\ where \def_hashmap_size_uint64 = (\_. 16)\ \ \same as @{typ nat}\ instance by standard (simp add: def_hashmap_size_uint64_def) end lemma word_nat_of_uint64_Rep_inject[simp]: \nat_of_uint64 ai = nat_of_uint64 bi \ ai = bi\ by transfer (simp add: word_unat_eq_iff) instance uint64 :: heap by standard (auto simp: inj_def exI[of _ nat_of_uint64]) instance uint64 :: semiring_numeral by standard lemma nat_of_uint64_012[simp]: \nat_of_uint64 0 = 0\ \nat_of_uint64 2 = 2\ \nat_of_uint64 1 = 1\ by (transfer, auto)+ definition uint64_of_nat_conv where [simp]: \uint64_of_nat_conv (x :: nat) = x\ -lemma less_upper_bintrunc_id: \n < 2 ^b \ n \ 0 \ bintrunc b n = n\ - unfolding uint32_of_nat_def - by (simp add: no_bintr_alt1) + +lemma less_upper_bintrunc_id: \n < 2 ^b \ n \ 0 \ take_bit b n = n\ for n :: int + by (rule take_bit_int_eq_self) lemma nat_of_uint64_uint64_of_nat_id: \n < 2^64 \ nat_of_uint64 (uint64_of_nat n) = n\ unfolding uint64_of_nat_def apply simp apply transfer apply (subst unat_eq_nat_uint) apply transfer by (auto simp: less_upper_bintrunc_id) lemma [sepref_fr_rules]: \(return o uint64_of_nat, RETURN o uint64_of_nat_conv) \ [\a. a < 2 ^64]\<^sub>a nat_assn\<^sup>k \ uint64_nat_assn\ by sepref_to_hoare (sep_auto simp: uint64_nat_rel_def br_def nat_of_uint64_uint64_of_nat_id) definition string_rel :: \(String.literal \ string) set\ where \string_rel = {(x, y). y = String.explode x}\ abbreviation string_assn :: \string \ String.literal \ assn\ where \string_assn \ pure string_rel\ lemma eq_string_eq: \((=), (=)) \ string_rel \ string_rel \ bool_rel\ by (auto intro!: frefI simp: string_rel_def String.less_literal_def less_than_char_def rel2p_def literal.explode_inject) lemmas eq_string_eq_hnr = eq_string_eq[sepref_import_param] definition string2_rel :: \(string \ string) set\ where \string2_rel \ \Id\list_rel\ abbreviation string2_assn :: \string \ string \ assn\ where \string2_assn \ pure string2_rel\ abbreviation monom_rel where \monom_rel \ \string_rel\list_rel\ abbreviation monom_assn where \monom_assn \ list_assn string_assn\ abbreviation monomial_rel where \monomial_rel \ monom_rel \\<^sub>r int_rel\ abbreviation monomial_assn where \monomial_assn \ monom_assn \\<^sub>a int_assn\ abbreviation poly_rel where \poly_rel \ \monomial_rel\list_rel\ abbreviation poly_assn where \poly_assn \ list_assn monomial_assn\ lemma poly_assn_alt_def: \poly_assn = pure poly_rel\ by (simp add: list_assn_pure_conv) abbreviation polys_assn where \polys_assn \ hm_fmap_assn uint64_nat_assn poly_assn\ lemma string_rel_string_assn: \(\ ((c, a) \ string_rel)) = string_assn a c\ by (auto simp: pure_app_eq) lemma single_valued_string_rel: \single_valued string_rel\ by (auto simp: single_valued_def string_rel_def) lemma IS_LEFT_UNIQUE_string_rel: \IS_LEFT_UNIQUE string_rel\ by (auto simp: IS_LEFT_UNIQUE_def single_valued_def string_rel_def literal.explode_inject) lemma IS_RIGHT_UNIQUE_string_rel: \IS_RIGHT_UNIQUE string_rel\ by (auto simp: single_valued_def string_rel_def literal.explode_inject) lemma single_valued_monom_rel: \single_valued monom_rel\ by (rule list_rel_sv) (auto intro!: frefI simp: string_rel_def rel2p_def single_valued_def p2rel_def) lemma single_valued_monomial_rel: \single_valued monomial_rel\ using single_valued_monom_rel by (auto intro!: frefI simp: rel2p_def single_valued_def p2rel_def) lemma single_valued_monom_rel': \IS_LEFT_UNIQUE monom_rel\ unfolding IS_LEFT_UNIQUE_def inv_list_rel_eq string2_rel_def by (rule list_rel_sv)+ (auto intro!: frefI simp: string_rel_def rel2p_def single_valued_def p2rel_def literal.explode_inject) lemma single_valued_monomial_rel': \IS_LEFT_UNIQUE monomial_rel\ using single_valued_monom_rel' unfolding IS_LEFT_UNIQUE_def inv_list_rel_eq by (auto intro!: frefI simp: rel2p_def single_valued_def p2rel_def) lemma [safe_constraint_rules]: \Sepref_Constraints.CONSTRAINT single_valued string_rel\ \Sepref_Constraints.CONSTRAINT IS_LEFT_UNIQUE string_rel\ by (auto simp: CONSTRAINT_def single_valued_def string_rel_def IS_LEFT_UNIQUE_def literal.explode_inject) lemma eq_string_monom_hnr[sepref_fr_rules]: \(uncurry (return oo (=)), uncurry (RETURN oo (=))) \ monom_assn\<^sup>k *\<^sub>a monom_assn\<^sup>k \\<^sub>a bool_assn\ using single_valued_monom_rel' single_valued_monom_rel unfolding list_assn_pure_conv by sepref_to_hoare (sep_auto simp: list_assn_pure_conv string_rel_string_assn single_valued_def IS_LEFT_UNIQUE_def dest!: mod_starD simp flip: inv_list_rel_eq) definition term_order_rel' where [simp]: \term_order_rel' x y = ((x, y) \ term_order_rel)\ lemma term_order_rel[def_pat_rules]: \(\)$(x,y)$term_order_rel \ term_order_rel'$x$y\ by auto lemma term_order_rel_alt_def: \term_order_rel = lexord (p2rel char.lexordp)\ by (auto simp: p2rel_def char.lexordp_conv_lexord var_order_rel_def intro!: arg_cong[of _ _ lexord]) instantiation char :: linorder begin definition less_char where [symmetric, simp]: "less_char = PAC_Polynomials_Term.less_char" definition less_eq_char where [symmetric, simp]: "less_eq_char = PAC_Polynomials_Term.less_eq_char" instance apply standard using char.linorder_axioms by (auto simp: class.linorder_def class.order_def class.preorder_def less_eq_char_def less_than_char_def class.order_axioms_def class.linorder_axioms_def p2rel_def less_char_def) end instantiation list :: (linorder) linorder begin definition less_list where "less_list = lexordp (<)" definition less_eq_list where "less_eq_list = lexordp_eq" instance proof standard have [dest]: \\x y :: 'a :: linorder list. (x, y) \ lexord {(x, y). x < y} \ lexordp_eq y x \ False\ by (metis lexordp_antisym lexordp_conv_lexord lexordp_eq_conv_lexord) have [simp]: \\x y :: 'a :: linorder list. lexordp_eq x y \ \ lexordp_eq y x \ (x, y) \ lexord {(x, y). x < y}\ using lexordp_conv_lexord lexordp_conv_lexordp_eq by blast show \(x < y) = Restricted_Predicates.strict (\) x y\ \x \ x\ \x \ y \ y \ z \ x \ z\ \x \ y \ y \ x \ x = y\ \x \ y \ y \ x\ for x y z :: \'a :: linorder list\ by (auto simp: less_list_def less_eq_list_def List.lexordp_def lexordp_conv_lexord lexordp_into_lexordp_eq lexordp_antisym antisym_def lexordp_eq_refl lexordp_eq_linear intro: lexordp_eq_trans dest: lexordp_eq_antisym) qed end lemma term_order_rel'_alt_def_lexord: \term_order_rel' x y = ord_class.lexordp x y\ and term_order_rel'_alt_def: \term_order_rel' x y \ x < y\ proof - show \term_order_rel' x y = ord_class.lexordp x y\ \term_order_rel' x y \ x < y\ unfolding less_than_char_of_char[symmetric, abs_def] by (auto simp: lexordp_conv_lexord less_eq_list_def less_list_def lexordp_def var_order_rel_def rel2p_def term_order_rel_alt_def p2rel_def) qed lemma list_rel_list_rel_order_iff: assumes \(a, b) \ \string_rel\list_rel\ \(a', b') \ \string_rel\list_rel\ shows \a < a' \ b < b'\ proof have H: \(a, b) \ \string_rel\list_rel \ (a, cs) \ \string_rel\list_rel \ b = cs\ for cs using single_valued_monom_rel' IS_RIGHT_UNIQUE_string_rel unfolding string2_rel_def by (subst (asm)list_rel_sv_iff[symmetric]) (auto simp: single_valued_def) assume \a < a'\ then consider u u' where \a' = a @ u # u'\ | u aa v w aaa where \a = u @ aa # v\ \a' = u @ aaa # w\ \aa < aaa\ by (subst (asm) less_list_def) (auto simp: lexord_def List.lexordp_def list_rel_append1 list_rel_split_right_iff) then show \b < b'\ proof cases case 1 then show \b < b'\ using assms by (subst less_list_def) (auto simp: lexord_def List.lexordp_def list_rel_append1 list_rel_split_right_iff dest: H) next case 2 then obtain u' aa' v' w' aaa' where \b = u' @ aa' # v'\ \b' = u' @ aaa' # w'\ \(aa, aa') \ string_rel\ \(aaa, aaa') \ string_rel\ using assms by (smt list_rel_append1 list_rel_split_right_iff single_valued_def single_valued_monom_rel) with \aa < aaa\ have \aa' < aaa'\ by (auto simp: string_rel_def less_literal.rep_eq less_list_def lexordp_conv_lexord lexordp_def char.lexordp_conv_lexord simp flip: less_char_def PAC_Polynomials_Term.less_char_def) then show \b < b'\ using \b = u' @ aa' # v'\ \b' = u' @ aaa' # w'\ by (subst less_list_def) (fastforce simp: lexord_def List.lexordp_def list_rel_append1 list_rel_split_right_iff) qed next have H: \(a, b) \ \string_rel\list_rel \ (a', b) \ \string_rel\list_rel \ a = a'\ for a a' b using single_valued_monom_rel' by (auto simp: single_valued_def IS_LEFT_UNIQUE_def simp flip: inv_list_rel_eq) assume \b < b'\ then consider u u' where \b' = b @ u # u'\ | u aa v w aaa where \b = u @ aa # v\ \b' = u @ aaa # w\ \aa < aaa\ by (subst (asm) less_list_def) (auto simp: lexord_def List.lexordp_def list_rel_append1 list_rel_split_right_iff) then show \a < a'\ proof cases case 1 then show \a < a'\ using assms by (subst less_list_def) (auto simp: lexord_def List.lexordp_def list_rel_append2 list_rel_split_left_iff dest: H) next case 2 then obtain u' aa' v' w' aaa' where \a = u' @ aa' # v'\ \a' = u' @ aaa' # w'\ \(aa', aa) \ string_rel\ \(aaa', aaa) \ string_rel\ using assms by (auto simp: lexord_def List.lexordp_def list_rel_append2 list_rel_split_left_iff dest: H) with \aa < aaa\ have \aa' < aaa'\ by (auto simp: string_rel_def less_literal.rep_eq less_list_def lexordp_conv_lexord lexordp_def char.lexordp_conv_lexord simp flip: less_char_def PAC_Polynomials_Term.less_char_def) then show \a < a'\ using \a = u' @ aa' # v'\ \a' = u' @ aaa' # w'\ by (subst less_list_def) (fastforce simp: lexord_def List.lexordp_def list_rel_append1 list_rel_split_right_iff) qed qed lemma string_rel_le[sepref_import_param]: shows \((<), (<)) \ \string_rel\list_rel \ \string_rel\list_rel \ bool_rel\ by (auto intro!: fun_relI simp: list_rel_list_rel_order_iff) (* TODO Move *) lemma [sepref_import_param]: assumes \CONSTRAINT IS_LEFT_UNIQUE R\ \CONSTRAINT IS_RIGHT_UNIQUE R\ shows \(remove1, remove1) \ R \ \R\list_rel \ \R\list_rel\ apply (intro fun_relI) subgoal premises p for x y xs ys using p(2) p(1) assms by (induction xs ys rule: list_rel_induct) (auto simp: IS_LEFT_UNIQUE_def single_valued_def) done instantiation pac_step :: (heap, heap, heap) heap begin instance proof standard obtain f :: \'a \ nat\ where f: \inj f\ by blast obtain g :: \nat \ nat \ nat \ nat \ nat \ nat\ where g: \inj g\ by blast obtain h :: \'b \ nat\ where h: \inj h\ by blast obtain i :: \'c \ nat\ where i: \inj i\ by blast have [iff]: \g a = g b \ a = b\\h a'' = h b'' \ a'' = b''\ \f a' = f b' \ a' = b'\ \i a''' = i b''' \ a''' = b'''\ for a b a' b' a'' b'' a''' b''' using f g h i unfolding inj_def by blast+ let ?f = \\x :: ('a, 'b, 'c) pac_step. g (case x of Add a b c d \ (0, i a, i b, i c, f d) | Del a \ (1, i a, 0, 0, 0) | Mult a b c d \ (2, i a, f b, i c, f d) | Extension a b c \ (3, i a, f c, 0, h b))\ have \inj ?f\ apply (auto simp: inj_def) apply (case_tac x; case_tac y) apply auto done then show \\f :: ('a, 'b, 'c) pac_step \ nat. inj f\ by blast qed end end \ No newline at end of file diff --git a/thys/Word_Lib/Ancient_Numeral.thy b/thys/Word_Lib/Ancient_Numeral.thy --- a/thys/Word_Lib/Ancient_Numeral.thy +++ b/thys/Word_Lib/Ancient_Numeral.thy @@ -1,235 +1,235 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) theory Ancient_Numeral - imports Main Reversed_Bit_Lists + imports Main Reversed_Bit_Lists Legacy_Aliases begin definition Bit :: "int \ bool \ int" (infixl "BIT" 90) where "k BIT b = (if b then 1 else 0) + k + k" lemma Bit_B0: "k BIT False = k + k" by (simp add: Bit_def) lemma Bit_B1: "k BIT True = k + k + 1" by (simp add: Bit_def) lemma Bit_B0_2t: "k BIT False = 2 * k" by (rule trans, rule Bit_B0) simp lemma Bit_B1_2t: "k BIT True = 2 * k + 1" by (rule trans, rule Bit_B1) simp lemma uminus_Bit_eq: "- k BIT b = (- k - of_bool b) BIT b" by (cases b) (simp_all add: Bit_def) lemma power_BIT: "2 ^ Suc n - 1 = (2 ^ n - 1) BIT True" by (simp add: Bit_B1) lemma bin_rl_simp [simp]: "bin_rest w BIT bin_last w = w" by (simp add: Bit_def) lemma bin_rest_BIT [simp]: "bin_rest (x BIT b) = x" by (simp add: Bit_def) lemma even_BIT [simp]: "even (x BIT b) \ \ b" by (simp add: Bit_def) lemma bin_last_BIT [simp]: "bin_last (x BIT b) = b" by simp lemma BIT_eq_iff [iff]: "u BIT b = v BIT c \ u = v \ b = c" by (auto simp: Bit_def) arith+ lemma BIT_bin_simps [simp]: "numeral k BIT False = numeral (Num.Bit0 k)" "numeral k BIT True = numeral (Num.Bit1 k)" "(- numeral k) BIT False = - numeral (Num.Bit0 k)" "(- numeral k) BIT True = - numeral (Num.BitM k)" by (simp_all only: Bit_B0 Bit_B1 numeral.simps numeral_BitM) lemma BIT_special_simps [simp]: shows "0 BIT False = 0" and "0 BIT True = 1" and "1 BIT False = 2" and "1 BIT True = 3" and "(- 1) BIT False = - 2" and "(- 1) BIT True = - 1" by (simp_all add: Bit_def) lemma Bit_eq_0_iff: "w BIT b = 0 \ w = 0 \ \ b" by (auto simp: Bit_def) arith lemma Bit_eq_m1_iff: "w BIT b = -1 \ w = -1 \ b" by (auto simp: Bit_def) arith lemma expand_BIT: "numeral (Num.Bit0 w) = numeral w BIT False" "numeral (Num.Bit1 w) = numeral w BIT True" "- numeral (Num.Bit0 w) = (- numeral w) BIT False" "- numeral (Num.Bit1 w) = (- numeral (w + Num.One)) BIT True" by (simp_all add: BitM_inc_eq add_One) lemma less_Bits: "v BIT b < w BIT c \ v < w \ v \ w \ \ b \ c" by (auto simp: Bit_def) lemma le_Bits: "v BIT b \ w BIT c \ v < w \ v \ w \ (\ b \ c)" by (auto simp: Bit_def) lemma pred_BIT_simps [simp]: "x BIT False - 1 = (x - 1) BIT True" "x BIT True - 1 = x BIT False" by (simp_all add: Bit_B0_2t Bit_B1_2t) lemma succ_BIT_simps [simp]: "x BIT False + 1 = x BIT True" "x BIT True + 1 = (x + 1) BIT False" by (simp_all add: Bit_B0_2t Bit_B1_2t) lemma add_BIT_simps [simp]: "x BIT False + y BIT False = (x + y) BIT False" "x BIT False + y BIT True = (x + y) BIT True" "x BIT True + y BIT False = (x + y) BIT True" "x BIT True + y BIT True = (x + y + 1) BIT False" by (simp_all add: Bit_B0_2t Bit_B1_2t) lemma mult_BIT_simps [simp]: "x BIT False * y = (x * y) BIT False" "x * y BIT False = (x * y) BIT False" "x BIT True * y = (x * y) BIT False + y" by (simp_all add: Bit_B0_2t Bit_B1_2t algebra_simps) lemma B_mod_2': "X = 2 \ (w BIT True) mod X = 1 \ (w BIT False) mod X = 0" by (simp add: Bit_B0 Bit_B1) lemma bin_ex_rl: "\w b. w BIT b = bin" by (metis bin_rl_simp) lemma bin_exhaust: "(\x b. bin = x BIT b \ Q) \ Q" by (metis bin_ex_rl) lemma bin_abs_lem: "bin = (w BIT b) \ bin \ -1 \ bin \ 0 \ nat \w\ < nat \bin\" apply clarsimp apply (unfold Bit_def) apply (cases b) apply (clarsimp, arith) apply (clarsimp, arith) done lemma bin_induct: assumes PPls: "P 0" and PMin: "P (- 1)" and PBit: "\bin bit. P bin \ P (bin BIT bit)" shows "P bin" apply (rule_tac P=P and a=bin and f1="nat \ abs" in wf_measure [THEN wf_induct]) apply (simp add: measure_def inv_image_def) apply (case_tac x rule: bin_exhaust) apply (frule bin_abs_lem) apply (auto simp add : PPls PMin PBit) done lemma Bit_div2: "(w BIT b) div 2 = w" by (fact bin_rest_BIT) lemma twice_conv_BIT: "2 * x = x BIT False" by (simp add: Bit_def) lemma BIT_lt0 [simp]: "x BIT b < 0 \ x < 0" by(cases b)(auto simp add: Bit_def) lemma BIT_ge0 [simp]: "x BIT b \ 0 \ x \ 0" by(cases b)(auto simp add: Bit_def) lemma bin_to_bl_aux_Bit_minus_simp [simp]: "0 < n \ bin_to_bl_aux n (w BIT b) bl = bin_to_bl_aux (n - 1) w (b # bl)" by (cases n) auto lemma bl_to_bin_BIT: "bl_to_bin bs BIT b = bl_to_bin (bs @ [b])" by (simp add: bl_to_bin_append Bit_def) lemma bin_nth_0_BIT: "bin_nth (w BIT b) 0 \ b" by simp lemma bin_nth_Suc_BIT: "bin_nth (w BIT b) (Suc n) = bin_nth w n" by (simp add: bit_Suc) lemma bin_nth_minus [simp]: "0 < n \ bin_nth (w BIT b) n = bin_nth w (n - 1)" by (cases n) (simp_all add: bit_Suc) lemma bin_sign_simps [simp]: "bin_sign (w BIT b) = bin_sign w" by (simp add: bin_sign_def Bit_def) lemma bin_nth_Bit: "bin_nth (w BIT b) n \ n = 0 \ b \ (\m. n = Suc m \ bin_nth w m)" by (cases n) auto lemmas sbintrunc_Suc_BIT [simp] = signed_take_bit_Suc [where a="w BIT b", simplified bin_last_BIT bin_rest_BIT] for w b lemmas sbintrunc_0_BIT_B0 [simp] = signed_take_bit_0 [where a="w BIT False", simplified bin_last_numeral_simps bin_rest_numeral_simps] for w lemmas sbintrunc_0_BIT_B1 [simp] = signed_take_bit_0 [where a="w BIT True", simplified bin_last_BIT bin_rest_numeral_simps] for w lemma sbintrunc_Suc_minus_Is: \0 < n \ sbintrunc (n - 1) w = y \ sbintrunc n (w BIT b) = y BIT b\ by (cases n) (simp_all add: Bit_def signed_take_bit_Suc) lemma bin_cat_Suc_Bit: "bin_cat w (Suc n) (v BIT b) = bin_cat w n v BIT b" by (auto simp add: Bit_def concat_bit_Suc) lemma int_not_BIT [simp]: "NOT (w BIT b) = (NOT w) BIT (\ b)" by (simp add: not_int_def Bit_def) lemma int_and_Bits [simp]: "(x BIT b) AND (y BIT c) = (x AND y) BIT (b \ c)" using and_int_rec [of \x BIT b\ \y BIT c\] by (auto simp add: Bit_B0_2t Bit_B1_2t) lemma int_or_Bits [simp]: "(x BIT b) OR (y BIT c) = (x OR y) BIT (b \ c)" using or_int_rec [of \x BIT b\ \y BIT c\] by (auto simp add: Bit_B0_2t Bit_B1_2t) lemma int_xor_Bits [simp]: "(x BIT b) XOR (y BIT c) = (x XOR y) BIT ((b \ c) \ \ (b \ c))" using xor_int_rec [of \x BIT b\ \y BIT c\] by (auto simp add: Bit_B0_2t Bit_B1_2t) lemma mod_BIT: "bin BIT bit mod 2 ^ Suc n = (bin mod 2 ^ n) BIT bit" for bit proof - have "2 * (bin mod 2 ^ n) + 1 = (2 * bin mod 2 ^ Suc n) + 1" by (simp add: mod_mult_mult1) also have "\ = ((2 * bin mod 2 ^ Suc n) + 1) mod 2 ^ Suc n" by (simp add: ac_simps pos_zmod_mult_2) also have "\ = (2 * bin + 1) mod 2 ^ Suc n" by (simp only: mod_simps) finally show ?thesis by (auto simp add: Bit_def) qed lemma minus_BIT_0: fixes x y :: int shows "x BIT b - y BIT False = (x - y) BIT b" by(simp add: Bit_def) lemma int_lsb_BIT [simp]: fixes x :: int shows "lsb (x BIT b) \ b" by(simp add: lsb_int_def) lemma int_shiftr_BIT [simp]: fixes x :: int shows int_shiftr0: "x >> 0 = x" and int_shiftr_Suc: "x BIT b >> Suc n = x >> n" proof - show "x >> 0 = x" by (simp add: shiftr_int_def) show "x BIT b >> Suc n = x >> n" by (cases b) (simp_all add: shiftr_int_def Bit_def add.commute pos_zdiv_mult_2) qed lemma msb_BIT [simp]: "msb (x BIT b) = msb x" by(simp add: msb_int_def) end \ No newline at end of file diff --git a/thys/Word_Lib/Bits_Int.thy b/thys/Word_Lib/Bits_Int.thy --- a/thys/Word_Lib/Bits_Int.thy +++ b/thys/Word_Lib/Bits_Int.thy @@ -1,1485 +1,1463 @@ (* * Copyright Brian Huffman, PSU; Jeremy Dawson and Gerwin Klein, NICTA * * SPDX-License-Identifier: BSD-2-Clause *) section \Bitwise Operations on integers\ theory Bits_Int imports "HOL-Library.Word" Traditional_Infix_Syntax begin subsection \Implicit bit representation of \<^typ>\int\\ -abbreviation (input) bin_last :: "int \ bool" - where "bin_last \ odd" - lemma bin_last_def: - "bin_last w \ w mod 2 = 1" + "(odd :: int \ bool) w \ w mod 2 = 1" by (fact odd_iff_mod_2_eq_one) -abbreviation (input) bin_rest :: "int \ int" - where "bin_rest w \ w div 2" - lemma bin_last_numeral_simps [simp]: "\ odd (0 :: int)" "odd (1 :: int)" "odd (- 1 :: int)" "odd (Numeral1 :: int)" "\ odd (numeral (Num.Bit0 w) :: int)" "odd (numeral (Num.Bit1 w) :: int)" "\ odd (- numeral (Num.Bit0 w) :: int)" "odd (- numeral (Num.Bit1 w) :: int)" by simp_all lemma bin_rest_numeral_simps [simp]: - "bin_rest 0 = 0" - "bin_rest 1 = 0" - "bin_rest (- 1) = - 1" - "bin_rest Numeral1 = 0" - "bin_rest (numeral (Num.Bit0 w)) = numeral w" - "bin_rest (numeral (Num.Bit1 w)) = numeral w" - "bin_rest (- numeral (Num.Bit0 w)) = - numeral w" - "bin_rest (- numeral (Num.Bit1 w)) = - numeral (w + Num.One)" + "(\k::int. k div 2) 0 = 0" + "(\k::int. k div 2) 1 = 0" + "(\k::int. k div 2) (- 1) = - 1" + "(\k::int. k div 2) Numeral1 = 0" + "(\k::int. k div 2) (numeral (Num.Bit0 w)) = numeral w" + "(\k::int. k div 2) (numeral (Num.Bit1 w)) = numeral w" + "(\k::int. k div 2) (- numeral (Num.Bit0 w)) = - numeral w" + "(\k::int. k div 2) (- numeral (Num.Bit1 w)) = - numeral (w + Num.One)" by simp_all -lemma bin_rl_eqI: "\bin_rest x = bin_rest y; odd x = odd y\ \ x = y" +lemma bin_rl_eqI: "\(\k::int. k div 2) x = (\k::int. k div 2) y; odd x = odd y\ \ x = y" by (auto elim: oddE) lemma [simp]: - shows bin_rest_lt0: "bin_rest i < 0 \ i < 0" - and bin_rest_ge_0: "bin_rest i \ 0 \ i \ 0" + shows bin_rest_lt0: "(\k::int. k div 2) i < 0 \ i < 0" + and bin_rest_ge_0: "(\k::int. k div 2) i \ 0 \ i \ 0" by auto -lemma bin_rest_gt_0 [simp]: "bin_rest x > 0 \ x > 1" +lemma bin_rest_gt_0 [simp]: "(\k::int. k div 2) x > 0 \ x > 1" by auto subsection \Bit projection\ -abbreviation (input) bin_nth :: \int \ nat \ bool\ - where \bin_nth \ bit\ - -lemma bin_nth_eq_iff: "bin_nth x = bin_nth y \ x = y" +lemma bin_nth_eq_iff: "(bit :: int \ nat \ bool) x = (bit :: int \ nat \ bool) y \ x = y" by (simp add: bit_eq_iff fun_eq_iff) lemma bin_eqI: - "x = y" if "\n. bin_nth x n \ bin_nth y n" + "x = y" if "\n. (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" using that bin_nth_eq_iff [of x y] by (simp add: fun_eq_iff) -lemma bin_eq_iff: "x = y \ (\n. bin_nth x n = bin_nth y n)" +lemma bin_eq_iff: "x = y \ (\n. (bit :: int \ nat \ bool) x n = (bit :: int \ nat \ bool) y n)" by (fact bit_eq_iff) -lemma bin_nth_zero [simp]: "\ bin_nth 0 n" +lemma bin_nth_zero [simp]: "\ (bit :: int \ nat \ bool) 0 n" by simp -lemma bin_nth_1 [simp]: "bin_nth 1 n \ n = 0" +lemma bin_nth_1 [simp]: "(bit :: int \ nat \ bool) 1 n \ n = 0" by (cases n) (simp_all add: bit_Suc) -lemma bin_nth_minus1 [simp]: "bin_nth (- 1) n" +lemma bin_nth_minus1 [simp]: "(bit :: int \ nat \ bool) (- 1) n" by (induction n) (simp_all add: bit_Suc) -lemma bin_nth_numeral: "bin_rest x = y \ bin_nth x (numeral n) = bin_nth y (pred_numeral n)" +lemma bin_nth_numeral: "(\k::int. k div 2) x = y \ (bit :: int \ nat \ bool) x (numeral n) = (bit :: int \ nat \ bool) y (pred_numeral n)" by (simp add: numeral_eq_Suc bit_Suc) lemmas bin_nth_numeral_simps [simp] = bin_nth_numeral [OF bin_rest_numeral_simps(8)] lemmas bin_nth_simps = bit_0 bit_Suc bin_nth_zero bin_nth_minus1 bin_nth_numeral_simps -lemma nth_2p_bin: "bin_nth (2 ^ n) m = (m = n)" \ \for use when simplifying with \bin_nth_Bit\\ +lemma nth_2p_bin: "(bit :: int \ nat \ bool) (2 ^ n) m = (m = n)" \ \for use when simplifying with \bin_nth_Bit\\ by (auto simp add: bit_exp_iff) -lemma nth_rest_power_bin: "bin_nth ((bin_rest ^^ k) w) n = bin_nth w (n + k)" +lemma nth_rest_power_bin: "(bit :: int \ nat \ bool) (((\k::int. k div 2) ^^ k) w) n = (bit :: int \ nat \ bool) w (n + k)" apply (induct k arbitrary: n) apply clarsimp apply clarsimp apply (simp only: bit_Suc [symmetric] add_Suc) done lemma bin_nth_numeral_unfold: - "bin_nth (numeral (num.Bit0 x)) n \ n > 0 \ bin_nth (numeral x) (n - 1)" - "bin_nth (numeral (num.Bit1 x)) n \ (n > 0 \ bin_nth (numeral x) (n - 1))" + "(bit :: int \ nat \ bool) (numeral (num.Bit0 x)) n \ n > 0 \ (bit :: int \ nat \ bool) (numeral x) (n - 1)" + "(bit :: int \ nat \ bool) (numeral (num.Bit1 x)) n \ (n > 0 \ (bit :: int \ nat \ bool) (numeral x) (n - 1))" by (cases n; simp)+ subsection \Truncating\ definition bin_sign :: "int \ int" where "bin_sign k = (if k \ 0 then 0 else - 1)" lemma bin_sign_simps [simp]: "bin_sign 0 = 0" "bin_sign 1 = 0" "bin_sign (- 1) = - 1" "bin_sign (numeral k) = 0" "bin_sign (- numeral k) = -1" by (simp_all add: bin_sign_def) -lemma bin_sign_rest [simp]: "bin_sign (bin_rest w) = bin_sign w" +lemma bin_sign_rest [simp]: "bin_sign ((\k::int. k div 2) w) = bin_sign w" by (simp add: bin_sign_def) -abbreviation (input) bintrunc :: \nat \ int \ int\ - where \bintrunc \ take_bit\ - -lemma bintrunc_mod2p: "bintrunc n w = w mod 2 ^ n" +lemma bintrunc_mod2p: "(take_bit :: nat \ int \ int) n w = w mod 2 ^ n" by (fact take_bit_eq_mod) -abbreviation (input) sbintrunc :: \nat \ int \ int\ - where \sbintrunc \ signed_take_bit\ - -abbreviation (input) norm_sint :: \nat \ int \ int\ - where \norm_sint n \ signed_take_bit (n - 1)\ - -lemma sbintrunc_mod2p: "sbintrunc n w = (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n" +lemma sbintrunc_mod2p: "(signed_take_bit :: nat \ int \ int) n w = (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n" by (simp add: bintrunc_mod2p signed_take_bit_eq_take_bit_shift) lemma sbintrunc_eq_take_bit: - \sbintrunc n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n\ + \(signed_take_bit :: nat \ int \ int) n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n\ by (fact signed_take_bit_eq_take_bit_shift) -lemma sign_bintr: "bin_sign (bintrunc n w) = 0" +lemma sign_bintr: "bin_sign ((take_bit :: nat \ int \ int) n w) = 0" by (simp add: bin_sign_def) -lemma bintrunc_n_0: "bintrunc n 0 = 0" +lemma bintrunc_n_0: "(take_bit :: nat \ int \ int) n 0 = 0" by (fact take_bit_of_0) -lemma sbintrunc_n_0: "sbintrunc n 0 = 0" +lemma sbintrunc_n_0: "(signed_take_bit :: nat \ int \ int) n 0 = 0" by (fact signed_take_bit_of_0) -lemma sbintrunc_n_minus1: "sbintrunc n (- 1) = -1" +lemma sbintrunc_n_minus1: "(signed_take_bit :: nat \ int \ int) n (- 1) = -1" by (fact signed_take_bit_of_minus_1) lemma bintrunc_Suc_numeral: - "bintrunc (Suc n) 1 = 1" - "bintrunc (Suc n) (- 1) = 1 + 2 * bintrunc n (- 1)" - "bintrunc (Suc n) (numeral (Num.Bit0 w)) = 2 * bintrunc n (numeral w)" - "bintrunc (Suc n) (numeral (Num.Bit1 w)) = 1 + 2 * bintrunc n (numeral w)" - "bintrunc (Suc n) (- numeral (Num.Bit0 w)) = 2 * bintrunc n (- numeral w)" - "bintrunc (Suc n) (- numeral (Num.Bit1 w)) = 1 + 2 * bintrunc n (- numeral (w + Num.One))" + "(take_bit :: nat \ int \ int) (Suc n) 1 = 1" + "(take_bit :: nat \ int \ int) (Suc n) (- 1) = 1 + 2 * (take_bit :: nat \ int \ int) n (- 1)" + "(take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) n (numeral w)" + "(take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) n (numeral w)" + "(take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) n (- numeral w)" + "(take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) n (- numeral (w + Num.One))" by (simp_all add: take_bit_Suc) lemma sbintrunc_0_numeral [simp]: - "sbintrunc 0 1 = -1" - "sbintrunc 0 (numeral (Num.Bit0 w)) = 0" - "sbintrunc 0 (numeral (Num.Bit1 w)) = -1" - "sbintrunc 0 (- numeral (Num.Bit0 w)) = 0" - "sbintrunc 0 (- numeral (Num.Bit1 w)) = -1" + "(signed_take_bit :: nat \ int \ int) 0 1 = -1" + "(signed_take_bit :: nat \ int \ int) 0 (numeral (Num.Bit0 w)) = 0" + "(signed_take_bit :: nat \ int \ int) 0 (numeral (Num.Bit1 w)) = -1" + "(signed_take_bit :: nat \ int \ int) 0 (- numeral (Num.Bit0 w)) = 0" + "(signed_take_bit :: nat \ int \ int) 0 (- numeral (Num.Bit1 w)) = -1" by simp_all lemma sbintrunc_Suc_numeral: - "sbintrunc (Suc n) 1 = 1" - "sbintrunc (Suc n) (numeral (Num.Bit0 w)) = 2 * sbintrunc n (numeral w)" - "sbintrunc (Suc n) (numeral (Num.Bit1 w)) = 1 + 2 * sbintrunc n (numeral w)" - "sbintrunc (Suc n) (- numeral (Num.Bit0 w)) = 2 * sbintrunc n (- numeral w)" - "sbintrunc (Suc n) (- numeral (Num.Bit1 w)) = 1 + 2 * sbintrunc n (- numeral (w + Num.One))" + "(signed_take_bit :: nat \ int \ int) (Suc n) 1 = 1" + "(signed_take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) n (numeral w)" + "(signed_take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) n (numeral w)" + "(signed_take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) n (- numeral w)" + "(signed_take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) n (- numeral (w + Num.One))" by (simp_all add: signed_take_bit_Suc) -lemma bin_sign_lem: "(bin_sign (sbintrunc n bin) = -1) = bit bin n" +lemma bin_sign_lem: "(bin_sign ((signed_take_bit :: nat \ int \ int) n bin) = -1) = bit bin n" by (simp add: bin_sign_def) -lemma nth_bintr: "bin_nth (bintrunc m w) n \ n < m \ bin_nth w n" +lemma nth_bintr: "(bit :: int \ nat \ bool) ((take_bit :: nat \ int \ int) m w) n \ n < m \ (bit :: int \ nat \ bool) w n" by (fact bit_take_bit_iff) -lemma nth_sbintr: "bin_nth (sbintrunc m w) n = (if n < m then bin_nth w n else bin_nth w m)" +lemma nth_sbintr: "(bit :: int \ nat \ bool) ((signed_take_bit :: nat \ int \ int) m w) n = (if n < m then (bit :: int \ nat \ bool) w n else (bit :: int \ nat \ bool) w m)" by (simp add: bit_signed_take_bit_iff min_def) lemma bin_nth_Bit0: - "bin_nth (numeral (Num.Bit0 w)) n \ - (\m. n = Suc m \ bin_nth (numeral w) m)" + "(bit :: int \ nat \ bool) (numeral (Num.Bit0 w)) n \ + (\m. n = Suc m \ (bit :: int \ nat \ bool) (numeral w) m)" using bit_double_iff [of \numeral w :: int\ n] by (auto intro: exI [of _ \n - 1\]) lemma bin_nth_Bit1: - "bin_nth (numeral (Num.Bit1 w)) n \ - n = 0 \ (\m. n = Suc m \ bin_nth (numeral w) m)" + "(bit :: int \ nat \ bool) (numeral (Num.Bit1 w)) n \ + n = 0 \ (\m. n = Suc m \ (bit :: int \ nat \ bool) (numeral w) m)" using even_bit_succ_iff [of \2 * numeral w :: int\ n] bit_double_iff [of \numeral w :: int\ n] by auto -lemma bintrunc_bintrunc_l: "n \ m \ bintrunc m (bintrunc n w) = bintrunc n w" +lemma bintrunc_bintrunc_l: "n \ m \ (take_bit :: nat \ int \ int) m ((take_bit :: nat \ int \ int) n w) = (take_bit :: nat \ int \ int) n w" by simp -lemma sbintrunc_sbintrunc_l: "n \ m \ sbintrunc m (sbintrunc n w) = sbintrunc n w" +lemma sbintrunc_sbintrunc_l: "n \ m \ (signed_take_bit :: nat \ int \ int) m ((signed_take_bit :: nat \ int \ int) n w) = (signed_take_bit :: nat \ int \ int) n w" by (simp add: min_def) -lemma bintrunc_bintrunc_ge: "n \ m \ bintrunc n (bintrunc m w) = bintrunc n w" +lemma bintrunc_bintrunc_ge: "n \ m \ (take_bit :: nat \ int \ int) n ((take_bit :: nat \ int \ int) m w) = (take_bit :: nat \ int \ int) n w" by (rule bin_eqI) (auto simp: nth_bintr) -lemma bintrunc_bintrunc_min [simp]: "bintrunc m (bintrunc n w) = bintrunc (min m n) w" +lemma bintrunc_bintrunc_min [simp]: "(take_bit :: nat \ int \ int) m ((take_bit :: nat \ int \ int) n w) = (take_bit :: nat \ int \ int) (min m n) w" by (rule take_bit_take_bit) -lemma sbintrunc_sbintrunc_min [simp]: "sbintrunc m (sbintrunc n w) = sbintrunc (min m n) w" +lemma sbintrunc_sbintrunc_min [simp]: "(signed_take_bit :: nat \ int \ int) m ((signed_take_bit :: nat \ int \ int) n w) = (signed_take_bit :: nat \ int \ int) (min m n) w" by (rule signed_take_bit_signed_take_bit) lemmas sbintrunc_Suc_Pls = signed_take_bit_Suc [where a="0::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Suc_Min = signed_take_bit_Suc [where a="-1::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Sucs = sbintrunc_Suc_Pls sbintrunc_Suc_Min sbintrunc_Suc_numeral lemmas sbintrunc_Pls = signed_take_bit_0 [where a="0::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Min = signed_take_bit_0 [where a="-1::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_0_simps = sbintrunc_Pls sbintrunc_Min lemmas sbintrunc_simps = sbintrunc_0_simps sbintrunc_Sucs -lemma bintrunc_minus: "0 < n \ bintrunc (Suc (n - 1)) w = bintrunc n w" +lemma bintrunc_minus: "0 < n \ (take_bit :: nat \ int \ int) (Suc (n - 1)) w = (take_bit :: nat \ int \ int) n w" by auto -lemma sbintrunc_minus: "0 < n \ sbintrunc (Suc (n - 1)) w = sbintrunc n w" +lemma sbintrunc_minus: "0 < n \ (signed_take_bit :: nat \ int \ int) (Suc (n - 1)) w = (signed_take_bit :: nat \ int \ int) n w" by auto lemmas sbintrunc_minus_simps = sbintrunc_Sucs [THEN [2] sbintrunc_minus [symmetric, THEN trans]] lemma sbintrunc_BIT_I: \0 < n \ - sbintrunc (n - 1) 0 = y \ - sbintrunc n 0 = 2 * y\ + (signed_take_bit :: nat \ int \ int) (n - 1) 0 = y \ + (signed_take_bit :: nat \ int \ int) n 0 = 2 * y\ by simp lemma sbintrunc_Suc_Is: - \sbintrunc n (- 1) = y \ - sbintrunc (Suc n) (- 1) = 1 + 2 * y\ + \(signed_take_bit :: nat \ int \ int) n (- 1) = y \ + (signed_take_bit :: nat \ int \ int) (Suc n) (- 1) = 1 + 2 * y\ by auto -lemma sbintrunc_Suc_lem: "sbintrunc (Suc n) x = y \ m = Suc n \ sbintrunc m x = y" +lemma sbintrunc_Suc_lem: "(signed_take_bit :: nat \ int \ int) (Suc n) x = y \ m = Suc n \ (signed_take_bit :: nat \ int \ int) m x = y" by (rule ssubst) lemmas sbintrunc_Suc_Ialts = sbintrunc_Suc_Is [THEN sbintrunc_Suc_lem] -lemma sbintrunc_bintrunc_lt: "m > n \ sbintrunc n (bintrunc m w) = sbintrunc n w" +lemma sbintrunc_bintrunc_lt: "m > n \ (signed_take_bit :: nat \ int \ int) n ((take_bit :: nat \ int \ int) m w) = (signed_take_bit :: nat \ int \ int) n w" by (rule bin_eqI) (auto simp: nth_sbintr nth_bintr) -lemma bintrunc_sbintrunc_le: "m \ Suc n \ bintrunc m (sbintrunc n w) = bintrunc m w" +lemma bintrunc_sbintrunc_le: "m \ Suc n \ (take_bit :: nat \ int \ int) m ((signed_take_bit :: nat \ int \ int) n w) = (take_bit :: nat \ int \ int) m w" by (rule take_bit_signed_take_bit) lemmas bintrunc_sbintrunc [simp] = order_refl [THEN bintrunc_sbintrunc_le] lemmas sbintrunc_bintrunc [simp] = lessI [THEN sbintrunc_bintrunc_lt] lemmas bintrunc_bintrunc [simp] = order_refl [THEN bintrunc_bintrunc_l] lemmas sbintrunc_sbintrunc [simp] = order_refl [THEN sbintrunc_sbintrunc_l] -lemma bintrunc_sbintrunc' [simp]: "0 < n \ bintrunc n (sbintrunc (n - 1) w) = bintrunc n w" +lemma bintrunc_sbintrunc' [simp]: "0 < n \ (take_bit :: nat \ int \ int) n ((signed_take_bit :: nat \ int \ int) (n - 1) w) = (take_bit :: nat \ int \ int) n w" by (cases n) simp_all -lemma sbintrunc_bintrunc' [simp]: "0 < n \ sbintrunc (n - 1) (bintrunc n w) = sbintrunc (n - 1) w" +lemma sbintrunc_bintrunc' [simp]: "0 < n \ (signed_take_bit :: nat \ int \ int) (n - 1) ((take_bit :: nat \ int \ int) n w) = (signed_take_bit :: nat \ int \ int) (n - 1) w" by (cases n) simp_all -lemma bin_sbin_eq_iff: "bintrunc (Suc n) x = bintrunc (Suc n) y \ sbintrunc n x = sbintrunc n y" +lemma bin_sbin_eq_iff: "(take_bit :: nat \ int \ int) (Suc n) x = (take_bit :: nat \ int \ int) (Suc n) y \ (signed_take_bit :: nat \ int \ int) n x = (signed_take_bit :: nat \ int \ int) n y" apply (rule iffI) apply (rule box_equals [OF _ sbintrunc_bintrunc sbintrunc_bintrunc]) apply simp apply (rule box_equals [OF _ bintrunc_sbintrunc bintrunc_sbintrunc]) apply simp done lemma bin_sbin_eq_iff': - "0 < n \ bintrunc n x = bintrunc n y \ sbintrunc (n - 1) x = sbintrunc (n - 1) y" + "0 < n \ (take_bit :: nat \ int \ int) n x = (take_bit :: nat \ int \ int) n y \ (signed_take_bit :: nat \ int \ int) (n - 1) x = (signed_take_bit :: nat \ int \ int) (n - 1) y" by (cases n) (simp_all add: bin_sbin_eq_iff) lemmas bintrunc_sbintruncS0 [simp] = bintrunc_sbintrunc' [unfolded One_nat_def] lemmas sbintrunc_bintruncS0 [simp] = sbintrunc_bintrunc' [unfolded One_nat_def] lemmas bintrunc_bintrunc_l' = le_add1 [THEN bintrunc_bintrunc_l] lemmas sbintrunc_sbintrunc_l' = le_add1 [THEN sbintrunc_sbintrunc_l] (* although bintrunc_minus_simps, if added to default simpset, tends to get applied where it's not wanted in developing the theories, we get a version for when the word length is given literally *) lemmas nat_non0_gr = trans [OF iszero_def [THEN Not_eq_iff [THEN iffD2]] refl] lemma bintrunc_numeral: - "bintrunc (numeral k) x = of_bool (odd x) + 2 * bintrunc (pred_numeral k) (x div 2)" + "(take_bit :: nat \ int \ int) (numeral k) x = of_bool (odd x) + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (x div 2)" by (simp add: numeral_eq_Suc take_bit_Suc mod_2_eq_odd) lemma sbintrunc_numeral: - "sbintrunc (numeral k) x = of_bool (odd x) + 2 * sbintrunc (pred_numeral k) (x div 2)" + "(signed_take_bit :: nat \ int \ int) (numeral k) x = of_bool (odd x) + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (x div 2)" by (simp add: numeral_eq_Suc signed_take_bit_Suc mod2_eq_if) lemma bintrunc_numeral_simps [simp]: - "bintrunc (numeral k) (numeral (Num.Bit0 w)) = - 2 * bintrunc (pred_numeral k) (numeral w)" - "bintrunc (numeral k) (numeral (Num.Bit1 w)) = - 1 + 2 * bintrunc (pred_numeral k) (numeral w)" - "bintrunc (numeral k) (- numeral (Num.Bit0 w)) = - 2 * bintrunc (pred_numeral k) (- numeral w)" - "bintrunc (numeral k) (- numeral (Num.Bit1 w)) = - 1 + 2 * bintrunc (pred_numeral k) (- numeral (w + Num.One))" - "bintrunc (numeral k) 1 = 1" + "(take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit0 w)) = + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" + "(take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit1 w)) = + 1 + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" + "(take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit0 w)) = + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (- numeral w)" + "(take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit1 w)) = + 1 + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (- numeral (w + Num.One))" + "(take_bit :: nat \ int \ int) (numeral k) 1 = 1" by (simp_all add: bintrunc_numeral) lemma sbintrunc_numeral_simps [simp]: - "sbintrunc (numeral k) (numeral (Num.Bit0 w)) = - 2 * sbintrunc (pred_numeral k) (numeral w)" - "sbintrunc (numeral k) (numeral (Num.Bit1 w)) = - 1 + 2 * sbintrunc (pred_numeral k) (numeral w)" - "sbintrunc (numeral k) (- numeral (Num.Bit0 w)) = - 2 * sbintrunc (pred_numeral k) (- numeral w)" - "sbintrunc (numeral k) (- numeral (Num.Bit1 w)) = - 1 + 2 * sbintrunc (pred_numeral k) (- numeral (w + Num.One))" - "sbintrunc (numeral k) 1 = 1" + "(signed_take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit0 w)) = + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" + "(signed_take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit1 w)) = + 1 + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" + "(signed_take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit0 w)) = + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (- numeral w)" + "(signed_take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit1 w)) = + 1 + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (- numeral (w + Num.One))" + "(signed_take_bit :: nat \ int \ int) (numeral k) 1 = 1" by (simp_all add: sbintrunc_numeral) -lemma no_bintr_alt1: "bintrunc n = (\w. w mod 2 ^ n :: int)" +lemma no_bintr_alt1: "(take_bit :: nat \ int \ int) n = (\w. w mod 2 ^ n :: int)" by (rule ext) (rule bintrunc_mod2p) -lemma range_bintrunc: "range (bintrunc n) = {i. 0 \ i \ i < 2 ^ n}" +lemma range_bintrunc: "range ((take_bit :: nat \ int \ int) n) = {i. 0 \ i \ i < 2 ^ n}" by (auto simp add: take_bit_eq_mod image_iff) (metis mod_pos_pos_trivial) -lemma no_sbintr_alt2: "sbintrunc n = (\w. (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n :: int)" +lemma no_sbintr_alt2: "(signed_take_bit :: nat \ int \ int) n = (\w. (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n :: int)" by (rule ext) (simp add : sbintrunc_mod2p) -lemma range_sbintrunc: "range (sbintrunc n) = {i. - (2 ^ n) \ i \ i < 2 ^ n}" +lemma range_sbintrunc: "range ((signed_take_bit :: nat \ int \ int) n) = {i. - (2 ^ n) \ i \ i < 2 ^ n}" proof - have \surj (\k::int. k + 2 ^ n)\ by (rule surjI [of _ \(\k. k - 2 ^ n)\]) simp - moreover have \sbintrunc n = ((\k. k - 2 ^ n) \ take_bit (Suc n) \ (\k. k + 2 ^ n))\ + moreover have \(signed_take_bit :: nat \ int \ int) n = ((\k. k - 2 ^ n) \ take_bit (Suc n) \ (\k. k + 2 ^ n))\ by (simp add: sbintrunc_eq_take_bit fun_eq_iff) ultimately show ?thesis apply (simp only: fun.set_map range_bintrunc) apply (auto simp add: image_iff) apply presburger done qed lemma sbintrunc_inc: - \k + 2 ^ Suc n \ sbintrunc n k\ if \k < - (2 ^ n)\ + \k + 2 ^ Suc n \ (signed_take_bit :: nat \ int \ int) n k\ if \k < - (2 ^ n)\ using that by (fact signed_take_bit_int_greater_eq) lemma sbintrunc_dec: - \sbintrunc n k \ k - 2 ^ (Suc n)\ if \k \ 2 ^ n\ + \(signed_take_bit :: nat \ int \ int) n k \ k - 2 ^ (Suc n)\ if \k \ 2 ^ n\ using that by (fact signed_take_bit_int_less_eq) -lemma bintr_ge0: "0 \ bintrunc n w" +lemma bintr_ge0: "0 \ (take_bit :: nat \ int \ int) n w" by (simp add: bintrunc_mod2p) -lemma bintr_lt2p: "bintrunc n w < 2 ^ n" +lemma bintr_lt2p: "(take_bit :: nat \ int \ int) n w < 2 ^ n" by (simp add: bintrunc_mod2p) -lemma bintr_Min: "bintrunc n (- 1) = 2 ^ n - 1" +lemma bintr_Min: "(take_bit :: nat \ int \ int) n (- 1) = 2 ^ n - 1" by (simp add: stable_imp_take_bit_eq) -lemma sbintr_ge: "- (2 ^ n) \ sbintrunc n w" +lemma sbintr_ge: "- (2 ^ n) \ (signed_take_bit :: nat \ int \ int) n w" by (simp add: sbintrunc_mod2p) -lemma sbintr_lt: "sbintrunc n w < 2 ^ n" +lemma sbintr_lt: "(signed_take_bit :: nat \ int \ int) n w < 2 ^ n" by (simp add: sbintrunc_mod2p) lemma sign_Pls_ge_0: "bin_sign bin = 0 \ bin \ 0" for bin :: int by (simp add: bin_sign_def) lemma sign_Min_lt_0: "bin_sign bin = -1 \ bin < 0" for bin :: int by (simp add: bin_sign_def) -lemma bin_rest_trunc: "bin_rest (bintrunc n bin) = bintrunc (n - 1) (bin_rest bin)" +lemma bin_rest_trunc: "(\k::int. k div 2) ((take_bit :: nat \ int \ int) n bin) = (take_bit :: nat \ int \ int) (n - 1) ((\k::int. k div 2) bin)" by (simp add: take_bit_rec [of n bin]) lemma bin_rest_power_trunc: - "(bin_rest ^^ k) (bintrunc n bin) = bintrunc (n - k) ((bin_rest ^^ k) bin)" + "((\k::int. k div 2) ^^ k) ((take_bit :: nat \ int \ int) n bin) = (take_bit :: nat \ int \ int) (n - k) (((\k::int. k div 2) ^^ k) bin)" by (induct k) (auto simp: bin_rest_trunc) -lemma bin_rest_trunc_i: "bintrunc n (bin_rest bin) = bin_rest (bintrunc (Suc n) bin)" +lemma bin_rest_trunc_i: "(take_bit :: nat \ int \ int) n ((\k::int. k div 2) bin) = (\k::int. k div 2) ((take_bit :: nat \ int \ int) (Suc n) bin)" by (auto simp add: take_bit_Suc) -lemma bin_rest_strunc: "bin_rest (sbintrunc (Suc n) bin) = sbintrunc n (bin_rest bin)" +lemma bin_rest_strunc: "(\k::int. k div 2) ((signed_take_bit :: nat \ int \ int) (Suc n) bin) = (signed_take_bit :: nat \ int \ int) n ((\k::int. k div 2) bin)" by (simp add: signed_take_bit_Suc) -lemma bintrunc_rest [simp]: "bintrunc n (bin_rest (bintrunc n bin)) = bin_rest (bintrunc n bin)" +lemma bintrunc_rest [simp]: "(take_bit :: nat \ int \ int) n ((\k::int. k div 2) ((take_bit :: nat \ int \ int) n bin)) = (\k::int. k div 2) ((take_bit :: nat \ int \ int) n bin)" by (induct n arbitrary: bin) (simp_all add: take_bit_Suc) -lemma sbintrunc_rest [simp]: "sbintrunc n (bin_rest (sbintrunc n bin)) = bin_rest (sbintrunc n bin)" +lemma sbintrunc_rest [simp]: "(signed_take_bit :: nat \ int \ int) n ((\k::int. k div 2) ((signed_take_bit :: nat \ int \ int) n bin)) = (\k::int. k div 2) ((signed_take_bit :: nat \ int \ int) n bin)" by (induct n arbitrary: bin) (simp_all add: signed_take_bit_Suc mod2_eq_if) -lemma bintrunc_rest': "bintrunc n \ bin_rest \ bintrunc n = bin_rest \ bintrunc n" +lemma bintrunc_rest': "(take_bit :: nat \ int \ int) n \ (\k::int. k div 2) \ (take_bit :: nat \ int \ int) n = (\k::int. k div 2) \ (take_bit :: nat \ int \ int) n" by (rule ext) auto -lemma sbintrunc_rest': "sbintrunc n \ bin_rest \ sbintrunc n = bin_rest \ sbintrunc n" +lemma sbintrunc_rest': "(signed_take_bit :: nat \ int \ int) n \ (\k::int. k div 2) \ (signed_take_bit :: nat \ int \ int) n = (\k::int. k div 2) \ (signed_take_bit :: nat \ int \ int) n" by (rule ext) auto lemma rco_lem: "f \ g \ f = g \ f \ f \ (g \ f) ^^ n = g ^^ n \ f" apply (rule ext) apply (induct_tac n) apply (simp_all (no_asm)) apply (drule fun_cong) apply (unfold o_def) apply (erule trans) apply simp done lemmas rco_bintr = bintrunc_rest' [THEN rco_lem [THEN fun_cong], unfolded o_def] lemmas rco_sbintr = sbintrunc_rest' [THEN rco_lem [THEN fun_cong], unfolded o_def] subsection \Splitting and concatenation\ definition bin_split :: \nat \ int \ int \ int\ where [simp]: \bin_split n k = (drop_bit n k, take_bit n k)\ lemma [code]: "bin_split (Suc n) w = (let (w1, w2) = bin_split n (w div 2) in (w1, of_bool (odd w) + 2 * w2))" "bin_split 0 w = (w, 0)" by (simp_all add: drop_bit_Suc take_bit_Suc mod_2_eq_odd) -abbreviation (input) bin_cat :: \int \ nat \ int \ int\ - where \bin_cat k n l \ concat_bit n l k\ - lemma bin_cat_eq_push_bit_add_take_bit: - \bin_cat k n l = push_bit n k + take_bit n l\ + \concat_bit n l k = push_bit n k + take_bit n l\ by (simp add: concat_bit_eq) -lemma bin_sign_cat: "bin_sign (bin_cat x n y) = bin_sign x" +lemma bin_sign_cat: "bin_sign ((\k n l. concat_bit n l k) x n y) = bin_sign x" proof - have \0 \ x\ if \0 \ x * 2 ^ n + y mod 2 ^ n\ proof - have \y mod 2 ^ n < 2 ^ n\ using pos_mod_bound [of \2 ^ n\ y] by simp then have \\ y mod 2 ^ n \ 2 ^ n\ by (simp add: less_le) with that have \x \ - 1\ by auto have *: \- 1 \ (- (y mod 2 ^ n)) div 2 ^ n\ by (simp add: zdiv_zminus1_eq_if) from that have \- (y mod 2 ^ n) \ x * 2 ^ n\ by simp then have \(- (y mod 2 ^ n)) div 2 ^ n \ (x * 2 ^ n) div 2 ^ n\ using zdiv_mono1 zero_less_numeral zero_less_power by blast with * have \- 1 \ x * 2 ^ n div 2 ^ n\ by simp with \x \ - 1\ show ?thesis by simp qed then show ?thesis by (simp add: bin_sign_def not_le not_less bin_cat_eq_push_bit_add_take_bit push_bit_eq_mult take_bit_eq_mod) qed -lemma bin_cat_assoc: "bin_cat (bin_cat x m y) n z = bin_cat x (m + n) (bin_cat y n z)" +lemma bin_cat_assoc: "(\k n l. concat_bit n l k) ((\k n l. concat_bit n l k) x m y) n z = (\k n l. concat_bit n l k) x (m + n) ((\k n l. concat_bit n l k) y n z)" by (fact concat_bit_assoc) -lemma bin_cat_assoc_sym: "bin_cat x m (bin_cat y n z) = bin_cat (bin_cat x (m - n) y) (min m n) z" +lemma bin_cat_assoc_sym: "(\k n l. concat_bit n l k) x m ((\k n l. concat_bit n l k) y n z) = (\k n l. concat_bit n l k) ((\k n l. concat_bit n l k) x (m - n) y) (min m n) z" by (fact concat_bit_assoc_sym) definition bin_rcat :: \nat \ int list \ int\ where \bin_rcat n = horner_sum (take_bit n) (2 ^ n) \ rev\ lemma bin_rcat_eq_foldl: - \bin_rcat n = foldl (\u v. bin_cat u n v) 0\ + \bin_rcat n = foldl (\u v. (\k n l. concat_bit n l k) u n v) 0\ proof fix ks :: \int list\ - show \bin_rcat n ks = foldl (\u v. bin_cat u n v) 0 ks\ + show \bin_rcat n ks = foldl (\u v. (\k n l. concat_bit n l k) u n v) 0 ks\ by (induction ks rule: rev_induct) (simp_all add: bin_rcat_def concat_bit_eq push_bit_eq_mult) qed fun bin_rsplit_aux :: "nat \ nat \ int \ int list \ int list" where "bin_rsplit_aux n m c bs = (if m = 0 \ n = 0 then bs else let (a, b) = bin_split n c in bin_rsplit_aux n (m - n) a (b # bs))" definition bin_rsplit :: "nat \ nat \ int \ int list" where "bin_rsplit n w = bin_rsplit_aux n (fst w) (snd w) []" fun bin_rsplitl_aux :: "nat \ nat \ int \ int list \ int list" where "bin_rsplitl_aux n m c bs = (if m = 0 \ n = 0 then bs else let (a, b) = bin_split (min m n) c in bin_rsplitl_aux n (m - n) a (b # bs))" definition bin_rsplitl :: "nat \ nat \ int \ int list" where "bin_rsplitl n w = bin_rsplitl_aux n (fst w) (snd w) []" declare bin_rsplit_aux.simps [simp del] declare bin_rsplitl_aux.simps [simp del] lemma bin_nth_cat: - "bin_nth (bin_cat x k y) n = - (if n < k then bin_nth y n else bin_nth x (n - k))" + "(bit :: int \ nat \ bool) ((\k n l. concat_bit n l k) x k y) n = + (if n < k then (bit :: int \ nat \ bool) y n else (bit :: int \ nat \ bool) x (n - k))" by (simp add: bit_concat_bit_iff) lemma bin_nth_drop_bit_iff: - \bin_nth (drop_bit n c) k \ bin_nth c (n + k)\ + \(bit :: int \ nat \ bool) (drop_bit n c) k \ (bit :: int \ nat \ bool) c (n + k)\ by (simp add: bit_drop_bit_eq) lemma bin_nth_take_bit_iff: - \bin_nth (take_bit n c) k \ k < n \ bin_nth c k\ + \(bit :: int \ nat \ bool) (take_bit n c) k \ k < n \ (bit :: int \ nat \ bool) c k\ by (fact bit_take_bit_iff) lemma bin_nth_split: "bin_split n c = (a, b) \ - (\k. bin_nth a k = bin_nth c (n + k)) \ - (\k. bin_nth b k = (k < n \ bin_nth c k))" + (\k. (bit :: int \ nat \ bool) a k = (bit :: int \ nat \ bool) c (n + k)) \ + (\k. (bit :: int \ nat \ bool) b k = (k < n \ (bit :: int \ nat \ bool) c k))" by (auto simp add: bin_nth_drop_bit_iff bin_nth_take_bit_iff) -lemma bin_cat_zero [simp]: "bin_cat 0 n w = bintrunc n w" +lemma bin_cat_zero [simp]: "(\k n l. concat_bit n l k) 0 n w = (take_bit :: nat \ int \ int) n w" by (simp add: bin_cat_eq_push_bit_add_take_bit) -lemma bintr_cat1: "bintrunc (k + n) (bin_cat a n b) = bin_cat (bintrunc k a) n b" +lemma bintr_cat1: "(take_bit :: nat \ int \ int) (k + n) ((\k n l. concat_bit n l k) a n b) = (\k n l. concat_bit n l k) ((take_bit :: nat \ int \ int) k a) n b" by (metis bin_cat_assoc bin_cat_zero) -lemma bintr_cat: "bintrunc m (bin_cat a n b) = - bin_cat (bintrunc (m - n) a) n (bintrunc (min m n) b)" - +lemma bintr_cat: "(take_bit :: nat \ int \ int) m ((\k n l. concat_bit n l k) a n b) = + (\k n l. concat_bit n l k) ((take_bit :: nat \ int \ int) (m - n) a) n ((take_bit :: nat \ int \ int) (min m n) b)" by (rule bin_eqI) (auto simp: bin_nth_cat nth_bintr) -lemma bintr_cat_same [simp]: "bintrunc n (bin_cat a n b) = bintrunc n b" +lemma bintr_cat_same [simp]: "(take_bit :: nat \ int \ int) n ((\k n l. concat_bit n l k) a n b) = (take_bit :: nat \ int \ int) n b" by (auto simp add : bintr_cat) -lemma cat_bintr [simp]: "bin_cat a n (bintrunc n b) = bin_cat a n b" +lemma cat_bintr [simp]: "(\k n l. concat_bit n l k) a n ((take_bit :: nat \ int \ int) n b) = (\k n l. concat_bit n l k) a n b" by (simp add: bin_cat_eq_push_bit_add_take_bit) -lemma split_bintrunc: "bin_split n c = (a, b) \ b = bintrunc n c" +lemma split_bintrunc: "bin_split n c = (a, b) \ b = (take_bit :: nat \ int \ int) n c" by simp -lemma bin_cat_split: "bin_split n w = (u, v) \ w = bin_cat u n v" +lemma bin_cat_split: "bin_split n w = (u, v) \ w = (\k n l. concat_bit n l k) u n v" by (auto simp add: bin_cat_eq_push_bit_add_take_bit bits_ident) lemma drop_bit_bin_cat_eq: - \drop_bit n (bin_cat v n w) = v\ + \drop_bit n ((\k n l. concat_bit n l k) v n w) = v\ by (rule bit_eqI) (simp add: bit_drop_bit_eq bit_concat_bit_iff) lemma take_bit_bin_cat_eq: - \take_bit n (bin_cat v n w) = take_bit n w\ + \take_bit n ((\k n l. concat_bit n l k) v n w) = take_bit n w\ by (rule bit_eqI) (simp add: bit_concat_bit_iff) -lemma bin_split_cat: "bin_split n (bin_cat v n w) = (v, bintrunc n w)" +lemma bin_split_cat: "bin_split n ((\k n l. concat_bit n l k) v n w) = (v, (take_bit :: nat \ int \ int) n w)" by (simp add: drop_bit_bin_cat_eq take_bit_bin_cat_eq) lemma bin_split_zero [simp]: "bin_split n 0 = (0, 0)" by simp lemma bin_split_minus1 [simp]: - "bin_split n (- 1) = (- 1, bintrunc n (- 1))" + "bin_split n (- 1) = (- 1, (take_bit :: nat \ int \ int) n (- 1))" by simp lemma bin_split_trunc: "bin_split (min m n) c = (a, b) \ - bin_split n (bintrunc m c) = (bintrunc (m - n) a, b)" + bin_split n ((take_bit :: nat \ int \ int) m c) = ((take_bit :: nat \ int \ int) (m - n) a, b)" apply (induct n arbitrary: m b c, clarsimp) apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) apply (case_tac m) apply (auto simp: Let_def drop_bit_Suc take_bit_Suc mod_2_eq_odd split: prod.split_asm) done lemma bin_split_trunc1: "bin_split n c = (a, b) \ - bin_split n (bintrunc m c) = (bintrunc (m - n) a, bintrunc m b)" + bin_split n ((take_bit :: nat \ int \ int) m c) = ((take_bit :: nat \ int \ int) (m - n) a, (take_bit :: nat \ int \ int) m b)" apply (induct n arbitrary: m b c, clarsimp) apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) apply (case_tac m) apply (auto simp: Let_def drop_bit_Suc take_bit_Suc mod_2_eq_odd split: prod.split_asm) done -lemma bin_cat_num: "bin_cat a n b = a * 2 ^ n + bintrunc n b" +lemma bin_cat_num: "(\k n l. concat_bit n l k) a n b = a * 2 ^ n + (take_bit :: nat \ int \ int) n b" by (simp add: bin_cat_eq_push_bit_add_take_bit push_bit_eq_mult) lemma bin_split_num: "bin_split n b = (b div 2 ^ n, b mod 2 ^ n)" by (simp add: drop_bit_eq_div take_bit_eq_mod) lemmas bin_rsplit_aux_simps = bin_rsplit_aux.simps bin_rsplitl_aux.simps lemmas rsplit_aux_simps = bin_rsplit_aux_simps lemmas th_if_simp1 = if_split [where P = "(=) l", THEN iffD1, THEN conjunct1, THEN mp] for l lemmas th_if_simp2 = if_split [where P = "(=) l", THEN iffD1, THEN conjunct2, THEN mp] for l lemmas rsplit_aux_simp1s = rsplit_aux_simps [THEN th_if_simp1] lemmas rsplit_aux_simp2ls = rsplit_aux_simps [THEN th_if_simp2] \ \these safe to \[simp add]\ as require calculating \m - n\\ lemmas bin_rsplit_aux_simp2s [simp] = rsplit_aux_simp2ls [unfolded Let_def] lemmas rbscl = bin_rsplit_aux_simp2s (2) lemmas rsplit_aux_0_simps [simp] = rsplit_aux_simp1s [OF disjI1] rsplit_aux_simp1s [OF disjI2] lemma bin_rsplit_aux_append: "bin_rsplit_aux n m c (bs @ cs) = bin_rsplit_aux n m c bs @ cs" apply (induct n m c bs rule: bin_rsplit_aux.induct) apply (subst bin_rsplit_aux.simps) apply (subst bin_rsplit_aux.simps) apply (clarsimp split: prod.split) done lemma bin_rsplitl_aux_append: "bin_rsplitl_aux n m c (bs @ cs) = bin_rsplitl_aux n m c bs @ cs" apply (induct n m c bs rule: bin_rsplitl_aux.induct) apply (subst bin_rsplitl_aux.simps) apply (subst bin_rsplitl_aux.simps) apply (clarsimp split: prod.split) done lemmas rsplit_aux_apps [where bs = "[]"] = bin_rsplit_aux_append bin_rsplitl_aux_append lemmas rsplit_def_auxs = bin_rsplit_def bin_rsplitl_def lemmas rsplit_aux_alts = rsplit_aux_apps [unfolded append_Nil rsplit_def_auxs [symmetric]] lemma bin_split_minus: "0 < n \ bin_split (Suc (n - 1)) w = bin_split n w" by auto lemma bin_split_pred_simp [simp]: "(0::nat) < numeral bin \ bin_split (numeral bin) w = - (let (w1, w2) = bin_split (numeral bin - 1) (bin_rest w) + (let (w1, w2) = bin_split (numeral bin - 1) ((\k::int. k div 2) w) in (w1, of_bool (odd w) + 2 * w2))" by (simp add: take_bit_rec drop_bit_rec mod_2_eq_odd) lemma bin_rsplit_aux_simp_alt: "bin_rsplit_aux n m c bs = (if m = 0 \ n = 0 then bs else let (a, b) = bin_split n c in bin_rsplit n (m - n, a) @ b # bs)" apply (simp add: bin_rsplit_aux.simps [of n m c bs]) apply (subst rsplit_aux_alts) apply (simp add: bin_rsplit_def) done lemmas bin_rsplit_simp_alt = trans [OF bin_rsplit_def bin_rsplit_aux_simp_alt] lemmas bthrs = bin_rsplit_simp_alt [THEN [2] trans] lemma bin_rsplit_size_sign' [rule_format]: - "n > 0 \ rev sw = bin_rsplit n (nw, w) \ \v\set sw. bintrunc n v = v" + "n > 0 \ rev sw = bin_rsplit n (nw, w) \ \v\set sw. (take_bit :: nat \ int \ int) n v = v" apply (induct sw arbitrary: nw w) apply clarsimp apply clarsimp apply (drule bthrs) apply (simp (no_asm_use) add: Let_def split: prod.split_asm if_split_asm) apply clarify apply simp done lemmas bin_rsplit_size_sign = bin_rsplit_size_sign' [OF asm_rl rev_rev_ident [THEN trans] set_rev [THEN equalityD2 [THEN subsetD]]] lemma bin_nth_rsplit [rule_format] : "n > 0 \ m < n \ \w k nw. rev sw = bin_rsplit n (nw, w) \ - k < size sw \ bin_nth (sw ! k) m = bin_nth w (k * n + m)" + k < size sw \ (bit :: int \ nat \ bool) (sw ! k) m = (bit :: int \ nat \ bool) w (k * n + m)" apply (induct sw) apply clarsimp apply clarsimp apply (drule bthrs) apply (simp (no_asm_use) add: Let_def split: prod.split_asm if_split_asm) apply (erule allE, erule impE, erule exI) apply (case_tac k) apply clarsimp prefer 2 apply clarsimp apply (erule allE) apply (erule (1) impE) apply (simp add: bit_drop_bit_eq ac_simps) apply (simp add: bit_take_bit_iff ac_simps) done -lemma bin_rsplit_all: "0 < nw \ nw \ n \ bin_rsplit n (nw, w) = [bintrunc n w]" +lemma bin_rsplit_all: "0 < nw \ nw \ n \ bin_rsplit n (nw, w) = [(take_bit :: nat \ int \ int) n w]" by (auto simp: bin_rsplit_def rsplit_aux_simp2ls split: prod.split dest!: split_bintrunc) lemma bin_rsplit_l [rule_format]: - "\bin. bin_rsplitl n (m, bin) = bin_rsplit n (m, bintrunc m bin)" + "\bin. bin_rsplitl n (m, bin) = bin_rsplit n (m, (take_bit :: nat \ int \ int) m bin)" apply (rule_tac a = "m" in wf_less_than [THEN wf_induct]) apply (simp (no_asm) add: bin_rsplitl_def bin_rsplit_def) apply (rule allI) apply (subst bin_rsplitl_aux.simps) apply (subst bin_rsplit_aux.simps) apply (clarsimp simp: Let_def split: prod.split) apply (simp add: ac_simps) apply (subst rsplit_aux_alts(1)) apply (subst rsplit_aux_alts(2)) apply clarsimp unfolding bin_rsplit_def bin_rsplitl_def apply (simp add: drop_bit_take_bit) apply (case_tac \x < n\) apply (simp_all add: not_less min_def) done lemma bin_rsplit_rcat [rule_format]: - "n > 0 \ bin_rsplit n (n * size ws, bin_rcat n ws) = map (bintrunc n) ws" + "n > 0 \ bin_rsplit n (n * size ws, bin_rcat n ws) = map ((take_bit :: nat \ int \ int) n) ws" apply (unfold bin_rsplit_def bin_rcat_eq_foldl) apply (rule_tac xs = ws in rev_induct) apply clarsimp apply clarsimp apply (subst rsplit_aux_alts) apply (simp add: drop_bit_bin_cat_eq take_bit_bin_cat_eq) done lemma bin_rsplit_aux_len_le [rule_format] : "\ws m. n \ 0 \ ws = bin_rsplit_aux n nw w bs \ length ws \ m \ nw + length bs * n \ m * n" proof - have *: R if d: "i \ j \ m < j'" and R1: "i * k \ j * k \ R" and R2: "Suc m * k' \ j' * k' \ R" for i j j' k k' m :: nat and R using d apply safe apply (rule R1, erule mult_le_mono1) apply (rule R2, erule Suc_le_eq [THEN iffD2 [THEN mult_le_mono1]]) done have **: "0 < sc \ sc - n + (n + lb * n) \ m * n \ sc + lb * n \ m * n" for sc m n lb :: nat apply safe apply arith apply (case_tac "sc \ n") apply arith apply (insert linorder_le_less_linear [of m lb]) apply (erule_tac k=n and k'=n in *) apply arith apply simp done show ?thesis apply (induct n nw w bs rule: bin_rsplit_aux.induct) apply (subst bin_rsplit_aux.simps) apply (simp add: ** Let_def split: prod.split) done qed lemma bin_rsplit_len_le: "n \ 0 \ ws = bin_rsplit n (nw, w) \ length ws \ m \ nw \ m * n" by (auto simp: bin_rsplit_def bin_rsplit_aux_len_le) lemma bin_rsplit_aux_len: "n \ 0 \ length (bin_rsplit_aux n nw w cs) = (nw + n - 1) div n + length cs" apply (induct n nw w cs rule: bin_rsplit_aux.induct) apply (subst bin_rsplit_aux.simps) apply (clarsimp simp: Let_def split: prod.split) apply (erule thin_rl) apply (case_tac m) apply simp apply (case_tac "m \ n") apply (auto simp add: div_add_self2) done lemma bin_rsplit_len: "n \ 0 \ length (bin_rsplit n (nw, w)) = (nw + n - 1) div n" by (auto simp: bin_rsplit_def bin_rsplit_aux_len) lemma bin_rsplit_aux_len_indep: "n \ 0 \ length bs = length cs \ length (bin_rsplit_aux n nw v bs) = length (bin_rsplit_aux n nw w cs)" proof (induct n nw w cs arbitrary: v bs rule: bin_rsplit_aux.induct) case (1 n m w cs v bs) show ?case proof (cases "m = 0") case True with \length bs = length cs\ show ?thesis by simp next case False from "1.hyps" [of \bin_split n w\ \drop_bit n w\ \take_bit n w\] \m \ 0\ \n \ 0\ have hyp: "\v bs. length bs = Suc (length cs) \ length (bin_rsplit_aux n (m - n) v bs) = length (bin_rsplit_aux n (m - n) (drop_bit n w) (take_bit n w # cs))" using bin_rsplit_aux_len by fastforce from \length bs = length cs\ \n \ 0\ show ?thesis by (auto simp add: bin_rsplit_aux_simp_alt Let_def bin_rsplit_len split: prod.split) qed qed lemma bin_rsplit_len_indep: "n \ 0 \ length (bin_rsplit n (nw, v)) = length (bin_rsplit n (nw, w))" apply (unfold bin_rsplit_def) apply (simp (no_asm)) apply (erule bin_rsplit_aux_len_indep) apply (rule refl) done subsection \Logical operations\ primrec bin_sc :: "nat \ bool \ int \ int" where - Z: "bin_sc 0 b w = of_bool b + 2 * bin_rest w" + Z: "bin_sc 0 b w = of_bool b + 2 * (\k::int. k div 2) w" | Suc: "bin_sc (Suc n) b w = of_bool (odd w) + 2 * bin_sc n b (w div 2)" lemma bin_nth_sc [simp]: "bit (bin_sc n b w) n \ b" by (induction n arbitrary: w) (simp_all add: bit_Suc) lemma bin_sc_sc_same [simp]: "bin_sc n c (bin_sc n b w) = bin_sc n c w" by (induction n arbitrary: w) (simp_all add: bit_Suc) lemma bin_sc_sc_diff: "m \ n \ bin_sc m c (bin_sc n b w) = bin_sc n b (bin_sc m c w)" apply (induct n arbitrary: w m) apply (case_tac [!] m) apply auto done -lemma bin_nth_sc_gen: "bin_nth (bin_sc n b w) m = (if m = n then b else bin_nth w m)" +lemma bin_nth_sc_gen: "(bit :: int \ nat \ bool) (bin_sc n b w) m = (if m = n then b else (bit :: int \ nat \ bool) w m)" apply (induct n arbitrary: w m) apply (case_tac m; simp add: bit_Suc) apply (case_tac m; simp add: bit_Suc) done lemma bin_sc_eq: \bin_sc n False = unset_bit n\ \bin_sc n True = Bit_Operations.set_bit n\ by (simp_all add: fun_eq_iff bit_eq_iff) (simp_all add: bin_nth_sc_gen bit_set_bit_iff bit_unset_bit_iff) -lemma bin_sc_nth [simp]: "bin_sc n (bin_nth w n) w = w" +lemma bin_sc_nth [simp]: "bin_sc n ((bit :: int \ nat \ bool) w n) w = w" by (rule bit_eqI) (simp add: bin_nth_sc_gen) lemma bin_sign_sc [simp]: "bin_sign (bin_sc n b w) = bin_sign w" proof (induction n arbitrary: w) case 0 then show ?case by (auto simp add: bin_sign_def) (use bin_rest_ge_0 in fastforce) next case (Suc n) from Suc [of \w div 2\] show ?case by (auto simp add: bin_sign_def split: if_splits) qed lemma bin_sc_bintr [simp]: - "bintrunc m (bin_sc n x (bintrunc m w)) = bintrunc m (bin_sc n x w)" + "(take_bit :: nat \ int \ int) m (bin_sc n x ((take_bit :: nat \ int \ int) m w)) = (take_bit :: nat \ int \ int) m (bin_sc n x w)" apply (cases x) apply (simp_all add: bin_sc_eq bit_eq_iff) apply (auto simp add: bit_take_bit_iff bit_set_bit_iff bit_unset_bit_iff) done lemma bin_clr_le: "bin_sc n False w \ w" by (simp add: bin_sc_eq unset_bit_less_eq) lemma bin_set_ge: "bin_sc n True w \ w" by (simp add: bin_sc_eq set_bit_greater_eq) -lemma bintr_bin_clr_le: "bintrunc n (bin_sc m False w) \ bintrunc n w" +lemma bintr_bin_clr_le: "(take_bit :: nat \ int \ int) n (bin_sc m False w) \ (take_bit :: nat \ int \ int) n w" by (simp add: bin_sc_eq take_bit_unset_bit_eq unset_bit_less_eq) -lemma bintr_bin_set_ge: "bintrunc n (bin_sc m True w) \ bintrunc n w" +lemma bintr_bin_set_ge: "(take_bit :: nat \ int \ int) n (bin_sc m True w) \ (take_bit :: nat \ int \ int) n w" by (simp add: bin_sc_eq take_bit_set_bit_eq set_bit_greater_eq) lemma bin_sc_FP [simp]: "bin_sc n False 0 = 0" by (induct n) auto lemma bin_sc_TM [simp]: "bin_sc n True (- 1) = - 1" by (induct n) auto lemmas bin_sc_simps = bin_sc.Z bin_sc.Suc bin_sc_TM bin_sc_FP lemma bin_sc_minus: "0 < n \ bin_sc (Suc (n - 1)) b w = bin_sc n b w" by auto lemmas bin_sc_Suc_minus = trans [OF bin_sc_minus [symmetric] bin_sc.Suc] lemma bin_sc_numeral [simp]: "bin_sc (numeral k) b w = of_bool (odd w) + 2 * bin_sc (pred_numeral k) b (w div 2)" by (simp add: numeral_eq_Suc) lemmas bin_sc_minus_simps = bin_sc_simps (2,3,4) [THEN [2] trans, OF bin_sc_minus [THEN sym]] instance int :: semiring_bit_syntax .. lemma test_bit_int_def [iff]: - "i !! n \ bin_nth i n" + "i !! n \ (bit :: int \ nat \ bool) i n" by (simp add: test_bit_eq_bit) lemma shiftl_int_def: "shiftl x n = x * 2 ^ n" for x :: int by (simp add: push_bit_int_def shiftl_eq_push_bit) lemma shiftr_int_def: "shiftr x n = x div 2 ^ n" for x :: int by (simp add: drop_bit_int_def shiftr_eq_drop_bit) subsubsection \Basic simplification rules\ lemmas int_not_def = not_int_def lemma int_not_simps [simp]: "NOT (0::int) = -1" "NOT (1::int) = -2" "NOT (- 1::int) = 0" "NOT (numeral w::int) = - numeral (w + Num.One)" "NOT (- numeral (Num.Bit0 w)::int) = numeral (Num.BitM w)" "NOT (- numeral (Num.Bit1 w)::int) = numeral (Num.Bit0 w)" by (simp_all add: not_int_def) lemma int_not_not: "NOT (NOT x) = x" for x :: int by (fact bit.double_compl) lemma int_and_0 [simp]: "0 AND x = 0" for x :: int by (fact bit.conj_zero_left) lemma int_and_m1 [simp]: "-1 AND x = x" for x :: int by (fact bit.conj_one_left) lemma int_or_zero [simp]: "0 OR x = x" for x :: int by (fact bit.disj_zero_left) lemma int_or_minus1 [simp]: "-1 OR x = -1" for x :: int by (fact bit.disj_one_left) lemma int_xor_zero [simp]: "0 XOR x = x" for x :: int by (fact bit.xor_zero_left) subsubsection \Binary destructors\ -lemma bin_rest_NOT [simp]: "bin_rest (NOT x) = NOT (bin_rest x)" +lemma bin_rest_NOT [simp]: "(\k::int. k div 2) (NOT x) = NOT ((\k::int. k div 2) x)" by (fact not_int_div_2) -lemma bin_last_NOT [simp]: "bin_last (NOT x) \ \ bin_last x" +lemma bin_last_NOT [simp]: "(odd :: int \ bool) (NOT x) \ \ (odd :: int \ bool) x" by simp -lemma bin_rest_AND [simp]: "bin_rest (x AND y) = bin_rest x AND bin_rest y" +lemma bin_rest_AND [simp]: "(\k::int. k div 2) (x AND y) = (\k::int. k div 2) x AND (\k::int. k div 2) y" by (subst and_int_rec) auto -lemma bin_last_AND [simp]: "bin_last (x AND y) \ bin_last x \ bin_last y" +lemma bin_last_AND [simp]: "(odd :: int \ bool) (x AND y) \ (odd :: int \ bool) x \ (odd :: int \ bool) y" by (subst and_int_rec) auto -lemma bin_rest_OR [simp]: "bin_rest (x OR y) = bin_rest x OR bin_rest y" +lemma bin_rest_OR [simp]: "(\k::int. k div 2) (x OR y) = (\k::int. k div 2) x OR (\k::int. k div 2) y" by (subst or_int_rec) auto -lemma bin_last_OR [simp]: "bin_last (x OR y) \ bin_last x \ bin_last y" +lemma bin_last_OR [simp]: "(odd :: int \ bool) (x OR y) \ (odd :: int \ bool) x \ (odd :: int \ bool) y" by (subst or_int_rec) auto -lemma bin_rest_XOR [simp]: "bin_rest (x XOR y) = bin_rest x XOR bin_rest y" +lemma bin_rest_XOR [simp]: "(\k::int. k div 2) (x XOR y) = (\k::int. k div 2) x XOR (\k::int. k div 2) y" by (subst xor_int_rec) auto -lemma bin_last_XOR [simp]: "bin_last (x XOR y) \ (bin_last x \ bin_last y) \ \ (bin_last x \ bin_last y)" +lemma bin_last_XOR [simp]: "(odd :: int \ bool) (x XOR y) \ ((odd :: int \ bool) x \ (odd :: int \ bool) y) \ \ ((odd :: int \ bool) x \ (odd :: int \ bool) y)" by (subst xor_int_rec) auto lemma bin_nth_ops: - "\x y. bin_nth (x AND y) n \ bin_nth x n \ bin_nth y n" - "\x y. bin_nth (x OR y) n \ bin_nth x n \ bin_nth y n" - "\x y. bin_nth (x XOR y) n \ bin_nth x n \ bin_nth y n" - "\x. bin_nth (NOT x) n \ \ bin_nth x n" + "\x y. (bit :: int \ nat \ bool) (x AND y) n \ (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" + "\x y. (bit :: int \ nat \ bool) (x OR y) n \ (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" + "\x y. (bit :: int \ nat \ bool) (x XOR y) n \ (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" + "\x. (bit :: int \ nat \ bool) (NOT x) n \ \ (bit :: int \ nat \ bool) x n" by (simp_all add: bit_and_iff bit_or_iff bit_xor_iff bit_not_iff) subsubsection \Derived properties\ lemma int_xor_minus1 [simp]: "-1 XOR x = NOT x" for x :: int by (fact bit.xor_one_left) lemma int_xor_extra_simps [simp]: "w XOR 0 = w" "w XOR -1 = NOT w" for w :: int by simp_all lemma int_or_extra_simps [simp]: "w OR 0 = w" "w OR -1 = -1" for w :: int by simp_all lemma int_and_extra_simps [simp]: "w AND 0 = 0" "w AND -1 = w" for w :: int by simp_all text \Commutativity of the above.\ lemma bin_ops_comm: fixes x y :: int shows int_and_comm: "x AND y = y AND x" and int_or_comm: "x OR y = y OR x" and int_xor_comm: "x XOR y = y XOR x" by (simp_all add: ac_simps) lemma bin_ops_same [simp]: "x AND x = x" "x OR x = x" "x XOR x = 0" for x :: int by simp_all lemmas bin_log_esimps = int_and_extra_simps int_or_extra_simps int_xor_extra_simps int_and_0 int_and_m1 int_or_zero int_or_minus1 int_xor_zero int_xor_minus1 subsubsection \Basic properties of logical (bit-wise) operations\ lemma bbw_ao_absorb: "x AND (y OR x) = x \ x OR (y AND x) = x" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_ao_absorbs_other: "x AND (x OR y) = x \ (y AND x) OR x = x" "(y OR x) AND x = x \ x OR (x AND y) = x" "(x OR y) AND x = x \ (x AND y) OR x = x" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemmas bbw_ao_absorbs [simp] = bbw_ao_absorb bbw_ao_absorbs_other lemma int_xor_not: "(NOT x) XOR y = NOT (x XOR y) \ x XOR (NOT y) = NOT (x XOR y)" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_and_assoc: "(x AND y) AND z = x AND (y AND z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_or_assoc: "(x OR y) OR z = x OR (y OR z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_xor_assoc: "(x XOR y) XOR z = x XOR (y XOR z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemmas bbw_assocs = int_and_assoc int_or_assoc int_xor_assoc (* BH: Why are these declared as simp rules??? *) lemma bbw_lcs [simp]: "y AND (x AND z) = x AND (y AND z)" "y OR (x OR z) = x OR (y OR z)" "y XOR (x XOR z) = x XOR (y XOR z)" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_not_dist: "NOT (x OR y) = (NOT x) AND (NOT y)" "NOT (x AND y) = (NOT x) OR (NOT y)" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_oa_dist: "(x AND y) OR z = (x OR z) AND (y OR z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_ao_dist: "(x OR y) AND z = (x AND z) OR (y AND z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) (* Why were these declared simp??? declare bin_ops_comm [simp] bbw_assocs [simp] *) subsubsection \Simplification with numerals\ text \Cases for \0\ and \-1\ are already covered by other simp rules.\ lemma bin_rest_neg_numeral_BitM [simp]: - "bin_rest (- numeral (Num.BitM w)) = - numeral w" + "(\k::int. k div 2) (- numeral (Num.BitM w)) = - numeral w" by simp lemma bin_last_neg_numeral_BitM [simp]: - "bin_last (- numeral (Num.BitM w))" + "(odd :: int \ bool) (- numeral (Num.BitM w))" by simp subsubsection \Interactions with arithmetic\ lemma le_int_or: "bin_sign y = 0 \ x \ x OR y" for x y :: int by (simp add: bin_sign_def or_greater_eq split: if_splits) lemmas int_and_le = xtrans(3) [OF bbw_ao_absorbs (2) [THEN conjunct2, symmetric] le_int_or] text \Interaction between bit-wise and arithmetic: good example of \bin_induction\.\ lemma bin_add_not: "x + NOT x = (-1::int)" by (simp add: not_int_def) lemma AND_mod: "x AND (2 ^ n - 1) = x mod 2 ^ n" for x :: int by (simp flip: take_bit_eq_mod add: take_bit_eq_mask mask_eq_exp_minus_1) subsubsection \Truncating results of bit-wise operations\ lemma bin_trunc_ao: - "bintrunc n x AND bintrunc n y = bintrunc n (x AND y)" - "bintrunc n x OR bintrunc n y = bintrunc n (x OR y)" + "(take_bit :: nat \ int \ int) n x AND (take_bit :: nat \ int \ int) n y = (take_bit :: nat \ int \ int) n (x AND y)" + "(take_bit :: nat \ int \ int) n x OR (take_bit :: nat \ int \ int) n y = (take_bit :: nat \ int \ int) n (x OR y)" by simp_all -lemma bin_trunc_xor: "bintrunc n (bintrunc n x XOR bintrunc n y) = bintrunc n (x XOR y)" +lemma bin_trunc_xor: "(take_bit :: nat \ int \ int) n ((take_bit :: nat \ int \ int) n x XOR (take_bit :: nat \ int \ int) n y) = (take_bit :: nat \ int \ int) n (x XOR y)" by simp -lemma bin_trunc_not: "bintrunc n (NOT (bintrunc n x)) = bintrunc n (NOT x)" +lemma bin_trunc_not: "(take_bit :: nat \ int \ int) n (NOT ((take_bit :: nat \ int \ int) n x)) = (take_bit :: nat \ int \ int) n (NOT x)" by (fact take_bit_not_take_bit) text \Want theorems of the form of \bin_trunc_xor\.\ -lemma bintr_bintr_i: "x = bintrunc n y \ bintrunc n x = bintrunc n y" +lemma bintr_bintr_i: "x = (take_bit :: nat \ int \ int) n y \ (take_bit :: nat \ int \ int) n x = (take_bit :: nat \ int \ int) n y" by auto lemmas bin_trunc_and = bin_trunc_ao(1) [THEN bintr_bintr_i] lemmas bin_trunc_or = bin_trunc_ao(2) [THEN bintr_bintr_i] subsubsection \More lemmas\ lemma not_int_cmp_0 [simp]: fixes i :: int shows "0 < NOT i \ i < -1" "0 \ NOT i \ i < 0" "NOT i < 0 \ i \ 0" "NOT i \ 0 \ i \ -1" by(simp_all add: int_not_def) arith+ lemma bbw_ao_dist2: "(x :: int) AND (y OR z) = x AND y OR x AND z" by (fact bit.conj_disj_distrib) lemmas int_and_ac = bbw_lcs(1) int_and_comm int_and_assoc lemma int_nand_same [simp]: fixes x :: int shows "x AND NOT x = 0" by simp lemma int_nand_same_middle: fixes x :: int shows "x AND y AND NOT x = 0" by (simp add: bit_eq_iff bit_and_iff bit_not_iff) lemma and_xor_dist: fixes x :: int shows "x AND (y XOR z) = (x AND y) XOR (x AND z)" by (fact bit.conj_xor_distrib) lemma int_and_lt0 [simp]: \x AND y < 0 \ x < 0 \ y < 0\ for x y :: int by (fact and_negative_int_iff) lemma int_and_ge0 [simp]: \x AND y \ 0 \ x \ 0 \ y \ 0\ for x y :: int by (fact and_nonnegative_int_iff) lemma int_and_1: fixes x :: int shows "x AND 1 = x mod 2" by (fact and_one_eq) lemma int_1_and: fixes x :: int shows "1 AND x = x mod 2" by (fact one_and_eq) lemma int_or_lt0 [simp]: \x OR y < 0 \ x < 0 \ y < 0\ for x y :: int by (fact or_negative_int_iff) lemma int_or_ge0 [simp]: \x OR y \ 0 \ x \ 0 \ y \ 0\ for x y :: int by (fact or_nonnegative_int_iff) lemma int_xor_lt0 [simp]: \x XOR y < 0 \ (x < 0) \ (y < 0)\ for x y :: int by (fact xor_negative_int_iff) lemma int_xor_ge0 [simp]: \x XOR y \ 0 \ (x \ 0 \ y \ 0)\ for x y :: int by (fact xor_nonnegative_int_iff) lemma even_conv_AND: \even i \ i AND 1 = 0\ for i :: int by (simp add: and_one_eq mod2_eq_if) lemma bin_last_conv_AND: - "bin_last i \ i AND 1 \ 0" + "(odd :: int \ bool) i \ i AND 1 \ 0" by (simp add: and_one_eq mod2_eq_if) lemma bitval_bin_last: - "of_bool (bin_last i) = i AND 1" + "of_bool ((odd :: int \ bool) i) = i AND 1" by (simp add: and_one_eq mod2_eq_if) lemma bin_sign_and: "bin_sign (i AND j) = - (bin_sign i * bin_sign j)" by(simp add: bin_sign_def) lemma int_not_neg_numeral: "NOT (- numeral n) = (Num.sub n num.One :: int)" by(simp add: int_not_def) lemma int_neg_numeral_pOne_conv_not: "- numeral (n + num.One) = (NOT (numeral n) :: int)" by(simp add: int_not_def) subsection \Setting and clearing bits\ lemma int_shiftl_BIT: fixes x :: int shows int_shiftl0 [simp]: "x << 0 = x" and int_shiftl_Suc [simp]: "x << Suc n = 2 * (x << n)" by (auto simp add: shiftl_int_def) lemma int_0_shiftl [simp]: "0 << n = (0 :: int)" by(induct n) simp_all -lemma bin_last_shiftl: "bin_last (x << n) \ n = 0 \ bin_last x" +lemma bin_last_shiftl: "(odd :: int \ bool) (x << n) \ n = 0 \ (odd :: int \ bool) x" by(cases n)(simp_all) -lemma bin_rest_shiftl: "bin_rest (x << n) = (if n > 0 then x << (n - 1) else bin_rest x)" +lemma bin_rest_shiftl: "(\k::int. k div 2) (x << n) = (if n > 0 then x << (n - 1) else (\k::int. k div 2) x)" by(cases n)(simp_all) -lemma bin_nth_shiftl [simp]: "bin_nth (x << n) m \ n \ m \ bin_nth x (m - n)" +lemma bin_nth_shiftl [simp]: "(bit :: int \ nat \ bool) (x << n) m \ n \ m \ (bit :: int \ nat \ bool) x (m - n)" by (simp add: bit_push_bit_iff_int shiftl_eq_push_bit) lemma bin_last_shiftr: "odd (x >> n) \ x !! n" for x :: int by (simp add: shiftr_eq_drop_bit bit_iff_odd_drop_bit) -lemma bin_rest_shiftr [simp]: "bin_rest (x >> n) = x >> Suc n" +lemma bin_rest_shiftr [simp]: "(\k::int. k div 2) (x >> n) = x >> Suc n" by (simp add: bit_eq_iff shiftr_eq_drop_bit drop_bit_Suc bit_drop_bit_eq drop_bit_half) -lemma bin_nth_shiftr [simp]: "bin_nth (x >> n) m = bin_nth x (n + m)" +lemma bin_nth_shiftr [simp]: "(bit :: int \ nat \ bool) (x >> n) m = (bit :: int \ nat \ bool) x (n + m)" by (simp add: shiftr_eq_drop_bit bit_drop_bit_eq) lemma bin_nth_conv_AND: fixes x :: int shows - "bin_nth x n \ x AND (1 << n) \ 0" + "(bit :: int \ nat \ bool) x n \ x AND (1 << n) \ 0" by (simp add: bit_eq_iff) (auto simp add: shiftl_eq_push_bit bit_and_iff bit_push_bit_iff bit_exp_iff) lemma int_shiftl_numeral [simp]: "(numeral w :: int) << numeral w' = numeral (num.Bit0 w) << pred_numeral w'" "(- numeral w :: int) << numeral w' = - numeral (num.Bit0 w) << pred_numeral w'" by(simp_all add: numeral_eq_Suc shiftl_int_def) (metis add_One mult_inc semiring_norm(11) semiring_norm(13) semiring_norm(2) semiring_norm(6) semiring_norm(87))+ lemma int_shiftl_One_numeral [simp]: "(1 :: int) << numeral w = 2 << pred_numeral w" using int_shiftl_numeral [of Num.One w] by simp lemma shiftl_ge_0 [simp]: fixes i :: int shows "i << n \ 0 \ i \ 0" by(induct n) simp_all lemma shiftl_lt_0 [simp]: fixes i :: int shows "i << n < 0 \ i < 0" by (metis not_le shiftl_ge_0) lemma int_shiftl_test_bit: "(n << i :: int) !! m \ m \ i \ n !! (m - i)" by simp lemma int_0shiftr [simp]: "(0 :: int) >> x = 0" by(simp add: shiftr_int_def) lemma int_minus1_shiftr [simp]: "(-1 :: int) >> x = -1" by(simp add: shiftr_int_def div_eq_minus1) lemma int_shiftr_ge_0 [simp]: fixes i :: int shows "i >> n \ 0 \ i \ 0" by (simp add: shiftr_eq_drop_bit) lemma int_shiftr_lt_0 [simp]: fixes i :: int shows "i >> n < 0 \ i < 0" by (metis int_shiftr_ge_0 not_less) lemma int_shiftr_numeral [simp]: "(1 :: int) >> numeral w' = 0" "(numeral num.One :: int) >> numeral w' = 0" "(numeral (num.Bit0 w) :: int) >> numeral w' = numeral w >> pred_numeral w'" "(numeral (num.Bit1 w) :: int) >> numeral w' = numeral w >> pred_numeral w'" "(- numeral (num.Bit0 w) :: int) >> numeral w' = - numeral w >> pred_numeral w'" "(- numeral (num.Bit1 w) :: int) >> numeral w' = - numeral (Num.inc w) >> pred_numeral w'" by (simp_all add: shiftr_eq_drop_bit numeral_eq_Suc add_One drop_bit_Suc) lemma int_shiftr_numeral_Suc0 [simp]: "(1 :: int) >> Suc 0 = 0" "(numeral num.One :: int) >> Suc 0 = 0" "(numeral (num.Bit0 w) :: int) >> Suc 0 = numeral w" "(numeral (num.Bit1 w) :: int) >> Suc 0 = numeral w" "(- numeral (num.Bit0 w) :: int) >> Suc 0 = - numeral w" "(- numeral (num.Bit1 w) :: int) >> Suc 0 = - numeral (Num.inc w)" by (simp_all add: shiftr_eq_drop_bit drop_bit_Suc add_One) lemma bin_nth_minus_p2: assumes sign: "bin_sign x = 0" and y: "y = 1 << n" and m: "m < n" and x: "x < y" - shows "bin_nth (x - y) m = bin_nth x m" + shows "(bit :: int \ nat \ bool) (x - y) m = (bit :: int \ nat \ bool) x m" proof - from sign y x have \x \ 0\ and \y = 2 ^ n\ and \x < 2 ^ n\ by (simp_all add: bin_sign_def shiftl_eq_push_bit push_bit_eq_mult split: if_splits) from \0 \ x\ \x < 2 ^ n\ \m < n\ have \bit x m \ bit (x - 2 ^ n) m\ proof (induction m arbitrary: x n) case 0 then show ?case by simp next case (Suc m) moreover define q where \q = n - 1\ ultimately have n: \n = Suc q\ by simp have \(x - 2 ^ Suc q) div 2 = x div 2 - 2 ^ q\ by simp moreover from Suc.IH [of \x div 2\ q] Suc.prems have \bit (x div 2) m \ bit (x div 2 - 2 ^ q) m\ by (simp add: n) ultimately show ?case by (simp add: bit_Suc n) qed with \y = 2 ^ n\ show ?thesis by simp qed lemma bin_clr_conv_NAND: "bin_sc n False i = i AND NOT (1 << n)" by (induct n arbitrary: i) (rule bin_rl_eqI; simp)+ lemma bin_set_conv_OR: "bin_sc n True i = i OR (1 << n)" by (induct n arbitrary: i) (rule bin_rl_eqI; simp)+ subsection \More lemmas on words\ lemma word_rcat_eq: \word_rcat ws = word_of_int (bin_rcat (LENGTH('a::len)) (map uint ws))\ for ws :: \'a::len word list\ apply (simp add: word_rcat_def bin_rcat_def rev_map) apply transfer apply (simp add: horner_sum_foldr foldr_map comp_def) done lemma sign_uint_Pls [simp]: "bin_sign (uint x) = 0" by (simp add: sign_Pls_ge_0) lemmas bin_log_bintrs = bin_trunc_not bin_trunc_xor bin_trunc_and bin_trunc_or \ \following definitions require both arithmetic and bit-wise word operations\ \ \to get \word_no_log_defs\ from \word_log_defs\, using \bin_log_bintrs\\ lemmas wils1 = bin_log_bintrs [THEN word_of_int_eq_iff [THEN iffD2], folded uint_word_of_int_eq, THEN eq_reflection] \ \the binary operations only\ (* BH: why is this needed? *) lemmas word_log_binary_defs = word_and_def word_or_def word_xor_def lemma setBit_no [simp]: "setBit (numeral bin) n = word_of_int (bin_sc n True (numeral bin))" by transfer (simp add: bin_sc_eq) lemma clearBit_no [simp]: "clearBit (numeral bin) n = word_of_int (bin_sc n False (numeral bin))" by transfer (simp add: bin_sc_eq) lemma eq_mod_iff: "0 < n \ b = b mod n \ 0 \ b \ b < n" for b n :: int by auto (metis pos_mod_conj)+ lemma split_uint_lem: "bin_split n (uint w) = (a, b) \ a = take_bit (LENGTH('a) - n) a \ b = take_bit (LENGTH('a)) b" for w :: "'a::len word" by transfer (simp add: drop_bit_take_bit ac_simps) \ \limited hom result\ lemma word_cat_hom: "LENGTH('a::len) \ LENGTH('b::len) + LENGTH('c::len) \ (word_cat (word_of_int w :: 'b word) (b :: 'c word) :: 'a word) = - word_of_int (bin_cat w (size b) (uint b))" + word_of_int ((\k n l. concat_bit n l k) w (size b) (uint b))" by transfer (simp add: take_bit_concat_bit_eq) lemma bintrunc_shiftl: "take_bit n (m << i) = take_bit (n - i) m << i" for m :: int by (rule bit_eqI) (auto simp add: bit_take_bit_iff) lemma uint_shiftl: "uint (n << i) = take_bit (size n) (uint n << i)" by transfer (simp add: push_bit_take_bit shiftl_eq_push_bit) lemma bin_mask_conv_pow2: "mask n = 2 ^ n - (1 :: int)" by (fact mask_eq_exp_minus_1) lemma bin_mask_ge0: "mask n \ (0 :: int)" by (fact mask_nonnegative_int) lemma and_bin_mask_conv_mod: "x AND mask n = x mod 2 ^ n" for x :: int by (simp flip: take_bit_eq_mod add: take_bit_eq_mask) lemma bin_mask_numeral: "mask (numeral n) = (1 :: int) + 2 * mask (pred_numeral n)" by (fact mask_numeral) lemma bin_nth_mask [simp]: "bit (mask n :: int) i \ i < n" by (simp add: bit_mask_iff) lemma bin_sign_mask [simp]: "bin_sign (mask n) = 0" by (simp add: bin_sign_def bin_mask_conv_pow2) lemma bin_mask_p1_conv_shift: "mask n + 1 = (1 :: int) << n" by (simp add: bin_mask_conv_pow2 shiftl_int_def) lemma sbintrunc_eq_in_range: - "(sbintrunc n x = x) = (x \ range (sbintrunc n))" - "(x = sbintrunc n x) = (x \ range (sbintrunc n))" + "((signed_take_bit :: nat \ int \ int) n x = x) = (x \ range ((signed_take_bit :: nat \ int \ int) n))" + "(x = (signed_take_bit :: nat \ int \ int) n x) = (x \ range ((signed_take_bit :: nat \ int \ int) n))" apply (simp_all add: image_def) apply (metis sbintrunc_sbintrunc)+ done lemma sbintrunc_If: "- 3 * (2 ^ n) \ x \ x < 3 * (2 ^ n) - \ sbintrunc n x = (if x < - (2 ^ n) then x + 2 * (2 ^ n) + \ (signed_take_bit :: nat \ int \ int) n x = (if x < - (2 ^ n) then x + 2 * (2 ^ n) else if x \ 2 ^ n then x - 2 * (2 ^ n) else x)" apply (simp add: no_sbintr_alt2, safe) apply (simp add: mod_pos_geq) apply (subst mod_add_self1[symmetric], simp) done lemma sint_range': \- (2 ^ (LENGTH('a) - Suc 0)) \ sint x \ sint x < 2 ^ (LENGTH('a) - Suc 0)\ for x :: \'a::len word\ apply transfer using sbintr_ge sbintr_lt apply auto done lemma signed_arith_eq_checks_to_ord: "(sint a + sint b = sint (a + b )) = ((a <=s a + b) = (0 <=s b))" "(sint a - sint b = sint (a - b )) = ((0 <=s a - b) = (b <=s a))" "(- sint a = sint (- a)) = (0 <=s (- a) = (a <=s 0))" using sint_range'[where x=a] sint_range'[where x=b] by (simp_all add: sint_word_ariths word_sle_eq word_sless_alt sbintrunc_If) lemma signed_mult_eq_checks_double_size: assumes mult_le: "(2 ^ (len_of TYPE ('a) - 1) + 1) ^ 2 \ (2 :: int) ^ (len_of TYPE ('b) - 1)" and le: "2 ^ (LENGTH('a) - 1) \ (2 :: int) ^ (len_of TYPE ('b) - 1)" shows "(sint (a :: 'a :: len word) * sint b = sint (a * b)) = (scast a * scast b = (scast (a * b) :: 'b :: len word))" proof - - have P: "sbintrunc (size a - 1) (sint a * sint b) \ range (sbintrunc (size a - 1))" + have P: "(signed_take_bit :: nat \ int \ int) (size a - 1) (sint a * sint b) \ range ((signed_take_bit :: nat \ int \ int) (size a - 1))" by simp have abs: "!! x :: 'a word. abs (sint x) < 2 ^ (size a - 1) + 1" apply (cut_tac x=x in sint_range') apply (simp add: abs_le_iff word_size) done have abs_ab: "abs (sint a * sint b) < 2 ^ (LENGTH('b) - 1)" using abs_mult_less[OF abs[where x=a] abs[where x=b]] mult_le by (simp add: abs_mult power2_eq_square word_size) define r s where \r = LENGTH('a) - 1\ \s = LENGTH('b) - 1\ then have \LENGTH('a) = Suc r\ \LENGTH('b) = Suc s\ \size a = Suc r\ \size b = Suc r\ by (simp_all add: word_size) then show ?thesis using P[unfolded range_sbintrunc] abs_ab le apply clarsimp apply (transfer fixing: r s) apply (auto simp add: signed_take_bit_int_eq_self simp flip: signed_take_bit_eq_iff_take_bit_eq) done qed code_identifier code_module Bits_Int \ (SML) Bit_Operations and (OCaml) Bit_Operations and (Haskell) Bit_Operations and (Scala) Bit_Operations end diff --git a/thys/Word_Lib/Bitwise.thy b/thys/Word_Lib/Bitwise.thy --- a/thys/Word_Lib/Bitwise.thy +++ b/thys/Word_Lib/Bitwise.thy @@ -1,506 +1,506 @@ (* * Copyright Thomas Sewell, NICTA and Sascha Boehme, TU Muenchen * * SPDX-License-Identifier: BSD-2-Clause *) theory Bitwise imports "HOL-Library.Word" More_Arithmetic Reversed_Bit_Lists begin text \Helper constants used in defining addition\ definition xor3 :: "bool \ bool \ bool \ bool" where "xor3 a b c = (a = (b = c))" definition carry :: "bool \ bool \ bool \ bool" where "carry a b c = ((a \ (b \ c)) \ (b \ c))" lemma carry_simps: "carry True a b = (a \ b)" "carry a True b = (a \ b)" "carry a b True = (a \ b)" "carry False a b = (a \ b)" "carry a False b = (a \ b)" "carry a b False = (a \ b)" by (auto simp add: carry_def) lemma xor3_simps: "xor3 True a b = (a = b)" "xor3 a True b = (a = b)" "xor3 a b True = (a = b)" "xor3 False a b = (a \ b)" "xor3 a False b = (a \ b)" "xor3 a b False = (a \ b)" by (simp_all add: xor3_def) text \Breaking up word equalities into equalities on their bit lists. Equalities are generated and manipulated in the reverse order to \<^const>\to_bl\.\ lemma bl_word_sub: "to_bl (x - y) = to_bl (x + (- y))" by simp lemma rbl_word_1: "rev (to_bl (1 :: 'a::len word)) = takefill False (LENGTH('a)) [True]" apply (rule_tac s="rev (to_bl (word_succ (0 :: 'a word)))" in trans) apply simp apply (simp only: rtb_rbl_ariths(1)[OF refl]) apply simp apply (case_tac "LENGTH('a)") apply simp apply (simp add: takefill_alt) done lemma rbl_word_if: "rev (to_bl (if P then x else y)) = map2 (If P) (rev (to_bl x)) (rev (to_bl y))" by (simp add: split_def) lemma rbl_add_carry_Cons: "(if car then rbl_succ else id) (rbl_add (x # xs) (y # ys)) = xor3 x y car # (if carry x y car then rbl_succ else id) (rbl_add xs ys)" by (simp add: carry_def xor3_def) lemma rbl_add_suc_carry_fold: "length xs = length ys \ \car. (if car then rbl_succ else id) (rbl_add xs ys) = (foldr (\(x, y) res car. xor3 x y car # res (carry x y car)) (zip xs ys) (\_. [])) car" apply (erule list_induct2) apply simp apply (simp only: rbl_add_carry_Cons) apply simp done lemma to_bl_plus_carry: "to_bl (x + y) = rev (foldr (\(x, y) res car. xor3 x y car # res (carry x y car)) (rev (zip (to_bl x) (to_bl y))) (\_. []) False)" using rbl_add_suc_carry_fold[where xs="rev (to_bl x)" and ys="rev (to_bl y)"] apply (simp add: word_add_rbl[OF refl refl]) apply (drule_tac x=False in spec) apply (simp add: zip_rev) done definition "rbl_plus cin xs ys = foldr (\(x, y) res car. xor3 x y car # res (carry x y car)) (zip xs ys) (\_. []) cin" lemma rbl_plus_simps: "rbl_plus cin (x # xs) (y # ys) = xor3 x y cin # rbl_plus (carry x y cin) xs ys" "rbl_plus cin [] ys = []" "rbl_plus cin xs [] = []" by (simp_all add: rbl_plus_def) lemma rbl_word_plus: "rev (to_bl (x + y)) = rbl_plus False (rev (to_bl x)) (rev (to_bl y))" by (simp add: rbl_plus_def to_bl_plus_carry zip_rev) definition "rbl_succ2 b xs = (if b then rbl_succ xs else xs)" lemma rbl_succ2_simps: "rbl_succ2 b [] = []" "rbl_succ2 b (x # xs) = (b \ x) # rbl_succ2 (x \ b) xs" by (simp_all add: rbl_succ2_def) lemma twos_complement: "- x = word_succ (NOT x)" using arg_cong[OF word_add_not[where x=x], where f="\a. a - x + 1"] by (simp add: word_succ_p1 word_sp_01[unfolded word_succ_p1] del: word_add_not) lemma rbl_word_neg: "rev (to_bl (- x)) = rbl_succ2 True (map Not (rev (to_bl x)))" for x :: \'a::len word\ by (simp add: twos_complement word_succ_rbl[OF refl] bl_word_not rev_map rbl_succ2_def) lemma rbl_word_cat: "rev (to_bl (word_cat x y :: 'a::len word)) = takefill False (LENGTH('a)) (rev (to_bl y) @ rev (to_bl x))" by (simp add: word_cat_bl word_rev_tf) lemma rbl_word_slice: "rev (to_bl (slice n w :: 'a::len word)) = takefill False (LENGTH('a)) (drop n (rev (to_bl w)))" apply (simp add: slice_take word_rev_tf rev_take) apply (cases "n < LENGTH('b)", simp_all) done lemma rbl_word_ucast: "rev (to_bl (ucast x :: 'a::len word)) = takefill False (LENGTH('a)) (rev (to_bl x))" apply (simp add: to_bl_ucast takefill_alt) apply (simp add: rev_drop) apply (cases "LENGTH('a) < LENGTH('b)") apply simp_all done lemma rbl_shiftl: "rev (to_bl (w << n)) = takefill False (size w) (replicate n False @ rev (to_bl w))" by (simp add: bl_shiftl takefill_alt word_size rev_drop) lemma rbl_shiftr: "rev (to_bl (w >> n)) = takefill False (size w) (drop n (rev (to_bl w)))" by (simp add: shiftr_slice rbl_word_slice word_size) definition "drop_nonempty v n xs = (if n < length xs then drop n xs else [last (v # xs)])" lemma drop_nonempty_simps: "drop_nonempty v (Suc n) (x # xs) = drop_nonempty x n xs" "drop_nonempty v 0 (x # xs) = (x # xs)" "drop_nonempty v n [] = [v]" by (simp_all add: drop_nonempty_def) definition "takefill_last x n xs = takefill (last (x # xs)) n xs" lemma takefill_last_simps: "takefill_last z (Suc n) (x # xs) = x # takefill_last x n xs" "takefill_last z 0 xs = []" "takefill_last z n [] = replicate n z" by (simp_all add: takefill_last_def) (simp_all add: takefill_alt) lemma rbl_sshiftr: "rev (to_bl (w >>> n)) = takefill_last False (size w) (drop_nonempty False n (rev (to_bl w)))" apply (cases "n < size w") apply (simp add: bl_sshiftr takefill_last_def word_size takefill_alt rev_take last_rev drop_nonempty_def) apply (subgoal_tac "(w >>> n) = of_bl (replicate (size w) (msb w))") apply (simp add: word_size takefill_last_def takefill_alt last_rev word_msb_alt word_rev_tf drop_nonempty_def take_Cons') apply (case_tac "LENGTH('a)", simp_all) apply (rule word_eqI) apply (simp add: nth_sshiftr word_size test_bit_of_bl msb_nth) done lemma nth_word_of_int: - "(word_of_int x :: 'a::len word) !! n = (n < LENGTH('a) \ bin_nth x n)" + "(word_of_int x :: 'a::len word) !! n = (n < LENGTH('a) \ bit x n)" apply (simp add: test_bit_bl word_size to_bl_of_bin) apply (subst conj_cong[OF refl], erule bin_nth_bl) apply auto done lemma nth_scast: "(scast (x :: 'a::len word) :: 'b::len word) !! n = (n < LENGTH('b) \ (if n < LENGTH('a) - 1 then x !! n else x !! (LENGTH('a) - 1)))" apply transfer apply (auto simp add: bit_signed_take_bit_iff min_def) done lemma rbl_word_scast: "rev (to_bl (scast x :: 'a::len word)) = takefill_last False (LENGTH('a)) (rev (to_bl x))" apply (rule nth_equalityI) apply (simp add: word_size takefill_last_def) apply (clarsimp simp: nth_scast takefill_last_def nth_takefill word_size rev_nth to_bl_nth) apply (cases "LENGTH('b)") apply simp apply (clarsimp simp: less_Suc_eq_le linorder_not_less last_rev word_msb_alt[symmetric] msb_nth) done definition rbl_mul :: "bool list \ bool list \ bool list" where "rbl_mul xs ys = foldr (\x sm. rbl_plus False (map ((\) x) ys) (False # sm)) xs []" lemma rbl_mul_simps: "rbl_mul (x # xs) ys = rbl_plus False (map ((\) x) ys) (False # rbl_mul xs ys)" "rbl_mul [] ys = []" by (simp_all add: rbl_mul_def) lemma takefill_le2: "length xs \ n \ takefill x m (takefill x n xs) = takefill x m xs" by (simp add: takefill_alt replicate_add[symmetric]) lemma take_rbl_plus: "\n b. take n (rbl_plus b xs ys) = rbl_plus b (take n xs) (take n ys)" apply (simp add: rbl_plus_def take_zip[symmetric]) apply (rule_tac list="zip xs ys" in list.induct) apply simp apply (clarsimp simp: split_def) apply (case_tac n, simp_all) done lemma word_rbl_mul_induct: "length xs \ size y \ rbl_mul xs (rev (to_bl y)) = take (length xs) (rev (to_bl (of_bl (rev xs) * y)))" for y :: "'a::len word" proof (induct xs) case Nil show ?case by (simp add: rbl_mul_simps) next case (Cons z zs) have rbl_word_plus': "to_bl (x + y) = rev (rbl_plus False (rev (to_bl x)) (rev (to_bl y)))" for x y :: "'a word" by (simp add: rbl_word_plus[symmetric]) have mult_bit: "to_bl (of_bl [z] * y) = map ((\) z) (to_bl y)" by (cases z) (simp cong: map_cong, simp add: map_replicate_const cong: map_cong) have shiftl: "of_bl xs * 2 * y = (of_bl xs * y) << 1" for xs by (simp add: shiftl_t2n) have zip_take_triv: "\xs ys n. n = length ys \ zip (take n xs) ys = zip xs ys" by (rule nth_equalityI) simp_all from Cons show ?case apply (simp add: trans [OF of_bl_append add.commute] rbl_mul_simps rbl_word_plus' distrib_right mult_bit shiftl rbl_shiftl) apply (simp add: takefill_alt word_size rev_map take_rbl_plus min_def) apply (simp add: rbl_plus_def zip_take_triv) done qed lemma rbl_word_mul: "rev (to_bl (x * y)) = rbl_mul (rev (to_bl x)) (rev (to_bl y))" for x :: "'a::len word" using word_rbl_mul_induct[where xs="rev (to_bl x)" and y=y] by (simp add: word_size) text \Breaking up inequalities into bitlist properties.\ definition "rev_bl_order F xs ys = (length xs = length ys \ ((xs = ys \ F) \ (\n < length xs. drop (Suc n) xs = drop (Suc n) ys \ \ xs ! n \ ys ! n)))" lemma rev_bl_order_simps: "rev_bl_order F [] [] = F" "rev_bl_order F (x # xs) (y # ys) = rev_bl_order ((y \ \ x) \ ((y \ \ x) \ F)) xs ys" apply (simp_all add: rev_bl_order_def) apply (rule conj_cong[OF refl]) apply (cases "xs = ys") apply (simp add: nth_Cons') apply blast apply (simp add: nth_Cons') apply safe apply (rule_tac x="n - 1" in exI) apply simp apply (rule_tac x="Suc n" in exI) apply simp done lemma rev_bl_order_rev_simp: "length xs = length ys \ rev_bl_order F (xs @ [x]) (ys @ [y]) = ((y \ \ x) \ ((y \ \ x) \ rev_bl_order F xs ys))" by (induct arbitrary: F rule: list_induct2) (auto simp: rev_bl_order_simps) lemma rev_bl_order_bl_to_bin: "length xs = length ys \ rev_bl_order True xs ys = (bl_to_bin (rev xs) \ bl_to_bin (rev ys)) \ rev_bl_order False xs ys = (bl_to_bin (rev xs) < bl_to_bin (rev ys))" apply (induct xs ys rule: list_induct2) apply (simp_all add: rev_bl_order_simps bl_to_bin_app_cat concat_bit_Suc) apply (auto simp add: bl_to_bin_def add1_zle_eq) done lemma word_le_rbl: "x \ y \ rev_bl_order True (rev (to_bl x)) (rev (to_bl y))" for x y :: "'a::len word" by (simp add: rev_bl_order_bl_to_bin word_le_def) lemma word_less_rbl: "x < y \ rev_bl_order False (rev (to_bl x)) (rev (to_bl y))" for x y :: "'a::len word" by (simp add: word_less_alt rev_bl_order_bl_to_bin) definition "map_last f xs = (if xs = [] then [] else butlast xs @ [f (last xs)])" lemma map_last_simps: "map_last f [] = []" "map_last f [x] = [f x]" "map_last f (x # y # zs) = x # map_last f (y # zs)" by (simp_all add: map_last_def) lemma word_sle_rbl: "x <=s y \ rev_bl_order True (map_last Not (rev (to_bl x))) (map_last Not (rev (to_bl y)))" using word_msb_alt[where w=x] word_msb_alt[where w=y] apply (simp add: word_sle_msb_le word_le_rbl) apply (subgoal_tac "length (to_bl x) = length (to_bl y)") apply (cases "to_bl x", simp) apply (cases "to_bl y", simp) apply (clarsimp simp: map_last_def rev_bl_order_rev_simp) apply auto done lemma word_sless_rbl: "x rev_bl_order False (map_last Not (rev (to_bl x))) (map_last Not (rev (to_bl y)))" using word_msb_alt[where w=x] word_msb_alt[where w=y] apply (simp add: word_sless_msb_less word_less_rbl) apply (subgoal_tac "length (to_bl x) = length (to_bl y)") apply (cases "to_bl x", simp) apply (cases "to_bl y", simp) apply (clarsimp simp: map_last_def rev_bl_order_rev_simp) apply auto done text \Lemmas for unpacking \<^term>\rev (to_bl n)\ for numerals n and also for irreducible values and expressions.\ lemma rev_bin_to_bl_simps: "rev (bin_to_bl 0 x) = []" "rev (bin_to_bl (Suc n) (numeral (num.Bit0 nm))) = False # rev (bin_to_bl n (numeral nm))" "rev (bin_to_bl (Suc n) (numeral (num.Bit1 nm))) = True # rev (bin_to_bl n (numeral nm))" "rev (bin_to_bl (Suc n) (numeral (num.One))) = True # replicate n False" "rev (bin_to_bl (Suc n) (- numeral (num.Bit0 nm))) = False # rev (bin_to_bl n (- numeral nm))" "rev (bin_to_bl (Suc n) (- numeral (num.Bit1 nm))) = True # rev (bin_to_bl n (- numeral (nm + num.One)))" "rev (bin_to_bl (Suc n) (- numeral (num.One))) = True # replicate n True" "rev (bin_to_bl (Suc n) (- numeral (num.Bit0 nm + num.One))) = True # rev (bin_to_bl n (- numeral (nm + num.One)))" "rev (bin_to_bl (Suc n) (- numeral (num.Bit1 nm + num.One))) = False # rev (bin_to_bl n (- numeral (nm + num.One)))" "rev (bin_to_bl (Suc n) (- numeral (num.One + num.One))) = False # rev (bin_to_bl n (- numeral num.One))" by (simp_all add: bin_to_bl_aux_append bin_to_bl_zero_aux bin_to_bl_minus1_aux replicate_append_same) lemma to_bl_upt: "to_bl x = rev (map ((!!) x) [0 ..< size x])" apply (rule nth_equalityI) apply (simp add: word_size) apply (auto simp: to_bl_nth word_size rev_nth) done lemma rev_to_bl_upt: "rev (to_bl x) = map ((!!) x) [0 ..< size x]" by (simp add: to_bl_upt) lemma upt_eq_list_intros: "j \ i \ [i ..< j] = []" "i = x \ x < j \ [x + 1 ..< j] = xs \ [i ..< j] = (x # xs)" by (simp_all add: upt_eq_Cons_conv) subsection \Tactic definition\ lemma if_bool_simps: "If p True y = (p \ y) \ If p False y = (\ p \ y) \ If p y True = (p \ y) \ If p y False = (p \ y)" by auto ML \ structure Word_Bitwise_Tac = struct val word_ss = simpset_of \<^theory_context>\Word\; fun mk_nat_clist ns = fold_rev (Thm.mk_binop \<^cterm>\Cons :: nat \ _\) ns \<^cterm>\[] :: nat list\; fun upt_conv ctxt ct = case Thm.term_of ct of (\<^const>\upt\ $ n $ m) => let val (i, j) = apply2 (snd o HOLogic.dest_number) (n, m); val ns = map (Numeral.mk_cnumber \<^ctyp>\nat\) (i upto (j - 1)) |> mk_nat_clist; val prop = Thm.mk_binop \<^cterm>\(=) :: nat list \ _\ ct ns |> Thm.apply \<^cterm>\Trueprop\; in try (fn () => Goal.prove_internal ctxt [] prop (K (REPEAT_DETERM (resolve_tac ctxt @{thms upt_eq_list_intros} 1 ORELSE simp_tac (put_simpset word_ss ctxt) 1))) |> mk_meta_eq) () end | _ => NONE; val expand_upt_simproc = Simplifier.make_simproc \<^context> "expand_upt" {lhss = [\<^term>\upt x y\], proc = K upt_conv}; fun word_len_simproc_fn ctxt ct = (case Thm.term_of ct of Const (\<^const_name>\len_of\, _) $ t => (let val T = fastype_of t |> dest_Type |> snd |> the_single val n = Numeral.mk_cnumber \<^ctyp>\nat\ (Word_Lib.dest_binT T); val prop = Thm.mk_binop \<^cterm>\(=) :: nat \ _\ ct n |> Thm.apply \<^cterm>\Trueprop\; in Goal.prove_internal ctxt [] prop (K (simp_tac (put_simpset word_ss ctxt) 1)) |> mk_meta_eq |> SOME end handle TERM _ => NONE | TYPE _ => NONE) | _ => NONE); val word_len_simproc = Simplifier.make_simproc \<^context> "word_len" {lhss = [\<^term>\len_of x\], proc = K word_len_simproc_fn}; (* convert 5 or nat 5 to Suc 4 when n_sucs = 1, Suc (Suc 4) when n_sucs = 2, or just 5 (discarding nat) when n_sucs = 0 *) fun nat_get_Suc_simproc_fn n_sucs ctxt ct = let val (f, arg) = dest_comb (Thm.term_of ct); val n = (case arg of \<^term>\nat\ $ n => n | n => n) |> HOLogic.dest_number |> snd; val (i, j) = if n > n_sucs then (n_sucs, n - n_sucs) else (n, 0); val arg' = funpow i HOLogic.mk_Suc (HOLogic.mk_number \<^typ>\nat\ j); val _ = if arg = arg' then raise TERM ("", []) else (); fun propfn g = HOLogic.mk_eq (g arg, g arg') |> HOLogic.mk_Trueprop |> Thm.cterm_of ctxt; val eq1 = Goal.prove_internal ctxt [] (propfn I) (K (simp_tac (put_simpset word_ss ctxt) 1)); in Goal.prove_internal ctxt [] (propfn (curry (op $) f)) (K (simp_tac (put_simpset HOL_ss ctxt addsimps [eq1]) 1)) |> mk_meta_eq |> SOME end handle TERM _ => NONE; fun nat_get_Suc_simproc n_sucs ts = Simplifier.make_simproc \<^context> "nat_get_Suc" {lhss = map (fn t => t $ \<^term>\n :: nat\) ts, proc = K (nat_get_Suc_simproc_fn n_sucs)}; val no_split_ss = simpset_of (put_simpset HOL_ss \<^context> |> Splitter.del_split @{thm if_split}); val expand_word_eq_sss = (simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms word_eq_rbl_eq word_le_rbl word_less_rbl word_sle_rbl word_sless_rbl}), map simpset_of [ put_simpset no_split_ss \<^context> addsimps @{thms rbl_word_plus rbl_word_and rbl_word_or rbl_word_not rbl_word_neg bl_word_sub rbl_word_xor rbl_word_cat rbl_word_slice rbl_word_scast rbl_word_ucast rbl_shiftl rbl_shiftr rbl_sshiftr rbl_word_if}, put_simpset no_split_ss \<^context> addsimps @{thms to_bl_numeral to_bl_neg_numeral to_bl_0 rbl_word_1}, put_simpset no_split_ss \<^context> addsimps @{thms rev_rev_ident rev_replicate rev_map to_bl_upt word_size} addsimprocs [word_len_simproc], put_simpset no_split_ss \<^context> addsimps @{thms list.simps split_conv replicate.simps list.map zip_Cons_Cons zip_Nil drop_Suc_Cons drop_0 drop_Nil foldr.simps list.map zip.simps(1) zip_Nil zip_Cons_Cons takefill_Suc_Cons takefill_Suc_Nil takefill.Z rbl_succ2_simps rbl_plus_simps rev_bin_to_bl_simps append.simps takefill_last_simps drop_nonempty_simps rev_bl_order_simps} addsimprocs [expand_upt_simproc, nat_get_Suc_simproc 4 [\<^term>\replicate\, \<^term>\takefill x\, \<^term>\drop\, \<^term>\bin_to_bl\, \<^term>\takefill_last x\, \<^term>\drop_nonempty x\]], put_simpset no_split_ss \<^context> addsimps @{thms xor3_simps carry_simps if_bool_simps} ]) fun tac ctxt = let val (ss, sss) = expand_word_eq_sss; in foldr1 (op THEN_ALL_NEW) ((CHANGED o safe_full_simp_tac (put_simpset ss ctxt)) :: map (fn ss => safe_full_simp_tac (put_simpset ss ctxt)) sss) end; end \ method_setup word_bitwise = \Scan.succeed (fn ctxt => Method.SIMPLE_METHOD (Word_Bitwise_Tac.tac ctxt 1))\ "decomposer for word equalities and inequalities into bit propositions" end diff --git a/thys/Word_Lib/Generic_set_bit.thy b/thys/Word_Lib/Generic_set_bit.thy --- a/thys/Word_Lib/Generic_set_bit.thy +++ b/thys/Word_Lib/Generic_set_bit.thy @@ -1,190 +1,190 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) (* Author: Jeremy Dawson, NICTA *) section \Operation variant for setting and unsetting bits\ theory Generic_set_bit imports "HOL-Library.Word" Bits_Int Most_significant_bit begin class set_bit = semiring_bits + fixes set_bit :: \'a \ nat \ bool \ 'a\ assumes bit_set_bit_iff [bit_simps]: \bit (set_bit a m b) n \ (if m = n then b else bit a n) \ 2 ^ n \ 0\ lemma set_bit_eq: \set_bit a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ for a :: \'a::{ring_bit_operations, set_bit}\ by (rule bit_eqI) (simp add: bit_simps) instantiation int :: set_bit begin definition set_bit_int :: \int \ nat \ bool \ int\ where \set_bit i n b = bin_sc n b i\ instance by standard (simp_all add: set_bit_int_def bin_nth_sc_gen bit_simps) end lemma int_set_bit_0 [simp]: fixes x :: int shows "set_bit x 0 b = of_bool b + 2 * (x div 2)" by (auto simp add: set_bit_int_def intro: bin_rl_eqI) lemma int_set_bit_Suc: fixes x :: int shows "set_bit x (Suc n) b = of_bool (odd x) + 2 * set_bit (x div 2) n b" by (auto simp add: set_bit_int_def intro: bin_rl_eqI) lemma bin_last_set_bit: - "bin_last (set_bit x n b) = (if n > 0 then bin_last x else b)" + "odd (set_bit x n b :: int) = (if n > 0 then odd x else b)" by (cases n) (simp_all add: int_set_bit_Suc) lemma bin_rest_set_bit: - "bin_rest (set_bit x n b) = (if n > 0 then set_bit (x div 2) (n - 1) b else x div 2)" + "(set_bit x n b :: int) div 2 = (if n > 0 then set_bit (x div 2) (n - 1) b else x div 2)" by (cases n) (simp_all add: int_set_bit_Suc) lemma int_set_bit_numeral: fixes x :: int shows "set_bit x (numeral w) b = of_bool (odd x) + 2 * set_bit (x div 2) (pred_numeral w) b" by (simp add: set_bit_int_def) lemmas int_set_bit_numerals [simp] = int_set_bit_numeral[where x="numeral w'"] int_set_bit_numeral[where x="- numeral w'"] int_set_bit_numeral[where x="Numeral1"] int_set_bit_numeral[where x="1"] int_set_bit_numeral[where x="0"] int_set_bit_Suc[where x="numeral w'"] int_set_bit_Suc[where x="- numeral w'"] int_set_bit_Suc[where x="Numeral1"] int_set_bit_Suc[where x="1"] int_set_bit_Suc[where x="0"] for w' lemma msb_set_bit [simp]: "msb (set_bit (x :: int) n b) \ msb x" by(simp add: msb_conv_bin_sign set_bit_int_def) instantiation word :: (len) set_bit begin definition set_bit_word :: \'a word \ nat \ bool \ 'a word\ where word_set_bit_def: \set_bit a n x = word_of_int (bin_sc n x (uint a))\ instance by standard (auto simp add: word_set_bit_def bin_nth_sc_gen bit_simps) end lemma set_bit_unfold: \set_bit w n b = (if b then Bit_Operations.set_bit n w else unset_bit n w)\ for w :: \'a::len word\ by (simp add: set_bit_eq) lemma bit_set_bit_word_iff [bit_simps]: \bit (set_bit w m b) n \ (if m = n then n < LENGTH('a) \ b else bit w n)\ for w :: \'a::len word\ by (auto simp add: bit_simps dest: bit_imp_le_length) lemma word_set_nth [simp]: "set_bit w n (test_bit w n) = w" for w :: "'a::len word" by (auto simp: word_test_bit_def word_set_bit_def) lemma test_bit_set: "(set_bit w n x) !! n \ n < size w \ x" for w :: "'a::len word" by (auto simp: word_size word_test_bit_def word_set_bit_def nth_bintr) lemma test_bit_set_gen: "test_bit (set_bit w n x) m = (if m = n then n < size w \ x else test_bit w m)" for w :: "'a::len word" apply (unfold word_size word_test_bit_def word_set_bit_def) apply (clarsimp simp add: nth_bintr bin_nth_sc_gen) apply (auto elim!: test_bit_size [unfolded word_size] simp add: word_test_bit_def [symmetric]) done lemma word_set_set_same [simp]: "set_bit (set_bit w n x) n y = set_bit w n y" for w :: "'a::len word" by (rule word_eqI) (simp add : test_bit_set_gen word_size) lemma word_set_set_diff: fixes w :: "'a::len word" assumes "m \ n" shows "set_bit (set_bit w m x) n y = set_bit (set_bit w n y) m x" by (rule word_eqI) (auto simp: test_bit_set_gen word_size assms) lemma set_bit_word_of_int: "set_bit (word_of_int x) n b = word_of_int (bin_sc n b x)" unfolding word_set_bit_def by (rule word_eqI)(simp add: word_size bin_nth_sc_gen nth_bintr) lemma word_set_numeral [simp]: "set_bit (numeral bin::'a::len word) n b = word_of_int (bin_sc n b (numeral bin))" unfolding word_numeral_alt by (rule set_bit_word_of_int) lemma word_set_neg_numeral [simp]: "set_bit (- numeral bin::'a::len word) n b = word_of_int (bin_sc n b (- numeral bin))" unfolding word_neg_numeral_alt by (rule set_bit_word_of_int) lemma word_set_bit_0 [simp]: "set_bit 0 n b = word_of_int (bin_sc n b 0)" unfolding word_0_wi by (rule set_bit_word_of_int) lemma word_set_bit_1 [simp]: "set_bit 1 n b = word_of_int (bin_sc n b 1)" unfolding word_1_wi by (rule set_bit_word_of_int) lemma word_set_nth_iff: "set_bit w n b = w \ w !! n = b \ n \ size w" for w :: "'a::len word" apply (rule iffI) apply (rule disjCI) apply (drule word_eqD) apply (erule sym [THEN trans]) apply (simp add: test_bit_set) apply (erule disjE) apply clarsimp apply (rule word_eqI) apply (clarsimp simp add : test_bit_set_gen) apply (drule test_bit_size) apply force done lemma word_clr_le: "w \ set_bit w n False" for w :: "'a::len word" apply (simp add: word_set_bit_def word_le_def) apply transfer apply (rule order_trans) apply (rule bintr_bin_clr_le) apply simp done lemma word_set_ge: "w \ set_bit w n True" for w :: "'a::len word" apply (simp add: word_set_bit_def word_le_def) apply transfer apply (rule order_trans [OF _ bintr_bin_set_ge]) apply simp done lemma set_bit_beyond: "size x \ n \ set_bit x n b = x" for x :: "'a :: len word" by (auto intro: word_eqI simp add: test_bit_set_gen word_size) lemma one_bit_shiftl: "set_bit 0 n True = (1 :: 'a :: len word) << n" apply (rule word_eqI) apply (auto simp add: test_bit_set_gen nth_shiftl word_size simp del: word_set_bit_0 shiftl_1) done lemmas one_bit_pow = trans [OF one_bit_shiftl shiftl_1] end diff --git a/thys/Word_Lib/Guide.thy b/thys/Word_Lib/Guide.thy --- a/thys/Word_Lib/Guide.thy +++ b/thys/Word_Lib/Guide.thy @@ -1,376 +1,380 @@ (* * Copyright Florian Haftmann * * SPDX-License-Identifier: BSD-2-Clause *) (*<*) theory Guide imports Word_Lib_Sumo Word_64 Ancient_Numeral begin hide_const (open) Generic_set_bit.set_bit (*>*) section \A short overview over bit operations and word types\ subsection \Basic theories and key ideas\ text \ When formalizing bit operations, it is tempting to represent bit values as explicit lists over a binary type. This however is a bad idea, mainly due to the inherent ambiguities in representation concerning repeating leading bits. Hence this approach avoids such explicit lists altogether following an algebraic path: \<^item> Bit values are represented by numeric types: idealized unbounded bit values can be represented by type \<^typ>\int\, bounded bit values by quotient types over \<^typ>\int\, aka \<^typ>\'a word\. \<^item> (A special case are idealized unbounded bit values ending in @{term [source] 0} which can be represented by type \<^typ>\nat\ but only support a restricted set of operations). The most fundamental ideas are developed in theory \<^theory>\HOL.Parity\ (which is part of \<^theory>\Main\): \<^item> Multiplication by \<^term>\2 :: int\ is a bit shift to the left and \<^item> Division by \<^term>\2 :: int\ is a bit shift to the right. \<^item> Concerning bounded bit values, iterated shifts to the left may result in eliminating all bits by shifting them all beyond the boundary. The property \<^prop>\(2 :: int) ^ n \ 0\ represents that \<^term>\n\ is \<^emph>\not\ beyond that boundary. \<^item> The projection on a single bit is then @{thm [mode=iff] bit_iff_odd [where ?'a = int, no_vars]}. \<^item> This leads to the most fundamental properties of bit values: \<^item> Equality rule: @{thm [display, mode=iff] bit_eq_iff [where ?'a = int, no_vars]} \<^item> Induction rule: @{thm [display, mode=iff] bits_induct [where ?'a = int, no_vars]} \<^item> Characteristic properties \<^prop>\bit (f x) n \ P x n\ are available in fact collection \<^text>\bit_simps\. On top of this, the following generic operations are provided after import of theory \<^theory>\HOL-Library.Bit_Operations\: \<^item> Singleton \<^term>\n\th bit: \<^term>\(2 :: int) ^ n\ \<^item> Bit mask upto bit \<^term>\n\: @{thm mask_eq_exp_minus_1 [where ?'a = int, no_vars]} \<^item> Left shift: @{thm push_bit_eq_mult [where ?'a = int, no_vars]} \<^item> Right shift: @{thm drop_bit_eq_div [where ?'a = int, no_vars]} \<^item> Truncation: @{thm take_bit_eq_mod [where ?'a = int, no_vars]} \<^item> Bitwise negation: @{thm [mode=iff] bit_not_iff [where ?'a = int, no_vars]} \<^item> Bitwise conjunction: @{thm [mode=iff] bit_and_iff [where ?'a = int, no_vars]} \<^item> Bitwise disjunction: @{thm [mode=iff] bit_or_iff [where ?'a = int, no_vars]} \<^item> Bitwise exclusive disjunction: @{thm [mode=iff] bit_xor_iff [where ?'a = int, no_vars]} \<^item> Setting a single bit: @{thm set_bit_def [where ?'a = int, no_vars]} \<^item> Unsetting a single bit: @{thm unset_bit_def [where ?'a = int, no_vars]} \<^item> Flipping a single bit: @{thm flip_bit_def [where ?'a = int, no_vars]} \<^item> Signed truncation, or modulus centered around \<^term>\0::int\: @{thm [display] signed_take_bit_def [where ?'a = int, no_vars]} \<^item> (Bounded) conversion from and to a list of bits: @{thm [display] horner_sum_bit_eq_take_bit [where ?'a = int, no_vars]} Bit concatenation on \<^typ>\int\ as given by @{thm [display] concat_bit_def [no_vars]} appears quite technical but is the logical foundation for the quite natural bit concatenation on \<^typ>\'a word\ (see below). Proper word types are introduced in theory \<^theory>\HOL-Library.Word\, with the following specific operations: \<^item> Standard arithmetic: @{term \(+) :: 'a::len word \ 'a word \ 'a word\}, @{term \uminus :: 'a::len word \ 'a word\}, @{term \(-) :: 'a::len word \ 'a word \ 'a word\}, @{term \(*) :: 'a::len word \ 'a word \ 'a word\}, @{term \0 :: 'a::len word\}, @{term \1 :: 'a::len word\}, numerals etc. \<^item> Standard bit operations: see above. \<^item> Conversion with unsigned interpretation of words: \<^item> @{term [source] \unsigned :: 'a::len word \ 'b::semiring_1\} \<^item> Important special cases as abbreviations: \<^item> @{term [source] \unat :: 'a::len word \ nat\} \<^item> @{term [source] \uint :: 'a::len word \ int\} \<^item> @{term [source] \ucast :: 'a::len word \ 'b::len word\} \<^item> Conversion with signed interpretation of words: \<^item> @{term [source] \signed :: 'a::len word \ 'b::ring_1\} \<^item> Important special cases as abbreviations: \<^item> @{term [source] \sint :: 'a::len word \ int\} \<^item> @{term [source] \scast :: 'a::len word \ 'b::len word\} \<^item> Operations with unsigned interpretation of words: \<^item> @{thm [mode=iff] word_le_nat_alt [no_vars]} \<^item> @{thm [mode=iff] word_less_nat_alt [no_vars]} \<^item> @{thm unat_div_distrib [no_vars]} \<^item> @{thm unat_drop_bit_eq [no_vars]} \<^item> @{thm unat_mod_distrib [no_vars]} \<^item> @{thm [mode=iff] udvd_iff_dvd [no_vars]} \<^item> Operations with signed interpretation of words: \<^item> @{thm [mode=iff] word_sle_eq [no_vars]} \<^item> @{thm [mode=iff] word_sless_alt [no_vars]} \<^item> @{thm sint_signed_drop_bit_eq [no_vars]} \<^item> Rotation and reversal: \<^item> @{term [source] \word_rotl :: nat \ 'a::len word \ 'a word\} \<^item> @{term [source] \word_rotr :: nat \ 'a::len word \ 'a word\} \<^item> @{term [source] \word_roti :: int \ 'a::len word \ 'a word\} \<^item> @{term [source] \word_reverse :: 'a::len word \ 'a word\} \<^item> Concatenation: @{term [source, display] \word_cat :: 'a::len word \ 'b::len word \ 'c::len word\} For proofs about words the following default strategies are applicable: \<^item> Using bit extensionality (facts \<^text>\bit_eq_iff\, \<^text>\bit_eqI\; fact collection \<^text>\bit_simps\). \<^item> Using the @{method transfer} method. \ subsection \More library theories\ text \ Note: currently, the theories listed here are hardly separate entities since they import each other in various ways. Always inspect them to understand what you pull in if you want to import one. \<^descr>[Syntax] \<^descr>[\<^theory>\Word_Lib.Hex_Words\] Printing word numerals as hexadecimal numerals. \<^descr>[\<^theory>\Word_Lib.Type_Syntax\] Pretty type-sensitive syntax for cast operations. \<^descr>[\<^theory>\Word_Lib.Word_Syntax\] Specific ASCII syntax for prominent bit operations on word. \<^descr>[Proof tools] \<^descr>[\<^theory>\Word_Lib.Norm_Words\] Rewriting word numerals to normal forms. \<^descr>[\<^theory>\Word_Lib.Bitwise\] Method @{method word_bitwise} decomposes word equalities and inequalities into bit propositions. \<^descr>[\<^theory>\Word_Lib.Word_EqI\] Method @{method word_eqI_solve} decomposes word equalities and inequalities into bit propositions. \<^descr>[Operations] \<^descr>[\<^theory>\Word_Lib.Signed_Division_Word\] Signed division on word: \<^item> @{term [source] \(sdiv) :: 'a::len word \ 'a word \ 'a word\} \<^item> @{term [source] \(smod) :: 'a::len word \ 'a word \ 'a word\} \<^descr>[\<^theory>\Word_Lib.Aligned\] \ \<^item> @{thm [mode=iff] is_aligned_iff_udvd [no_vars]} \<^descr>[\<^theory>\Word_Lib.Least_significant_bit\] The least significant bit as an alias: @{thm [mode=iff] lsb_odd [where ?'a = int, no_vars]} \<^descr>[\<^theory>\Word_Lib.Most_significant_bit\] The most significant bit: \<^item> @{thm [mode=iff] msb_int_def [of k]} \<^item> @{thm [mode=iff] word_msb_sint [no_vars]} \<^item> @{thm [mode=iff] msb_word_iff_sless_0 [no_vars]} \<^item> @{thm [mode=iff] msb_word_iff_bit [no_vars]} \<^descr>[\<^theory>\Word_Lib.Traditional_Infix_Syntax\] Clones of existing operations decorated with traditional syntax: \<^item> @{thm test_bit_eq_bit [no_vars]} \<^item> @{thm shiftl_eq_push_bit [no_vars]} \<^item> @{thm shiftr_eq_drop_bit [no_vars]} \<^item> @{thm sshiftr_eq [no_vars]} \<^descr>[\<^theory>\Word_Lib.Next_and_Prev\] \ \<^item> @{thm word_next_unfold [no_vars]} \<^item> @{thm word_prev_unfold [no_vars]} \<^descr>[\<^theory>\Word_Lib.Enumeration_Word\] More on explicit enumeration of word types. \<^descr>[\<^theory>\Word_Lib.More_Word_Operations\] Even more operations on word. \<^descr>[Types] \<^descr>[\<^theory>\Word_Lib.Signed_Words\] Formal tagging of word types with a \<^text>\signed\ marker. \<^descr>[Lemmas] \<^descr>[\<^theory>\Word_Lib.More_Word\] More lemmas on words. \<^descr>[\<^theory>\Word_Lib.Word_Lemmas\] More lemmas on words, covering many other theories mentioned here. \<^descr>[Words of popular lengths]. \<^descr>[\<^theory>\Word_Lib.Word_8\] for 8-bit words. \<^descr>[\<^theory>\Word_Lib.Word_16\] for 16-bit words. \<^descr>[\<^theory>\Word_Lib.Word_32\] for 32-bit words. \<^descr>[\<^theory>\Word_Lib.Word_64\] for 64-bit words. This theory is not part of \<^text>\Word_Lib_Sumo\, because it shadows names from \<^theory>\Word_Lib.Word_32\. They can be used together, but then require to use qualified names in applications. \ subsection \More library sessions\ text \ \<^descr>[\<^text>\Native_Word\] Makes machine words and machine arithmetic available for code generation. It provides a common abstraction that hides the differences between the different target languages. The code generator maps these operations to the APIs of the target languages. \ subsection \Legacy theories\ text \ The following theories contain material which has been factored out since it is not recommended to use it in new applications, mostly because matters can be expressed succinctly using already existing operations. This section gives some indication how to migrate away from those theories. However theorem coverage may still be terse in some cases. \<^descr>[\<^theory>\Word_Lib.Word_Lib_Sumo\] An entry point importing any relevant theory in that session. Intended for backward compatibility: start importing this theory when migrating applications to Isabelle2021, and later sort out what you really need. You may need to include \<^theory>\Word_Lib.Word_64\ separately. \<^descr>[\<^theory>\Word_Lib.Generic_set_bit\] Kind of an alias: @{thm set_bit_eq [no_vars]} \<^descr>[\<^theory>\Word_Lib.Typedef_Morphisms\] A low-level extension to HOL typedef providing conversions along type morphisms. The @{method transfer} method seems to be sufficient for most applications though. \<^descr>[\<^theory>\Word_Lib.Bit_Comprehension\] Comprehension syntax for bit values over predicates \<^typ>\nat \ bool\. For \<^typ>\'a::len word\, straightforward alternatives exist; difficult to handle for \<^typ>\int\. \<^descr>[\<^theory>\Word_Lib.Reversed_Bit_Lists\] Representation of bit values as explicit list in \<^emph>\reversed\ order. This should rarely be necessary: the \<^const>\bit\ projection should be sufficient in most cases. In case explicit lists are needed, existing operations can be used: @{thm [display] horner_sum_bit_eq_take_bit [where ?'a = int, no_vars]} \<^descr>[\<^theory>\Word_Lib.Many_More\] Collection of operations and theorems which are kept for backward compatibility and not used in other theories in session \<^text>\Word_Lib\. They are used in applications of \<^text>\Word_Lib\, but should be migrated to there. \ section \Changelog\ text \ \<^descr>[Changes since AFP 2021] ~ \<^item> Theory \<^theory>\Word_Lib.Ancient_Numeral\ is no part of \<^theory>\Word_Lib.Word_Lib_Sumo\ any longer. \<^item> Abbreviation \<^abbrev>\max_word\ moved from distribution into theory \<^theory>\Word_Lib.Legacy_Aliases\. + + \<^item> Abbreviations \<^abbrev>\bin_nth\, \<^abbrev>\bin_last\, \<^abbrev>\bin_rest\, + \<^abbrev>\bintrunc\, \<^abbrev>\sbintrunc\, \<^abbrev>\norm_sint\, + \<^abbrev>\bin_cat\ moved into theory \<^theory>\Word_Lib.Legacy_Aliases\. \ (*<*) end (*>*) diff --git a/thys/Word_Lib/Least_significant_bit.thy b/thys/Word_Lib/Least_significant_bit.thy --- a/thys/Word_Lib/Least_significant_bit.thy +++ b/thys/Word_Lib/Least_significant_bit.thy @@ -1,94 +1,94 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) (* Author: Jeremy Dawson, NICTA *) section \Operation variant for the least significant bit\ theory Least_significant_bit imports "HOL-Library.Word" Bits_Int begin class lsb = semiring_bits + fixes lsb :: \'a \ bool\ assumes lsb_odd: \lsb = odd\ instantiation int :: lsb begin definition lsb_int :: \int \ bool\ where \lsb i = i !! 0\ for i :: int instance by standard (simp add: fun_eq_iff lsb_int_def) end -lemma bin_last_conv_lsb: "bin_last = lsb" +lemma bin_last_conv_lsb: "odd = (lsb :: int \ bool)" by (simp add: lsb_odd) lemma int_lsb_numeral [simp]: "lsb (0 :: int) = False" "lsb (1 :: int) = True" "lsb (Numeral1 :: int) = True" "lsb (- 1 :: int) = True" "lsb (- Numeral1 :: int) = True" "lsb (numeral (num.Bit0 w) :: int) = False" "lsb (numeral (num.Bit1 w) :: int) = True" "lsb (- numeral (num.Bit0 w) :: int) = False" "lsb (- numeral (num.Bit1 w) :: int) = True" by (simp_all add: lsb_int_def) instantiation word :: (len) lsb begin definition lsb_word :: \'a word \ bool\ where word_lsb_def: \lsb a \ odd (uint a)\ for a :: \'a word\ instance apply standard apply (simp add: fun_eq_iff word_lsb_def) apply transfer apply simp done end lemma lsb_word_eq: \lsb = (odd :: 'a word \ bool)\ for w :: \'a::len word\ by (fact lsb_odd) lemma word_lsb_alt: "lsb w = test_bit w 0" for w :: "'a::len word" by (auto simp: word_test_bit_def word_lsb_def) lemma word_lsb_1_0 [simp]: "lsb (1::'a::len word) \ \ lsb (0::'b::len word)" unfolding word_lsb_def by simp lemma word_lsb_int: "lsb w \ uint w mod 2 = 1" apply (simp add: lsb_odd flip: odd_iff_mod_2_eq_one) apply transfer apply simp done lemmas word_ops_lsb = lsb0 [unfolded word_lsb_alt] lemma word_lsb_numeral [simp]: - "lsb (numeral bin :: 'a::len word) \ bin_last (numeral bin)" + "lsb (numeral bin :: 'a::len word) \ odd (numeral bin :: int)" unfolding word_lsb_alt test_bit_numeral by simp lemma word_lsb_neg_numeral [simp]: - "lsb (- numeral bin :: 'a::len word) \ bin_last (- numeral bin)" + "lsb (- numeral bin :: 'a::len word) \ odd (- numeral bin :: int)" by (simp add: word_lsb_alt) lemma word_lsb_nat:"lsb w = (unat w mod 2 = 1)" apply (simp add: word_lsb_def Groebner_Basis.algebra(31)) apply transfer apply (simp add: even_nat_iff) done end diff --git a/thys/Word_Lib/Legacy_Aliases.thy b/thys/Word_Lib/Legacy_Aliases.thy --- a/thys/Word_Lib/Legacy_Aliases.thy +++ b/thys/Word_Lib/Legacy_Aliases.thy @@ -1,27 +1,48 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section \Legacy aliases\ theory Legacy_Aliases imports "HOL-Library.Word" begin +abbreviation (input) bin_nth :: \int \ nat \ bool\ + where \bin_nth \ bit\ + +abbreviation (input) bin_last :: \int \ bool\ + where \bin_last \ odd\ + +abbreviation (input) bin_rest :: \int \ int\ + where \bin_rest w \ w div 2\ + +abbreviation (input) bintrunc :: \nat \ int \ int\ + where \bintrunc \ take_bit\ + +abbreviation (input) sbintrunc :: \nat \ int \ int\ + where \sbintrunc \ signed_take_bit\ + +abbreviation (input) bin_cat :: \int \ nat \ int \ int\ + where \bin_cat k n l \ concat_bit n l k\ + +abbreviation (input) norm_sint :: \nat \ int \ int\ + where \norm_sint n \ signed_take_bit (n - 1)\ + abbreviation (input) max_word :: \'a::len word\ - where "max_word \ - 1" + where \max_word \ - 1\ definition complement :: "'a :: len word \ 'a word" where "complement x \ NOT x" lemma complement_mask: "complement (2 ^ n - 1) = NOT (mask n)" unfolding complement_def mask_eq_decr_exp by simp lemmas less_def = less_eq [symmetric] lemmas le_def = not_less [symmetric, where ?'a = nat] end diff --git a/thys/Word_Lib/Most_significant_bit.thy b/thys/Word_Lib/Most_significant_bit.thy --- a/thys/Word_Lib/Most_significant_bit.thy +++ b/thys/Word_Lib/Most_significant_bit.thy @@ -1,203 +1,203 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) (* Author: Jeremy Dawson, NICTA *) section \Dedicated operation for the most significant bit\ theory Most_significant_bit imports "HOL-Library.Word" Bits_Int Traditional_Infix_Syntax More_Arithmetic begin class msb = fixes msb :: \'a \ bool\ instantiation int :: msb begin definition \msb x \ x < 0\ for x :: int instance .. end lemma msb_conv_bin_sign: "msb x \ bin_sign x = -1" by(simp add: bin_sign_def not_le msb_int_def) lemma msb_bin_rest [simp]: "msb (x div 2) = msb x" for x :: int by (simp add: msb_int_def) lemma int_msb_and [simp]: "msb ((x :: int) AND y) \ msb x \ msb y" by(simp add: msb_int_def) lemma int_msb_or [simp]: "msb ((x :: int) OR y) \ msb x \ msb y" by(simp add: msb_int_def) lemma int_msb_xor [simp]: "msb ((x :: int) XOR y) \ msb x \ msb y" by(simp add: msb_int_def) lemma int_msb_not [simp]: "msb (NOT (x :: int)) \ \ msb x" by(simp add: msb_int_def not_less) lemma msb_shiftl [simp]: "msb ((x :: int) << n) \ msb x" by(simp add: msb_int_def) lemma msb_shiftr [simp]: "msb ((x :: int) >> r) \ msb x" by(simp add: msb_int_def) lemma msb_bin_sc [simp]: "msb (bin_sc n b x) \ msb x" by(simp add: msb_conv_bin_sign) lemma msb_0 [simp]: "msb (0 :: int) = False" by(simp add: msb_int_def) lemma msb_1 [simp]: "msb (1 :: int) = False" by(simp add: msb_int_def) lemma msb_numeral [simp]: "msb (numeral n :: int) = False" "msb (- numeral n :: int) = True" by(simp_all add: msb_int_def) instantiation word :: (len) msb begin definition msb_word :: \'a word \ bool\ - where \msb a \ bin_sign (sbintrunc (LENGTH('a) - 1) (uint a)) = - 1\ + where \msb a \ bin_sign (signed_take_bit (LENGTH('a) - 1) (uint a)) = - 1\ lemma msb_word_eq: \msb w \ bit w (LENGTH('a) - 1)\ for w :: \'a::len word\ by (simp add: msb_word_def bin_sign_lem bit_uint_iff) instance .. end lemma msb_word_iff_bit: \msb w \ bit w (LENGTH('a) - Suc 0)\ for w :: \'a::len word\ by (simp add: msb_word_def bin_sign_def bit_uint_iff) lemma word_msb_def: "msb a \ bin_sign (sint a) = - 1" by (simp add: msb_word_def sint_uint) lemma word_msb_sint: "msb w \ sint w < 0" by (simp add: msb_word_eq bit_last_iff) lemma msb_word_iff_sless_0: \msb w \ w by (simp add: word_msb_sint word_sless_alt) -lemma msb_word_of_int: "msb (word_of_int x::'a::len word) = bin_nth x (LENGTH('a) - 1)" +lemma msb_word_of_int: "msb (word_of_int x::'a::len word) = bit x (LENGTH('a) - 1)" by (simp add: word_msb_def bin_sign_lem) lemma word_msb_numeral [simp]: - "msb (numeral w::'a::len word) = bin_nth (numeral w) (LENGTH('a) - 1)" + "msb (numeral w::'a::len word) = bit (numeral w :: int) (LENGTH('a) - 1)" unfolding word_numeral_alt by (rule msb_word_of_int) lemma word_msb_neg_numeral [simp]: - "msb (- numeral w::'a::len word) = bin_nth (- numeral w) (LENGTH('a) - 1)" + "msb (- numeral w::'a::len word) = bit (- numeral w :: int) (LENGTH('a) - 1)" unfolding word_neg_numeral_alt by (rule msb_word_of_int) lemma word_msb_0 [simp]: "\ msb (0::'a::len word)" by (simp add: word_msb_def bin_sign_def sint_uint sbintrunc_eq_take_bit) lemma word_msb_1 [simp]: "msb (1::'a::len word) \ LENGTH('a) = 1" unfolding word_1_wi msb_word_of_int eq_iff [where 'a=nat] by (simp add: Suc_le_eq) -lemma word_msb_nth: "msb w = bin_nth (uint w) (LENGTH('a) - 1)" +lemma word_msb_nth: "msb w = bit (uint w) (LENGTH('a) - 1)" for w :: "'a::len word" by (simp add: word_msb_def sint_uint bin_sign_lem) lemma msb_nth: "msb w = w !! (LENGTH('a) - 1)" for w :: "'a::len word" by (simp add: word_msb_nth word_test_bit_def) lemma word_msb_n1 [simp]: "msb (-1::'a::len word)" by (simp add: msb_word_eq not_le) lemma msb_shift: "msb w \ w >> (LENGTH('a) - 1) \ 0" for w :: "'a::len word" by (simp add: msb_word_eq shiftr_word_eq bit_iff_odd_drop_bit drop_bit_eq_zero_iff_not_bit_last) lemmas word_ops_msb = msb1 [unfolded msb_nth [symmetric, unfolded One_nat_def]] lemma word_sint_msb_eq: "sint x = uint x - (if msb x then 2 ^ size x else 0)" apply (cases \LENGTH('a)\) apply (simp_all add: msb_word_def bin_sign_def bit_simps word_size) apply transfer apply (auto simp add: take_bit_Suc_from_most signed_take_bit_eq_if_positive signed_take_bit_eq_if_negative minus_exp_eq_not_mask ac_simps) apply (subst disjunctive_add) apply (simp_all add: bit_simps) done lemma word_sle_msb_le: "x <=s y \ (msb y \ msb x) \ ((msb x \ \ msb y) \ x \ y)" apply (simp add: word_sle_eq word_sint_msb_eq word_size word_le_def) apply safe apply (rule order_trans[OF _ uint_ge_0]) apply (simp add: order_less_imp_le) apply (erule notE[OF leD]) apply (rule order_less_le_trans[OF _ uint_ge_0]) apply simp done lemma word_sless_msb_less: "x (msb y \ msb x) \ ((msb x \ \ msb y) \ x < y)" by (auto simp add: word_sless_eq word_sle_msb_le) lemma not_msb_from_less: "(v :: 'a word) < 2 ^ (LENGTH('a :: len) - 1) \ \ msb v" apply (clarsimp simp add: msb_nth) apply (drule less_mask_eq) apply (drule word_eqD, drule(1) iffD2) apply simp done lemma sint_eq_uint: "\ msb x \ sint x = uint x" apply (simp add: msb_word_eq) apply transfer apply auto apply (smt One_nat_def bintrunc_bintrunc_l bintrunc_sbintrunc' diff_le_self len_gt_0 signed_take_bit_eq_if_positive) done lemma scast_eq_ucast: "\ msb x \ scast x = ucast x" apply (cases \LENGTH('a)\) apply simp apply (rule bit_word_eqI) apply (auto simp add: bit_signed_iff bit_unsigned_iff min_def msb_word_eq) apply (erule notE) apply (metis le_less_Suc_eq test_bit_bin test_bit_word_eq) done lemma msb_ucast_eq: "LENGTH('a) = LENGTH('b) \ msb (ucast x :: ('a::len) word) = msb (x :: ('b::len) word)" by (simp add: msb_word_eq bit_simps) lemma msb_big: "msb (a :: ('a::len) word) = (a \ 2 ^ (LENGTH('a) - Suc 0))" apply (rule iffI) apply (clarsimp simp: msb_nth) apply (drule bang_is_le) apply simp apply (rule ccontr) apply (subgoal_tac "a = a AND mask (LENGTH('a) - Suc 0)") apply (cut_tac and_mask_less' [where w=a and n="LENGTH('a) - Suc 0"]) apply (clarsimp simp: word_not_le [symmetric]) apply clarsimp apply (rule sym, subst and_mask_eq_iff_shiftr_0) apply (clarsimp simp: msb_shift) done end diff --git a/thys/Word_Lib/Norm_Words.thy b/thys/Word_Lib/Norm_Words.thy --- a/thys/Word_Lib/Norm_Words.thy +++ b/thys/Word_Lib/Norm_Words.thy @@ -1,110 +1,109 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Normalising Word Numerals" theory Norm_Words imports Bits_Int Signed_Words begin text \ Normalise word numerals, including negative ones apart from @{term "-1"}, to the interval \[0..2^len_of 'a)\. Only for concrete word lengths. \ lemma neg_num_bintr: - "(- numeral x :: 'a::len word) = - word_of_int (bintrunc (LENGTH('a)) (-numeral x))" + "(- numeral x :: 'a::len word) = word_of_int (take_bit LENGTH('a) (- numeral x))" by transfer simp ML \ fun is_refl (Const (@{const_name Pure.eq}, _) $ x $ y) = (x = y) | is_refl _ = false; fun signed_dest_wordT (Type (@{type_name word}, [Type (@{type_name signed}, [T])])) = Word_Lib.dest_binT T | signed_dest_wordT T = Word_Lib.dest_wordT T fun typ_size_of t = signed_dest_wordT (type_of (Thm.term_of t)); fun num_len (Const (@{const_name Num.Bit0}, _) $ n) = num_len n + 1 | num_len (Const (@{const_name Num.Bit1}, _) $ n) = num_len n + 1 | num_len (Const (@{const_name Num.One}, _)) = 1 | num_len (Const (@{const_name numeral}, _) $ t) = num_len t | num_len (Const (@{const_name uminus}, _) $ t) = num_len t | num_len t = raise TERM ("num_len", [t]) fun unsigned_norm is_neg _ ctxt ct = (if is_neg orelse num_len (Thm.term_of ct) > typ_size_of ct then let val btr = if is_neg then @{thm neg_num_bintr} else @{thm num_abs_bintr} val th = [Thm.reflexive ct, mk_eq btr] MRS transitive_thm (* will work in context of theory Word as well *) val ss = simpset_of (@{context} addsimps @{thms bintrunc_numeral}) val cnv = simplify (put_simpset ss ctxt) th in if is_refl (Thm.prop_of cnv) then NONE else SOME cnv end else NONE) handle TERM ("num_len", _) => NONE | TYPE ("dest_binT", _, _) => NONE \ simproc_setup unsigned_norm ("numeral n::'a::len word") = \unsigned_norm false\ simproc_setup unsigned_norm_neg0 ("-numeral (num.Bit0 num)::'a::len word") = \unsigned_norm true\ simproc_setup unsigned_norm_neg1 ("-numeral (num.Bit1 num)::'a::len word") = \unsigned_norm true\ lemma minus_one_norm: "(-1 :: 'a :: len word) = of_nat (2 ^ LENGTH('a) - 1)" by (simp add:of_nat_diff) lemmas minus_one_norm_num = minus_one_norm [where 'a="'b::len bit0"] minus_one_norm [where 'a="'b::len0 bit1"] lemma "f (7 :: 2 word) = f 3" by simp lemma "f 7 = f (3 :: 2 word)" by simp lemma "f (-2) = f (21 + 1 :: 3 word)" by simp lemma "f (-2) = f (13 + 1 :: 'a::len word)" apply simp (* does not touch generic word length *) oops lemma "f (-2) = f (0xFFFFFFFE :: 32 word)" by simp lemma "(-1 :: 2 word) = 3" by simp lemma "f (-2) = f (0xFFFFFFFE :: 32 signed word)" by simp text \ We leave @{term "-1"} untouched by default, because it is often useful and its normal form can be large. To include it in the normalisation, add @{thm [source] minus_one_norm_num}. The additional normalisation is restricted to concrete numeral word lengths, like the rest. \ context notes minus_one_norm_num [simp] begin lemma "f (-1) = f (15 :: 4 word)" by simp lemma "f (-1) = f (7 :: 3 word)" by simp lemma "f (-1) = f (0xFFFF :: 16 word)" by simp lemma "f (-1) = f (0xFFFF + 1 :: 'a::len word)" apply simp (* does not touch generic -1 *) oops end end diff --git a/thys/Word_Lib/Reversed_Bit_Lists.thy b/thys/Word_Lib/Reversed_Bit_Lists.thy --- a/thys/Word_Lib/Reversed_Bit_Lists.thy +++ b/thys/Word_Lib/Reversed_Bit_Lists.thy @@ -1,2318 +1,2318 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) (* Author: Jeremy Dawson, NICTA *) section \Bit values as reversed lists of bools\ theory Reversed_Bit_Lists imports "HOL-Library.Word" Typedef_Morphisms Least_significant_bit Most_significant_bit Even_More_List "HOL-Library.Sublist" Aligned begin lemma horner_sum_of_bool_2_concat: \horner_sum of_bool 2 (concat (map (\x. map (bit x) [0.. for ws :: \'a::len word list\ proof (induction ws) case Nil then show ?case by simp next case (Cons w ws) moreover have \horner_sum of_bool 2 (map (bit w) [0.. proof transfer fix k :: int have \map (\n. n < LENGTH('a) \ bit k n) [0.. by simp then show \horner_sum of_bool 2 (map (\n. n < LENGTH('a) \ bit k n) [0.. by (simp only: horner_sum_bit_eq_take_bit) qed ultimately show ?case by (simp add: horner_sum_append) qed subsection \Implicit augmentation of list prefixes\ primrec takefill :: "'a \ nat \ 'a list \ 'a list" where Z: "takefill fill 0 xs = []" | Suc: "takefill fill (Suc n) xs = (case xs of [] \ fill # takefill fill n xs | y # ys \ y # takefill fill n ys)" lemma nth_takefill: "m < n \ takefill fill n l ! m = (if m < length l then l ! m else fill)" apply (induct n arbitrary: m l) apply clarsimp apply clarsimp apply (case_tac m) apply (simp split: list.split) apply (simp split: list.split) done lemma takefill_alt: "takefill fill n l = take n l @ replicate (n - length l) fill" by (induct n arbitrary: l) (auto split: list.split) lemma takefill_replicate [simp]: "takefill fill n (replicate m fill) = replicate n fill" by (simp add: takefill_alt replicate_add [symmetric]) lemma takefill_le': "n = m + k \ takefill x m (takefill x n l) = takefill x m l" by (induct m arbitrary: l n) (auto split: list.split) lemma length_takefill [simp]: "length (takefill fill n l) = n" by (simp add: takefill_alt) lemma take_takefill': "n = k + m \ take k (takefill fill n w) = takefill fill k w" by (induct k arbitrary: w n) (auto split: list.split) lemma drop_takefill: "drop k (takefill fill (m + k) w) = takefill fill m (drop k w)" by (induct k arbitrary: w) (auto split: list.split) lemma takefill_le [simp]: "m \ n \ takefill x m (takefill x n l) = takefill x m l" by (auto simp: le_iff_add takefill_le') lemma take_takefill [simp]: "m \ n \ take m (takefill fill n w) = takefill fill m w" by (auto simp: le_iff_add take_takefill') lemma takefill_append: "takefill fill (m + length xs) (xs @ w) = xs @ (takefill fill m w)" by (induct xs) auto lemma takefill_same': "l = length xs \ takefill fill l xs = xs" by (induct xs arbitrary: l) auto lemmas takefill_same [simp] = takefill_same' [OF refl] lemma tf_rev: "n + k = m + length bl \ takefill x m (rev (takefill y n bl)) = rev (takefill y m (rev (takefill x k (rev bl))))" apply (rule nth_equalityI) apply (auto simp add: nth_takefill rev_nth) apply (rule_tac f = "\n. bl ! n" in arg_cong) apply arith done lemma takefill_minus: "0 < n \ takefill fill (Suc (n - 1)) w = takefill fill n w" by auto lemmas takefill_Suc_cases = list.cases [THEN takefill.Suc [THEN trans]] lemmas takefill_Suc_Nil = takefill_Suc_cases (1) lemmas takefill_Suc_Cons = takefill_Suc_cases (2) lemmas takefill_minus_simps = takefill_Suc_cases [THEN [2] takefill_minus [symmetric, THEN trans]] lemma takefill_numeral_Nil [simp]: "takefill fill (numeral k) [] = fill # takefill fill (pred_numeral k) []" by (simp add: numeral_eq_Suc) lemma takefill_numeral_Cons [simp]: "takefill fill (numeral k) (x # xs) = x # takefill fill (pred_numeral k) xs" by (simp add: numeral_eq_Suc) subsection \Range projection\ definition bl_of_nth :: "nat \ (nat \ 'a) \ 'a list" where "bl_of_nth n f = map f (rev [0.. rev (bl_of_nth n f) ! m = f m" by (simp add: bl_of_nth_def rev_map) lemma bl_of_nth_inj: "(\k. k < n \ f k = g k) \ bl_of_nth n f = bl_of_nth n g" by (simp add: bl_of_nth_def) lemma bl_of_nth_nth_le: "n \ length xs \ bl_of_nth n (nth (rev xs)) = drop (length xs - n) xs" apply (induct n arbitrary: xs) apply clarsimp apply clarsimp apply (rule trans [OF _ hd_Cons_tl]) apply (frule Suc_le_lessD) apply (simp add: rev_nth trans [OF drop_Suc drop_tl, symmetric]) apply (subst hd_drop_conv_nth) apply force apply simp_all apply (rule_tac f = "\n. drop n xs" in arg_cong) apply simp done lemma bl_of_nth_nth [simp]: "bl_of_nth (length xs) ((!) (rev xs)) = xs" by (simp add: bl_of_nth_nth_le) subsection \More\ definition rotater1 :: "'a list \ 'a list" where "rotater1 ys = (case ys of [] \ [] | x # xs \ last ys # butlast ys)" definition rotater :: "nat \ 'a list \ 'a list" where "rotater n = rotater1 ^^ n" lemmas rotater_0' [simp] = rotater_def [where n = "0", simplified] lemma rotate1_rl': "rotater1 (l @ [a]) = a # l" by (cases l) (auto simp: rotater1_def) lemma rotate1_rl [simp] : "rotater1 (rotate1 l) = l" apply (unfold rotater1_def) apply (cases "l") apply (case_tac [2] "list") apply auto done lemma rotate1_lr [simp] : "rotate1 (rotater1 l) = l" by (cases l) (auto simp: rotater1_def) lemma rotater1_rev': "rotater1 (rev xs) = rev (rotate1 xs)" by (cases "xs") (simp add: rotater1_def, simp add: rotate1_rl') lemma rotater_rev': "rotater n (rev xs) = rev (rotate n xs)" by (induct n) (auto simp: rotater_def intro: rotater1_rev') lemma rotater_rev: "rotater n ys = rev (rotate n (rev ys))" using rotater_rev' [where xs = "rev ys"] by simp lemma rotater_drop_take: "rotater n xs = drop (length xs - n mod length xs) xs @ take (length xs - n mod length xs) xs" by (auto simp: rotater_rev rotate_drop_take rev_take rev_drop) lemma rotater_Suc [simp]: "rotater (Suc n) xs = rotater1 (rotater n xs)" unfolding rotater_def by auto lemma nth_rotater: \rotater m xs ! n = xs ! ((n + (length xs - m mod length xs)) mod length xs)\ if \n < length xs\ using that by (simp add: rotater_drop_take nth_append not_less less_diff_conv ac_simps le_mod_geq) lemma nth_rotater1: \rotater1 xs ! n = xs ! ((n + (length xs - 1)) mod length xs)\ if \n < length xs\ using that nth_rotater [of n xs 1] by simp lemma rotate_inv_plus [rule_format]: "\k. k = m + n \ rotater k (rotate n xs) = rotater m xs \ rotate k (rotater n xs) = rotate m xs \ rotater n (rotate k xs) = rotate m xs \ rotate n (rotater k xs) = rotater m xs" by (induct n) (auto simp: rotater_def rotate_def intro: funpow_swap1 [THEN trans]) lemmas rotate_inv_rel = le_add_diff_inverse2 [symmetric, THEN rotate_inv_plus] lemmas rotate_inv_eq = order_refl [THEN rotate_inv_rel, simplified] lemmas rotate_lr [simp] = rotate_inv_eq [THEN conjunct1] lemmas rotate_rl [simp] = rotate_inv_eq [THEN conjunct2, THEN conjunct1] lemma rotate_gal: "rotater n xs = ys \ rotate n ys = xs" by auto lemma rotate_gal': "ys = rotater n xs \ xs = rotate n ys" by auto lemma length_rotater [simp]: "length (rotater n xs) = length xs" by (simp add : rotater_rev) lemma rotate_eq_mod: "m mod length xs = n mod length xs \ rotate m xs = rotate n xs" apply (rule box_equals) defer apply (rule rotate_conv_mod [symmetric])+ apply simp done lemma restrict_to_left: "x = y \ x = z \ y = z" by simp lemmas rotate_eqs = trans [OF rotate0 [THEN fun_cong] id_apply] rotate_rotate [symmetric] rotate_id rotate_conv_mod rotate_eq_mod lemmas rrs0 = rotate_eqs [THEN restrict_to_left, simplified rotate_gal [symmetric] rotate_gal' [symmetric]] lemmas rrs1 = rrs0 [THEN refl [THEN rev_iffD1]] lemmas rotater_eqs = rrs1 [simplified length_rotater] lemmas rotater_0 = rotater_eqs (1) lemmas rotater_add = rotater_eqs (2) lemma butlast_map: "xs \ [] \ butlast (map f xs) = map f (butlast xs)" by (induct xs) auto lemma rotater1_map: "rotater1 (map f xs) = map f (rotater1 xs)" by (cases xs) (auto simp: rotater1_def last_map butlast_map) lemma rotater_map: "rotater n (map f xs) = map f (rotater n xs)" by (induct n) (auto simp: rotater_def rotater1_map) lemma but_last_zip [rule_format] : "\ys. length xs = length ys \ xs \ [] \ last (zip xs ys) = (last xs, last ys) \ butlast (zip xs ys) = zip (butlast xs) (butlast ys)" apply (induct xs) apply auto apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+ done lemma but_last_map2 [rule_format] : "\ys. length xs = length ys \ xs \ [] \ last (map2 f xs ys) = f (last xs) (last ys) \ butlast (map2 f xs ys) = map2 f (butlast xs) (butlast ys)" apply (induct xs) apply auto apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+ done lemma rotater1_zip: "length xs = length ys \ rotater1 (zip xs ys) = zip (rotater1 xs) (rotater1 ys)" apply (unfold rotater1_def) apply (cases xs) apply auto apply ((case_tac ys, auto simp: neq_Nil_conv but_last_zip)[1])+ done lemma rotater1_map2: "length xs = length ys \ rotater1 (map2 f xs ys) = map2 f (rotater1 xs) (rotater1 ys)" by (simp add: rotater1_map rotater1_zip) lemmas lrth = box_equals [OF asm_rl length_rotater [symmetric] length_rotater [symmetric], THEN rotater1_map2] lemma rotater_map2: "length xs = length ys \ rotater n (map2 f xs ys) = map2 f (rotater n xs) (rotater n ys)" by (induct n) (auto intro!: lrth) lemma rotate1_map2: "length xs = length ys \ rotate1 (map2 f xs ys) = map2 f (rotate1 xs) (rotate1 ys)" by (cases xs; cases ys) auto lemmas lth = box_equals [OF asm_rl length_rotate [symmetric] length_rotate [symmetric], THEN rotate1_map2] lemma rotate_map2: "length xs = length ys \ rotate n (map2 f xs ys) = map2 f (rotate n xs) (rotate n ys)" by (induct n) (auto intro!: lth) subsection \Explicit bit representation of \<^typ>\int\\ primrec bl_to_bin_aux :: "bool list \ int \ int" where Nil: "bl_to_bin_aux [] w = w" | Cons: "bl_to_bin_aux (b # bs) w = bl_to_bin_aux bs (of_bool b + 2 * w)" definition bl_to_bin :: "bool list \ int" where "bl_to_bin bs = bl_to_bin_aux bs 0" primrec bin_to_bl_aux :: "nat \ int \ bool list \ bool list" where Z: "bin_to_bl_aux 0 w bl = bl" - | Suc: "bin_to_bl_aux (Suc n) w bl = bin_to_bl_aux n (bin_rest w) ((bin_last w) # bl)" + | Suc: "bin_to_bl_aux (Suc n) w bl = bin_to_bl_aux n (w div 2) (odd w # bl)" definition bin_to_bl :: "nat \ int \ bool list" where "bin_to_bl n w = bin_to_bl_aux n w []" lemma bin_to_bl_aux_zero_minus_simp [simp]: "0 < n \ bin_to_bl_aux n 0 bl = bin_to_bl_aux (n - 1) 0 (False # bl)" by (cases n) auto lemma bin_to_bl_aux_minus1_minus_simp [simp]: "0 < n \ bin_to_bl_aux n (- 1) bl = bin_to_bl_aux (n - 1) (- 1) (True # bl)" by (cases n) auto lemma bin_to_bl_aux_one_minus_simp [simp]: "0 < n \ bin_to_bl_aux n 1 bl = bin_to_bl_aux (n - 1) 0 (True # bl)" by (cases n) auto lemma bin_to_bl_aux_Bit0_minus_simp [simp]: "0 < n \ bin_to_bl_aux n (numeral (Num.Bit0 w)) bl = bin_to_bl_aux (n - 1) (numeral w) (False # bl)" by (cases n) simp_all lemma bin_to_bl_aux_Bit1_minus_simp [simp]: "0 < n \ bin_to_bl_aux n (numeral (Num.Bit1 w)) bl = bin_to_bl_aux (n - 1) (numeral w) (True # bl)" by (cases n) simp_all lemma bl_to_bin_aux_append: "bl_to_bin_aux (bs @ cs) w = bl_to_bin_aux cs (bl_to_bin_aux bs w)" by (induct bs arbitrary: w) auto lemma bin_to_bl_aux_append: "bin_to_bl_aux n w bs @ cs = bin_to_bl_aux n w (bs @ cs)" by (induct n arbitrary: w bs) auto lemma bl_to_bin_append: "bl_to_bin (bs @ cs) = bl_to_bin_aux cs (bl_to_bin bs)" unfolding bl_to_bin_def by (rule bl_to_bin_aux_append) lemma bin_to_bl_aux_alt: "bin_to_bl_aux n w bs = bin_to_bl n w @ bs" by (simp add: bin_to_bl_def bin_to_bl_aux_append) lemma bin_to_bl_0 [simp]: "bin_to_bl 0 bs = []" by (auto simp: bin_to_bl_def) lemma size_bin_to_bl_aux: "length (bin_to_bl_aux n w bs) = n + length bs" by (induct n arbitrary: w bs) auto lemma size_bin_to_bl [simp]: "length (bin_to_bl n w) = n" by (simp add: bin_to_bl_def size_bin_to_bl_aux) lemma bl_bin_bl': "bin_to_bl (n + length bs) (bl_to_bin_aux bs w) = bin_to_bl_aux n w bs" apply (induct bs arbitrary: w n) apply auto apply (simp_all only: add_Suc [symmetric]) apply (auto simp add: bin_to_bl_def) done lemma bl_bin_bl [simp]: "bin_to_bl (length bs) (bl_to_bin bs) = bs" unfolding bl_to_bin_def apply (rule box_equals) apply (rule bl_bin_bl') prefer 2 apply (rule bin_to_bl_aux.Z) apply simp done lemma bl_to_bin_inj: "bl_to_bin bs = bl_to_bin cs \ length bs = length cs \ bs = cs" apply (rule_tac box_equals) defer apply (rule bl_bin_bl) apply (rule bl_bin_bl) apply simp done lemma bl_to_bin_False [simp]: "bl_to_bin (False # bl) = bl_to_bin bl" by (auto simp: bl_to_bin_def) lemma bl_to_bin_Nil [simp]: "bl_to_bin [] = 0" by (auto simp: bl_to_bin_def) lemma bin_to_bl_zero_aux: "bin_to_bl_aux n 0 bl = replicate n False @ bl" by (induct n arbitrary: bl) (auto simp: replicate_app_Cons_same) lemma bin_to_bl_zero: "bin_to_bl n 0 = replicate n False" by (simp add: bin_to_bl_def bin_to_bl_zero_aux) lemma bin_to_bl_minus1_aux: "bin_to_bl_aux n (- 1) bl = replicate n True @ bl" by (induct n arbitrary: bl) (auto simp: replicate_app_Cons_same) lemma bin_to_bl_minus1: "bin_to_bl n (- 1) = replicate n True" by (simp add: bin_to_bl_def bin_to_bl_minus1_aux) subsection \Semantic interpretation of \<^typ>\bool list\ as \<^typ>\int\\ -lemma bin_bl_bin': "bl_to_bin (bin_to_bl_aux n w bs) = bl_to_bin_aux bs (bintrunc n w)" +lemma bin_bl_bin': "bl_to_bin (bin_to_bl_aux n w bs) = bl_to_bin_aux bs (take_bit n w)" by (induct n arbitrary: w bs) (auto simp: bl_to_bin_def take_bit_Suc ac_simps mod_2_eq_odd) -lemma bin_bl_bin [simp]: "bl_to_bin (bin_to_bl n w) = bintrunc n w" +lemma bin_bl_bin [simp]: "bl_to_bin (bin_to_bl n w) = take_bit n w" by (auto simp: bin_to_bl_def bin_bl_bin') lemma bl_to_bin_rep_F: "bl_to_bin (replicate n False @ bl) = bl_to_bin bl" by (simp add: bin_to_bl_zero_aux [symmetric] bin_bl_bin') (simp add: bl_to_bin_def) -lemma bin_to_bl_trunc [simp]: "n \ m \ bin_to_bl n (bintrunc m w) = bin_to_bl n w" +lemma bin_to_bl_trunc [simp]: "n \ m \ bin_to_bl n (take_bit m w) = bin_to_bl n w" by (auto intro: bl_to_bin_inj) lemma bin_to_bl_aux_bintr: - "bin_to_bl_aux n (bintrunc m bin) bl = + "bin_to_bl_aux n (take_bit m bin) bl = replicate (n - m) False @ bin_to_bl_aux (min n m) bin bl" apply (induct n arbitrary: m bin bl) apply clarsimp apply clarsimp apply (case_tac "m") apply (clarsimp simp: bin_to_bl_zero_aux) apply (erule thin_rl) apply (induct_tac n) apply (auto simp add: take_bit_Suc) done lemma bin_to_bl_bintr: - "bin_to_bl n (bintrunc m bin) = replicate (n - m) False @ bin_to_bl (min n m) bin" + "bin_to_bl n (take_bit m bin) = replicate (n - m) False @ bin_to_bl (min n m) bin" unfolding bin_to_bl_def by (rule bin_to_bl_aux_bintr) lemma bl_to_bin_rep_False: "bl_to_bin (replicate n False) = 0" by (induct n) auto lemma len_bin_to_bl_aux: "length (bin_to_bl_aux n w bs) = n + length bs" by (fact size_bin_to_bl_aux) lemma len_bin_to_bl: "length (bin_to_bl n w) = n" by (fact size_bin_to_bl) (* FIXME: duplicate *) lemma sign_bl_bin': "bin_sign (bl_to_bin_aux bs w) = bin_sign w" by (induction bs arbitrary: w) (simp_all add: bin_sign_def) lemma sign_bl_bin: "bin_sign (bl_to_bin bs) = 0" by (simp add: bl_to_bin_def sign_bl_bin') -lemma bl_sbin_sign_aux: "hd (bin_to_bl_aux (Suc n) w bs) = (bin_sign (sbintrunc n w) = -1)" +lemma bl_sbin_sign_aux: "hd (bin_to_bl_aux (Suc n) w bs) = (bin_sign (signed_take_bit n w) = -1)" by (induction n arbitrary: w bs) (auto simp add: bin_sign_def even_iff_mod_2_eq_zero bit_Suc) -lemma bl_sbin_sign: "hd (bin_to_bl (Suc n) w) = (bin_sign (sbintrunc n w) = -1)" +lemma bl_sbin_sign: "hd (bin_to_bl (Suc n) w) = (bin_sign (signed_take_bit n w) = -1)" unfolding bin_to_bl_def by (rule bl_sbin_sign_aux) lemma bin_nth_of_bl_aux: - "bin_nth (bl_to_bin_aux bl w) n = - (n < size bl \ rev bl ! n \ n \ length bl \ bin_nth w (n - size bl))" + "bit (bl_to_bin_aux bl w) n = + (n < size bl \ rev bl ! n \ n \ length bl \ bit w (n - size bl))" apply (induction bl arbitrary: w) apply simp_all apply safe apply (simp_all add: not_le nth_append bit_double_iff even_bit_succ_iff split: if_splits) done -lemma bin_nth_of_bl: "bin_nth (bl_to_bin bl) n = (n < length bl \ rev bl ! n)" +lemma bin_nth_of_bl: "bit (bl_to_bin bl) n = (n < length bl \ rev bl ! n)" by (simp add: bl_to_bin_def bin_nth_of_bl_aux) -lemma bin_nth_bl: "n < m \ bin_nth w n = nth (rev (bin_to_bl m w)) n" +lemma bin_nth_bl: "n < m \ bit w n = nth (rev (bin_to_bl m w)) n" apply (induct n arbitrary: m w) apply clarsimp apply (case_tac m, clarsimp) apply (clarsimp simp: bin_to_bl_def) apply (simp add: bin_to_bl_aux_alt) apply (case_tac m, clarsimp) apply (clarsimp simp: bin_to_bl_def) apply (simp add: bin_to_bl_aux_alt bit_Suc) done lemma nth_bin_to_bl_aux: "n < m + length bl \ (bin_to_bl_aux m w bl) ! n = (if n < m then bit w (m - 1 - n) else bl ! (n - m))" apply (induction bl arbitrary: w) apply simp_all apply (simp add: bin_nth_bl [of \m - Suc n\ m] rev_nth flip: bin_to_bl_def) apply (metis One_nat_def Suc_pred add_diff_cancel_left' add_diff_cancel_right' bin_to_bl_aux_alt bin_to_bl_def diff_Suc_Suc diff_is_0_eq diff_zero less_Suc_eq_0_disj less_antisym less_imp_Suc_add list.size(3) nat_less_le nth_append size_bin_to_bl_aux) done -lemma nth_bin_to_bl: "n < m \ (bin_to_bl m w) ! n = bin_nth w (m - Suc n)" +lemma nth_bin_to_bl: "n < m \ (bin_to_bl m w) ! n = bit w (m - Suc n)" by (simp add: bin_to_bl_def nth_bin_to_bl_aux) lemma takefill_bintrunc: "takefill False n bl = rev (bin_to_bl n (bl_to_bin (rev bl)))" apply (rule nth_equalityI) apply simp apply (clarsimp simp: nth_takefill rev_nth nth_bin_to_bl bin_nth_of_bl) done lemma bl_bin_bl_rtf: "bin_to_bl n (bl_to_bin bl) = rev (takefill False n (rev bl))" by (simp add: takefill_bintrunc) lemma bl_to_bin_lt2p_aux: "bl_to_bin_aux bs w < (w + 1) * (2 ^ length bs)" proof (induction bs arbitrary: w) case Nil then show ?case by simp next case (Cons b bs) from Cons.IH [of \1 + 2 * w\] Cons.IH [of \2 * w\] show ?case apply (auto simp add: algebra_simps) apply (subst mult_2 [of \2 ^ length bs\]) apply (simp only: add.assoc) apply (rule pos_add_strict) apply simp_all done qed lemma bl_to_bin_lt2p_drop: "bl_to_bin bs < 2 ^ length (dropWhile Not bs)" proof (induct bs) case Nil then show ?case by simp next case (Cons b bs) with bl_to_bin_lt2p_aux[where w=1] show ?case by (simp add: bl_to_bin_def) qed lemma bl_to_bin_lt2p: "bl_to_bin bs < 2 ^ length bs" by (metis bin_bl_bin bintr_lt2p bl_bin_bl) lemma bl_to_bin_ge2p_aux: "bl_to_bin_aux bs w \ w * (2 ^ length bs)" proof (induction bs arbitrary: w) case Nil then show ?case by simp next case (Cons b bs) from Cons.IH [of \1 + 2 * w\] Cons.IH [of \2 * w\] show ?case apply (auto simp add: algebra_simps) apply (rule add_le_imp_le_left [of \2 ^ length bs\]) apply (rule add_increasing) apply simp_all done qed lemma bl_to_bin_ge0: "bl_to_bin bs \ 0" apply (unfold bl_to_bin_def) apply (rule xtrans(4)) apply (rule bl_to_bin_ge2p_aux) apply simp done -lemma butlast_rest_bin: "butlast (bin_to_bl n w) = bin_to_bl (n - 1) (bin_rest w)" +lemma butlast_rest_bin: "butlast (bin_to_bl n w) = bin_to_bl (n - 1) (w div 2)" apply (unfold bin_to_bl_def) apply (cases n, clarsimp) apply clarsimp apply (auto simp add: bin_to_bl_aux_alt) done -lemma butlast_bin_rest: "butlast bl = bin_to_bl (length bl - Suc 0) (bin_rest (bl_to_bin bl))" +lemma butlast_bin_rest: "butlast bl = bin_to_bl (length bl - Suc 0) (bl_to_bin bl div 2)" using butlast_rest_bin [where w="bl_to_bin bl" and n="length bl"] by simp lemma butlast_rest_bl2bin_aux: - "bl \ [] \ bl_to_bin_aux (butlast bl) w = bin_rest (bl_to_bin_aux bl w)" + "bl \ [] \ bl_to_bin_aux (butlast bl) w = bl_to_bin_aux bl w div 2" by (induct bl arbitrary: w) auto -lemma butlast_rest_bl2bin: "bl_to_bin (butlast bl) = bin_rest (bl_to_bin bl)" +lemma butlast_rest_bl2bin: "bl_to_bin (butlast bl) = bl_to_bin bl div 2" by (cases bl) (auto simp: bl_to_bin_def butlast_rest_bl2bin_aux) lemma trunc_bl2bin_aux: - "bintrunc m (bl_to_bin_aux bl w) = - bl_to_bin_aux (drop (length bl - m) bl) (bintrunc (m - length bl) w)" + "take_bit m (bl_to_bin_aux bl w) = + bl_to_bin_aux (drop (length bl - m) bl) (take_bit (m - length bl) w)" proof (induct bl arbitrary: w) case Nil show ?case by simp next case (Cons b bl) show ?case proof (cases "m - length bl") case 0 then have "Suc (length bl) - m = Suc (length bl - m)" by simp with Cons show ?thesis by simp next case (Suc n) then have "m - Suc (length bl) = n" by simp with Cons Suc show ?thesis by (simp add: take_bit_Suc ac_simps) qed qed -lemma trunc_bl2bin: "bintrunc m (bl_to_bin bl) = bl_to_bin (drop (length bl - m) bl)" +lemma trunc_bl2bin: "take_bit m (bl_to_bin bl) = bl_to_bin (drop (length bl - m) bl)" by (simp add: bl_to_bin_def trunc_bl2bin_aux) -lemma trunc_bl2bin_len [simp]: "bintrunc (length bl) (bl_to_bin bl) = bl_to_bin bl" +lemma trunc_bl2bin_len [simp]: "take_bit (length bl) (bl_to_bin bl) = bl_to_bin bl" by (simp add: trunc_bl2bin) -lemma bl2bin_drop: "bl_to_bin (drop k bl) = bintrunc (length bl - k) (bl_to_bin bl)" +lemma bl2bin_drop: "bl_to_bin (drop k bl) = take_bit (length bl - k) (bl_to_bin bl)" apply (rule trans) prefer 2 apply (rule trunc_bl2bin [symmetric]) apply (cases "k \ length bl") apply auto done -lemma take_rest_power_bin: "m \ n \ take m (bin_to_bl n w) = bin_to_bl m ((bin_rest ^^ (n - m)) w)" +lemma take_rest_power_bin: "m \ n \ take m (bin_to_bl n w) = bin_to_bl m (((\w. w div 2) ^^ (n - m)) w)" apply (rule nth_equalityI) apply simp apply (clarsimp simp add: nth_bin_to_bl nth_rest_power_bin) done -lemma last_bin_last': "size xs > 0 \ last xs \ bin_last (bl_to_bin_aux xs w)" +lemma last_bin_last': "size xs > 0 \ last xs \ odd (bl_to_bin_aux xs w)" by (induct xs arbitrary: w) auto -lemma last_bin_last: "size xs > 0 \ last xs \ bin_last (bl_to_bin xs)" +lemma last_bin_last: "size xs > 0 \ last xs \ odd (bl_to_bin xs)" unfolding bl_to_bin_def by (erule last_bin_last') -lemma bin_last_last: "bin_last w \ last (bin_to_bl (Suc n) w)" +lemma bin_last_last: "odd w \ last (bin_to_bl (Suc n) w)" by (simp add: bin_to_bl_def) (auto simp: bin_to_bl_aux_alt) lemma drop_bin2bl_aux: "drop m (bin_to_bl_aux n bin bs) = bin_to_bl_aux (n - m) bin (drop (m - n) bs)" apply (induction n arbitrary: m bin bs) apply auto apply (case_tac "m \ n") apply (auto simp add: not_le Suc_diff_le) apply (case_tac "m - n") apply auto apply (use Suc_diff_Suc in fastforce) done lemma drop_bin2bl: "drop m (bin_to_bl n bin) = bin_to_bl (n - m) bin" by (simp add: bin_to_bl_def drop_bin2bl_aux) lemma take_bin2bl_lem1: "take m (bin_to_bl_aux m w bs) = bin_to_bl m w" apply (induct m arbitrary: w bs) apply clarsimp apply clarsimp apply (simp add: bin_to_bl_aux_alt) apply (simp add: bin_to_bl_def) apply (simp add: bin_to_bl_aux_alt) done lemma take_bin2bl_lem: "take m (bin_to_bl_aux (m + n) w bs) = take m (bin_to_bl (m + n) w)" by (induct n arbitrary: w bs) (simp_all (no_asm) add: bin_to_bl_def take_bin2bl_lem1, simp) lemma bin_split_take: "bin_split n c = (a, b) \ bin_to_bl m a = take m (bin_to_bl (m + n) c)" apply (induct n arbitrary: b c) apply clarsimp apply (clarsimp simp: Let_def split: prod.split_asm) apply (simp add: bin_to_bl_def) apply (simp add: take_bin2bl_lem drop_bit_Suc) done lemma bin_to_bl_drop_bit: "k = m + n \ bin_to_bl m (drop_bit n c) = take m (bin_to_bl k c)" using bin_split_take by simp lemma bin_split_take1: "k = m + n \ bin_split n c = (a, b) \ bin_to_bl m a = take m (bin_to_bl k c)" using bin_split_take by simp lemma bl_bin_bl_rep_drop: "bin_to_bl n (bl_to_bin bl) = replicate (n - length bl) False @ drop (length bl - n) bl" by (simp add: bl_to_bin_inj bl_to_bin_rep_F trunc_bl2bin) lemma bl_to_bin_aux_cat: - "bl_to_bin_aux bs (bin_cat w nv v) = - bin_cat w (nv + length bs) (bl_to_bin_aux bs v)" + "bl_to_bin_aux bs (concat_bit nv v w) = + concat_bit (nv + length bs) (bl_to_bin_aux bs v) w" by (rule bit_eqI) (auto simp add: bin_nth_of_bl_aux bin_nth_cat algebra_simps) lemma bin_to_bl_aux_cat: - "bin_to_bl_aux (nv + nw) (bin_cat v nw w) bs = + "bin_to_bl_aux (nv + nw) (concat_bit nw w v) bs = bin_to_bl_aux nv v (bin_to_bl_aux nw w bs)" by (induction nw arbitrary: w bs) (simp_all add: concat_bit_Suc) -lemma bl_to_bin_aux_alt: "bl_to_bin_aux bs w = bin_cat w (length bs) (bl_to_bin bs)" +lemma bl_to_bin_aux_alt: "bl_to_bin_aux bs w = concat_bit (length bs) (bl_to_bin bs) w" using bl_to_bin_aux_cat [where nv = "0" and v = "0"] by (simp add: bl_to_bin_def [symmetric]) lemma bin_to_bl_cat: - "bin_to_bl (nv + nw) (bin_cat v nw w) = + "bin_to_bl (nv + nw) (concat_bit nw w v) = bin_to_bl_aux nv v (bin_to_bl nw w)" by (simp add: bin_to_bl_def bin_to_bl_aux_cat) lemmas bl_to_bin_aux_app_cat = trans [OF bl_to_bin_aux_append bl_to_bin_aux_alt] lemmas bin_to_bl_aux_cat_app = trans [OF bin_to_bl_aux_cat bin_to_bl_aux_alt] lemma bl_to_bin_app_cat: - "bl_to_bin (bsa @ bs) = bin_cat (bl_to_bin bsa) (length bs) (bl_to_bin bs)" + "bl_to_bin (bsa @ bs) = concat_bit (length bs) (bl_to_bin bs) (bl_to_bin bsa)" by (simp only: bl_to_bin_aux_app_cat bl_to_bin_def) lemma bin_to_bl_cat_app: - "bin_to_bl (n + nw) (bin_cat w nw wa) = bin_to_bl n w @ bin_to_bl nw wa" + "bin_to_bl (n + nw) (concat_bit nw wa w) = bin_to_bl n w @ bin_to_bl nw wa" by (simp only: bin_to_bl_def bin_to_bl_aux_cat_app) text \\bl_to_bin_app_cat_alt\ and \bl_to_bin_app_cat\ are easily interderivable.\ -lemma bl_to_bin_app_cat_alt: "bin_cat (bl_to_bin cs) n w = bl_to_bin (cs @ bin_to_bl n w)" +lemma bl_to_bin_app_cat_alt: "concat_bit n w (bl_to_bin cs) = bl_to_bin (cs @ bin_to_bl n w)" by (simp add: bl_to_bin_app_cat) lemma mask_lem: "(bl_to_bin (True # replicate n False)) = bl_to_bin (replicate n True) + 1" apply (unfold bl_to_bin_def) apply (induct n) apply simp apply (simp only: Suc_eq_plus1 replicate_add append_Cons [symmetric] bl_to_bin_aux_append) apply simp done lemma bin_exhaust: "(\x b. bin = of_bool b + 2 * x \ Q) \ Q" for bin :: int apply (cases \even bin\) apply (auto elim!: evenE oddE) apply fastforce apply fastforce done primrec rbl_succ :: "bool list \ bool list" where Nil: "rbl_succ Nil = Nil" | Cons: "rbl_succ (x # xs) = (if x then False # rbl_succ xs else True # xs)" primrec rbl_pred :: "bool list \ bool list" where Nil: "rbl_pred Nil = Nil" | Cons: "rbl_pred (x # xs) = (if x then False # xs else True # rbl_pred xs)" primrec rbl_add :: "bool list \ bool list \ bool list" where \ \result is length of first arg, second arg may be longer\ Nil: "rbl_add Nil x = Nil" | Cons: "rbl_add (y # ys) x = (let ws = rbl_add ys (tl x) in (y \ hd x) # (if hd x \ y then rbl_succ ws else ws))" primrec rbl_mult :: "bool list \ bool list \ bool list" where \ \result is length of first arg, second arg may be longer\ Nil: "rbl_mult Nil x = Nil" | Cons: "rbl_mult (y # ys) x = (let ws = False # rbl_mult ys x in if y then rbl_add ws x else ws)" lemma size_rbl_pred: "length (rbl_pred bl) = length bl" by (induct bl) auto lemma size_rbl_succ: "length (rbl_succ bl) = length bl" by (induct bl) auto lemma size_rbl_add: "length (rbl_add bl cl) = length bl" by (induct bl arbitrary: cl) (auto simp: Let_def size_rbl_succ) lemma size_rbl_mult: "length (rbl_mult bl cl) = length bl" by (induct bl arbitrary: cl) (auto simp add: Let_def size_rbl_add) lemmas rbl_sizes [simp] = size_rbl_pred size_rbl_succ size_rbl_add size_rbl_mult lemmas rbl_Nils = rbl_pred.Nil rbl_succ.Nil rbl_add.Nil rbl_mult.Nil lemma rbl_add_app2: "length blb \ length bla \ rbl_add bla (blb @ blc) = rbl_add bla blb" apply (induct bla arbitrary: blb) apply simp apply clarsimp apply (case_tac blb, clarsimp) apply (clarsimp simp: Let_def) done lemma rbl_add_take2: "length blb \ length bla \ rbl_add bla (take (length bla) blb) = rbl_add bla blb" apply (induct bla arbitrary: blb) apply simp apply clarsimp apply (case_tac blb, clarsimp) apply (clarsimp simp: Let_def) done lemma rbl_mult_app2: "length blb \ length bla \ rbl_mult bla (blb @ blc) = rbl_mult bla blb" apply (induct bla arbitrary: blb) apply simp apply clarsimp apply (case_tac blb, clarsimp) apply (clarsimp simp: Let_def rbl_add_app2) done lemma rbl_mult_take2: "length blb \ length bla \ rbl_mult bla (take (length bla) blb) = rbl_mult bla blb" apply (rule trans) apply (rule rbl_mult_app2 [symmetric]) apply simp apply (rule_tac f = "rbl_mult bla" in arg_cong) apply (rule append_take_drop_id) done lemma rbl_add_split: "P (rbl_add (y # ys) (x # xs)) = (\ws. length ws = length ys \ ws = rbl_add ys xs \ (y \ ((x \ P (False # rbl_succ ws)) \ (\ x \ P (True # ws)))) \ (\ y \ P (x # ws)))" by (cases y) (auto simp: Let_def) lemma rbl_mult_split: "P (rbl_mult (y # ys) xs) = (\ws. length ws = Suc (length ys) \ ws = False # rbl_mult ys xs \ (y \ P (rbl_add ws xs)) \ (\ y \ P ws))" by (auto simp: Let_def) lemma rbl_pred: "rbl_pred (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin - 1))" proof (unfold bin_to_bl_def, induction n arbitrary: bin) case 0 then show ?case by simp next case (Suc n) obtain b k where \bin = of_bool b + 2 * k\ using bin_exhaust by blast moreover have \(2 * k - 1) div 2 = k - 1\ using even_succ_div_2 [of \2 * (k - 1)\] by simp ultimately show ?case using Suc [of \bin div 2\] by simp (auto simp add: bin_to_bl_aux_alt) qed lemma rbl_succ: "rbl_succ (rev (bin_to_bl n bin)) = rev (bin_to_bl n (bin + 1))" apply (unfold bin_to_bl_def) apply (induction n arbitrary: bin) apply simp_all apply (case_tac bin rule: bin_exhaust) apply (simp_all add: bin_to_bl_aux_alt ac_simps) done lemma rbl_add: "\bina binb. rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = rev (bin_to_bl n (bina + binb))" apply (unfold bin_to_bl_def) apply (induct n) apply simp apply clarsimp apply (case_tac bina rule: bin_exhaust) apply (case_tac binb rule: bin_exhaust) apply (case_tac b) apply (case_tac [!] "ba") apply (auto simp: rbl_succ bin_to_bl_aux_alt Let_def ac_simps) done lemma rbl_add_long: "m \ n \ rbl_add (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = rev (bin_to_bl n (bina + binb))" apply (rule box_equals [OF _ rbl_add_take2 rbl_add]) apply (rule_tac f = "rbl_add (rev (bin_to_bl n bina))" in arg_cong) apply (rule rev_swap [THEN iffD1]) apply (simp add: rev_take drop_bin2bl) apply simp done lemma rbl_mult_gt1: "m \ length bl \ rbl_mult bl (rev (bin_to_bl m binb)) = rbl_mult bl (rev (bin_to_bl (length bl) binb))" apply (rule trans) apply (rule rbl_mult_take2 [symmetric]) apply simp_all apply (rule_tac f = "rbl_mult bl" in arg_cong) apply (rule rev_swap [THEN iffD1]) apply (simp add: rev_take drop_bin2bl) done lemma rbl_mult_gt: "m > n \ rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl m binb)) = rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb))" by (auto intro: trans [OF rbl_mult_gt1]) lemmas rbl_mult_Suc = lessI [THEN rbl_mult_gt] lemma rbbl_Cons: "b # rev (bin_to_bl n x) = rev (bin_to_bl (Suc n) (of_bool b + 2 * x))" by (simp add: bin_to_bl_def) (simp add: bin_to_bl_aux_alt) lemma rbl_mult: "rbl_mult (rev (bin_to_bl n bina)) (rev (bin_to_bl n binb)) = rev (bin_to_bl n (bina * binb))" apply (induct n arbitrary: bina binb) apply simp_all apply (unfold bin_to_bl_def) apply clarsimp apply (case_tac bina rule: bin_exhaust) apply (case_tac binb rule: bin_exhaust) apply (simp_all add: bin_to_bl_aux_alt) apply (simp_all add: rbbl_Cons rbl_mult_Suc rbl_add algebra_simps) done lemma sclem: "size (concat (map (bin_to_bl n) xs)) = length xs * n" by (simp add: length_concat comp_def sum_list_triv) lemma bin_cat_foldl_lem: - "foldl (\u. bin_cat u n) x xs = - bin_cat x (size xs * n) (foldl (\u. bin_cat u n) y xs)" + "foldl (\u k. concat_bit n k u) x xs = + concat_bit (size xs * n) (foldl (\u k. concat_bit n k u) y xs) x" apply (induct xs arbitrary: x) apply simp apply (simp (no_asm)) apply (frule asm_rl) apply (drule meta_spec) apply (erule trans) - apply (drule_tac x = "bin_cat y n a" in meta_spec) + apply (drule_tac x = "concat_bit n a y" in meta_spec) apply (simp add: bin_cat_assoc_sym) done lemma bin_rcat_bl: "bin_rcat n wl = bl_to_bin (concat (map (bin_to_bl n) wl))" apply (unfold bin_rcat_eq_foldl) apply (rule sym) apply (induct wl) apply (auto simp add: bl_to_bin_append) apply (simp add: bl_to_bin_aux_alt sclem) apply (simp add: bin_cat_foldl_lem [symmetric]) done -lemma bin_last_bl_to_bin: "bin_last (bl_to_bin bs) \ bs \ [] \ last bs" +lemma bin_last_bl_to_bin: "odd (bl_to_bin bs) \ bs \ [] \ last bs" by(cases "bs = []")(auto simp add: bl_to_bin_def last_bin_last'[where w=0]) -lemma bin_rest_bl_to_bin: "bin_rest (bl_to_bin bs) = bl_to_bin (butlast bs)" +lemma bin_rest_bl_to_bin: "bl_to_bin bs div 2 = bl_to_bin (butlast bs)" by(cases "bs = []")(simp_all add: bl_to_bin_def butlast_rest_bl2bin_aux) lemma bl_xor_aux_bin: "map2 (\x y. x \ y) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = bin_to_bl_aux n (v XOR w) (map2 (\x y. x \ y) bs cs)" apply (induction n arbitrary: v w bs cs) apply auto apply (case_tac v rule: bin_exhaust) apply (case_tac w rule: bin_exhaust) apply clarsimp done lemma bl_or_aux_bin: "map2 (\) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = bin_to_bl_aux n (v OR w) (map2 (\) bs cs)" by (induct n arbitrary: v w bs cs) simp_all lemma bl_and_aux_bin: "map2 (\) (bin_to_bl_aux n v bs) (bin_to_bl_aux n w cs) = bin_to_bl_aux n (v AND w) (map2 (\) bs cs)" by (induction n arbitrary: v w bs cs) simp_all lemma bl_not_aux_bin: "map Not (bin_to_bl_aux n w cs) = bin_to_bl_aux n (NOT w) (map Not cs)" by (induct n arbitrary: w cs) auto lemma bl_not_bin: "map Not (bin_to_bl n w) = bin_to_bl n (NOT w)" by (simp add: bin_to_bl_def bl_not_aux_bin) lemma bl_and_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v AND w)" by (simp add: bin_to_bl_def bl_and_aux_bin) lemma bl_or_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v OR w)" by (simp add: bin_to_bl_def bl_or_aux_bin) lemma bl_xor_bin: "map2 (\) (bin_to_bl n v) (bin_to_bl n w) = bin_to_bl n (v XOR w)" using bl_xor_aux_bin by (simp add: bin_to_bl_def) subsection \Type \<^typ>\'a word\\ lift_definition of_bl :: \bool list \ 'a::len word\ is bl_to_bin . lift_definition to_bl :: \'a::len word \ bool list\ is \bin_to_bl LENGTH('a)\ by (simp add: bl_to_bin_inj) lemma to_bl_eq: \to_bl w = bin_to_bl (LENGTH('a)) (uint w)\ for w :: \'a::len word\ by transfer simp lemma bit_of_bl_iff [bit_simps]: \bit (of_bl bs :: 'a word) n \ rev bs ! n \ n < LENGTH('a::len) \ n < length bs\ by transfer (simp add: bin_nth_of_bl ac_simps) lemma rev_to_bl_eq: \rev (to_bl w) = map (bit w) [0.. for w :: \'a::len word\ apply (rule nth_equalityI) apply (simp add: to_bl.rep_eq) apply (simp add: bin_nth_bl bit_word.rep_eq to_bl.rep_eq) done lemma to_bl_eq_rev: \to_bl w = map (bit w) (rev [0.. for w :: \'a::len word\ using rev_to_bl_eq [of w] apply (subst rev_is_rev_conv [symmetric]) apply (simp add: rev_map) done lemma of_bl_rev_eq: \of_bl (rev bs) = horner_sum of_bool 2 bs\ apply (rule bit_word_eqI) apply (simp add: bit_of_bl_iff) apply transfer apply (simp add: bit_horner_sum_bit_iff ac_simps) done lemma of_bl_eq: \of_bl bs = horner_sum of_bool 2 (rev bs)\ using of_bl_rev_eq [of \rev bs\] by simp lemma bshiftr1_eq: \bshiftr1 b w = of_bl (b # butlast (to_bl w))\ apply transfer apply auto apply (subst bl_to_bin_app_cat [of \[True]\, simplified]) apply simp apply (metis One_nat_def add.commute bin_bl_bin bin_last_bl_to_bin bin_rest_bl_to_bin butlast_bin_rest concat_bit_eq last.simps list.distinct(1) list.size(3) list.size(4) odd_iff_mod_2_eq_one plus_1_eq_Suc power_Suc0_right push_bit_of_1 size_bin_to_bl take_bit_eq_mod trunc_bl2bin_len) apply (simp add: butlast_rest_bl2bin) done lemma length_to_bl_eq: \length (to_bl w) = LENGTH('a)\ for w :: \'a::len word\ by transfer simp lemma word_rotr_eq: \word_rotr n w = of_bl (rotater n (to_bl w))\ apply (rule bit_word_eqI) subgoal for n apply (cases \n < LENGTH('a)\) apply (simp_all add: bit_word_rotr_iff bit_of_bl_iff rotater_rev length_to_bl_eq nth_rotate rev_to_bl_eq ac_simps) done done lemma word_rotl_eq: \word_rotl n w = of_bl (rotate n (to_bl w))\ proof - have \rotate n (to_bl w) = rev (rotater n (rev (to_bl w)))\ by (simp add: rotater_rev') then show ?thesis apply (simp add: word_rotl_eq_word_rotr bit_of_bl_iff length_to_bl_eq rev_to_bl_eq) apply (rule bit_word_eqI) subgoal for n apply (cases \n < LENGTH('a)\) apply (simp_all add: bit_word_rotr_iff bit_of_bl_iff nth_rotater) done done qed lemma to_bl_def': "(to_bl :: 'a::len word \ bool list) = bin_to_bl (LENGTH('a)) \ uint" by transfer (simp add: fun_eq_iff) \ \type definitions theorem for in terms of equivalent bool list\ lemma td_bl: "type_definition (to_bl :: 'a::len word \ bool list) of_bl {bl. length bl = LENGTH('a)}" apply (standard; transfer) apply (auto dest: sym) done interpretation word_bl: type_definition "to_bl :: 'a::len word \ bool list" of_bl "{bl. length bl = LENGTH('a::len)}" by (fact td_bl) lemmas word_bl_Rep' = word_bl.Rep [unfolded mem_Collect_eq, iff] lemma word_size_bl: "size w = size (to_bl w)" by (auto simp: word_size) lemma to_bl_use_of_bl: "to_bl w = bl \ w = of_bl bl \ length bl = length (to_bl w)" by (fastforce elim!: word_bl.Abs_inverse [unfolded mem_Collect_eq]) lemma length_bl_gt_0 [iff]: "0 < length (to_bl x)" for x :: "'a::len word" unfolding word_bl_Rep' by (rule len_gt_0) lemma bl_not_Nil [iff]: "to_bl x \ []" for x :: "'a::len word" by (fact length_bl_gt_0 [unfolded length_greater_0_conv]) lemma length_bl_neq_0 [iff]: "length (to_bl x) \ 0" for x :: "'a::len word" by (fact length_bl_gt_0 [THEN gr_implies_not0]) lemma hd_bl_sign_sint: "hd (to_bl w) = (bin_sign (sint w) = -1)" apply transfer apply (auto simp add: bin_sign_def) using bin_sign_lem bl_sbin_sign apply fastforce using bin_sign_lem bl_sbin_sign apply force done lemma of_bl_drop': "lend = length bl - LENGTH('a::len) \ of_bl (drop lend bl) = (of_bl bl :: 'a word)" by transfer (simp flip: trunc_bl2bin) lemma test_bit_of_bl: "(of_bl bl::'a::len word) !! n = (rev bl ! n \ n < LENGTH('a) \ n < length bl)" by transfer (simp add: bin_nth_of_bl ac_simps) lemma no_of_bl: "(numeral bin ::'a::len word) = of_bl (bin_to_bl (LENGTH('a)) (numeral bin))" by transfer simp lemma uint_bl: "to_bl w = bin_to_bl (size w) (uint w)" by transfer simp lemma to_bl_bin: "bl_to_bin (to_bl w) = uint w" by (simp add: uint_bl word_size) lemma to_bl_of_bin: "to_bl (word_of_int bin::'a::len word) = bin_to_bl (LENGTH('a)) bin" by (auto simp: uint_bl word_ubin.eq_norm word_size) lemma to_bl_numeral [simp]: "to_bl (numeral bin::'a::len word) = bin_to_bl (LENGTH('a)) (numeral bin)" unfolding word_numeral_alt by (rule to_bl_of_bin) lemma to_bl_neg_numeral [simp]: "to_bl (- numeral bin::'a::len word) = bin_to_bl (LENGTH('a)) (- numeral bin)" unfolding word_neg_numeral_alt by (rule to_bl_of_bin) lemma to_bl_to_bin [simp] : "bl_to_bin (to_bl w) = uint w" by (simp add: uint_bl word_size) lemma uint_bl_bin: "bl_to_bin (bin_to_bl (LENGTH('a)) (uint x)) = uint x" for x :: "'a::len word" by (rule trans [OF bin_bl_bin word_ubin.norm_Rep]) lemma ucast_bl: "ucast w = of_bl (to_bl w)" by transfer simp lemma ucast_down_bl: \(ucast :: 'a::len word \ 'b::len word) (of_bl bl) = of_bl bl\ if \is_down (ucast :: 'a::len word \ 'b::len word)\ using that by transfer simp lemma of_bl_append_same: "of_bl (X @ to_bl w) = w" by transfer (simp add: bl_to_bin_app_cat) lemma ucast_of_bl_up: \ucast (of_bl bl :: 'a::len word) = of_bl bl\ if \size bl \ size (of_bl bl :: 'a::len word)\ using that apply transfer apply (rule bit_eqI) apply (auto simp add: bit_take_bit_iff) apply (subst (asm) trunc_bl2bin_len [symmetric]) apply (auto simp only: bit_take_bit_iff) done lemma word_rev_tf: "to_bl (of_bl bl::'a::len word) = rev (takefill False (LENGTH('a)) (rev bl))" by transfer (simp add: bl_bin_bl_rtf) lemma word_rep_drop: "to_bl (of_bl bl::'a::len word) = replicate (LENGTH('a) - length bl) False @ drop (length bl - LENGTH('a)) bl" by (simp add: word_rev_tf takefill_alt rev_take) lemma to_bl_ucast: "to_bl (ucast (w::'b::len word) ::'a::len word) = replicate (LENGTH('a) - LENGTH('b)) False @ drop (LENGTH('b) - LENGTH('a)) (to_bl w)" apply (unfold ucast_bl) apply (rule trans) apply (rule word_rep_drop) apply simp done lemma ucast_up_app: \to_bl (ucast w :: 'b::len word) = replicate n False @ (to_bl w)\ if \source_size (ucast :: 'a word \ 'b word) + n = target_size (ucast :: 'a word \ 'b word)\ for w :: \'a::len word\ using that by (auto simp add : source_size target_size to_bl_ucast) lemma ucast_down_drop [OF refl]: "uc = ucast \ source_size uc = target_size uc + n \ to_bl (uc w) = drop n (to_bl w)" by (auto simp add : source_size target_size to_bl_ucast) lemma scast_down_drop [OF refl]: "sc = scast \ source_size sc = target_size sc + n \ to_bl (sc w) = drop n (to_bl w)" apply (subgoal_tac "sc = ucast") apply safe apply simp apply (erule ucast_down_drop) apply (rule down_cast_same [symmetric]) apply (simp add : source_size target_size is_down) done lemma word_0_bl [simp]: "of_bl [] = 0" by transfer simp lemma word_1_bl: "of_bl [True] = 1" by transfer (simp add: bl_to_bin_def) lemma of_bl_0 [simp]: "of_bl (replicate n False) = 0" by transfer (simp add: bl_to_bin_rep_False) lemma to_bl_0 [simp]: "to_bl (0::'a::len word) = replicate (LENGTH('a)) False" by (simp add: uint_bl word_size bin_to_bl_zero) \ \links with \rbl\ operations\ lemma word_succ_rbl: "to_bl w = bl \ to_bl (word_succ w) = rev (rbl_succ (rev bl))" by transfer (simp add: rbl_succ) lemma word_pred_rbl: "to_bl w = bl \ to_bl (word_pred w) = rev (rbl_pred (rev bl))" by transfer (simp add: rbl_pred) lemma word_add_rbl: "to_bl v = vbl \ to_bl w = wbl \ to_bl (v + w) = rev (rbl_add (rev vbl) (rev wbl))" apply transfer apply (drule sym) apply (drule sym) apply (simp add: rbl_add) done lemma word_mult_rbl: "to_bl v = vbl \ to_bl w = wbl \ to_bl (v * w) = rev (rbl_mult (rev vbl) (rev wbl))" apply transfer apply (drule sym) apply (drule sym) apply (simp add: rbl_mult) done lemma rtb_rbl_ariths: "rev (to_bl w) = ys \ rev (to_bl (word_succ w)) = rbl_succ ys" "rev (to_bl w) = ys \ rev (to_bl (word_pred w)) = rbl_pred ys" "rev (to_bl v) = ys \ rev (to_bl w) = xs \ rev (to_bl (v * w)) = rbl_mult ys xs" "rev (to_bl v) = ys \ rev (to_bl w) = xs \ rev (to_bl (v + w)) = rbl_add ys xs" by (auto simp: rev_swap [symmetric] word_succ_rbl word_pred_rbl word_mult_rbl word_add_rbl) lemma of_bl_length_less: \(of_bl x :: 'a::len word) < 2 ^ k\ if \length x = k\ \k < LENGTH('a)\ proof - from that have \length x < LENGTH('a)\ by simp then have \(of_bl x :: 'a::len word) < 2 ^ length x\ apply (simp add: of_bl_eq) apply transfer apply (simp add: take_bit_horner_sum_bit_eq) apply (subst length_rev [symmetric]) apply (simp only: horner_sum_of_bool_2_less) done with that show ?thesis by simp qed lemma word_eq_rbl_eq: "x = y \ rev (to_bl x) = rev (to_bl y)" by simp lemma bl_word_not: "to_bl (NOT w) = map Not (to_bl w)" by transfer (simp add: bl_not_bin) lemma bl_word_xor: "to_bl (v XOR w) = map2 (\) (to_bl v) (to_bl w)" by transfer (simp flip: bl_xor_bin) lemma bl_word_or: "to_bl (v OR w) = map2 (\) (to_bl v) (to_bl w)" by transfer (simp flip: bl_or_bin) lemma bl_word_and: "to_bl (v AND w) = map2 (\) (to_bl v) (to_bl w)" by transfer (simp flip: bl_and_bin) -lemma bin_nth_uint': "bin_nth (uint w) n \ rev (bin_to_bl (size w) (uint w)) ! n \ n < size w" +lemma bin_nth_uint': "bit (uint w) n \ rev (bin_to_bl (size w) (uint w)) ! n \ n < size w" apply (unfold word_size) apply (safe elim!: bin_nth_uint_imp) apply (frule bin_nth_uint_imp) apply (fast dest!: bin_nth_bl)+ done lemmas bin_nth_uint = bin_nth_uint' [unfolded word_size] lemma test_bit_bl: "w !! n \ rev (to_bl w) ! n \ n < size w" by transfer (auto simp add: bin_nth_bl) lemma to_bl_nth: "n < size w \ to_bl w ! n = w !! (size w - Suc n)" by (simp add: word_size rev_nth test_bit_bl) lemma map_bit_interval_eq: \map (bit w) [0.. for w :: \'a::len word\ proof (rule nth_equalityI) show \length (map (bit w) [0.. by simp fix m assume \m < length (map (bit w) [0.. then have \m < n\ by simp then have \bit w m \ takefill False n (rev (to_bl w)) ! m\ by (auto simp add: nth_takefill not_less rev_nth to_bl_nth word_size test_bit_word_eq dest: bit_imp_le_length) with \m < n \show \map (bit w) [0.. takefill False n (rev (to_bl w)) ! m\ by simp qed lemma to_bl_unfold: \to_bl w = rev (map (bit w) [0.. for w :: \'a::len word\ by (simp add: map_bit_interval_eq takefill_bintrunc to_bl_def flip: bin_to_bl_def) lemma nth_rev_to_bl: \rev (to_bl w) ! n \ bit w n\ if \n < LENGTH('a)\ for w :: \'a::len word\ using that by (simp add: to_bl_unfold) lemma nth_to_bl: \to_bl w ! n \ bit w (LENGTH('a) - Suc n)\ if \n < LENGTH('a)\ for w :: \'a::len word\ using that by (simp add: to_bl_unfold rev_nth) lemma of_bl_rep_False: "of_bl (replicate n False @ bs) = of_bl bs" by (auto simp: of_bl_def bl_to_bin_rep_F) lemma [code abstract]: \Word.the_int (of_bl bs :: 'a word) = horner_sum of_bool 2 (take LENGTH('a::len) (rev bs))\ apply (simp add: of_bl_eq flip: take_bit_horner_sum_bit_eq) apply transfer apply simp done lemma [code]: \to_bl w = map (bit w) (rev [0.. for w :: \'a::len word\ by (fact to_bl_eq_rev) lemma word_reverse_eq_of_bl_rev_to_bl: \word_reverse w = of_bl (rev (to_bl w))\ by (rule bit_word_eqI) (auto simp add: bit_word_reverse_iff bit_of_bl_iff nth_to_bl) lemmas word_reverse_no_def [simp] = word_reverse_eq_of_bl_rev_to_bl [of "numeral w"] for w lemma to_bl_word_rev: "to_bl (word_reverse w) = rev (to_bl w)" by (rule nth_equalityI) (simp_all add: nth_rev_to_bl word_reverse_def word_rep_drop flip: of_bl_eq) lemma to_bl_n1 [simp]: "to_bl (-1::'a::len word) = replicate (LENGTH('a)) True" apply (rule word_bl.Abs_inverse') apply simp apply (rule word_eqI) apply (clarsimp simp add: word_size) apply (auto simp add: word_bl.Abs_inverse test_bit_bl word_size) done lemma rbl_word_or: "rev (to_bl (x OR y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" by (simp add: zip_rev bl_word_or rev_map) lemma rbl_word_and: "rev (to_bl (x AND y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" by (simp add: zip_rev bl_word_and rev_map) lemma rbl_word_xor: "rev (to_bl (x XOR y)) = map2 (\) (rev (to_bl x)) (rev (to_bl y))" by (simp add: zip_rev bl_word_xor rev_map) lemma rbl_word_not: "rev (to_bl (NOT x)) = map Not (rev (to_bl x))" by (simp add: bl_word_not rev_map) lemma bshiftr1_numeral [simp]: \bshiftr1 b (numeral w :: 'a word) = of_bl (b # butlast (bin_to_bl LENGTH('a::len) (numeral w)))\ by (simp add: bshiftr1_eq) lemma bshiftr1_bl: "to_bl (bshiftr1 b w) = b # butlast (to_bl w)" unfolding bshiftr1_eq by (rule word_bl.Abs_inverse) simp lemma shiftl1_of_bl: "shiftl1 (of_bl bl) = of_bl (bl @ [False])" by transfer (simp add: bl_to_bin_append) lemma shiftl1_bl: "shiftl1 w = of_bl (to_bl w @ [False])" for w :: "'a::len word" proof - have "shiftl1 w = shiftl1 (of_bl (to_bl w))" by simp also have "\ = of_bl (to_bl w @ [False])" by (rule shiftl1_of_bl) finally show ?thesis . qed lemma bl_shiftl1: "to_bl (shiftl1 w) = tl (to_bl w) @ [False]" for w :: "'a::len word" by (simp add: shiftl1_bl word_rep_drop drop_Suc drop_Cons') (fast intro!: Suc_leI) \ \Generalized version of \bl_shiftl1\. Maybe this one should replace it?\ lemma bl_shiftl1': "to_bl (shiftl1 w) = tl (to_bl w @ [False])" by (simp add: shiftl1_bl word_rep_drop drop_Suc del: drop_append) lemma shiftr1_bl: \shiftr1 w = of_bl (butlast (to_bl w))\ proof (rule bit_word_eqI) fix n assume \n < LENGTH('a)\ show \bit (shiftr1 w) n \ bit (of_bl (butlast (to_bl w)) :: 'a word) n\ proof (cases \n = LENGTH('a) - 1\) case True then show ?thesis by (simp add: bit_shiftr1_iff bit_of_bl_iff) next case False with \n < LENGTH('a)\ have \n < LENGTH('a) - 1\ by simp with \n < LENGTH('a)\ show ?thesis by (simp add: bit_shiftr1_iff bit_of_bl_iff rev_nth nth_butlast word_size test_bit_word_eq to_bl_nth) qed qed lemma bl_shiftr1: "to_bl (shiftr1 w) = False # butlast (to_bl w)" for w :: "'a::len word" by (simp add: shiftr1_bl word_rep_drop len_gt_0 [THEN Suc_leI]) \ \Generalized version of \bl_shiftr1\. Maybe this one should replace it?\ lemma bl_shiftr1': "to_bl (shiftr1 w) = butlast (False # to_bl w)" apply (rule word_bl.Abs_inverse') apply (simp del: butlast.simps) apply (simp add: shiftr1_bl of_bl_def) done lemma bl_sshiftr1: "to_bl (sshiftr1 w) = hd (to_bl w) # butlast (to_bl w)" for w :: "'a::len word" proof (rule nth_equalityI) fix n assume \n < length (to_bl (sshiftr1 w))\ then have \n < LENGTH('a)\ by simp then show \to_bl (sshiftr1 w) ! n \ (hd (to_bl w) # butlast (to_bl w)) ! n\ apply (cases n) apply (simp_all add: to_bl_nth word_size hd_conv_nth test_bit_eq_bit bit_sshiftr1_iff nth_butlast Suc_diff_Suc nth_to_bl) done qed simp lemma drop_shiftr: "drop n (to_bl (w >> n)) = take (size w - n) (to_bl w)" for w :: "'a::len word" apply (unfold shiftr_def) apply (induct n) prefer 2 apply (simp add: drop_Suc bl_shiftr1 butlast_drop [symmetric]) apply (rule butlast_take [THEN trans]) apply (auto simp: word_size) done lemma drop_sshiftr: "drop n (to_bl (w >>> n)) = take (size w - n) (to_bl w)" for w :: "'a::len word" apply (simp_all add: word_size sshiftr_eq) apply (rule nth_equalityI) apply (simp_all add: word_size nth_to_bl bit_signed_drop_bit_iff) done lemma take_shiftr: "n \ size w \ take n (to_bl (w >> n)) = replicate n False" apply (unfold shiftr_def) apply (induct n) prefer 2 apply (simp add: bl_shiftr1' length_0_conv [symmetric] word_size) apply (rule take_butlast [THEN trans]) apply (auto simp: word_size) done lemma take_sshiftr': "n \ size w \ hd (to_bl (w >>> n)) = hd (to_bl w) \ take n (to_bl (w >>> n)) = replicate n (hd (to_bl w))" for w :: "'a::len word" apply (auto simp add: sshiftr_eq hd_bl_sign_sint bin_sign_def not_le word_size sint_signed_drop_bit_eq) apply (rule nth_equalityI) apply (auto simp add: nth_to_bl bit_signed_drop_bit_iff bit_last_iff) apply (rule nth_equalityI) apply (auto simp add: nth_to_bl bit_signed_drop_bit_iff bit_last_iff) done lemmas hd_sshiftr = take_sshiftr' [THEN conjunct1] lemmas take_sshiftr = take_sshiftr' [THEN conjunct2] lemma atd_lem: "take n xs = t \ drop n xs = d \ xs = t @ d" by (auto intro: append_take_drop_id [symmetric]) lemmas bl_shiftr = atd_lem [OF take_shiftr drop_shiftr] lemmas bl_sshiftr = atd_lem [OF take_sshiftr drop_sshiftr] lemma shiftl_of_bl: "of_bl bl << n = of_bl (bl @ replicate n False)" by (induct n) (auto simp: shiftl_def shiftl1_of_bl replicate_app_Cons_same) lemma shiftl_bl: "w << n = of_bl (to_bl w @ replicate n False)" for w :: "'a::len word" proof - have "w << n = of_bl (to_bl w) << n" by simp also have "\ = of_bl (to_bl w @ replicate n False)" by (rule shiftl_of_bl) finally show ?thesis . qed lemma bl_shiftl: "to_bl (w << n) = drop n (to_bl w) @ replicate (min (size w) n) False" by (simp add: shiftl_bl word_rep_drop word_size) lemma shiftr1_bl_of: "length bl \ LENGTH('a) \ shiftr1 (of_bl bl::'a::len word) = of_bl (butlast bl)" by transfer (simp add: butlast_rest_bl2bin trunc_bl2bin) lemma shiftr_bl_of: "length bl \ LENGTH('a) \ (of_bl bl::'a::len word) >> n = of_bl (take (length bl - n) bl)" apply (unfold shiftr_def) apply (induct n) apply clarsimp apply clarsimp apply (subst shiftr1_bl_of) apply simp apply (simp add: butlast_take) done lemma shiftr_bl: "x >> n \ of_bl (take (LENGTH('a) - n) (to_bl x))" for x :: "'a::len word" using shiftr_bl_of [where 'a='a, of "to_bl x"] by simp lemma aligned_bl_add_size [OF refl]: "size x - n = m \ n \ size x \ drop m (to_bl x) = replicate n False \ take m (to_bl y) = replicate m False \ to_bl (x + y) = take m (to_bl x) @ drop m (to_bl y)" for x :: \'a::len word\ apply (subgoal_tac "x AND y = 0") prefer 2 apply (rule word_bl.Rep_eqD) apply (simp add: bl_word_and) apply (rule align_lem_and [THEN trans]) apply (simp_all add: word_size)[5] apply simp apply (subst word_plus_and_or [symmetric]) apply (simp add : bl_word_or) apply (rule align_lem_or) apply (simp_all add: word_size) done lemma mask_bl: "mask n = of_bl (replicate n True)" by (auto simp add : test_bit_of_bl word_size intro: word_eqI) lemma bl_and_mask': "to_bl (w AND mask n :: 'a::len word) = replicate (LENGTH('a) - n) False @ drop (LENGTH('a) - n) (to_bl w)" apply (rule nth_equalityI) apply simp apply (clarsimp simp add: to_bl_nth word_size) apply (auto simp add: word_size test_bit_bl nth_append rev_nth) done lemma slice1_eq_of_bl: \(slice1 n w :: 'b::len word) = of_bl (takefill False n (to_bl w))\ for w :: \'a::len word\ proof (rule bit_word_eqI) fix m assume \m < LENGTH('b)\ show \bit (slice1 n w :: 'b::len word) m \ bit (of_bl (takefill False n (to_bl w)) :: 'b word) m\ by (cases \m \ n\; cases \LENGTH('a) \ n\) (auto simp add: bit_slice1_iff bit_of_bl_iff not_less rev_nth not_le nth_takefill nth_to_bl algebra_simps) qed lemma slice1_no_bin [simp]: "slice1 n (numeral w :: 'b word) = of_bl (takefill False n (bin_to_bl (LENGTH('b::len)) (numeral w)))" by (simp add: slice1_eq_of_bl) (* TODO: neg_numeral *) lemma slice_no_bin [simp]: "slice n (numeral w :: 'b word) = of_bl (takefill False (LENGTH('b::len) - n) (bin_to_bl (LENGTH('b::len)) (numeral w)))" by (simp add: slice_def) (* TODO: neg_numeral *) lemma slice_take': "slice n w = of_bl (take (size w - n) (to_bl w))" by (simp add: slice_def word_size slice1_eq_of_bl takefill_alt) lemmas slice_take = slice_take' [unfolded word_size] \ \shiftr to a word of the same size is just slice, slice is just shiftr then ucast\ lemmas shiftr_slice = trans [OF shiftr_bl [THEN meta_eq_to_obj_eq] slice_take [symmetric]] lemma slice1_down_alt': "sl = slice1 n w \ fs = size sl \ fs + k = n \ to_bl sl = takefill False fs (drop k (to_bl w))" apply (simp add: slice1_eq_of_bl) apply transfer apply (simp add: bl_bin_bl_rep_drop) using drop_takefill apply force done lemma slice1_up_alt': "sl = slice1 n w \ fs = size sl \ fs = n + k \ to_bl sl = takefill False fs (replicate k False @ (to_bl w))" apply (simp add: slice1_eq_of_bl) apply transfer apply (simp add: bl_bin_bl_rep_drop flip: takefill_append) apply (metis diff_add_inverse) done lemmas sd1 = slice1_down_alt' [OF refl refl, unfolded word_size] lemmas su1 = slice1_up_alt' [OF refl refl, unfolded word_size] lemmas slice1_down_alt = le_add_diff_inverse [THEN sd1] lemmas slice1_up_alts = le_add_diff_inverse [symmetric, THEN su1] le_add_diff_inverse2 [symmetric, THEN su1] lemma slice1_tf_tf': "to_bl (slice1 n w :: 'a::len word) = rev (takefill False (LENGTH('a)) (rev (takefill False n (to_bl w))))" unfolding slice1_eq_of_bl by (rule word_rev_tf) lemmas slice1_tf_tf = slice1_tf_tf' [THEN word_bl.Rep_inverse', symmetric] lemma revcast_eq_of_bl: \(revcast w :: 'b::len word) = of_bl (takefill False (LENGTH('b)) (to_bl w))\ for w :: \'a::len word\ by (simp add: revcast_def slice1_eq_of_bl) lemmas revcast_no_def [simp] = revcast_eq_of_bl [where w="numeral w", unfolded word_size] for w lemma to_bl_revcast: "to_bl (revcast w :: 'a::len word) = takefill False (LENGTH('a)) (to_bl w)" apply (rule nth_equalityI) apply simp apply (cases \LENGTH('a) \ LENGTH('b)\) apply (auto simp add: nth_to_bl nth_takefill bit_revcast_iff) done lemma word_cat_bl: "word_cat a b = of_bl (to_bl a @ to_bl b)" apply (rule bit_word_eqI) apply (simp add: bit_word_cat_iff bit_of_bl_iff nth_append not_less nth_rev_to_bl) apply (meson bit_word.rep_eq less_diff_conv2 nth_rev_to_bl) done lemma of_bl_append: "(of_bl (xs @ ys) :: 'a::len word) = of_bl xs * 2^(length ys) + of_bl ys" apply transfer apply (simp add: bl_to_bin_app_cat bin_cat_num) done lemma of_bl_False [simp]: "of_bl (False#xs) = of_bl xs" by (rule word_eqI) (auto simp: test_bit_of_bl nth_append) lemma of_bl_True [simp]: "(of_bl (True # xs) :: 'a::len word) = 2^length xs + of_bl xs" by (subst of_bl_append [where xs="[True]", simplified]) (simp add: word_1_bl) lemma of_bl_Cons: "of_bl (x#xs) = of_bool x * 2^length xs + of_bl xs" by (cases x) simp_all lemma word_split_bl': "std = size c - size b \ (word_split c = (a, b)) \ (a = of_bl (take std (to_bl c)) \ b = of_bl (drop std (to_bl c)))" apply (simp add: word_split_def) apply transfer apply (cases \LENGTH('b) \ LENGTH('a)\) apply (auto simp add: drop_bit_take_bit drop_bin2bl bin_to_bl_drop_bit [symmetric, of \LENGTH('a)\ \LENGTH('a) - LENGTH('b)\ \LENGTH('b)\] min_absorb2) done lemma word_split_bl: "std = size c - size b \ (a = of_bl (take std (to_bl c)) \ b = of_bl (drop std (to_bl c))) \ word_split c = (a, b)" apply (rule iffI) defer apply (erule (1) word_split_bl') apply (case_tac "word_split c") apply (auto simp add: word_size) apply (frule word_split_bl' [rotated]) apply (auto simp add: word_size) done lemma word_split_bl_eq: "(word_split c :: ('c::len word \ 'd::len word)) = (of_bl (take (LENGTH('a::len) - LENGTH('d::len)) (to_bl c)), of_bl (drop (LENGTH('a) - LENGTH('d)) (to_bl c)))" for c :: "'a::len word" apply (rule word_split_bl [THEN iffD1]) apply (unfold word_size) apply (rule refl conjI)+ done lemma word_rcat_bl: \word_rcat wl = of_bl (concat (map to_bl wl))\ proof - define ws where \ws = rev wl\ moreover have \word_rcat (rev ws) = of_bl (concat (map to_bl (rev ws)))\ apply (simp add: word_rcat_def of_bl_eq rev_concat rev_map comp_def rev_to_bl_eq flip: horner_sum_of_bool_2_concat) apply transfer apply simp done ultimately show ?thesis by simp qed lemma size_rcat_lem': "size (concat (map to_bl wl)) = length wl * size (hd wl)" by (induct wl) (auto simp: word_size) lemmas size_rcat_lem = size_rcat_lem' [unfolded word_size] lemma nth_rcat_lem: "n < length (wl::'a word list) * LENGTH('a::len) \ rev (concat (map to_bl wl)) ! n = rev (to_bl (rev wl ! (n div LENGTH('a)))) ! (n mod LENGTH('a))" apply (induct wl) apply clarsimp apply (clarsimp simp add : nth_append size_rcat_lem) apply (simp flip: mult_Suc minus_div_mult_eq_mod add: less_Suc_eq_le not_less) apply (metis (no_types, lifting) diff_is_0_eq div_le_mono len_not_eq_0 less_Suc_eq less_mult_imp_div_less nonzero_mult_div_cancel_right not_le nth_Cons_0) done lemma foldl_eq_foldr: "foldl (+) x xs = foldr (+) (x # xs) 0" for x :: "'a::comm_monoid_add" by (induct xs arbitrary: x) (auto simp: add.assoc) lemmas word_cat_bl_no_bin [simp] = word_cat_bl [where a="numeral a" and b="numeral b", unfolded to_bl_numeral] for a b (* FIXME: negative numerals, 0 and 1 *) lemmas word_split_bl_no_bin [simp] = word_split_bl_eq [where c="numeral c", unfolded to_bl_numeral] for c lemmas word_rot_defs = word_roti_eq_word_rotr_word_rotl word_rotr_eq word_rotl_eq lemma to_bl_rotl: "to_bl (word_rotl n w) = rotate n (to_bl w)" by (simp add: word_rotl_eq to_bl_use_of_bl) lemmas blrs0 = rotate_eqs [THEN to_bl_rotl [THEN trans]] lemmas word_rotl_eqs = blrs0 [simplified word_bl_Rep' word_bl.Rep_inject to_bl_rotl [symmetric]] lemma to_bl_rotr: "to_bl (word_rotr n w) = rotater n (to_bl w)" by (simp add: word_rotr_eq to_bl_use_of_bl) lemmas brrs0 = rotater_eqs [THEN to_bl_rotr [THEN trans]] lemmas word_rotr_eqs = brrs0 [simplified word_bl_Rep' word_bl.Rep_inject to_bl_rotr [symmetric]] declare word_rotr_eqs (1) [simp] declare word_rotl_eqs (1) [simp] lemmas abl_cong = arg_cong [where f = "of_bl"] locale word_rotate begin lemmas word_rot_defs' = to_bl_rotl to_bl_rotr lemmas blwl_syms [symmetric] = bl_word_not bl_word_and bl_word_or bl_word_xor lemmas lbl_lbl = trans [OF word_bl_Rep' word_bl_Rep' [symmetric]] lemmas ths_map2 [OF lbl_lbl] = rotate_map2 rotater_map2 lemmas ths_map [where xs = "to_bl v"] = rotate_map rotater_map for v lemmas th1s [simplified word_rot_defs' [symmetric]] = ths_map2 ths_map end lemmas bl_word_rotl_dt = trans [OF to_bl_rotl rotate_drop_take, simplified word_bl_Rep'] lemmas bl_word_rotr_dt = trans [OF to_bl_rotr rotater_drop_take, simplified word_bl_Rep'] lemma bl_word_roti_dt': "n = nat ((- i) mod int (size (w :: 'a::len word))) \ to_bl (word_roti i w) = drop n (to_bl w) @ take n (to_bl w)" apply (unfold word_roti_eq_word_rotr_word_rotl) apply (simp add: bl_word_rotl_dt bl_word_rotr_dt word_size) apply safe apply (simp add: zmod_zminus1_eq_if) apply safe apply (simp add: nat_mult_distrib) apply (simp add: nat_diff_distrib [OF pos_mod_sign pos_mod_conj [THEN conjunct2, THEN order_less_imp_le]] nat_mod_distrib) apply (simp add: nat_mod_distrib) done lemmas bl_word_roti_dt = bl_word_roti_dt' [unfolded word_size] lemmas word_rotl_dt = bl_word_rotl_dt [THEN word_bl.Rep_inverse' [symmetric]] lemmas word_rotr_dt = bl_word_rotr_dt [THEN word_bl.Rep_inverse' [symmetric]] lemmas word_roti_dt = bl_word_roti_dt [THEN word_bl.Rep_inverse' [symmetric]] lemmas word_rotr_dt_no_bin' [simp] = word_rotr_dt [where w="numeral w", unfolded to_bl_numeral] for w (* FIXME: negative numerals, 0 and 1 *) lemmas word_rotl_dt_no_bin' [simp] = word_rotl_dt [where w="numeral w", unfolded to_bl_numeral] for w (* FIXME: negative numerals, 0 and 1 *) lemma max_word_bl: "to_bl (- 1::'a::len word) = replicate LENGTH('a) True" by (fact to_bl_n1) lemma to_bl_mask: "to_bl (mask n :: 'a::len word) = replicate (LENGTH('a) - n) False @ replicate (min (LENGTH('a)) n) True" by (simp add: mask_bl word_rep_drop min_def) lemma map_replicate_True: "n = length xs \ map (\(x,y). x \ y) (zip xs (replicate n True)) = xs" by (induct xs arbitrary: n) auto lemma map_replicate_False: "n = length xs \ map (\(x,y). x \ y) (zip xs (replicate n False)) = replicate n False" by (induct xs arbitrary: n) auto lemma bl_and_mask: fixes w :: "'a::len word" and n :: nat defines "n' \ LENGTH('a) - n" shows "to_bl (w AND mask n) = replicate n' False @ drop n' (to_bl w)" proof - note [simp] = map_replicate_True map_replicate_False have "to_bl (w AND mask n) = map2 (\) (to_bl w) (to_bl (mask n::'a::len word))" by (simp add: bl_word_and) also have "to_bl w = take n' (to_bl w) @ drop n' (to_bl w)" by simp also have "map2 (\) \ (to_bl (mask n::'a::len word)) = replicate n' False @ drop n' (to_bl w)" unfolding to_bl_mask n'_def by (subst zip_append) auto finally show ?thesis . qed lemma drop_rev_takefill: "length xs \ n \ drop (n - length xs) (rev (takefill False n (rev xs))) = xs" by (simp add: takefill_alt rev_take) declare bin_to_bl_def [simp] lemmas of_bl_reasoning = to_bl_use_of_bl of_bl_append lemma uint_of_bl_is_bl_to_bin_drop: "length (dropWhile Not l) \ LENGTH('a) \ uint (of_bl l :: 'a::len word) = bl_to_bin l" apply transfer apply (simp add: take_bit_eq_mod) apply (rule Divides.mod_less) apply (rule bl_to_bin_ge0) using bl_to_bin_lt2p_drop apply (rule order.strict_trans2) apply simp done corollary uint_of_bl_is_bl_to_bin: "length l\LENGTH('a) \ uint ((of_bl::bool list\ ('a :: len) word) l) = bl_to_bin l" apply(rule uint_of_bl_is_bl_to_bin_drop) using le_trans length_dropWhile_le by blast lemma bin_to_bl_or: "bin_to_bl n (a OR b) = map2 (\) (bin_to_bl n a) (bin_to_bl n b)" using bl_or_aux_bin[where n=n and v=a and w=b and bs="[]" and cs="[]"] by simp lemma word_and_1_bl: fixes x::"'a::len word" shows "(x AND 1) = of_bl [x !! 0]" by (simp add: mod_2_eq_odd test_bit_word_eq and_one_eq) lemma word_1_and_bl: fixes x::"'a::len word" shows "(1 AND x) = of_bl [x !! 0]" by (simp add: mod_2_eq_odd test_bit_word_eq one_and_eq) lemma of_bl_drop: "of_bl (drop n xs) = (of_bl xs AND mask (length xs - n))" apply (clarsimp simp: bang_eq test_bit_of_bl rev_nth cong: rev_conj_cong) apply (safe; simp add: word_size to_bl_nth) done lemma to_bl_1: "to_bl (1::'a::len word) = replicate (LENGTH('a) - 1) False @ [True]" by (rule nth_equalityI) (auto simp add: to_bl_unfold nth_append rev_nth bit_1_iff not_less not_le) lemma eq_zero_set_bl: "(w = 0) = (True \ set (to_bl w))" apply (auto simp add: to_bl_unfold) apply (rule bit_word_eqI) apply auto done lemma of_drop_to_bl: "of_bl (drop n (to_bl x)) = (x AND mask (size x - n))" by (simp add: of_bl_drop word_size_bl) lemma unat_of_bl_length: "unat (of_bl xs :: 'a::len word) < 2 ^ (length xs)" proof (cases "length xs < LENGTH('a)") case True then have "(of_bl xs::'a::len word) < 2 ^ length xs" by (simp add: of_bl_length_less) with True show ?thesis by (simp add: word_less_nat_alt unat_of_nat) next case False have "unat (of_bl xs::'a::len word) < 2 ^ LENGTH('a)" by (simp split: unat_split) also from False have "LENGTH('a) \ length xs" by simp then have "2 ^ LENGTH('a) \ (2::nat) ^ length xs" by (rule power_increasing) simp finally show ?thesis . qed lemma word_msb_alt: "msb w \ hd (to_bl w)" for w :: "'a::len word" apply (simp add: msb_word_eq) apply (subst hd_conv_nth) apply simp apply (subst nth_to_bl) apply simp apply simp done lemma word_lsb_last: \lsb w \ last (to_bl w)\ for w :: \'a::len word\ using nth_to_bl [of \LENGTH('a) - Suc 0\ w] by (simp add: lsb_odd last_conv_nth) lemma is_aligned_to_bl: "is_aligned (w :: 'a :: len word) n = (True \ set (drop (size w - n) (to_bl w)))" apply (simp add: is_aligned_mask eq_zero_set_bl) apply (clarsimp simp: in_set_conv_nth word_size) apply (simp add: to_bl_nth word_size cong: conj_cong) apply (simp add: diff_diff_less) apply safe apply (case_tac "n \ LENGTH('a)") prefer 2 apply (rule_tac x=i in exI) apply clarsimp apply (subgoal_tac "\j < LENGTH('a). j < n \ LENGTH('a) - n + j = i") apply (erule exE) apply (rule_tac x=j in exI) apply clarsimp apply (thin_tac "w !! t" for t) apply (rule_tac x="i + n - LENGTH('a)" in exI) apply clarsimp apply arith apply (rule_tac x="LENGTH('a) - n + i" in exI) apply clarsimp apply arith done lemma is_aligned_replicate: fixes w::"'a::len word" assumes aligned: "is_aligned w n" and nv: "n \ LENGTH('a)" shows "to_bl w = (take (LENGTH('a) - n) (to_bl w)) @ replicate n False" proof - from nv have rl: "\q. q < 2 ^ (LENGTH('a) - n) \ to_bl (2 ^ n * (of_nat q :: 'a word)) = drop n (to_bl (of_nat q :: 'a word)) @ replicate n False" by (metis bl_shiftl le_antisym min_def shiftl_t2n wsst_TYs(3)) show ?thesis using aligned by (auto simp: rl elim: is_alignedE) qed lemma is_aligned_drop: fixes w::"'a::len word" assumes "is_aligned w n" "n \ LENGTH('a)" shows "drop (LENGTH('a) - n) (to_bl w) = replicate n False" proof - have "to_bl w = take (LENGTH('a) - n) (to_bl w) @ replicate n False" by (rule is_aligned_replicate) fact+ then have "drop (LENGTH('a) - n) (to_bl w) = drop (LENGTH('a) - n) \" by simp also have "\ = replicate n False" by simp finally show ?thesis . qed lemma less_is_drop_replicate: fixes x::"'a::len word" assumes lt: "x < 2 ^ n" shows "to_bl x = replicate (LENGTH('a) - n) False @ drop (LENGTH('a) - n) (to_bl x)" by (metis assms bl_and_mask' less_mask_eq) lemma is_aligned_add_conv: fixes off::"'a::len word" assumes aligned: "is_aligned w n" and offv: "off < 2 ^ n" shows "to_bl (w + off) = (take (LENGTH('a) - n) (to_bl w)) @ (drop (LENGTH('a) - n) (to_bl off))" proof cases assume nv: "n \ LENGTH('a)" show ?thesis proof (subst aligned_bl_add_size, simp_all only: word_size) show "drop (LENGTH('a) - n) (to_bl w) = replicate n False" by (subst is_aligned_replicate [OF aligned nv]) (simp add: word_size) from offv show "take (LENGTH('a) - n) (to_bl off) = replicate (LENGTH('a) - n) False" by (subst less_is_drop_replicate, assumption) simp qed fact next assume "\ n \ LENGTH('a)" with offv show ?thesis by (simp add: power_overflow) qed lemma is_aligned_replicateI: "to_bl p = addr @ replicate n False \ is_aligned (p::'a::len word) n" apply (simp add: is_aligned_to_bl word_size) apply (subgoal_tac "length addr = LENGTH('a) - n") apply (simp add: replicate_not_True) apply (drule arg_cong [where f=length]) apply simp done lemma to_bl_2p: "n < LENGTH('a) \ to_bl ((2::'a::len word) ^ n) = replicate (LENGTH('a) - Suc n) False @ True # replicate n False" apply (subst shiftl_1 [symmetric]) apply (subst bl_shiftl) apply (simp add: to_bl_1 min_def word_size) done lemma xor_2p_to_bl: fixes x::"'a::len word" shows "to_bl (x XOR 2^n) = (if n < LENGTH('a) then take (LENGTH('a)-Suc n) (to_bl x) @ (\rev (to_bl x)!n) # drop (LENGTH('a)-n) (to_bl x) else to_bl x)" proof - have x: "to_bl x = take (LENGTH('a)-Suc n) (to_bl x) @ drop (LENGTH('a)-Suc n) (to_bl x)" by simp show ?thesis apply simp apply (rule conjI) apply (clarsimp simp: word_size) apply (simp add: bl_word_xor to_bl_2p) apply (subst x) apply (subst zip_append) apply simp apply (simp add: map_zip_replicate_False_xor drop_minus) apply (auto simp add: word_size nth_w2p intro!: word_eqI) done qed lemma is_aligned_replicateD: "\ is_aligned (w::'a::len word) n; n \ LENGTH('a) \ \ \xs. to_bl w = xs @ replicate n False \ length xs = size w - n" apply (subst is_aligned_replicate, assumption+) apply (rule exI, rule conjI, rule refl) apply (simp add: word_size) done text \right-padding a word to a certain length\ definition "bl_pad_to bl sz \ bl @ (replicate (sz - length bl) False)" lemma bl_pad_to_length: assumes lbl: "length bl \ sz" shows "length (bl_pad_to bl sz) = sz" using lbl by (simp add: bl_pad_to_def) lemma bl_pad_to_prefix: "prefix bl (bl_pad_to bl sz)" by (simp add: bl_pad_to_def) lemma of_bl_length: "length xs < LENGTH('a) \ of_bl xs < (2 :: 'a::len word) ^ length xs" by (simp add: of_bl_length_less) lemma of_bl_mult_and_not_mask_eq: "\is_aligned (a :: 'a::len word) n; length b + m \ n\ \ a + of_bl b * (2^m) AND NOT(mask n) = a" apply (simp flip: push_bit_eq_mult subtract_mask(1) take_bit_eq_mask) apply (subst disjunctive_add) apply (auto simp add: bit_simps not_le not_less) apply (meson is_aligned_imp_not_bit is_aligned_weaken less_diff_conv2) apply (erule is_alignedE') apply (simp add: take_bit_push_bit) apply (rule bit_word_eqI) apply (auto simp add: bit_simps) done lemma bin_to_bl_of_bl_eq: "\is_aligned (a::'a::len word) n; length b + c \ n; length b + c < LENGTH('a)\ \ bin_to_bl (length b) (uint ((a + of_bl b * 2^c) >> c)) = b" apply (simp flip: push_bit_eq_mult take_bit_eq_mask add: shiftr_eq_drop_bit) apply (subst disjunctive_add) apply (auto simp add: bit_simps not_le not_less unsigned_or_eq unsigned_drop_bit_eq unsigned_push_bit_eq bin_to_bl_or simp flip: bin_to_bl_def) apply (meson is_aligned_imp_not_bit is_aligned_weaken less_diff_conv2) apply (erule is_alignedE') apply (rule nth_equalityI) apply (auto simp add: nth_bin_to_bl bit_simps rev_nth simp flip: bin_to_bl_def) done (* FIXME: move to Word distribution *) lemma bin_nth_minus_Bit0[simp]: - "0 < n \ bin_nth (numeral (num.Bit0 w)) n = bin_nth (numeral w) (n - 1)" + "0 < n \ bit (numeral (num.Bit0 w) :: int) n = bit (numeral w :: int) (n - 1)" by (cases n; simp) lemma bin_nth_minus_Bit1[simp]: - "0 < n \ bin_nth (numeral (num.Bit1 w)) n = bin_nth (numeral w) (n - 1)" + "0 < n \ bit (numeral (num.Bit1 w) :: int) n = bit (numeral w :: int) (n - 1)" by (cases n; simp) (* casting a long word to a shorter word and casting back to the long word is equal to the original long word -- if the word is small enough. 'l is the longer word. 's is the shorter word. *) lemma bl_cast_long_short_long_ingoreLeadingZero_generic: "\ length (dropWhile Not (to_bl w)) \ LENGTH('s); LENGTH('s) \ LENGTH('l) \ \ (of_bl :: _ \ 'l::len word) (to_bl ((of_bl::_ \ 's::len word) (to_bl w))) = w" by (rule word_uint_eqI) (simp add: uint_of_bl_is_bl_to_bin uint_of_bl_is_bl_to_bin_drop) (* Casting between longer and shorter word. 'l is the longer word. 's is the shorter word. For example: 'l::len word is 128 word (full ipv6 address) 's::len word is 16 word (address piece of ipv6 address in colon-text-representation) *) corollary ucast_short_ucast_long_ingoreLeadingZero: "\ length (dropWhile Not (to_bl w)) \ LENGTH('s); LENGTH('s) \ LENGTH('l) \ \ (ucast:: 's::len word \ 'l::len word) ((ucast:: 'l::len word \ 's::len word) w) = w" apply (subst ucast_bl)+ apply (rule bl_cast_long_short_long_ingoreLeadingZero_generic; simp) done lemma length_drop_mask: fixes w::"'a::len word" shows "length (dropWhile Not (to_bl (w AND mask n))) \ n" proof - have "length (takeWhile Not (replicate n False @ ls)) = n + length (takeWhile Not ls)" for ls n by(subst takeWhile_append2) simp+ then show ?thesis unfolding bl_and_mask by (simp add: dropWhile_eq_drop) qed lemma map_bits_rev_to_bl: "map ((!!) x) [0.. of_bl xs * 2^c < (2::'a::len word) ^ (length xs + c)" by (simp add: of_bl_length word_less_power_trans2) lemma of_bl_max: "(of_bl xs :: 'a::len word) \ mask (length xs)" proof - define ys where \ys = rev xs\ have \take_bit (length ys) (horner_sum of_bool 2 ys :: 'a word) = horner_sum of_bool 2 ys\ by transfer (simp add: take_bit_horner_sum_bit_eq min_def) then have \(of_bl (rev ys) :: 'a word) \ mask (length ys)\ by (simp only: of_bl_rev_eq less_eq_mask_iff_take_bit_eq_self) with ys_def show ?thesis by simp qed text\Some auxiliaries for sign-shifting by the entire word length or more\ lemma sshiftr_clamp_pos: assumes "LENGTH('a) \ n" "0 \ sint x" shows "(x::'a::len word) >>> n = 0" apply (rule word_sint.Rep_eqD) apply (unfold sshiftr_div_2n Word.sint_0) apply (rule div_pos_pos_trivial) subgoal using assms(2) . apply (rule order.strict_trans[where b="2 ^ (LENGTH('a) - 1)"]) using sint_lt assms(1) by auto lemma sshiftr_clamp_neg: assumes "LENGTH('a) \ n" "sint x < 0" shows "(x::'a::len word) >>> n = -1" proof - have *: "- (2 ^ n) < sint x" apply (rule order.strict_trans2[where b="- (2 ^ (LENGTH('a) - 1))"]) using assms(1) sint_ge by auto show ?thesis apply (rule word_sint.Rep_eqD) apply (unfold sshiftr_div_2n Word.sint_n1) apply (subst div_minus_minus[symmetric]) apply (rule div_pos_neg_trivial) subgoal using assms(2) by linarith using * by simp qed lemma sshiftr_clamp: assumes "LENGTH('a) \ n" shows "(x::'a::len word) >>> n = x >>> LENGTH('a)" apply (cases "0 \ sint x") subgoal apply (subst sshiftr_clamp_pos[OF assms]) defer apply (subst sshiftr_clamp_pos) by auto apply (subst sshiftr_clamp_neg[OF assms]) defer apply (subst sshiftr_clamp_neg) by auto text\ Like @{thm shiftr1_bl_of}, but the precondition is stronger because we need to pick the msb out of the list. \ lemma sshiftr1_bl_of: "length bl = LENGTH('a) \ sshiftr1 (of_bl bl::'a::len word) = of_bl (hd bl # butlast bl)" apply (rule word_bl.Rep_eqD) apply (subst bl_sshiftr1[of "of_bl bl :: 'a word"]) by (simp add: word_bl.Abs_inverse) text\ Like @{thm sshiftr1_bl_of}, with a weaker precondition. We still get a direct equation for @{term \sshiftr1 (of_bl bl)\}, it's just uglier. \ lemma sshiftr1_bl_of': "LENGTH('a) \ length bl \ sshiftr1 (of_bl bl::'a::len word) = of_bl (hd (drop (length bl - LENGTH('a)) bl) # butlast (drop (length bl - LENGTH('a)) bl))" apply (subst of_bl_drop'[symmetric, of "length bl - LENGTH('a)"]) using sshiftr1_bl_of[of "drop (length bl - LENGTH('a)) bl"] by auto text\ Like @{thm shiftr_bl_of}. \ lemma sshiftr_bl_of: assumes "length bl = LENGTH('a)" shows "(of_bl bl::'a::len word) >>> n = of_bl (replicate n (hd bl) @ take (length bl - n) bl)" proof - { fix n assume "n \ LENGTH('a)" hence "(of_bl bl::'a::len word) >>> n = of_bl (replicate n (hd bl) @ take (length bl - n) bl)" proof (induction n) case (Suc n) hence "n < length bl" by (simp add: assms) hence ne: "\take (length bl - n) bl = []" by auto have left: "hd (replicate n (hd bl) @ take (length bl - n) bl) = (hd bl)" by (cases "0 < n") auto have right: "butlast (take (length bl - n) bl) = take (length bl - Suc n) bl" by (subst butlast_take) auto have "(of_bl bl::'a::len word) >>> Suc n = sshiftr1 ((of_bl bl::'a::len word) >>> n)" unfolding sshiftr_eq_funpow_sshiftr1 by simp also have "\ = of_bl (replicate (Suc n) (hd bl) @ take (length bl - Suc n) bl)" apply (subst Suc.IH[OF Suc_leD[OF Suc.prems]]) apply (subst sshiftr1_bl_of) subgoal using assms Suc.prems by simp apply (rule arg_cong[where f=of_bl]) apply (subst butlast_append) unfolding left right using ne by simp finally show ?case . qed (transfer, simp) } note pos = this { assume n: "LENGTH('a) \ n" have "(of_bl bl::'a::len word) >>> n = (of_bl bl::'a::len word) >>> LENGTH('a)" by (rule sshiftr_clamp[OF n]) also have "\ = of_bl (replicate LENGTH('a) (hd bl) @ take (length bl - LENGTH('a)) bl)" apply (rule pos) .. also have "\ = of_bl (replicate n (hd bl) @ take (length bl - n) bl)" proof - have "(of_bl (replicate LENGTH('a) (hd bl)) :: 'a word) = of_bl (replicate n (hd bl))" apply (subst of_bl_drop'[symmetric, of "n - LENGTH('a)" "replicate n (hd bl)"]) unfolding length_replicate by (auto simp: n) thus ?thesis by (simp add: assms n) qed finally have "(of_bl bl::'a::len word) >>> n = of_bl (replicate n (hd bl) @ take (length bl - n) bl)" . } thus ?thesis using pos by fastforce qed text\Like @{thm shiftr_bl}\ lemma sshiftr_bl: "x >>> n \ of_bl (replicate n (msb x) @ take (LENGTH('a) - n) (to_bl x))" for x :: "'a::len word" unfolding word_msb_alt by (smt (z3) length_to_bl_eq sshiftr_bl_of word_bl.Rep_inverse) end diff --git a/thys/Word_Lib/Rsplit.thy b/thys/Word_Lib/Rsplit.thy --- a/thys/Word_Lib/Rsplit.thy +++ b/thys/Word_Lib/Rsplit.thy @@ -1,168 +1,168 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) (* Author: Jeremy Dawson and Gerwin Klein, NICTA *) section \Splitting words into lists\ theory Rsplit imports "HOL-Library.Word" Bits_Int begin definition word_rsplit :: "'a::len word \ 'b::len word list" where "word_rsplit w = map word_of_int (bin_rsplit LENGTH('b) (LENGTH('a), uint w))" lemma word_rsplit_no: "(word_rsplit (numeral bin :: 'b::len word) :: 'a word list) = map word_of_int (bin_rsplit (LENGTH('a::len)) (LENGTH('b), take_bit (LENGTH('b)) (numeral bin)))" by (simp add: word_rsplit_def of_nat_take_bit) lemmas word_rsplit_no_cl [simp] = word_rsplit_no [unfolded bin_rsplitl_def bin_rsplit_l [symmetric]] text \ This odd result arises from the fact that the statement of the result implies that the decoded words are of the same type, and therefore of the same length, as the original word.\ lemma word_rsplit_same: "word_rsplit w = [w]" apply (simp add: word_rsplit_def bin_rsplit_all) apply transfer apply simp done lemma word_rsplit_empty_iff_size: "word_rsplit w = [] \ size w = 0" by (simp add: word_rsplit_def bin_rsplit_def word_size bin_rsplit_aux_simp_alt Let_def split: prod.split) lemma test_bit_rsplit: "sw = word_rsplit w \ m < size (hd sw) \ k < length sw \ (rev sw ! k) !! m = w !! (k * size (hd sw) + m)" for sw :: "'a::len word list" apply (unfold word_rsplit_def word_test_bit_def) apply (rule trans) - apply (rule_tac f = "\x. bin_nth x m" in arg_cong) + apply (rule_tac f = "\x. bit x m" in arg_cong) apply (rule nth_map [symmetric]) apply simp apply (rule bin_nth_rsplit) apply simp_all apply (simp add : word_size rev_map) apply (rule trans) defer apply (rule map_ident [THEN fun_cong]) apply (rule refl [THEN map_cong]) apply simp using bin_rsplit_size_sign take_bit_int_eq_self_iff by blast lemma test_bit_rsplit_alt: \(word_rsplit w :: 'b::len word list) ! i !! m \ w !! ((length (word_rsplit w :: 'b::len word list) - Suc i) * size (hd (word_rsplit w :: 'b::len word list)) + m)\ if \i < length (word_rsplit w :: 'b::len word list)\ \m < size (hd (word_rsplit w :: 'b::len word list))\ \0 < length (word_rsplit w :: 'b::len word list)\ for w :: \'a::len word\ apply (rule trans) apply (rule test_bit_cong) apply (rule rev_nth [of _ \rev (word_rsplit w)\, simplified rev_rev_ident]) apply simp apply (rule that(1)) apply simp apply (rule test_bit_rsplit) apply (rule refl) apply (rule asm_rl) apply (rule that(2)) apply (rule diff_Suc_less) apply (rule that(3)) done lemma word_rsplit_len_indep [OF refl refl refl refl]: "[u,v] = p \ [su,sv] = q \ word_rsplit u = su \ word_rsplit v = sv \ length su = length sv" by (auto simp: word_rsplit_def bin_rsplit_len_indep) lemma length_word_rsplit_size: "n = LENGTH('a::len) \ length (word_rsplit w :: 'a word list) \ m \ size w \ m * n" by (auto simp: word_rsplit_def word_size bin_rsplit_len_le) lemmas length_word_rsplit_lt_size = length_word_rsplit_size [unfolded Not_eq_iff linorder_not_less [symmetric]] lemma length_word_rsplit_exp_size: "n = LENGTH('a::len) \ length (word_rsplit w :: 'a word list) = (size w + n - 1) div n" by (auto simp: word_rsplit_def word_size bin_rsplit_len) lemma length_word_rsplit_even_size: "n = LENGTH('a::len) \ size w = m * n \ length (word_rsplit w :: 'a word list) = m" by (cases \LENGTH('a)\) (simp_all add: length_word_rsplit_exp_size div_nat_eqI) lemmas length_word_rsplit_exp_size' = refl [THEN length_word_rsplit_exp_size] \ \alternative proof of \word_rcat_rsplit\\ lemmas tdle = times_div_less_eq_dividend lemmas dtle = xtrans(4) [OF tdle mult.commute] lemma word_rcat_rsplit: "word_rcat (word_rsplit w) = w" apply (rule word_eqI) apply (clarsimp simp: test_bit_rcat word_size) apply (subst refl [THEN test_bit_rsplit]) apply (simp_all add: word_size refl [THEN length_word_rsplit_size [simplified not_less [symmetric], simplified]]) apply safe apply (erule xtrans(7), rule dtle)+ done lemma size_word_rsplit_rcat_size: "word_rcat ws = frcw \ size frcw = length ws * LENGTH('a) \ length (word_rsplit frcw::'a word list) = length ws" for ws :: "'a::len word list" and frcw :: "'b::len word" by (cases \LENGTH('a)\) (simp_all add: word_size length_word_rsplit_exp_size' div_nat_eqI) lemma msrevs: "0 < n \ (k * n + m) div n = m div n + k" "(k * n + m) mod n = m mod n" for n :: nat by (auto simp: add.commute) lemma word_rsplit_rcat_size [OF refl]: "word_rcat ws = frcw \ size frcw = length ws * LENGTH('a) \ word_rsplit frcw = ws" for ws :: "'a::len word list" apply (frule size_word_rsplit_rcat_size, assumption) apply (clarsimp simp add : word_size) apply (rule nth_equalityI, assumption) apply clarsimp apply (rule word_eqI [rule_format]) apply (rule trans) apply (rule test_bit_rsplit_alt) apply (clarsimp simp: word_size)+ apply (rule trans) apply (rule test_bit_rcat [OF refl refl]) apply (simp add: word_size) apply (subst rev_nth) apply arith apply (simp add: le0 [THEN [2] xtrans(7), THEN diff_Suc_less]) apply safe apply (simp add: diff_mult_distrib) apply (cases "size ws") apply simp_all done lemma word_rsplit_upt: "\ size x = LENGTH('a :: len) * n; n \ 0 \ \ word_rsplit x = map (\i. ucast (x >> i * len_of TYPE ('a)) :: 'a word) (rev [0 ..< n])" apply (subgoal_tac "length (word_rsplit x :: 'a word list) = n") apply (rule nth_equalityI, simp) apply (intro allI word_eqI impI) apply (simp add: test_bit_rsplit_alt word_size) apply (simp add: nth_ucast nth_shiftr rev_nth field_simps) apply (simp add: length_word_rsplit_exp_size) apply transfer apply (metis (no_types, lifting) Nat.add_diff_assoc Suc_leI add_0_left diff_Suc_less div_less len_gt_0 msrevs(1) mult.commute) done end \ No newline at end of file diff --git a/thys/Word_Lib/Word_Lemmas.thy b/thys/Word_Lib/Word_Lemmas.thy --- a/thys/Word_Lib/Word_Lemmas.thy +++ b/thys/Word_Lib/Word_Lemmas.thy @@ -1,1901 +1,1901 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Lemmas with Generic Word Length" theory Word_Lemmas imports Type_Syntax Signed_Division_Word Signed_Words More_Word Most_significant_bit Enumeration_Word Aligned begin lemma max_word_not_less [simp]: "\ - 1 < x" for x :: \'a::len word\ by (fact word_order.extremum_strict) (* The seL4 bitfield generator produces functions containing mask and shift operations, such that * invoking two of them consecutively can produce something like the following. *) lemma bitfield_op_twice: "(x AND NOT (mask n << m) OR ((y AND mask n) << m)) AND NOT (mask n << m) = x AND NOT (mask n << m)" for x :: \'a::len word\ by (induct n arbitrary: m) (auto simp: word_ao_dist) lemma bitfield_op_twice'': "\NOT a = b << c; \x. b = mask x\ \ (x AND a OR (y AND b << c)) AND a = x AND a" for a b :: \'a::len word\ apply clarsimp apply (cut_tac n=xa and m=c and x=x and y=y in bitfield_op_twice) apply (clarsimp simp:mask_eq_decr_exp) apply (drule not_switch) apply clarsimp done lemma bit_twiddle_min: "(y::'a::len word) XOR (((x::'a::len word) XOR y) AND (if x < y then -1 else 0)) = min x y" by (auto simp add: Parity.bit_eq_iff bit_xor_iff min_def) lemma bit_twiddle_max: "(x::'a::len word) XOR (((x::'a::len word) XOR y) AND (if x < y then -1 else 0)) = max x y" by (auto simp add: Parity.bit_eq_iff bit_xor_iff max_def) lemma swap_with_xor: "\(x::'a::len word) = a XOR b; y = b XOR x; z = x XOR y\ \ z = b \ y = a" by (auto simp add: Parity.bit_eq_iff bit_xor_iff max_def) lemma scast_nop1 [simp]: "((scast ((of_int x)::('a::len) word))::'a sword) = of_int x" apply (simp only: scast_eq) by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse) lemma scast_nop2 [simp]: "((scast ((of_int x)::('a::len) sword))::'a word) = of_int x" apply (simp only: scast_eq) by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse) lemmas scast_nop = scast_nop1 scast_nop2 scast_id lemma le_mask_imp_and_mask: "(x::'a::len word) \ mask n \ x AND mask n = x" by (metis and_mask_eq_iff_le_mask) lemma or_not_mask_nop: "((x::'a::len word) OR NOT (mask n)) AND mask n = x AND mask n" by (metis word_and_not word_ao_dist2 word_bw_comms(1) word_log_esimps(3)) lemma mask_subsume: "\n \ m\ \ ((x::'a::len word) OR y AND mask n) AND NOT (mask m) = x AND NOT (mask m)" by (auto simp add: Parity.bit_eq_iff bit_not_iff bit_or_iff bit_and_iff bit_mask_iff) lemma and_mask_0_iff_le_mask: fixes w :: "'a::len word" shows "(w AND NOT(mask n) = 0) = (w \ mask n)" by (simp add: mask_eq_0_eq_x le_mask_imp_and_mask and_mask_eq_iff_le_mask) lemma mask_twice2: "n \ m \ ((x::'a::len word) AND mask m) AND mask n = x AND mask n" by (metis mask_twice min_def) lemma uint_2_id: "LENGTH('a) \ 2 \ uint (2::('a::len) word) = 2" by simp lemma bintrunc_id: - "\m \ of_nat n; 0 < m\ \ bintrunc n m = m" + "\m \ int n; 0 < m\ \ take_bit n m = m" by (simp add: bintrunc_mod2p le_less_trans) lemma shiftr1_unfold: "shiftr1 x = x >> 1" by (metis One_nat_def comp_apply funpow.simps(1) funpow.simps(2) id_apply shiftr_def) lemma shiftr1_is_div_2: "(x::('a::len) word) >> 1 = x div 2" by transfer (simp add: drop_bit_Suc) lemma shiftl1_is_mult: "(x << 1) = (x :: 'a::len word) * 2" by (metis One_nat_def mult_2 mult_2_right one_add_one power_0 power_Suc shiftl_t2n) lemma div_of_0_id[simp]:"(0::('a::len) word) div n = 0" by (simp add: word_div_def) lemma degenerate_word:"LENGTH('a) = 1 \ (x::('a::len) word) = 0 \ x = 1" by (metis One_nat_def less_irrefl_nat sint_1_cases) lemma div_by_0_word:"(x::('a::len) word) div 0 = 0" by (metis div_0 div_by_0 unat_0 word_arith_nat_defs(6) word_div_1) lemma div_less_dividend_word:"\x \ 0; n \ 1\ \ (x::('a::len) word) div n < x" apply (cases \n = 0\) apply clarsimp apply (simp add:word_neq_0_conv) apply (subst word_arith_nat_div) apply (rule word_of_nat_less) apply (rule div_less_dividend) using unat_eq_zero word_unat_Rep_inject1 apply force apply (simp add:unat_gt_0) done lemma shiftr1_lt:"x \ 0 \ (x::('a::len) word) >> 1 < x" apply (subst shiftr1_is_div_2) apply (rule div_less_dividend_word) apply simp+ done lemma word_less_div: fixes x :: "('a::len) word" and y :: "('a::len) word" shows "x div y = 0 \ y = 0 \ x < y" apply (case_tac "y = 0", clarsimp+) by (metis One_nat_def Suc_le_mono le0 le_div_geq not_less unat_0 unat_div unat_gt_0 word_less_nat_alt zero_less_one) lemma not_degenerate_imp_2_neq_0:"LENGTH('a) > 1 \ (2::('a::len) word) \ 0" by (metis numerals(1) power_not_zero power_zero_numeral) lemma shiftr1_0_or_1:"(x::('a::len) word) >> 1 = 0 \ x = 0 \ x = 1" apply (subst (asm) shiftr1_is_div_2) apply (drule word_less_div) apply (case_tac "LENGTH('a) = 1") apply (simp add:degenerate_word) apply (erule disjE) apply (subgoal_tac "(2::'a word) \ 0") apply simp apply (rule not_degenerate_imp_2_neq_0) apply (subgoal_tac "LENGTH('a) \ 0") apply arith apply simp apply (rule x_less_2_0_1', simp+) done lemma word_overflow:"(x::('a::len) word) + 1 > x \ x + 1 = 0" apply clarsimp by (metis diff_0 eq_diff_eq less_x_plus_1) lemma word_overflow_unat:"unat ((x::('a::len) word) + 1) = unat x + 1 \ x + 1 = 0" by (metis Suc_eq_plus1 add.commute unatSuc) lemma even_word_imp_odd_next:"even (unat (x::('a::len) word)) \ x + 1 = 0 \ odd (unat (x + 1))" apply (cut_tac x=x in word_overflow_unat) apply clarsimp done lemma odd_word_imp_even_next:"odd (unat (x::('a::len) word)) \ x + 1 = 0 \ even (unat (x + 1))" apply (cut_tac x=x in word_overflow_unat) apply clarsimp done lemma overflow_imp_lsb:"(x::('a::len) word) + 1 = 0 \ x !! 0" using even_plus_one_iff [of x] by (simp add: test_bit_word_eq) lemma odd_iff_lsb:"odd (unat (x::('a::len) word)) = x !! 0" by transfer (simp add: even_nat_iff) lemma of_nat_neq_iff_word: "x mod 2 ^ LENGTH('a) \ y mod 2 ^ LENGTH('a) \ (((of_nat x)::('a::len) word) \ of_nat y) = (x \ y)" apply (rule iffI) apply (case_tac "x = y") apply (subst (asm) of_nat_eq_iff[symmetric]) apply simp+ apply (case_tac "((of_nat x)::('a::len) word) = of_nat y") apply (subst (asm) word_unat.norm_eq_iff[symmetric]) apply simp+ done lemma shiftr1_irrelevant_lsb:"(x::('a::len) word) !! 0 \ x >> 1 = (x + 1) >> 1" apply (cases \LENGTH('a)\; transfer) apply (simp_all add: take_bit_drop_bit) apply (simp add: drop_bit_take_bit drop_bit_Suc) done lemma shiftr1_0_imp_only_lsb:"((x::('a::len) word) + 1) >> 1 = 0 \ x = 0 \ x + 1 = 0" by (metis One_nat_def shiftr1_0_or_1 word_less_1 word_overflow) lemma shiftr1_irrelevant_lsb':"\((x::('a::len) word) !! 0) \ x >> 1 = (x + 1) >> 1" by (metis shiftr1_irrelevant_lsb) lemma lsb_this_or_next:"\(((x::('a::len) word) + 1) !! 0) \ x !! 0" by (metis (poly_guards_query) even_word_imp_odd_next odd_iff_lsb overflow_imp_lsb) (* Perhaps this one should be a simp lemma, but it seems a little dangerous. *) lemma cast_chunk_assemble_id: "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ (((ucast ((ucast (x::'b word))::'a word))::'b word) OR (((ucast ((ucast (x >> n))::'a word))::'b word) << n)) = x" apply (subgoal_tac "((ucast ((ucast (x >> n))::'a word))::'b word) = x >> n") apply clarsimp apply (subst and_not_mask[symmetric]) apply (subst ucast_ucast_mask) apply (subst word_ao_dist2[symmetric]) apply clarsimp apply (rule ucast_ucast_len) apply (rule shiftr_less_t2n') apply (subst and_mask_eq_iff_le_mask) apply (simp_all add: mask_eq_decr_exp flip: mult_2_right) apply (metis add_diff_cancel_left' len_gt_0 mult_2_right zero_less_diff) done lemma cast_chunk_scast_assemble_id: "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ (((ucast ((scast (x::'b word))::'a word))::'b word) OR (((ucast ((scast (x >> n))::'a word))::'b word) << n)) = x" apply (subgoal_tac "((scast x)::'a word) = ((ucast x)::'a word)") apply (subgoal_tac "((scast (x >> n))::'a word) = ((ucast (x >> n))::'a word)") apply (simp add:cast_chunk_assemble_id) apply (subst down_cast_same[symmetric], subst is_down, arith, simp)+ done lemma mask_or_not_mask: "x AND mask n OR x AND NOT (mask n) = x" for x :: \'a::len word\ apply (subst word_oa_dist, simp) apply (subst word_oa_dist2, simp) done lemma is_aligned_add_not_aligned: "\is_aligned (p::'a::len word) n; \ is_aligned (q::'a::len word) n\ \ \ is_aligned (p + q) n" by (metis is_aligned_addD1) lemma word_gr0_conv_Suc: "(m::'a::len word) > 0 \ \n. m = n + 1" by (metis add.commute add_minus_cancel) lemma neg_mask_add_aligned: "\ is_aligned p n; q < 2 ^ n \ \ (p + q) AND NOT (mask n) = p AND NOT (mask n)" by (metis is_aligned_add_helper is_aligned_neg_mask_eq) lemma word_sless_sint_le:"x sint x \ sint y - 1" by (metis word_sless_alt zle_diff1_eq) lemma upper_trivial: fixes x :: "'a::len word" shows "x \ 2 ^ LENGTH('a) - 1 \ x < 2 ^ LENGTH('a) - 1" by (simp add: less_le) lemma constraint_expand: fixes x :: "'a::len word" shows "x \ {y. lower \ y \ y \ upper} = (lower \ x \ x \ upper)" by (rule mem_Collect_eq) lemma card_map_elide: "card ((of_nat :: nat \ 'a::len word) ` {0.. CARD('a::len word)" proof - let ?of_nat = "of_nat :: nat \ 'a word" from word_unat.Abs_inj_on have "inj_on ?of_nat {i. i < CARD('a word)}" by (simp add: unats_def card_word) moreover have "{0.. {i. i < CARD('a word)}" using that by auto ultimately have "inj_on ?of_nat {0.. CARD('a::len word) \ card ((of_nat::nat \ 'a::len word) ` {0..UCAST('b \ 'a) (UCAST('a \ 'b) x) = x\ if \x \ UCAST('b::len \ 'a) (- 1)\ for x :: \'a::len word\ proof - from that have a1: \x \ word_of_int (uint (word_of_int (2 ^ LENGTH('b) - 1) :: 'b word))\ by simp have f2: "((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ \ (0::int) \ - 1 + 2 ^ LENGTH('b) \ (0::int) \ - 1 + 2 ^ LENGTH('b) + - 1 * 2 ^ LENGTH('b) \ (- (1::int) + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b) = - 1 + 2 ^ LENGTH('b)) = ((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ \ (1::int) \ 2 ^ LENGTH('b) \ 2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)) = 1)" by force have f3: "\i ia. \ (0::int) \ i \ 0 \ i + - 1 * ia \ i mod ia = i" using mod_pos_pos_trivial by force have "(1::int) \ 2 ^ LENGTH('b)" by simp then have "2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ len_of TYPE ('b)) = 1" using f3 f2 by blast then have f4: "- (1::int) + 2 ^ LENGTH('b) = (- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)" by linarith have f5: "x \ word_of_int (uint (word_of_int (- 1 + 2 ^ LENGTH('b))::'b word))" using a1 by force have f6: "2 ^ LENGTH('b) + - (1::int) = - 1 + 2 ^ LENGTH('b)" by force have f7: "- (1::int) * 1 = - 1" by auto have "\x0 x1. (x1::int) - x0 = x1 + - 1 * x0" by force then have "x \ 2 ^ LENGTH('b) - 1" using f7 f6 f5 f4 by (metis uint_word_of_int wi_homs(2) word_arith_wis(8) word_of_int_2p) then have \uint x \ uint (2 ^ LENGTH('b) - (1 :: 'a word))\ by (simp add: word_le_def) then have \uint x \ 2 ^ LENGTH('b) - 1\ by (simp add: uint_word_ariths) (metis \1 \ 2 ^ LENGTH('b)\ \uint x \ uint (2 ^ LENGTH('b) - 1)\ linorder_not_less lt2p_lem uint_1 uint_minus_simple_alt uint_power_lower word_le_def zle_diff1_eq) then show ?thesis apply (simp add: word_ubin.eq_norm bintrunc_mod2p unsigned_ucast_eq) apply (metis \x \ 2 ^ LENGTH('b) - 1\ and_mask_eq_iff_le_mask mask_eq_decr_exp take_bit_eq_mask) done qed lemma remdups_enum_upto: fixes s::"'a::len word" shows "remdups [s .e. e] = [s .e. e]" by simp lemma card_enum_upto: fixes s::"'a::len word" shows "card (set [s .e. e]) = Suc (unat e) - unat s" by (subst List.card_set) (simp add: remdups_enum_upto) lemma complement_nth_w2p: shows "n' < LENGTH('a) \ (NOT (2 ^ n :: 'a::len word)) !! n' = (n' \ n)" by (fastforce simp: word_ops_nth_size word_size nth_w2p) lemma word_unat_and_lt: "unat x < n \ unat y < n \ unat (x AND y) < n" by (meson le_less_trans word_and_le1 word_and_le2 word_le_nat_alt) lemma word_unat_mask_lt: "m \ size w \ unat ((w::'a::len word) AND mask m) < 2 ^ m" by (rule word_unat_and_lt) (simp add: unat_mask word_size) lemma unat_shiftr_less_t2n: fixes x :: "'a :: len word" shows "unat x < 2 ^ (n + m) \ unat (x >> n) < 2 ^ m" by (simp add: shiftr_div_2n' power_add mult.commute less_mult_imp_div_less) lemma le_or_mask: "w \ w' \ w OR mask x \ w' OR mask x" for w w' :: \'a::len word\ by (metis neg_mask_add_mask add.commute le_word_or1 mask_2pm1 neg_mask_mono_le word_plus_mono_left) lemma le_shiftr1': "\ shiftr1 u \ shiftr1 v ; shiftr1 u \ shiftr1 v \ \ u \ v" apply transfer apply simp done lemma le_shiftr': "\ u >> n \ v >> n ; u >> n \ v >> n \ \ (u::'a::len word) \ v" apply (induct n; simp add: shiftr_def) apply (case_tac "(shiftr1 ^^ n) u = (shiftr1 ^^ n) v", simp) apply (fastforce dest: le_shiftr1') done lemma word_add_no_overflow:"(x::'a::len word) < - 1 \ x < x + 1" using less_x_plus_1 order_less_le by blast lemma lt_plus_1_le_word: fixes x :: "'a::len word" assumes bound:"n < unat (maxBound::'a word)" shows "x < 1 + of_nat n = (x \ of_nat n)" by (metis add.commute bound max_word_max word_Suc_leq word_not_le word_of_nat_less) lemma unat_ucast_up_simp: fixes x :: "'a::len word" assumes "LENGTH('a) \ LENGTH('b)" shows "unat (ucast x :: 'b::len word) = unat x" unfolding ucast_eq unat_eq_nat_uint apply (subst int_word_uint) apply (subst mod_pos_pos_trivial; simp?) apply (rule lt2p_lem) apply (simp add: assms) done lemma unat_ucast_less_no_overflow: "\n < 2 ^ LENGTH('a); unat f < n\ \ (f::('a::len) word) < of_nat n" by (erule (1) order_le_less_trans[OF _ of_nat_mono_maybe,rotated]) simp lemma unat_ucast_less_no_overflow_simp: "n < 2 ^ LENGTH('a) \ (unat f < n) = ((f::('a::len) word) < of_nat n)" using unat_less_helper unat_ucast_less_no_overflow by blast lemma unat_ucast_no_overflow_le: assumes no_overflow: "unat b < (2 :: nat) ^ LENGTH('a)" and upward_cast: "LENGTH('a) < LENGTH('b)" shows "(ucast (f::'a::len word) < (b :: 'b :: len word)) = (unat f < unat b)" proof - have LR: "ucast f < b \ unat f < unat b" apply (rule unat_less_helper) apply (simp add:ucast_nat_def) apply (rule_tac 'b1 = 'b in ucast_less_ucast[OF order.strict_implies_order, THEN iffD1]) apply (rule upward_cast) apply (simp add: ucast_ucast_mask less_mask_eq word_less_nat_alt unat_power_lower[OF upward_cast] no_overflow) done have RL: "unat f < unat b \ ucast f < b" proof- assume ineq: "unat f < unat b" have "ucast (f::'a::len word) < ((ucast (ucast b ::'a::len word)) :: 'b :: len word)" apply (simp add: ucast_less_ucast[OF order.strict_implies_order] upward_cast) apply (simp only: flip: ucast_nat_def) apply (rule unat_ucast_less_no_overflow[OF no_overflow ineq]) done then show ?thesis apply (rule order_less_le_trans) apply (simp add:ucast_ucast_mask word_and_le2) done qed then show ?thesis by (simp add:RL LR iffI) qed lemmas ucast_up_mono = ucast_less_ucast[THEN iffD2] lemma minus_one_word: "(-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1" by simp lemma mask_exceed: "n \ LENGTH('a) \ (x::'a::len word) AND NOT (mask n) = 0" by (simp add: and_not_mask shiftr_eq_0) lemma word_shift_by_2: "x * 4 = (x::'a::len word) << 2" by (simp add: shiftl_t2n) lemma le_2p_upper_bits: "\ (p::'a::len word) \ 2^n - 1; n < LENGTH('a) \ \ \n'\n. n' < LENGTH('a) \ \ p !! n'" by (subst upper_bits_unset_is_l2p; simp) lemma le2p_bits_unset: "p \ 2 ^ n - 1 \ \n'\n. n' < LENGTH('a) \ \ (p::'a::len word) !! n'" using upper_bits_unset_is_l2p [where p=p] by (cases "n < LENGTH('a)") auto lemma ucast_less_shiftl_helper: "\ LENGTH('b) + 2 < LENGTH('a); 2 ^ (LENGTH('b) + 2) \ n\ \ (ucast (x :: 'b::len word) << 2) < (n :: 'a::len word)" apply (erule order_less_le_trans[rotated]) using ucast_less[where x=x and 'a='a] apply (simp only: shiftl_t2n field_simps) apply (rule word_less_power_trans2; simp) done lemma word_power_nonzero: "\ (x :: 'a::len word) < 2 ^ (LENGTH('a) - n); n < LENGTH('a); x \ 0 \ \ x * 2 ^ n \ 0" by (metis and_mask_eq_iff_shiftr_0 less_mask_eq p2_gt_0 semiring_normalization_rules(7) shiftl_shiftr_id shiftl_t2n) lemma less_1_helper: "n \ m \ (n - 1 :: int) < m" by arith lemma div_power_helper: "\ x \ y; y < LENGTH('a) \ \ (2 ^ y - 1) div (2 ^ x :: 'a::len word) = 2 ^ (y - x) - 1" apply (rule word_uint.Rep_eqD) apply (simp only: uint_word_ariths uint_div uint_power_lower) apply (subst mod_pos_pos_trivial, fastforce, fastforce)+ apply (subst mod_pos_pos_trivial) apply (simp add: le_diff_eq uint_2p_alt) apply (rule less_1_helper) apply (rule power_increasing; simp) apply (subst mod_pos_pos_trivial) apply (simp add: uint_2p_alt) apply (rule less_1_helper) apply (rule power_increasing; simp) apply (subst int_div_sub_1; simp add: uint_2p_alt) apply (subst power_0[symmetric]) apply (simp add: uint_2p_alt le_imp_power_dvd power_diff_power_eq) done lemma word_add_power_off: fixes a :: "'a :: len word" assumes ak: "a < k" and kw: "k < 2 ^ (LENGTH('a) - m)" and mw: "m < LENGTH('a)" and off: "off < 2 ^ m" shows "(a * 2 ^ m) + off < k * 2 ^ m" proof (cases "m = 0") case True then show ?thesis using off ak by simp next case False from ak have ak1: "a + 1 \ k" by (rule inc_le) then have "(a + 1) * 2 ^ m \ 0" apply - apply (rule word_power_nonzero) apply (erule order_le_less_trans [OF _ kw]) apply (rule mw) apply (rule less_is_non_zero_p1 [OF ak]) done then have "(a * 2 ^ m) + off < ((a + 1) * 2 ^ m)" using kw mw apply - apply (simp add: distrib_right) apply (rule word_plus_strict_mono_right [OF off]) apply (rule is_aligned_no_overflow'') apply (rule is_aligned_mult_triv2) apply assumption done also have "\ \ k * 2 ^ m" using ak1 mw kw False apply - apply (erule word_mult_le_mono1) apply (simp add: p2_gt_0) apply (simp add: word_less_nat_alt) apply (meson nat_mult_power_less_eq zero_less_numeral) done finally show ?thesis . qed lemma offset_not_aligned: "\ is_aligned (p::'a::len word) n; i > 0; i < 2 ^ n; n < LENGTH('a)\ \ \ is_aligned (p + of_nat i) n" apply (erule is_aligned_add_not_aligned) apply transfer apply (auto simp add: is_aligned_iff_udvd) apply (metis bintrunc_bintrunc_ge int_ops(1) nat_int_comparison(1) nat_less_le take_bit_eq_0_iff take_bit_nat_eq_self_iff take_bit_of_nat) done lemma length_upto_enum_one: fixes x :: "'a :: len word" assumes lt1: "x < y" and lt2: "z < y" and lt3: "x \ z" shows "[x , y .e. z] = [x]" unfolding upto_enum_step_def proof (subst upto_enum_red, subst if_not_P [OF leD [OF lt3]], clarsimp, rule conjI) show "unat ((z - x) div (y - x)) = 0" proof (subst unat_div, rule div_less) have syx: "unat (y - x) = unat y - unat x" by (rule unat_sub [OF order_less_imp_le]) fact moreover have "unat (z - x) = unat z - unat x" by (rule unat_sub) fact ultimately show "unat (z - x) < unat (y - x)" using lt2 lt3 unat_mono word_less_minus_mono_left by blast qed then show "(z - x) div (y - x) * (y - x) = 0" by (metis mult_zero_left unat_0 word_unat.Rep_eqD) qed lemma max_word_mask: "(- 1 :: 'a::len word) = mask LENGTH('a)" by (fact minus_1_eq_mask) lemmas mask_len_max = max_word_mask[symmetric] lemma mask_out_first_mask_some: "\ x AND NOT (mask n) = y; n \ m \ \ x AND NOT (mask m) = y AND NOT (mask m)" for x y :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma mask_lower_twice: "n \ m \ (x AND NOT (mask n)) AND NOT (mask m) = x AND NOT (mask m)" for x :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma mask_lower_twice2: "(a AND NOT (mask n)) AND NOT (mask m) = a AND NOT (mask (max n m))" for a :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma ucast_and_neg_mask: "ucast (x AND NOT (mask n)) = ucast x AND NOT (mask n)" apply (rule bit_word_eqI) apply (auto simp add: bit_simps dest: bit_imp_le_length) done lemma ucast_and_mask: "ucast (x AND mask n) = ucast x AND mask n" apply (rule bit_word_eqI) apply (auto simp add: bit_simps dest: bit_imp_le_length) done lemma ucast_mask_drop: "LENGTH('a :: len) \ n \ (ucast (x AND mask n) :: 'a word) = ucast x" apply (rule bit_word_eqI) apply (auto simp add: bit_simps dest: bit_imp_le_length) done (* negating a mask which has been shifted to the very left *) lemma NOT_mask_shifted_lenword: "NOT (mask len << (LENGTH('a) - len) ::'a::len word) = mask (LENGTH('a) - len)" by (rule bit_word_eqI) (auto simp add: shiftl_word_eq word_size bit_not_iff bit_push_bit_iff bit_mask_iff) (* Comparisons between different word sizes. *) lemma eq_ucast_ucast_eq: "LENGTH('b) \ LENGTH('a) \ x = ucast y \ ucast x = y" for x :: "'a::len word" and y :: "'b::len word" by transfer simp lemma le_ucast_ucast_le: "x \ ucast y \ ucast x \ y" for x :: "'a::len word" and y :: "'b::len word" by (smt le_unat_uoi linorder_not_less order_less_imp_le ucast_nat_def unat_arith_simps(1)) lemma less_ucast_ucast_less: "LENGTH('b) \ LENGTH('a) \ x < ucast y \ ucast x < y" for x :: "'a::len word" and y :: "'b::len word" by (metis ucast_nat_def unat_mono unat_ucast_up_simp word_of_nat_less) lemma ucast_le_ucast: "LENGTH('a) \ LENGTH('b) \ (ucast x \ (ucast y::'b::len word)) = (x \ y)" for x :: "'a::len word" by (simp add: unat_arith_simps(1) unat_ucast_up_simp) lemmas ucast_up_mono_le = ucast_le_ucast[THEN iffD2] lemma ucast_le_ucast_eq: fixes x y :: "'a::len word" assumes x: "x < 2 ^ n" assumes y: "y < 2 ^ n" assumes n: "n = LENGTH('b::len)" shows "(UCAST('a \ 'b) x \ UCAST('a \ 'b) y) = (x \ y)" apply (rule iffI) apply (cases "LENGTH('b) < LENGTH('a)") apply (subst less_mask_eq[OF x, symmetric]) apply (subst less_mask_eq[OF y, symmetric]) apply (unfold n) apply (subst ucast_ucast_mask[symmetric])+ apply (simp add: ucast_le_ucast)+ apply (erule ucast_mono_le[OF _ y[unfolded n]]) done lemma ucast_or_distrib: fixes x :: "'a::len word" fixes y :: "'a::len word" shows "(ucast (x OR y) :: ('b::len) word) = ucast x OR ucast y" by transfer simp lemma shiftr_less: "(w::'a::len word) < k \ w >> n < k" by (metis div_le_dividend le_less_trans shiftr_div_2n' unat_arith_simps(2)) lemma word_and_notzeroD: "w AND w' \ 0 \ w \ 0 \ w' \ 0" by auto lemma word_exists_nth: "(w::'a::len word) \ 0 \ \i. w !! i" by (simp add: bit_eq_iff test_bit_word_eq) lemma shiftr_le_0: "unat (w::'a::len word) < 2 ^ n \ w >> n = (0::'a::len word)" by (rule word_unat.Rep_eqD) (simp add: shiftr_div_2n') lemma of_nat_shiftl: "(of_nat x << n) = (of_nat (x * 2 ^ n) :: ('a::len) word)" proof - have "(of_nat x::'a word) << n = of_nat (2 ^ n) * of_nat x" using shiftl_t2n by (metis word_unat_power) thus ?thesis by simp qed lemma shiftl_1_not_0: "n < LENGTH('a) \ (1::'a::len word) << n \ 0" by (simp add: shiftl_t2n) lemma max_word_not_0 [simp]: "- 1 \ (0 :: 'a::len word)" by simp lemma ucast_zero_is_aligned: "UCAST('a::len \ 'b::len) w = 0 \ n \ LENGTH('b) \ is_aligned w n" by (clarsimp simp: is_aligned_mask word_eq_iff word_size nth_ucast) lemma unat_ucast_eq_unat_and_mask: "unat (UCAST('b::len \ 'a::len) w) = unat (w AND mask LENGTH('a))" apply (simp flip: take_bit_eq_mask) apply transfer apply (simp add: ac_simps) done lemma unat_max_word_pos[simp]: "0 < unat (- 1 :: 'a::len word)" using unat_gt_0 [of \- 1 :: 'a::len word\] by simp (* Miscellaneous conditional injectivity rules. *) lemma mult_pow2_inj: assumes ws: "m + n \ LENGTH('a)" assumes le: "x \ mask m" "y \ mask m" assumes eq: "x * 2 ^ n = y * (2 ^ n::'a::len word)" shows "x = y" proof (rule bit_word_eqI) fix q assume \q < LENGTH('a)\ from eq have \push_bit n x = push_bit n y\ by (simp add: push_bit_eq_mult) moreover from le have \take_bit m x = x\ \take_bit m y = y\ by (simp_all add: less_eq_mask_iff_take_bit_eq_self) ultimately have \push_bit n (take_bit m x) = push_bit n (take_bit m y)\ by simp_all with \q < LENGTH('a)\ ws show \bit x q \ bit y q\ apply (simp add: push_bit_take_bit) unfolding bit_eq_iff apply (simp add: bit_simps not_le) apply (metis (full_types) \take_bit m x = x\ \take_bit m y = y\ add.commute add_diff_cancel_right' add_less_cancel_right bit_take_bit_iff le_add2 less_le_trans) done qed lemma word_of_nat_inj: assumes bounded: "x < 2 ^ LENGTH('a)" "y < 2 ^ LENGTH('a)" assumes of_nats: "of_nat x = (of_nat y :: 'a::len word)" shows "x = y" by (rule contrapos_pp[OF of_nats]; cases "x < y"; cases "y < x") (auto dest: bounded[THEN of_nat_mono_maybe]) (* Uints *) lemma uints_mono_iff: "uints l \ uints m \ l \ m" using power_increasing_iff[of "2::int" l m] apply (auto simp: uints_num subset_iff simp del: power_increasing_iff) apply (meson less_irrefl not_le zero_le_numeral zero_le_power) done lemmas uints_monoI = uints_mono_iff[THEN iffD2] lemma Bit_in_uints_Suc: "of_bool c + 2 * w \ uints (Suc m)" if "w \ uints m" using that by (auto simp: uints_num) lemma Bit_in_uintsI: "of_bool c + 2 * w \ uints m" if "w \ uints (m - 1)" "m > 0" using Bit_in_uints_Suc[OF that(1)] that(2) by auto lemma bin_cat_in_uintsI: - \bin_cat a n b \ uints m\ if \a \ uints l\ \m \ l + n\ + \concat_bit n b a \ uints m\ if \a \ uints l\ \m \ l + n\ proof - from \m \ l + n\ obtain q where \m = l + n + q\ using le_Suc_ex by blast then have \(2::int) ^ m = 2 ^ n * 2 ^ (l + q)\ by (simp add: ac_simps power_add) moreover have \a mod 2 ^ (l + q) = a\ using \a \ uints l\ by (auto simp add: uints_def take_bit_eq_mod power_add Divides.mod_mult2_eq) ultimately have \concat_bit n b a = take_bit m (concat_bit n b a)\ by (simp add: concat_bit_eq take_bit_eq_mod push_bit_eq_mult Divides.mod_mult2_eq) then show ?thesis by (simp add: uints_def) qed -lemma bin_cat_cong: "bin_cat a n b = bin_cat c m d" - if "n = m" "a = c" "bintrunc m b = bintrunc m d" +lemma bin_cat_cong: "concat_bit n b a = concat_bit m d c" + if "n = m" "a = c" "take_bit m b = take_bit m d" using that(3) unfolding that(1,2) by (simp add: bin_cat_eq_push_bit_add_take_bit) -lemma bin_cat_eqD1: "bin_cat a n b = bin_cat c n d \ a = c" +lemma bin_cat_eqD1: "concat_bit n b a = concat_bit n d c \ a = c" by (metis drop_bit_bin_cat_eq) -lemma bin_cat_eqD2: "bin_cat a n b = bin_cat c n d \ bintrunc n b = bintrunc n d" +lemma bin_cat_eqD2: "concat_bit n b a = concat_bit n d c \ take_bit n b = take_bit n d" by (metis take_bit_bin_cat_eq) -lemma bin_cat_inj: "(bin_cat a n b) = bin_cat c n d \ a = c \ bintrunc n b = bintrunc n d" +lemma bin_cat_inj: "(concat_bit n b a) = concat_bit n d c \ a = c \ take_bit n b = take_bit n d" by (auto intro: bin_cat_cong bin_cat_eqD1 bin_cat_eqD2) lemma word_of_int_bin_cat_eq_iff: - "(word_of_int (bin_cat (uint a) LENGTH('b) (uint b))::'c::len word) = - word_of_int (bin_cat (uint c) LENGTH('b) (uint d)) \ b = d \ a = c" + "(word_of_int (concat_bit LENGTH('b) (uint b) (uint a))::'c::len word) = + word_of_int (concat_bit LENGTH('b) (uint d) (uint c)) \ b = d \ a = c" if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" for a::"'a::len word" and b::"'b::len word" by (subst word_uint.Abs_inject) (auto simp: bin_cat_inj intro!: that bin_cat_in_uintsI) lemma word_cat_inj: "(word_cat a b::'c::len word) = word_cat c d \ a = c \ b = d" if "LENGTH('a) + LENGTH('b) \ LENGTH('c)" for a::"'a::len word" and b::"'b::len word" using word_of_int_bin_cat_eq_iff [OF that, of b a d c] by transfer auto lemma p2_eq_1: "2 ^ n = (1::'a::len word) \ n = 0" proof - have "2 ^ n = (1::'a word) \ n = 0" by (metis One_nat_def not_less one_less_numeral_iff p2_eq_0 p2_gt_0 power_0 power_0 power_inject_exp semiring_norm(76) unat_power_lower zero_neq_one) then show ?thesis by auto qed (* usually: x,y = (len_of TYPE ('a)) *) lemma bitmagic_zeroLast_leq_or1Last: "(a::('a::len) word) AND (mask len << x - len) \ a OR mask (y - len)" by (meson le_word_or2 order_trans word_and_le2) lemma zero_base_lsb_imp_set_eq_as_bit_operation: fixes base ::"'a::len word" assumes valid_prefix: "mask (LENGTH('a) - len) AND base = 0" shows "(base = NOT (mask (LENGTH('a) - len)) AND a) \ (a \ {base .. base OR mask (LENGTH('a) - len)})" proof have helper3: "x OR y = x OR y AND NOT x" for x y ::"'a::len word" by (simp add: word_oa_dist2) from assms show "base = NOT (mask (LENGTH('a) - len)) AND a \ a \ {base..base OR mask (LENGTH('a) - len)}" apply(simp add: word_and_le1) apply(metis helper3 le_word_or2 word_bw_comms(1) word_bw_comms(2)) done next assume "a \ {base..base OR mask (LENGTH('a) - len)}" hence a: "base \ a \ a \ base OR mask (LENGTH('a) - len)" by simp show "base = NOT (mask (LENGTH('a) - len)) AND a" proof - have f2: "\x\<^sub>0. base AND NOT (mask x\<^sub>0) \ a AND NOT (mask x\<^sub>0)" using a neg_mask_mono_le by blast have f3: "\x\<^sub>0. a AND NOT (mask x\<^sub>0) \ (base OR mask (LENGTH('a) - len)) AND NOT (mask x\<^sub>0)" using a neg_mask_mono_le by blast have f4: "base = base AND NOT (mask (LENGTH('a) - len))" using valid_prefix by (metis mask_eq_0_eq_x word_bw_comms(1)) hence f5: "\x\<^sub>6. (base OR x\<^sub>6) AND NOT (mask (LENGTH('a) - len)) = base OR x\<^sub>6 AND NOT (mask (LENGTH('a) - len))" using word_ao_dist by (metis) have f6: "\x\<^sub>2 x\<^sub>3. a AND NOT (mask x\<^sub>2) \ x\<^sub>3 \ \ (base OR mask (LENGTH('a) - len)) AND NOT (mask x\<^sub>2) \ x\<^sub>3" using f3 dual_order.trans by auto have "base = (base OR mask (LENGTH('a) - len)) AND NOT (mask (LENGTH('a) - len))" using f5 by auto hence "base = a AND NOT (mask (LENGTH('a) - len))" using f2 f4 f6 by (metis eq_iff) thus "base = NOT (mask (LENGTH('a) - len)) AND a" by (metis word_bw_comms(1)) qed qed lemma of_nat_eq_signed_scast: "(of_nat x = (y :: ('a::len) signed word)) = (of_nat x = (scast y :: 'a word))" by (metis scast_of_nat scast_scast_id(2)) lemma word_aligned_add_no_wrap_bounded: "\ w + 2^n \ x; w + 2^n \ 0; is_aligned w n \ \ (w::'a::len word) < x" by (blast dest: is_aligned_no_overflow le_less_trans word_leq_le_minus_one) lemma mask_Suc: "mask (Suc n) = (2 :: 'a::len word) ^ n + mask n" by (simp add: mask_eq_decr_exp) lemma mask_mono: "sz' \ sz \ mask sz' \ (mask sz :: 'a::len word)" by (simp add: le_mask_iff shiftr_mask_le) lemma aligned_mask_disjoint: "\ is_aligned (a :: 'a :: len word) n; b \ mask n \ \ a AND b = 0" by (metis and_zero_eq is_aligned_mask le_mask_imp_and_mask word_bw_lcs(1)) lemma word_and_or_mask_aligned: "\ is_aligned a n; b \ mask n \ \ a + b = a OR b" by (simp add: aligned_mask_disjoint word_plus_and_or_coroll) lemma word_and_or_mask_aligned2: \is_aligned b n \ a \ mask n \ a + b = a OR b\ using word_and_or_mask_aligned [of b n a] by (simp add: ac_simps) lemma is_aligned_ucastI: "is_aligned w n \ is_aligned (ucast w) n" apply transfer apply (auto simp add: min_def) apply (metis bintrunc_bintrunc_ge bintrunc_n_0 nat_less_le not_le take_bit_eq_0_iff) done lemma ucast_le_maskI: "a \ mask n \ UCAST('a::len \ 'b::len) a \ mask n" by (metis and_mask_eq_iff_le_mask ucast_and_mask) lemma ucast_add_mask_aligned: "\ a \ mask n; is_aligned b n \ \ UCAST ('a::len \ 'b::len) (a + b) = ucast a + ucast b" by (metis add.commute is_aligned_ucastI ucast_le_maskI ucast_or_distrib word_and_or_mask_aligned) lemma ucast_shiftl: "LENGTH('b) \ LENGTH ('a) \ UCAST ('a::len \ 'b::len) x << n = ucast (x << n)" by word_eqI_solve lemma ucast_leq_mask: "LENGTH('a) \ n \ ucast (x::'a::len word) \ mask n" apply (simp add: less_eq_mask_iff_take_bit_eq_self) apply transfer apply (simp add: ac_simps) done lemma shiftl_inj: "\ x << n = y << n; x \ mask (LENGTH('a)-n); y \ mask (LENGTH('a)-n) \ \ x = (y :: 'a :: len word)" apply word_eqI apply (rename_tac n') apply (case_tac "LENGTH('a) - n \ n'", simp) by (metis add.commute add.right_neutral diff_add_inverse le_diff_conv linorder_not_less zero_order(1)) lemma distinct_word_add_ucast_shift_inj: "\ p + (UCAST('a::len \ 'b::len) off << n) = p' + (ucast off' << n); is_aligned p n'; is_aligned p' n'; n' = n + LENGTH('a); n' < LENGTH('b) \ \ p' = p \ off' = off" apply (simp add: word_and_or_mask_aligned le_mask_shiftl_le_mask[where n="LENGTH('a)"] ucast_leq_mask) apply (simp add: is_aligned_nth) apply (rule conjI; word_eqI) apply (metis add.commute test_bit_conj_lt diff_add_inverse le_diff_conv nat_less_le) apply (rename_tac i) apply (erule_tac x="i+n" in allE) apply simp done lemma word_upto_Nil: "y < x \ [x .e. y ::'a::len word] = []" by (simp add: upto_enum_red not_le word_less_nat_alt) lemma word_enum_decomp_elem: assumes "[x .e. (y ::'a::len word)] = as @ a # bs" shows "x \ a \ a \ y" proof - have "set as \ set [x .e. y] \ a \ set [x .e. y]" using assms by (auto dest: arg_cong[where f=set]) then show ?thesis by auto qed lemma word_enum_prefix: "[x .e. (y ::'a::len word)] = as @ a # bs \ as = (if x < a then [x .e. a - 1] else [])" apply (induct as arbitrary: x; clarsimp) apply (case_tac "x < y") prefer 2 apply (case_tac "x = y", simp) apply (simp add: not_less) apply (drule (1) dual_order.not_eq_order_implies_strict) apply (simp add: word_upto_Nil) apply (simp add: word_upto_Cons_eq) apply (case_tac "x < y") prefer 2 apply (case_tac "x = y", simp) apply (simp add: not_less) apply (drule (1) dual_order.not_eq_order_implies_strict) apply (simp add: word_upto_Nil) apply (clarsimp simp: word_upto_Cons_eq) apply (frule word_enum_decomp_elem) apply clarsimp apply (rule conjI) prefer 2 apply (subst word_Suc_le[symmetric]; clarsimp) apply (drule meta_spec) apply (drule (1) meta_mp) apply clarsimp apply (rule conjI; clarsimp) apply (subst (2) word_upto_Cons_eq) apply unat_arith apply simp done lemma word_enum_decomp_set: "[x .e. (y ::'a::len word)] = as @ a # bs \ a \ set as" by (metis distinct_append distinct_enum_upto' not_distinct_conv_prefix) lemma word_enum_decomp: assumes "[x .e. (y ::'a::len word)] = as @ a # bs" shows "x \ a \ a \ y \ a \ set as \ (\z \ set as. x \ z \ z \ y)" proof - from assms have "set as \ set [x .e. y] \ a \ set [x .e. y]" by (auto dest: arg_cong[where f=set]) with word_enum_decomp_set[OF assms] show ?thesis by auto qed lemma of_nat_unat_le_mask_ucast: "\of_nat (unat t) = w; t \ mask LENGTH('a)\ \ t = UCAST('a::len \ 'b::len) w" by (clarsimp simp: ucast_nat_def ucast_ucast_mask simp flip: and_mask_eq_iff_le_mask) lemma less_diff_gt0: "a < b \ (0 :: 'a :: len word) < b - a" by unat_arith lemma unat_plus_gt: "unat ((a :: 'a :: len word) + b) \ unat a + unat b" by (clarsimp simp: unat_plus_if_size) lemma const_less: "\ (a :: 'a :: len word) - 1 < b; a \ b \ \ a < b" by (metis less_1_simp word_le_less_eq) lemma add_mult_aligned_neg_mask: \(x + y * m) AND NOT(mask n) = (x AND NOT(mask n)) + y * m\ if \m AND (2 ^ n - 1) = 0\ for x y m :: \'a::len word\ by (metis (no_types, hide_lams) add.assoc add.commute add.right_neutral add_uminus_conv_diff mask_eq_decr_exp mask_eqs(2) mask_eqs(6) mult.commute mult_zero_left subtract_mask(1) that) lemma unat_of_nat_minus_1: "\ n < 2 ^ LENGTH('a); n \ 0 \ \ unat ((of_nat n:: 'a :: len word) - 1) = n - 1" by (simp add: of_nat_diff unat_eq_of_nat) lemma word_eq_zeroI: "a \ a - 1 \ a = 0" for a :: "'a :: len word" by (simp add: word_must_wrap) lemma word_add_format: "(-1 :: 'a :: len word) + b + c = b + (c - 1)" by simp lemma upto_enum_word_nth: "\ i \ j; k \ unat (j - i) \ \ [i .e. j] ! k = i + of_nat k" apply (clarsimp simp: upto_enum_def nth_append) apply (clarsimp simp: word_le_nat_alt[symmetric]) apply (rule conjI, clarsimp) apply (subst toEnum_of_nat, unat_arith) apply unat_arith apply (clarsimp simp: not_less unat_sub[symmetric]) apply unat_arith done lemma upto_enum_step_nth: "\ a \ c; n \ unat ((c - a) div (b - a)) \ \ [a, b .e. c] ! n = a + of_nat n * (b - a)" by (clarsimp simp: upto_enum_step_def not_less[symmetric] upto_enum_word_nth) lemma upto_enum_inc_1_len: "a < - 1 \ [(0 :: 'a :: len word) .e. 1 + a] = [0 .e. a] @ [1 + a]" apply (simp add: upto_enum_word) apply (subgoal_tac "unat (1+a) = 1 + unat a") apply simp apply (subst unat_plus_simple[THEN iffD1]) apply (metis add.commute no_plus_overflow_neg olen_add_eqv) apply unat_arith done lemma neg_mask_add: "y AND mask n = 0 \ x + y AND NOT(mask n) = (x AND NOT(mask n)) + y" for x y :: \'a::len word\ by (clarsimp simp: mask_out_sub_mask mask_eqs(7)[symmetric] mask_twice) lemma shiftr_shiftl_shiftr[simp]: "(x :: 'a :: len word) >> a << a >> a = x >> a" by word_eqI_solve lemma add_right_shift: "\ x AND mask n = 0; y AND mask n = 0; x \ x + y \ \ (x + y :: ('a :: len) word) >> n = (x >> n) + (y >> n)" apply (simp add: no_olen_add_nat is_aligned_mask[symmetric]) apply (simp add: unat_arith_simps shiftr_div_2n' split del: if_split) apply (subst if_P) apply (erule order_le_less_trans[rotated]) apply (simp add: add_mono) apply (simp add: shiftr_div_2n' is_aligned_iff_dvd_nat) done lemma sub_right_shift: "\ x AND mask n = 0; y AND mask n = 0; y \ x \ \ (x - y) >> n = (x >> n :: 'a :: len word) - (y >> n)" using add_right_shift[where x="x - y" and y=y and n=n] by (simp add: aligned_sub_aligned is_aligned_mask[symmetric] word_sub_le) lemma and_and_mask_simple: "y AND mask n = mask n \ (x AND y) AND mask n = x AND mask n" by (simp add: ac_simps) lemma and_and_mask_simple_not: "y AND mask n = 0 \ (x AND y) AND mask n = 0" by (simp add: ac_simps) lemma word_and_le': "b \ c \ (a :: 'a :: len word) AND b \ c" by (metis word_and_le1 order_trans) lemma word_and_less': "b < c \ (a :: 'a :: len word) AND b < c" by transfer simp lemma shiftr_w2p: "x < LENGTH('a) \ 2 ^ x = (2 ^ (LENGTH('a) - 1) >> (LENGTH('a) - 1 - x) :: 'a :: len word)" by word_eqI_solve lemma t2p_shiftr: "\ b \ a; a < LENGTH('a) \ \ (2 :: 'a :: len word) ^ a >> b = 2 ^ (a - b)" by word_eqI_solve lemma scast_1[simp]: "scast (1 :: 'a :: len signed word) = (1 :: 'a word)" by simp lemma unsigned_uminus1 [simp]: \(unsigned (-1::'b::len word)::'c::len word) = mask LENGTH('b)\ by word_eqI lemma ucast_ucast_mask_eq: "\ UCAST('a::len \ 'b::len) x = y; x AND mask LENGTH('b) = x \ \ x = ucast y" by word_eqI_solve lemma ucast_up_eq: "\ ucast x = (ucast y::'b::len word); LENGTH('a) \ LENGTH ('b) \ \ ucast x = (ucast y::'a::len word)" by word_eqI_solve lemma ucast_up_neq: "\ ucast x \ (ucast y::'b::len word); LENGTH('b) \ LENGTH ('a) \ \ ucast x \ (ucast y::'a::len word)" by (fastforce dest: ucast_up_eq) lemma mask_AND_less_0: "\ x AND mask n = 0; m \ n \ \ x AND mask m = 0" for x :: \'a::len word\ by (metis mask_twice2 word_and_notzeroD) lemma mask_len_id [simp]: "(x :: 'a :: len word) AND mask LENGTH('a) = x" using uint_lt2p [of x] by (simp add: mask_eq_iff) lemma scast_ucast_down_same: "LENGTH('b) \ LENGTH('a) \ SCAST('a \ 'b) = UCAST('a::len \ 'b::len)" by (simp add: down_cast_same is_down) lemma word_aligned_0_sum: "\ a + b = 0; is_aligned (a :: 'a :: len word) n; b \ mask n; n < LENGTH('a) \ \ a = 0 \ b = 0" by (simp add: word_plus_and_or_coroll aligned_mask_disjoint word_or_zero) lemma mask_eq1_nochoice: "\ LENGTH('a) > 1; (x :: 'a :: len word) AND 1 = x \ \ x = 0 \ x = 1" by (metis word_and_1) lemma shiftr_and_eq_shiftl: "(w >> n) AND x = y \ w AND (x << n) = (y << n)" for y :: "'a:: len word" by (metis (no_types, lifting) and_not_mask bit.conj_ac(1) bit.conj_ac(2) mask_eq_0_eq_x shiftl_mask_is_0 shiftl_over_and_dist) lemma add_mask_lower_bits': "\ len = LENGTH('a); is_aligned (x :: 'a :: len word) n; \n' \ n. n' < len \ \ p !! n' \ \ x + p AND NOT(mask n) = x" using add_mask_lower_bits by auto lemma leq_mask_shift: "(x :: 'a :: len word) \ mask (low_bits + high_bits) \ (x >> low_bits) \ mask high_bits" by (simp add: le_mask_iff shiftr_shiftr) lemma ucast_ucast_eq_mask_shift: "(x :: 'a :: len word) \ mask (low_bits + LENGTH('b)) \ ucast((ucast (x >> low_bits)) :: 'b :: len word) = x >> low_bits" by (meson and_mask_eq_iff_le_mask eq_ucast_ucast_eq not_le_imp_less shiftr_less_t2n' ucast_ucast_len) lemma const_le_unat: "\ b < 2 ^ LENGTH('a); of_nat b \ a \ \ b \ unat (a :: 'a :: len word)" apply (simp add: word_le_def) apply (simp only: uint_nat zle_int) apply transfer apply (simp add: take_bit_nat_eq_self) done lemma upt_enum_offset_trivial: "\ x < 2 ^ LENGTH('a) - 1 ; n \ unat x \ \ ([(0 :: 'a :: len word) .e. x] ! n) = of_nat n" apply (induct x arbitrary: n) apply simp by (simp add: upto_enum_word_nth) lemma word_le_mask_out_plus_2sz: "x \ (x AND NOT(mask sz)) + 2 ^ sz - 1" for x :: \'a::len word\ by (metis add_diff_eq word_neg_and_le) lemma ucast_add: "ucast (a + (b :: 'a :: len word)) = ucast a + (ucast b :: ('a signed word))" by transfer (simp add: take_bit_add) lemma ucast_minus: "ucast (a - (b :: 'a :: len word)) = ucast a - (ucast b :: ('a signed word))" apply (insert ucast_add[where a=a and b="-b"]) apply (metis (no_types, hide_lams) add_diff_eq diff_add_cancel ucast_add) done lemma scast_ucast_add_one [simp]: "scast (ucast (x :: 'a::len word) + (1 :: 'a signed word)) = x + 1" apply (subst ucast_1[symmetric]) apply (subst ucast_add[symmetric]) apply clarsimp done lemma word_and_le_plus_one: "a > 0 \ (x :: 'a :: len word) AND (a - 1) < a" by (simp add: gt0_iff_gem1 word_and_less') lemma unat_of_ucast_then_shift_eq_unat_of_shift[simp]: "LENGTH('b) \ LENGTH('a) \ unat ((ucast (x :: 'a :: len word) :: 'b :: len word) >> n) = unat (x >> n)" by (simp add: shiftr_div_2n' unat_ucast_up_simp) lemma unat_of_ucast_then_mask_eq_unat_of_mask[simp]: "LENGTH('b) \ LENGTH('a) \ unat ((ucast (x :: 'a :: len word) :: 'b :: len word) AND mask m) = unat (x AND mask m)" by (metis ucast_and_mask unat_ucast_up_simp) lemma shiftr_less_t2n3: "\ (2 :: 'a word) ^ (n + m) = 0; m < LENGTH('a) \ \ (x :: 'a :: len word) >> n < 2 ^ m" by (fastforce intro: shiftr_less_t2n' simp: mask_eq_decr_exp power_overflow) lemma unat_shiftr_le_bound: "\ 2 ^ (LENGTH('a :: len) - n) - 1 \ bnd; 0 < n \ \ unat ((x :: 'a word) >> n) \ bnd" apply transfer apply (simp add: take_bit_drop_bit) apply (simp add: drop_bit_take_bit) apply (rule order_trans) defer apply assumption apply (simp add: nat_le_iff of_nat_diff) done lemma shiftr_eqD: "\ x >> n = y >> n; is_aligned x n; is_aligned y n \ \ x = y" by (metis is_aligned_shiftr_shiftl) lemma word_shiftr_shiftl_shiftr_eq_shiftr: "a \ b \ (x :: 'a :: len word) >> a << b >> b = x >> a" by (simp add: mask_shift multi_shift_simps(5) shiftr_shiftr) lemma of_int_uint_ucast: "of_int (uint (x :: 'a::len word)) = (ucast x :: 'b::len word)" by (fact Word.of_int_uint) lemma mod_mask_drop: "\ m = 2 ^ n; 0 < m; mask n AND msk = mask n \ \ (x mod m) AND msk = x mod m" for x :: \'a::len word\ by (simp add: word_mod_2p_is_mask word_bw_assocs) lemma mask_eq_ucast_eq: "\ x AND mask LENGTH('a) = (x :: ('c :: len word)); LENGTH('a) \ LENGTH('b)\ \ ucast (ucast x :: ('a :: len word)) = (ucast x :: ('b :: len word))" by (metis ucast_and_mask ucast_id ucast_ucast_mask ucast_up_eq) lemma of_nat_less_t2n: "of_nat i < (2 :: ('a :: len) word) ^ n \ n < LENGTH('a) \ unat (of_nat i :: 'a word) < 2 ^ n" by (metis order_less_trans p2_gt_0 unat_less_power word_neq_0_conv) lemma two_power_increasing_less_1: "\ n \ m; m \ LENGTH('a) \ \ (2 :: 'a :: len word) ^ n - 1 \ 2 ^ m - 1" by (metis diff_diff_cancel le_m1_iff_lt less_imp_diff_less p2_gt_0 two_power_increasing word_1_le_power word_le_minus_mono_left word_less_sub_1) lemma word_sub_mono4: "\ y + x \ z + x; y \ y + x; z \ z + x \ \ y \ z" for y :: "'a :: len word" by (simp add: word_add_le_iff2) lemma eq_or_less_helperD: "\ n = unat (2 ^ m - 1 :: 'a :: len word) \ n < unat (2 ^ m - 1 :: 'a word); m < LENGTH('a) \ \ n < 2 ^ m" by (meson le_less_trans nat_less_le unat_less_power word_power_less_1) lemma mask_sub: "n \ m \ mask m - mask n = mask m AND NOT(mask n :: 'a::len word)" by (metis (full_types) and_mask_eq_iff_shiftr_0 mask_out_sub_mask shiftr_mask_le word_bw_comms(1)) lemma neg_mask_diff_bound: "sz'\ sz \ (ptr AND NOT(mask sz')) - (ptr AND NOT(mask sz)) \ 2 ^ sz - 2 ^ sz'" (is "_ \ ?lhs \ ?rhs") for ptr :: \'a::len word\ proof - assume lt: "sz' \ sz" hence "?lhs = ptr AND (mask sz AND NOT(mask sz'))" by (metis add_diff_cancel_left' multiple_mask_trivia) also have "\ \ ?rhs" using lt by (metis (mono_tags) add_diff_eq diff_eq_eq eq_iff mask_2pm1 mask_sub word_and_le') finally show ?thesis by simp qed lemma mask_out_eq_0: "\ idx < 2 ^ sz; sz < LENGTH('a) \ \ (of_nat idx :: 'a :: len word) AND NOT(mask sz) = 0" by (simp add: of_nat_power less_mask_eq mask_eq_0_eq_x) lemma is_aligned_neg_mask_eq': "is_aligned ptr sz = (ptr AND NOT(mask sz) = ptr)" using is_aligned_mask mask_eq_0_eq_x by blast lemma neg_mask_mask_unat: "sz < LENGTH('a) \ unat ((ptr :: 'a :: len word) AND NOT(mask sz)) + unat (ptr AND mask sz) = unat ptr" by (metis AND_NOT_mask_plus_AND_mask_eq unat_plus_simple word_and_le2) lemma unat_pow_le_intro: "LENGTH('a) \ n \ unat (x :: 'a :: len word) < 2 ^ n" by (metis lt2p_lem not_le of_nat_le_iff of_nat_numeral semiring_1_class.of_nat_power uint_nat) lemma unat_shiftl_less_t2n: "\ unat (x :: 'a :: len word) < 2 ^ (m - n); m < LENGTH('a) \ \ unat (x << n) < 2 ^ m" by (metis (no_types) of_nat_power diff_le_self le_less_trans shiftl_less_t2n unat_less_power word_unat.Rep_inverse) lemma unat_is_aligned_add: "\ is_aligned p n; unat d < 2 ^ n \ \ unat (p + d AND mask n) = unat d \ unat (p + d AND NOT(mask n)) = unat p" by (metis add.right_neutral and_mask_eq_iff_le_mask and_not_mask le_mask_iff mask_add_aligned mask_out_add_aligned mult_zero_right shiftl_t2n shiftr_le_0) lemma unat_shiftr_shiftl_mask_zero: "\ c + a \ LENGTH('a) + b ; c < LENGTH('a) \ \ unat (((q :: 'a :: len word) >> a << b) AND NOT(mask c)) = 0" by (fastforce intro: unat_is_aligned_add[where p=0 and n=c, simplified, THEN conjunct2] unat_shiftl_less_t2n unat_shiftr_less_t2n unat_pow_le_intro) lemmas of_nat_ucast = ucast_of_nat[symmetric] lemma shift_then_mask_eq_shift_low_bits: "x \ mask (low_bits + high_bits) \ (x >> low_bits) AND mask high_bits = x >> low_bits" for x :: \'a::len word\ by (simp add: leq_mask_shift le_mask_imp_and_mask) lemma leq_low_bits_iff_zero: "\ x \ mask (low bits + high bits); x >> low_bits = 0 \ \ (x AND mask low_bits = 0) = (x = 0)" for x :: \'a::len word\ using and_mask_eq_iff_shiftr_0 by force lemma unat_less_iff: "\ unat (a :: 'a :: len word) = b; c < 2 ^ LENGTH('a) \ \ (a < of_nat c) = (b < c)" using unat_ucast_less_no_overflow_simp by blast lemma is_aligned_no_overflow3: "\ is_aligned (a :: 'a :: len word) n; n < LENGTH('a); b < 2 ^ n; c \ 2 ^ n; b < c \ \ a + b \ a + (c - 1)" by (meson is_aligned_no_wrap' le_m1_iff_lt not_le word_less_sub_1 word_plus_mono_right) lemma mask_add_aligned_right: "is_aligned p n \ (q + p) AND mask n = q AND mask n" by (simp add: mask_add_aligned add.commute) lemma leq_high_bits_shiftr_low_bits_leq_bits_mask: "x \ mask high_bits \ (x :: 'a :: len word) << low_bits \ mask (low_bits + high_bits)" by (metis le_mask_shiftl_le_mask) lemma word_two_power_neg_ineq: "2 ^ m \ (0 :: 'a word) \ 2 ^ n \ - (2 ^ m :: 'a :: len word)" apply (cases "n < LENGTH('a)"; simp add: power_overflow) apply (cases "m < LENGTH('a)"; simp add: power_overflow) apply (simp add: word_le_nat_alt unat_minus word_size) apply (cases "LENGTH('a)"; simp) apply (simp add: less_Suc_eq_le) apply (drule power_increasing[where a=2 and n=n] power_increasing[where a=2 and n=m], simp)+ apply (drule(1) add_le_mono) apply simp done lemma unat_shiftl_absorb: "\ x \ 2 ^ p; p + k < LENGTH('a) \ \ unat (x :: 'a :: len word) * 2 ^ k = unat (x * 2 ^ k)" by (smt add_diff_cancel_right' add_lessD1 le_add2 le_less_trans mult.commute nat_le_power_trans unat_lt2p unat_mult_lem unat_power_lower word_le_nat_alt) lemma word_plus_mono_right_split: "\ unat ((x :: 'a :: len word) AND mask sz) + unat z < 2 ^ sz; sz < LENGTH('a) \ \ x \ x + z" apply (subgoal_tac "(x AND NOT(mask sz)) + (x AND mask sz) \ (x AND NOT(mask sz)) + ((x AND mask sz) + z)") apply (simp add:word_plus_and_or_coroll2 field_simps) apply (rule word_plus_mono_right) apply (simp add: less_le_trans no_olen_add_nat) using of_nat_power is_aligned_no_wrap' by force lemma mul_not_mask_eq_neg_shiftl: "NOT(mask n :: 'a::len word) = -1 << n" by (simp add: NOT_mask shiftl_t2n) lemma shiftr_mul_not_mask_eq_and_not_mask: "(x >> n) * NOT(mask n) = - (x AND NOT(mask n))" for x :: \'a::len word\ by (metis NOT_mask and_not_mask mult_minus_left semiring_normalization_rules(7) shiftl_t2n) lemma mask_eq_n1_shiftr: "n \ LENGTH('a) \ (mask n :: 'a :: len word) = -1 >> (LENGTH('a) - n)" by (metis diff_diff_cancel eq_refl mask_full shiftr_mask2) lemma is_aligned_mask_out_add_eq: "is_aligned p n \ (p + x) AND NOT(mask n) = p + (x AND NOT(mask n))" by (simp add: mask_out_sub_mask mask_add_aligned) lemmas is_aligned_mask_out_add_eq_sub = is_aligned_mask_out_add_eq[where x="a - b" for a b, simplified field_simps] lemma aligned_bump_down: "is_aligned x n \ (x - 1) AND NOT(mask n) = x - 2 ^ n" by (drule is_aligned_mask_out_add_eq[where x="-1"]) (simp add: NOT_mask) lemma unat_2tp_if: "unat (2 ^ n :: ('a :: len) word) = (if n < LENGTH ('a) then 2 ^ n else 0)" by (split if_split, simp_all add: power_overflow) lemma mask_of_mask: "mask (n::nat) AND mask (m::nat) = (mask (min m n) :: 'a::len word)" by word_eqI_solve lemma unat_signed_ucast_less_ucast: "LENGTH('a) \ LENGTH('b) \ unat (ucast (x :: 'a :: len word) :: 'b :: len signed word) = unat x" by (simp add: unat_ucast_up_simp) lemma toEnum_of_ucast: "LENGTH('b) \ LENGTH('a) \ (toEnum (unat (b::'b :: len word))::'a :: len word) = of_nat (unat b)" by (simp add: unat_pow_le_intro) lemmas unat_ucast_mask = unat_ucast_eq_unat_and_mask[where w=a for a] lemma t2n_mask_eq_if: "2 ^ n AND mask m = (if n < m then 2 ^ n else (0 :: 'a::len word))" by (rule word_eqI, auto simp add: word_size nth_w2p split: if_split) lemma unat_ucast_le: "unat (ucast (x :: 'a :: len word) :: 'b :: len word) \ unat x" by (simp add: ucast_nat_def word_unat_less_le) lemma ucast_le_up_down_iff: "\ LENGTH('a) \ LENGTH('b); (x :: 'b :: len word) \ ucast (- 1 :: 'a :: len word) \ \ (ucast x \ (y :: 'a word)) = (x \ ucast y)" using le_max_word_ucast_id ucast_le_ucast by metis lemma ucast_ucast_mask_shift: "a \ LENGTH('a) + b \ ucast (ucast (p AND mask a >> b) :: 'a :: len word) = p AND mask a >> b" by (metis add.commute le_mask_iff shiftr_mask_le ucast_ucast_eq_mask_shift word_and_le') lemma unat_ucast_mask_shift: "a \ LENGTH('a) + b \ unat (ucast (p AND mask a >> b) :: 'a :: len word) = unat (p AND mask a >> b)" by (metis linear ucast_ucast_mask_shift unat_ucast_up_simp) lemma mask_overlap_zero: "a \ b \ (p AND mask a) AND NOT(mask b) = 0" for p :: \'a::len word\ by (metis NOT_mask_AND_mask mask_lower_twice2 max_def) lemma mask_shifl_overlap_zero: "a + c \ b \ (p AND mask a << c) AND NOT(mask b) = 0" for p :: \'a::len word\ by (metis and_mask_0_iff_le_mask mask_mono mask_shiftl_decompose order_trans shiftl_over_and_dist word_and_le' word_and_le2) lemma mask_overlap_zero': "a \ b \ (p AND NOT(mask a)) AND mask b = 0" for p :: \'a::len word\ using mask_AND_NOT_mask mask_AND_less_0 by blast lemma mask_rshift_mult_eq_rshift_lshift: "((a :: 'a :: len word) >> b) * (1 << c) = (a >> b << c)" by (simp add: shiftl_t2n) lemma shift_alignment: "a \ b \ is_aligned (p >> a << a) b" using is_aligned_shift is_aligned_weaken by blast lemma mask_split_sum_twice: "a \ b \ (p AND NOT(mask a)) + ((p AND mask a) AND NOT(mask b)) + (p AND mask b) = p" for p :: \'a::len word\ by (simp add: add.commute multiple_mask_trivia word_bw_comms(1) word_bw_lcs(1) word_plus_and_or_coroll2) lemma mask_shift_eq_mask_mask: "(p AND mask a >> b << b) = (p AND mask a) AND NOT(mask b)" for p :: \'a::len word\ by (simp add: and_not_mask) lemma mask_shift_sum: "\ a \ b; unat n = unat (p AND mask b) \ \ (p AND NOT(mask a)) + (p AND mask a >> b) * (1 << b) + n = (p :: 'a :: len word)" by (metis and_not_mask mask_rshift_mult_eq_rshift_lshift mask_split_sum_twice word_unat.Rep_eqD) lemma is_up_compose: "\ is_up uc; is_up uc' \ \ is_up (uc' \ uc)" unfolding is_up_def by (simp add: Word.target_size Word.source_size) lemma of_int_sint_scast: "of_int (sint (x :: 'a :: len word)) = (scast x :: 'b :: len word)" by (fact Word.of_int_sint) lemma scast_of_nat_to_signed [simp]: "scast (of_nat x :: 'a :: len word) = (of_nat x :: 'a signed word)" by transfer simp lemma scast_of_nat_signed_to_unsigned_add: "scast (of_nat x + of_nat y :: 'a :: len signed word) = (of_nat x + of_nat y :: 'a :: len word)" by (metis of_nat_add scast_of_nat) lemma scast_of_nat_unsigned_to_signed_add: "(scast (of_nat x + of_nat y :: 'a :: len word)) = (of_nat x + of_nat y :: 'a :: len signed word)" by (metis Abs_fnat_hom_add scast_of_nat_to_signed) lemma and_mask_cases: fixes x :: "'a :: len word" assumes len: "n < LENGTH('a)" shows "x AND mask n \ of_nat ` set [0 ..< 2 ^ n]" apply (simp flip: take_bit_eq_mask) apply (rule image_eqI [of _ _ \unat (take_bit n x)\]) using len apply simp_all apply transfer apply simp done lemma sint_eq_uint_2pl: "\ (a :: 'a :: len word) < 2 ^ (LENGTH('a) - 1) \ \ sint a = uint a" by (simp add: not_msb_from_less sint_eq_uint word_2p_lem word_size) lemma pow_sub_less: "\ a + b \ LENGTH('a); unat (x :: 'a :: len word) = 2 ^ a \ \ unat (x * 2 ^ b - 1) < 2 ^ (a + b)" by (metis (mono_tags) eq_or_less_helperD not_less of_nat_numeral power_add semiring_1_class.of_nat_power unat_pow_le_intro word_unat.Rep_inverse) lemma sle_le_2pl: "\ (b :: 'a :: len word) < 2 ^ (LENGTH('a) - 1); a \ b \ \ a <=s b" by (simp add: not_msb_from_less word_sle_msb_le) lemma sless_less_2pl: "\ (b :: 'a :: len word) < 2 ^ (LENGTH('a) - 1); a < b \ \ a > n = w AND mask (size w - n)" for w :: \'a::len word\ by (cases "n \ size w"; clarsimp simp: word_and_le2 and_mask shiftl_zero_size) lemma aligned_sub_aligned_simple: "\ is_aligned a n; is_aligned b n \ \ is_aligned (a - b) n" by (simp add: aligned_sub_aligned) lemma minus_one_shift: "- (1 << n) = (-1 << n :: 'a::len word)" by (simp add: mask_eq_decr_exp NOT_eq flip: mul_not_mask_eq_neg_shiftl) lemma ucast_eq_mask: "(UCAST('a::len \ 'b::len) x = UCAST('a \ 'b) y) = (x AND mask LENGTH('b) = y AND mask LENGTH('b))" by (rule iffI; word_eqI_solve) context fixes w :: "'a::len word" begin private lemma sbintrunc_uint_ucast: assumes "Suc n = LENGTH('b::len)" - shows "sbintrunc n (uint (ucast w :: 'b word)) = sbintrunc n (uint w)" - by (metis assms sbintrunc_bintrunc ucast_eq word_ubin.eq_norm) + shows "signed_take_bit n (uint (ucast w :: 'b word)) = signed_take_bit n (uint w)" + by (rule bit_eqI) (use assms in \simp add: bit_simps\) private lemma test_bit_sbintrunc: assumes "i < LENGTH('a)" - shows "(word_of_int (sbintrunc n (uint w)) :: 'a word) !! i + shows "(word_of_int (signed_take_bit n (uint w)) :: 'a word) !! i = (if n < i then w !! n else w !! i)" using assms by (simp add: nth_sbintr) (simp add: test_bit_bin) private lemma test_bit_sbintrunc_ucast: assumes len_a: "i < LENGTH('a)" - shows "(word_of_int (sbintrunc (LENGTH('b) - 1) (uint (ucast w :: 'b word))) :: 'a word) !! i + shows "(word_of_int (signed_take_bit (LENGTH('b) - 1) (uint (ucast w :: 'b word))) :: 'a word) !! i = (if LENGTH('b::len) \ i then w !! (LENGTH('b) - 1) else w !! i)" apply (subst sbintrunc_uint_ucast) apply simp apply (subst test_bit_sbintrunc) apply (rule len_a) apply (rule if_cong[OF _ refl refl]) using leD less_linear by fastforce lemma scast_ucast_high_bits: \scast (ucast w :: 'b::len word) = w \ (\ i \ {LENGTH('b) ..< size w}. w !! i = w !! (LENGTH('b) - 1))\ proof (cases \LENGTH('a) \ LENGTH('b)\) case True moreover define m where \m = LENGTH('b) - LENGTH('a)\ ultimately have \LENGTH('b) = m + LENGTH('a)\ by simp then show ?thesis apply (simp_all add: signed_ucast_eq word_size) apply (rule bit_word_eqI) apply (simp add: bit_signed_take_bit_iff) done next case False define q where \q = LENGTH('b) - 1\ then have \LENGTH('b) = Suc q\ by simp moreover define m where \m = Suc LENGTH('a) - LENGTH('b)\ with False \LENGTH('b) = Suc q\ have \LENGTH('a) = m + q\ by (simp add: not_le) ultimately show ?thesis apply (simp_all add: signed_ucast_eq word_size) apply (transfer fixing: m q) apply (simp add: signed_take_bit_take_bit) apply rule apply (subst bit_eq_iff) apply (simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def) apply (auto simp add: Suc_le_eq) using less_imp_le_nat apply blast using less_imp_le_nat apply blast done qed lemma scast_ucast_mask_compare: "scast (ucast w :: 'b::len word) = w \ (w \ mask (LENGTH('b) - 1) \ NOT(mask (LENGTH('b) - 1)) \ w)" apply (clarsimp simp: le_mask_high_bits neg_mask_le_high_bits scast_ucast_high_bits word_size) apply (rule iffI; clarsimp) apply (rename_tac i j; case_tac "i = LENGTH('b) - 1"; case_tac "j = LENGTH('b) - 1") by auto lemma ucast_less_shiftl_helper': "\ LENGTH('b) + (a::nat) < LENGTH('a); 2 ^ (LENGTH('b) + a) \ n\ \ (ucast (x :: 'b::len word) << a) < (n :: 'a::len word)" apply (erule order_less_le_trans[rotated]) using ucast_less[where x=x and 'a='a] apply (simp only: shiftl_t2n field_simps) apply (rule word_less_power_trans2; simp) done end lemma ucast_ucast_mask2: "is_down (UCAST ('a \ 'b)) \ UCAST ('b::len \ 'c::len) (UCAST ('a::len \ 'b::len) x) = UCAST ('a \ 'c) (x AND mask LENGTH('b))" by word_eqI_solve lemma ucast_NOT: "ucast (NOT x) = NOT(ucast x) AND mask (LENGTH('a))" for x::"'a::len word" by word_eqI lemma ucast_NOT_down: "is_down UCAST('a::len \ 'b::len) \ UCAST('a \ 'b) (NOT x) = NOT(UCAST('a \ 'b) x)" by word_eqI lemma upto_enum_step_shift: "\ is_aligned p n \ \ ([p , p + 2 ^ m .e. p + 2 ^ n - 1]) = map ((+) p) [0, 2 ^ m .e. 2 ^ n - 1]" apply (erule is_aligned_get_word_bits) prefer 2 apply (simp add: map_idI) apply (clarsimp simp: upto_enum_step_def) apply (frule is_aligned_no_overflow) apply (simp add: linorder_not_le [symmetric]) done lemma upto_enum_step_shift_red: "\ is_aligned p sz; sz < LENGTH('a); us \ sz \ \ [p :: 'a :: len word, p + 2 ^ us .e. p + 2 ^ sz - 1] = map (\x. p + of_nat x * 2 ^ us) [0 ..< 2 ^ (sz - us)]" apply (subst upto_enum_step_shift, assumption) apply (simp add: upto_enum_step_red) done lemma upto_enum_step_subset: "set [x, y .e. z] \ {x .. z}" apply (clarsimp simp: upto_enum_step_def linorder_not_less) apply (drule div_to_mult_word_lt) apply (rule conjI) apply (erule word_random[rotated]) apply simp apply (rule order_trans) apply (erule word_plus_mono_right) apply simp apply simp done lemma ucast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" fixes L :: "int \ int \ int" assumes lift_M: "\x y. uint (M x y) = L (uint x) (uint y) mod 2 ^ LENGTH('a)" assumes lift_M': "\x y. uint (M' x y) = L (uint x) (uint y) mod 2 ^ LENGTH('b)" assumes distrib: "\x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b)) = (L x y) mod (2 ^ LENGTH('b))" assumes is_down: "is_down (ucast :: 'a word \ 'b word)" shows "ucast (M a b) = M' (ucast a) (ucast b)" apply (simp only: ucast_eq) apply (subst lift_M) apply (subst of_int_uint [symmetric], subst lift_M') apply (subst (1 2) int_word_uint) apply (subst word_ubin.norm_eq_iff [symmetric]) apply (subst (1 2) bintrunc_mod2p) apply (insert is_down) apply (unfold is_down_def) apply (clarsimp simp: target_size source_size) apply (clarsimp simp: mod_exp_eq min_def) apply (rule distrib [symmetric]) done lemma ucast_down_add: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) + b) = (ucast a + ucast b :: 'b::len word)" by (rule ucast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp) lemma ucast_down_minus: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) - b) = (ucast a - ucast b :: 'b::len word)" apply (rule ucast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_diff_left_eq mod_diff_right_eq) apply simp done lemma ucast_down_mult: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) * b) = (ucast a * ucast b :: 'b::len word)" apply (rule ucast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_mult_eq) apply simp done lemma scast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" fixes L :: "int \ int \ int" assumes lift_M: "\x y. uint (M x y) = L (uint x) (uint y) mod 2 ^ LENGTH('a)" assumes lift_M': "\x y. uint (M' x y) = L (uint x) (uint y) mod 2 ^ LENGTH('b)" assumes distrib: "\x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b)) = (L x y) mod (2 ^ LENGTH('b))" assumes is_down: "is_down (scast :: 'a word \ 'b word)" shows "scast (M a b) = M' (scast a) (scast b)" apply (subst (1 2 3) down_cast_same [symmetric]) apply (insert is_down) apply (clarsimp simp: is_down_def target_size source_size is_down) apply (rule ucast_distrib [where L=L, OF lift_M lift_M' distrib]) apply (insert is_down) apply (clarsimp simp: is_down_def target_size source_size is_down) done lemma scast_down_add: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) + b) = (scast a + scast b :: 'b::len word)" by (rule scast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp) lemma scast_down_minus: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) - b) = (scast a - scast b :: 'b::len word)" apply (rule scast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_diff_left_eq mod_diff_right_eq) apply simp done lemma scast_down_mult: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) * b) = (scast a * scast b :: 'b::len word)" apply (rule scast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_mult_eq) apply simp done lemma scast_ucast_1: "\ is_down (ucast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma scast_ucast_3: "\ is_down (ucast :: 'a word \ 'c word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma scast_ucast_4: "\ is_up (ucast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma scast_scast_b: "\ is_up (scast :: 'a word \ 'b word) \ \ (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_eq sint_up_scast) lemma ucast_scast_1: "\ is_down (scast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_eq ucast_down_wi) lemma ucast_scast_3: "\ is_down (scast :: 'a word \ 'c word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_eq ucast_down_wi) lemma ucast_scast_4: "\ is_up (scast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis down_cast_same scast_eq sint_up_scast) lemma ucast_ucast_a: "\ is_down (ucast :: 'b word \ 'c word) \ \ (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma ucast_ucast_b: "\ is_up (ucast :: 'a word \ 'b word) \ \ (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis ucast_up_ucast) lemma scast_scast_a: "\ is_down (scast :: 'b word \ 'c word) \ \ (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" apply (simp only: scast_eq) apply (metis down_cast_same is_up_down scast_eq ucast_down_wi) done lemma scast_down_wi [OF refl]: "uc = scast \ is_down uc \ uc (word_of_int x) = word_of_int x" by (metis down_cast_same is_up_down ucast_down_wi) lemmas cast_simps = is_down is_up scast_down_add scast_down_minus scast_down_mult ucast_down_add ucast_down_minus ucast_down_mult scast_ucast_1 scast_ucast_3 scast_ucast_4 ucast_scast_1 ucast_scast_3 ucast_scast_4 ucast_ucast_a ucast_ucast_b scast_scast_a scast_scast_b ucast_down_wi scast_down_wi ucast_of_nat scast_of_nat uint_up_ucast sint_up_scast up_scast_surj up_ucast_surj lemma sdiv_word_max: "(sint (a :: ('a::len) word) sdiv sint (b :: ('a::len) word) < (2 ^ (size a - 1))) = ((a \ - (2 ^ (size a - 1)) \ (b \ -1)))" (is "?lhs = (\ ?a_int_min \ \ ?b_minus1)") proof (rule classical) assume not_thesis: "\ ?thesis" have not_zero: "b \ 0" using not_thesis by (clarsimp) have result_range: "sint a sdiv sint b \ (sints (size a)) \ {2 ^ (size a - 1)}" apply (cut_tac sdiv_int_range [where a="sint a" and b="sint b"]) apply (erule rev_subsetD) using sint_range' [where x=a] sint_range' [where x=b] apply (auto simp: max_def abs_if word_size sints_num) done have result_range_overflow: "(sint a sdiv sint b = 2 ^ (size a - 1)) = (?a_int_min \ ?b_minus1)" apply (rule iffI [rotated]) apply (clarsimp simp: signed_divide_int_def sgn_if word_size sint_int_min) apply (rule classical) apply (case_tac "?a_int_min") apply (clarsimp simp: word_size sint_int_min) apply (metis diff_0_right int_sdiv_negated_is_minus1 minus_diff_eq minus_int_code(2) power_eq_0_iff sint_minus1 zero_neq_numeral) apply (subgoal_tac "abs (sint a) < 2 ^ (size a - 1)") apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] apply (clarsimp simp: word_size) apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] apply (insert word_sint.Rep [where x="a"])[1] apply (clarsimp simp: minus_le_iff word_size abs_if sints_num split: if_split_asm) apply (metis minus_minus sint_int_min word_sint.Rep_inject) done have result_range_simple: "(sint a sdiv sint b \ (sints (size a))) \ ?thesis" apply (insert sdiv_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: word_size sints_num sint_int_min) done show ?thesis apply (rule UnE [OF result_range result_range_simple]) apply simp apply (clarsimp simp: word_size) using result_range_overflow apply (clarsimp simp: word_size) done qed lemmas sdiv_word_min' = sdiv_word_min [simplified word_size, simplified] lemmas sdiv_word_max' = sdiv_word_max [simplified word_size, simplified] lemma signed_arith_ineq_checks_to_eq: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) = (sint a + sint b = sint (a + b ))" "((- (2 ^ (size a - 1)) \ (sint a - sint b)) \ (sint a - sint b \ (2 ^ (size a - 1) - 1))) = (sint a - sint b = sint (a - b))" "((- (2 ^ (size a - 1)) \ (- sint a)) \ (- sint a) \ (2 ^ (size a - 1) - 1)) = ((- sint a) = sint (- a))" "((- (2 ^ (size a - 1)) \ (sint a * sint b)) \ (sint a * sint b \ (2 ^ (size a - 1) - 1))) = (sint a * sint b = sint (a * b))" "((- (2 ^ (size a - 1)) \ (sint a sdiv sint b)) \ (sint a sdiv sint b \ (2 ^ (size a - 1) - 1))) = (sint a sdiv sint b = sint (a sdiv b))" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) = (sint a smod sint b = sint (a smod b))" by (auto simp: sint_word_ariths word_size signed_div_arith signed_mod_arith sbintrunc_eq_in_range range_sbintrunc) lemma signed_arith_sint: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) \ sint (a + b) = (sint a + sint b)" "((- (2 ^ (size a - 1)) \ (sint a - sint b)) \ (sint a - sint b \ (2 ^ (size a - 1) - 1))) \ sint (a - b) = (sint a - sint b)" "((- (2 ^ (size a - 1)) \ (- sint a)) \ (- sint a) \ (2 ^ (size a - 1) - 1)) \ sint (- a) = (- sint a)" "((- (2 ^ (size a - 1)) \ (sint a * sint b)) \ (sint a * sint b \ (2 ^ (size a - 1) - 1))) \ sint (a * b) = (sint a * sint b)" "((- (2 ^ (size a - 1)) \ (sint a sdiv sint b)) \ (sint a sdiv sint b \ (2 ^ (size a - 1) - 1))) \ sint (a sdiv b) = (sint a sdiv sint b)" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) \ sint (a smod b) = (sint a smod sint b)" by (subst (asm) signed_arith_ineq_checks_to_eq; simp)+ end