# HG changeset patch # User haftmann # Date 1603965783 0 # Node ID c7038c397ae38a1d70b2f2cc684a67856b1669be # Parent d8661799afb25ef1be44ca5605849707d8854813 moved most material from session HOL-Word to Word_Lib in the AFP diff -r d8661799afb2 -r c7038c397ae3 NEWS --- a/NEWS Thu Oct 29 09:59:40 2020 +0000 +++ b/NEWS Thu Oct 29 10:03:03 2020 +0000 @@ -109,55 +109,56 @@ * Library theory "Signed_Division" provides operations for signed division, instantiated for type int. -* Session HOL-Word: Type word is restricted to bit strings consisting +* Self-contained library theory "Word" taken over from former session +"HOL-Word". + +* Theory "Word": Type word is restricted to bit strings consisting of at least one bit. INCOMPATIBILITY. -* Session HOL-Word: Various operations dealing with bit values -represented as reversed lists of bools are separated into theory -Reversed_Bit_Lists, included by theory More_Word rather than theory Word. -INCOMPATIBILITY. - -* Session HOL-Word: Bit operations NOT, AND, OR, XOR are based on +* Theory "Word": Bit operations NOT, AND, OR, XOR are based on generic algebraic bit operations from HOL-Library.Bit_Operations. INCOMPATIBILITY. -* Session HOL-Word: Most operations on type word are set up +* Theory "Word": Most operations on type word are set up for transfer and lifting. INCOMPATIBILITY. -* Session HOL-Word: Generic type conversions. INCOMPATIBILITY, +* Theory "Word": Generic type conversions. INCOMPATIBILITY, sometimes additional rewrite rules must be added to applications to get a confluent system again. -* Session HOL-Word: Theory "Word_Bitwise" has been moved to AFP entry -Word_Lib as theory "Bitwise". INCOMPATIBILITY. - -* Session HOL-Word: Compound operation "bin_split" simplifies by default -into its components "drop_bit" and "take_bit". INCOMPATIBILITY. - -* Session HOL-Word: Uniform polymorphic "mask" operation for both +* Theory "Word": Uniform polymorphic "mask" operation for both types int and word. INCOMPATIBILITY. -* Session HOL-Word: Syntax for signed compare operators has been +* Theory "Word": Syntax for signed compare operators has been consolidated with syntax of regular compare operators. Minor INCOMPATIBILITY. -* Session HOL-Word: Operations lsb, msb and set_bit are separated -into theories Misc_lsb, Misc_msb and Misc_set_bit respectively. -INCOMPATIBILITY. - -* Session HOL-Word: Misc ancient material has been factored out into -separate theories. INCOMPATIBILITY, prefer theory "More_Word" -over "Word" to use it. - -* Session HOL-Word: Ancient int numeral representation has been -factored out in separate theory "Ancient_Numeral". INCOMPATIBILITY. - -* Session HOL-Word: Operations "bin_last", "bin_rest", "bin_nth", -"bintrunc", "sbintrunc", "norm_sint", "bin_cat" and "max_word" are now -mere input abbreviations. Minor INCOMPATIBILITY. - -* Session HOL-Word: Theory HOL-Library.Z2 is not imported any longer. -Minor INCOMPATIBILITY. +* Former session "HOL-Word": Various operations dealing with bit values +represented as reversed lists of bools are separated into theory +Reversed_Bit_Lists in session Word_Lib in the AFP. INCOMPATIBILITY. + +* Former session "HOL-Word": Theory "Word_Bitwise" has been moved to AFP +entry Word_Lib as theory "Bitwise". INCOMPATIBILITY. + +* Former session "HOL-Word": Compound operation "bin_split" simplifies +by default into its components "drop_bit" and "take_bit". +INCOMPATIBILITY. + +* Former session "HOL-Word": Operations lsb, msb and set_bit are +separated into theories Misc_lsb, Misc_msb and Misc_set_bit respectively +in session Word_Lib in the AFP. INCOMPATIBILITY. + +* Former session "HOL-Word": Ancient int numeral representation has been +factored out in separate theory "Ancient_Numeral" in session Word_Lib +in the AFP. INCOMPATIBILITY. + +* Former session "HOL-Word": Operations "bin_last", "bin_rest", +"bin_nth", "bintrunc", "sbintrunc", "norm_sint", "bin_cat" and +"max_word" are now mere input abbreviations. Minor INCOMPATIBILITY. + +* Former session "HOL-Word": Misc ancient material has been factored out +into separate theories and moved to session Word_Lib in the AFP. See +theory "Guide" there for further information. INCOMPATIBILITY. * Session HOL-TPTP: The "tptp_isabelle" and "tptp_sledgehammer" commands are in working order again, as opposed to outputting "GaveUp" on nearly diff -r d8661799afb2 -r c7038c397ae3 src/Doc/System/Server.thy --- a/src/Doc/System/Server.thy Thu Oct 29 09:59:40 2020 +0000 +++ b/src/Doc/System/Server.thy Thu Oct 29 10:03:03 2020 +0000 @@ -766,7 +766,7 @@ text \ Build of a session image from the Isabelle distribution: - @{verbatim [display] \session_build {"session": "HOL-Word"}\} + @{verbatim [display] \session_build {"session": "HOL-Algebra"}\} Build a session image from the Archive of Formal Proofs: @{verbatim [display] \session_build {"session": "Coinductive", "dirs": ["$AFP_BASE/thys"]}\} diff -r d8661799afb2 -r c7038c397ae3 src/Doc/System/Sessions.thy --- a/src/Doc/System/Sessions.thy Thu Oct 29 09:59:40 2020 +0000 +++ b/src/Doc/System/Sessions.thy Thu Oct 29 10:03:03 2020 +0000 @@ -466,7 +466,7 @@ \<^smallskip> Build some session images with cleanup of their descendants, while retaining their ancestry: - @{verbatim [display] \isabelle build -b -c HOL-Algebra HOL-Word\} + @{verbatim [display] \isabelle build -b -c HOL-Library HOL-Algebra\} \<^smallskip> Clean all sessions without building anything: diff -r d8661799afb2 -r c7038c397ae3 src/HOL/Codegenerator_Test/Candidates.thy --- a/src/HOL/Codegenerator_Test/Candidates.thy Thu Oct 29 09:59:40 2020 +0000 +++ b/src/HOL/Codegenerator_Test/Candidates.thy Thu Oct 29 10:03:03 2020 +0000 @@ -16,7 +16,7 @@ "HOL-Computational_Algebra.Polynomial_Factorial" "HOL-Number_Theory.Eratosthenes" "HOL-Examples.Records" - "HOL-Word.Word" + "HOL-Library.Word" begin text \Drop technical stuff from \<^theory>\HOL.Quickcheck_Narrowing\ which is tailored towards Haskell\ diff -r d8661799afb2 -r c7038c397ae3 src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy Thu Oct 29 09:59:40 2020 +0000 +++ b/src/HOL/Library/Library.thy Thu Oct 29 10:03:03 2020 +0000 @@ -99,6 +99,7 @@ Type_Length Uprod While_Combinator + Word Z2 begin end diff -r d8661799afb2 -r c7038c397ae3 src/HOL/Library/Tools/smt_word.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Tools/smt_word.ML Thu Oct 29 10:03:03 2020 +0000 @@ -0,0 +1,158 @@ +(* Title: HOL/Word/Tools/smt_word.ML + Author: Sascha Boehme, TU Muenchen + +SMT setup for words. +*) + +signature SMT_WORD = +sig + val add_word_shift': term * string -> Context.generic -> Context.generic +end + +structure SMT_Word : SMT_WORD = +struct + +open Word_Lib + + +(* SMT-LIB logic *) + +(* "QF_AUFBV" is too restrictive for Isabelle's problems, which contain aritmetic and quantifiers. + Better set the logic to "" and make at least Z3 happy. *) +fun smtlib_logic ts = + if exists (Term.exists_type (Term.exists_subtype is_wordT)) ts then SOME "" else NONE + + +(* SMT-LIB builtins *) + +local + val smtlibC = SMTLIB_Interface.smtlibC + + val wordT = \<^typ>\'a::len word\ + + fun index1 s i = "(_ " ^ s ^ " " ^ string_of_int i ^ ")" + fun index2 s i j = "(_ " ^ s ^ " " ^ string_of_int i ^ " " ^ string_of_int j ^ ")" + + fun word_typ (Type (\<^type_name>\word\, [T])) = + Option.map (rpair [] o index1 "BitVec") (try dest_binT T) + | word_typ _ = NONE + + fun word_num (Type (\<^type_name>\word\, [T])) k = + Option.map (index1 ("bv" ^ string_of_int k)) (try dest_binT T) + | word_num _ _ = NONE + + fun if_fixed pred m n T ts = + let val (Us, U) = Term.strip_type T + in + if pred (U, Us) then SOME (n, length Us, ts, Term.list_comb o pair (Const (m, T))) else NONE + end + + fun if_fixed_all m = if_fixed (forall (can dest_wordT) o (op ::)) m + fun if_fixed_args m = if_fixed (forall (can dest_wordT) o snd) m + + fun add_word_fun f (t, n) = + let val (m, _) = Term.dest_Const t + in SMT_Builtin.add_builtin_fun smtlibC (Term.dest_Const t, K (f m n)) end + + val mk_nat = HOLogic.mk_number \<^typ>\nat\ + + fun mk_shift c [u, t] = Const c $ mk_nat (snd (HOLogic.dest_number u)) $ t + | mk_shift c ts = raise TERM ("bad arguments", Const c :: ts) + + fun shift m n T ts = + let val U = Term.domain_type (Term.range_type T) + in + (case (can dest_wordT U, try (snd o HOLogic.dest_number o hd) ts) of + (true, SOME i) => + SOME (n, 2, [hd (tl ts), HOLogic.mk_number U i], mk_shift (m, T)) + | _ => NONE) (* FIXME: also support non-numerical shifts *) + end + + fun mk_shift' c [t, u] = Const c $ t $ mk_nat (snd (HOLogic.dest_number u)) + | mk_shift' c ts = raise TERM ("bad arguments", Const c :: ts) + + fun shift' m n T ts = + let val U = Term.domain_type T + in + (case (can dest_wordT U, try (snd o HOLogic.dest_number o hd o tl) ts) of + (true, SOME i) => + SOME (n, 2, [hd ts, HOLogic.mk_number U i], mk_shift' (m, T)) + | _ => NONE) (* FIXME: also support non-numerical shifts *) + end + + fun mk_extract c i ts = Term.list_comb (Const c, mk_nat i :: ts) + + fun extract m n T ts = + let val U = Term.range_type (Term.range_type T) + in + (case (try (snd o HOLogic.dest_number o hd) ts, try dest_wordT U) of + (SOME lb, SOME i) => + SOME (index2 n (i + lb - 1) lb, 1, tl ts, mk_extract (m, T) lb) + | _ => NONE) + end + + fun mk_extend c ts = Term.list_comb (Const c, ts) + + fun extend m n T ts = + let val (U1, U2) = Term.dest_funT T + in + (case (try dest_wordT U1, try dest_wordT U2) of + (SOME i, SOME j) => + if j-i >= 0 then SOME (index1 n (j-i), 1, ts, mk_extend (m, T)) + else NONE + | _ => NONE) + end + + fun mk_rotate c i ts = Term.list_comb (Const c, mk_nat i :: ts) + + fun rotate m n T ts = + let val U = Term.domain_type (Term.range_type T) + in + (case (can dest_wordT U, try (snd o HOLogic.dest_number o hd) ts) of + (true, SOME i) => SOME (index1 n i, 1, tl ts, mk_rotate (m, T) i) + | _ => NONE) + end +in + +val setup_builtins = + SMT_Builtin.add_builtin_typ smtlibC (wordT, word_typ, word_num) #> + fold (add_word_fun if_fixed_all) [ + (\<^term>\uminus :: 'a::len word \ _\, "bvneg"), + (\<^term>\plus :: 'a::len word \ _\, "bvadd"), + (\<^term>\minus :: 'a::len word \ _\, "bvsub"), + (\<^term>\times :: 'a::len word \ _\, "bvmul"), + (\<^term>\NOT :: 'a::len word \ _\, "bvnot"), + (\<^term>\(AND) :: 'a::len word \ _\, "bvand"), + (\<^term>\(OR) :: 'a::len word \ _\, "bvor"), + (\<^term>\(XOR) :: 'a::len word \ _\, "bvxor"), + (\<^term>\word_cat :: 'a::len word \ _\, "concat") ] #> + fold (add_word_fun shift) [ + (\<^term>\push_bit :: nat \ 'a::len word \ _ \, "bvshl"), + (\<^term>\drop_bit :: nat \ 'a::len word \ _\, "bvlshr"), + (\<^term>\signed_drop_bit :: nat \ 'a::len word \ _\, "bvashr") ] #> + add_word_fun extract + (\<^term>\slice :: _ \ 'a::len word \ _\, "extract") #> + fold (add_word_fun extend) [ + (\<^term>\ucast :: 'a::len word \ _\, "zero_extend"), + (\<^term>\scast :: 'a::len word \ _\, "sign_extend") ] #> + fold (add_word_fun rotate) [ + (\<^term>\word_rotl\, "rotate_left"), + (\<^term>\word_rotr\, "rotate_right") ] #> + fold (add_word_fun if_fixed_args) [ + (\<^term>\less :: 'a::len word \ _\, "bvult"), + (\<^term>\less_eq :: 'a::len word \ _\, "bvule"), + (\<^term>\word_sless\, "bvslt"), + (\<^term>\word_sle\, "bvsle") ] + +val add_word_shift' = add_word_fun shift' + +end + + +(* setup *) + +val _ = Theory.setup (Context.theory_map ( + SMTLIB_Interface.add_logic (20, smtlib_logic) #> + setup_builtins)) + +end; diff -r d8661799afb2 -r c7038c397ae3 src/HOL/Library/Tools/word_lib.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Tools/word_lib.ML Thu Oct 29 10:03:03 2020 +0000 @@ -0,0 +1,47 @@ +(* Title: HOL/Word/Tools/word_lib.ML + Author: Sascha Boehme, TU Muenchen and Thomas Sewell, NICTA + +Helper routines for operating on the word type. +*) + +signature WORD_LIB = +sig + val dest_binT: typ -> int + val is_wordT: typ -> bool + val dest_wordT: typ -> int + val mk_wordT: int -> typ +end + +structure Word_Lib: WORD_LIB = +struct + +fun dest_binT T = + (case T of + Type (\<^type_name>\Numeral_Type.num0\, _) => 0 + | Type (\<^type_name>\Numeral_Type.num1\, _) => 1 + | Type (\<^type_name>\Numeral_Type.bit0\, [T]) => 2 * dest_binT T + | Type (\<^type_name>\Numeral_Type.bit1\, [T]) => 1 + 2 * dest_binT T + | _ => raise TYPE ("dest_binT", [T], [])) + +fun is_wordT (Type (\<^type_name>\word\, _)) = true + | is_wordT _ = false + +fun dest_wordT (Type (\<^type_name>\word\, [T])) = dest_binT T + | dest_wordT T = raise TYPE ("dest_wordT", [T], []) + + +fun mk_bitT i T = + if i = 0 + then Type (\<^type_name>\Numeral_Type.bit0\, [T]) + else Type (\<^type_name>\Numeral_Type.bit1\, [T]) + +fun mk_binT size = + if size = 0 then \<^typ>\Numeral_Type.num0\ + else if size = 1 then \<^typ>\Numeral_Type.num1\ + else let val (q, r) = Integer.div_mod size 2 in mk_bitT r (mk_binT q) end + +fun mk_wordT size = + if size >= 0 then Type (\<^type_name>\word\, [mk_binT size]) + else raise TYPE ("mk_wordT: " ^ string_of_int size, [], []) + +end diff -r d8661799afb2 -r c7038c397ae3 src/HOL/Library/Word.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Word.thy Thu Oct 29 10:03:03 2020 +0000 @@ -0,0 +1,4588 @@ +(* Title: HOL/Library/Word.thy + Author: Jeremy Dawson and Gerwin Klein, NICTA, et. al. +*) + +section \A type of finite bit strings\ + +theory Word +imports + "HOL-Library.Type_Length" + "HOL-Library.Boolean_Algebra" + "HOL-Library.Bit_Operations" +begin + +subsection \Preliminaries\ + +lemma signed_take_bit_decr_length_iff: + \signed_take_bit (LENGTH('a::len) - Suc 0) k = signed_take_bit (LENGTH('a) - Suc 0) l + \ take_bit LENGTH('a) k = take_bit LENGTH('a) l\ + by (cases \LENGTH('a)\) + (simp_all add: signed_take_bit_eq_iff_take_bit_eq) + + +subsection \Fundamentals\ + +subsubsection \Type definition\ + +quotient_type (overloaded) 'a word = int / \\k l. take_bit LENGTH('a) k = take_bit LENGTH('a::len) l\ + morphisms rep Word by (auto intro!: equivpI reflpI sympI transpI) + +hide_const (open) rep \ \only for foundational purpose\ +hide_const (open) Word \ \only for code generation\ + + +subsubsection \Basic arithmetic\ + +instantiation word :: (len) comm_ring_1 +begin + +lift_definition zero_word :: \'a word\ + is 0 . + +lift_definition one_word :: \'a word\ + is 1 . + +lift_definition plus_word :: \'a word \ 'a word \ 'a word\ + is \(+)\ + by (auto simp add: take_bit_eq_mod intro: mod_add_cong) + +lift_definition minus_word :: \'a word \ 'a word \ 'a word\ + is \(-)\ + by (auto simp add: take_bit_eq_mod intro: mod_diff_cong) + +lift_definition uminus_word :: \'a word \ 'a word\ + is uminus + by (auto simp add: take_bit_eq_mod intro: mod_minus_cong) + +lift_definition times_word :: \'a word \ 'a word \ 'a word\ + is \(*)\ + by (auto simp add: take_bit_eq_mod intro: mod_mult_cong) + +instance + by (standard; transfer) (simp_all add: algebra_simps) + +end + +context + includes lifting_syntax + notes + power_transfer [transfer_rule] + transfer_rule_of_bool [transfer_rule] + transfer_rule_numeral [transfer_rule] + transfer_rule_of_nat [transfer_rule] + transfer_rule_of_int [transfer_rule] +begin + +lemma power_transfer_word [transfer_rule]: + \(pcr_word ===> (=) ===> pcr_word) (^) (^)\ + by transfer_prover + +lemma [transfer_rule]: + \((=) ===> pcr_word) of_bool of_bool\ + by transfer_prover + +lemma [transfer_rule]: + \((=) ===> pcr_word) numeral numeral\ + by transfer_prover + +lemma [transfer_rule]: + \((=) ===> pcr_word) int of_nat\ + by transfer_prover + +lemma [transfer_rule]: + \((=) ===> pcr_word) (\k. k) of_int\ +proof - + have \((=) ===> pcr_word) of_int of_int\ + by transfer_prover + then show ?thesis by (simp add: id_def) +qed + +lemma [transfer_rule]: + \(pcr_word ===> (\)) even ((dvd) 2 :: 'a::len word \ bool)\ +proof - + have even_word_unfold: "even k \ (\l. take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l))" (is "?P \ ?Q") + for k :: int + proof + assume ?P + then show ?Q + by auto + next + assume ?Q + then obtain l where "take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l)" .. + then have "even (take_bit LENGTH('a) k)" + by simp + then show ?P + by simp + qed + show ?thesis by (simp only: even_word_unfold [abs_def] dvd_def [where ?'a = "'a word", abs_def]) + transfer_prover +qed + +end + +lemma exp_eq_zero_iff [simp]: + \2 ^ n = (0 :: 'a::len word) \ n \ LENGTH('a)\ + by transfer simp + +lemma word_exp_length_eq_0 [simp]: + \(2 :: 'a::len word) ^ LENGTH('a) = 0\ + by simp + + +subsubsection \Basic tool setup\ + +ML_file \Tools/word_lib.ML\ + + +subsubsection \Basic code generation setup\ + +context +begin + +qualified lift_definition the_int :: \'a::len word \ int\ + is \take_bit LENGTH('a)\ . + +end + +lemma [code abstype]: + \Word.Word (Word.the_int w) = w\ + by transfer simp + +lemma Word_eq_word_of_int [code_post, simp]: + \Word.Word = of_int\ + by (rule; transfer) simp + +quickcheck_generator word + constructors: + \0 :: 'a::len word\, + \numeral :: num \ 'a::len word\ + +instantiation word :: (len) equal +begin + +lift_definition equal_word :: \'a word \ 'a word \ bool\ + is \\k l. take_bit LENGTH('a) k = take_bit LENGTH('a) l\ + by simp + +instance + by (standard; transfer) rule + +end + +lemma [code]: + \HOL.equal v w \ HOL.equal (Word.the_int v) (Word.the_int w)\ + by transfer (simp add: equal) + +lemma [code]: + \Word.the_int 0 = 0\ + by transfer simp + +lemma [code]: + \Word.the_int 1 = 1\ + by transfer simp + +lemma [code]: + \Word.the_int (v + w) = take_bit LENGTH('a) (Word.the_int v + Word.the_int w)\ + for v w :: \'a::len word\ + by transfer (simp add: take_bit_add) + +lemma [code]: + \Word.the_int (- w) = (let k = Word.the_int w in if w = 0 then 0 else 2 ^ LENGTH('a) - k)\ + for w :: \'a::len word\ + by transfer (auto simp add: take_bit_eq_mod zmod_zminus1_eq_if) + +lemma [code]: + \Word.the_int (v - w) = take_bit LENGTH('a) (Word.the_int v - Word.the_int w)\ + for v w :: \'a::len word\ + by transfer (simp add: take_bit_diff) + +lemma [code]: + \Word.the_int (v * w) = take_bit LENGTH('a) (Word.the_int v * Word.the_int w)\ + for v w :: \'a::len word\ + by transfer (simp add: take_bit_mult) + + +subsubsection \Basic conversions\ + +abbreviation word_of_nat :: \nat \ 'a::len word\ + where \word_of_nat \ of_nat\ + +abbreviation word_of_int :: \int \ 'a::len word\ + where \word_of_int \ of_int\ + +lemma word_of_nat_eq_iff: + \word_of_nat m = (word_of_nat n :: 'a::len word) \ take_bit LENGTH('a) m = take_bit LENGTH('a) n\ + by transfer (simp add: take_bit_of_nat) + +lemma word_of_int_eq_iff: + \word_of_int k = (word_of_int l :: 'a::len word) \ take_bit LENGTH('a) k = take_bit LENGTH('a) l\ + by transfer rule + +lemma word_of_nat_eq_0_iff [simp]: + \word_of_nat n = (0 :: 'a::len word) \ 2 ^ LENGTH('a) dvd n\ + using word_of_nat_eq_iff [where ?'a = 'a, of n 0] by (simp add: take_bit_eq_0_iff) + +lemma word_of_int_eq_0_iff [simp]: + \word_of_int k = (0 :: 'a::len word) \ 2 ^ LENGTH('a) dvd k\ + using word_of_int_eq_iff [where ?'a = 'a, of k 0] by (simp add: take_bit_eq_0_iff) + +context semiring_1 +begin + +lift_definition unsigned :: \'b::len word \ 'a\ + is \of_nat \ nat \ take_bit LENGTH('b)\ + by simp + +lemma unsigned_0 [simp]: + \unsigned 0 = 0\ + by transfer simp + +lemma unsigned_1 [simp]: + \unsigned 1 = 1\ + by transfer simp + +lemma unsigned_numeral [simp]: + \unsigned (numeral n :: 'b::len word) = of_nat (take_bit LENGTH('b) (numeral n))\ + by transfer (simp add: nat_take_bit_eq) + +lemma unsigned_neg_numeral [simp]: + \unsigned (- numeral n :: 'b::len word) = of_nat (nat (take_bit LENGTH('b) (- numeral n)))\ + by transfer simp + +end + +context semiring_1 +begin + +lemma unsigned_of_nat [simp]: + \unsigned (word_of_nat n :: 'b::len word) = of_nat (take_bit LENGTH('b) n)\ + by transfer (simp add: nat_eq_iff take_bit_of_nat) + +lemma unsigned_of_int [simp]: + \unsigned (word_of_int k :: 'b::len word) = of_nat (nat (take_bit LENGTH('b) k))\ + by transfer simp + +end + +context semiring_char_0 +begin + +lemma unsigned_word_eqI: + \v = w\ if \unsigned v = unsigned w\ + using that by transfer (simp add: eq_nat_nat_iff) + +lemma word_eq_iff_unsigned: + \v = w \ unsigned v = unsigned w\ + by (auto intro: unsigned_word_eqI) + +lemma inj_unsigned [simp]: + \inj unsigned\ + by (rule injI) (simp add: unsigned_word_eqI) + +lemma unsigned_eq_0_iff: + \unsigned w = 0 \ w = 0\ + using word_eq_iff_unsigned [of w 0] by simp + +end + +context ring_1 +begin + +lift_definition signed :: \'b::len word \ 'a\ + is \of_int \ signed_take_bit (LENGTH('b) - Suc 0)\ + by (simp flip: signed_take_bit_decr_length_iff) + +lemma signed_0 [simp]: + \signed 0 = 0\ + by transfer simp + +lemma signed_1 [simp]: + \signed (1 :: 'b::len word) = (if LENGTH('b) = 1 then - 1 else 1)\ + by (transfer fixing: uminus; cases \LENGTH('b)\) (auto dest: gr0_implies_Suc) + +lemma signed_minus_1 [simp]: + \signed (- 1 :: 'b::len word) = - 1\ + by (transfer fixing: uminus) simp + +lemma signed_numeral [simp]: + \signed (numeral n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - 1) (numeral n))\ + by transfer simp + +lemma signed_neg_numeral [simp]: + \signed (- numeral n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - 1) (- numeral n))\ + by transfer simp + +lemma signed_of_nat [simp]: + \signed (word_of_nat n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - Suc 0) (int n))\ + by transfer simp + +lemma signed_of_int [simp]: + \signed (word_of_int n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - Suc 0) n)\ + by transfer simp + +end + +context ring_char_0 +begin + +lemma signed_word_eqI: + \v = w\ if \signed v = signed w\ + using that by transfer (simp flip: signed_take_bit_decr_length_iff) + +lemma word_eq_iff_signed: + \v = w \ signed v = signed w\ + by (auto intro: signed_word_eqI) + +lemma inj_signed [simp]: + \inj signed\ + by (rule injI) (simp add: signed_word_eqI) + +lemma signed_eq_0_iff: + \signed w = 0 \ w = 0\ + using word_eq_iff_signed [of w 0] by simp + +end + +abbreviation unat :: \'a::len word \ nat\ + where \unat \ unsigned\ + +abbreviation uint :: \'a::len word \ int\ + where \uint \ unsigned\ + +abbreviation sint :: \'a::len word \ int\ + where \sint \ signed\ + +abbreviation ucast :: \'a::len word \ 'b::len word\ + where \ucast \ unsigned\ + +abbreviation scast :: \'a::len word \ 'b::len word\ + where \scast \ signed\ + +context + includes lifting_syntax +begin + +lemma [transfer_rule]: + \(pcr_word ===> (=)) (nat \ take_bit LENGTH('a)) (unat :: 'a::len word \ nat)\ + using unsigned.transfer [where ?'a = nat] by simp + +lemma [transfer_rule]: + \(pcr_word ===> (=)) (take_bit LENGTH('a)) (uint :: 'a::len word \ int)\ + using unsigned.transfer [where ?'a = int] by (simp add: comp_def) + +lemma [transfer_rule]: + \(pcr_word ===> (=)) (signed_take_bit (LENGTH('a) - Suc 0)) (sint :: 'a::len word \ int)\ + using signed.transfer [where ?'a = int] by simp + +lemma [transfer_rule]: + \(pcr_word ===> pcr_word) (take_bit LENGTH('a)) (ucast :: 'a::len word \ 'b::len word)\ +proof (rule rel_funI) + fix k :: int and w :: \'a word\ + assume \pcr_word k w\ + then have \w = word_of_int k\ + by (simp add: pcr_word_def cr_word_def relcompp_apply) + moreover have \pcr_word (take_bit LENGTH('a) k) (ucast (word_of_int k :: 'a word))\ + by transfer (simp add: pcr_word_def cr_word_def relcompp_apply) + ultimately show \pcr_word (take_bit LENGTH('a) k) (ucast w)\ + by simp +qed + +lemma [transfer_rule]: + \(pcr_word ===> pcr_word) (signed_take_bit (LENGTH('a) - Suc 0)) (scast :: 'a::len word \ 'b::len word)\ +proof (rule rel_funI) + fix k :: int and w :: \'a word\ + assume \pcr_word k w\ + then have \w = word_of_int k\ + by (simp add: pcr_word_def cr_word_def relcompp_apply) + moreover have \pcr_word (signed_take_bit (LENGTH('a) - Suc 0) k) (scast (word_of_int k :: 'a word))\ + by transfer (simp add: pcr_word_def cr_word_def relcompp_apply) + ultimately show \pcr_word (signed_take_bit (LENGTH('a) - Suc 0) k) (scast w)\ + by simp +qed + +end + +lemma of_nat_unat [simp]: + \of_nat (unat w) = unsigned w\ + by transfer simp + +lemma of_int_uint [simp]: + \of_int (uint w) = unsigned w\ + by transfer simp + +lemma of_int_sint [simp]: + \of_int (sint a) = signed a\ + by transfer (simp_all add: take_bit_signed_take_bit) + +lemma nat_uint_eq [simp]: + \nat (uint w) = unat w\ + by transfer simp + +lemma sgn_uint_eq [simp]: + \sgn (uint w) = of_bool (w \ 0)\ + by transfer (simp add: less_le) + +text \Aliasses only for code generation\ + +context +begin + +qualified lift_definition of_int :: \int \ 'a::len word\ + is \take_bit LENGTH('a)\ . + +qualified lift_definition of_nat :: \nat \ 'a::len word\ + is \int \ take_bit LENGTH('a)\ . + +qualified lift_definition the_nat :: \'a::len word \ nat\ + is \nat \ take_bit LENGTH('a)\ by simp + +qualified lift_definition the_signed_int :: \'a::len word \ int\ + is \signed_take_bit (LENGTH('a) - Suc 0)\ by (simp add: signed_take_bit_decr_length_iff) + +qualified lift_definition cast :: \'a::len word \ 'b::len word\ + is \take_bit LENGTH('a)\ by simp + +qualified lift_definition signed_cast :: \'a::len word \ 'b::len word\ + is \signed_take_bit (LENGTH('a) - Suc 0)\ by (metis signed_take_bit_decr_length_iff) + +end + +lemma [code_abbrev, simp]: + \Word.the_int = uint\ + by transfer rule + +lemma [code]: + \Word.the_int (Word.of_int k :: 'a::len word) = take_bit LENGTH('a) k\ + by transfer simp + +lemma [code_abbrev, simp]: + \Word.of_int = word_of_int\ + by (rule; transfer) simp + +lemma [code]: + \Word.the_int (Word.of_nat n :: 'a::len word) = take_bit LENGTH('a) (int n)\ + by transfer (simp add: take_bit_of_nat) + +lemma [code_abbrev, simp]: + \Word.of_nat = word_of_nat\ + by (rule; transfer) (simp add: take_bit_of_nat) + +lemma [code]: + \Word.the_nat w = nat (Word.the_int w)\ + by transfer simp + +lemma [code_abbrev, simp]: + \Word.the_nat = unat\ + by (rule; transfer) simp + +lemma [code]: + \Word.the_signed_int w = signed_take_bit (LENGTH('a) - Suc 0) (Word.the_int w)\ + for w :: \'a::len word\ + by transfer (simp add: signed_take_bit_take_bit) + +lemma [code_abbrev, simp]: + \Word.the_signed_int = sint\ + by (rule; transfer) simp + +lemma [code]: + \Word.the_int (Word.cast w :: 'b::len word) = take_bit LENGTH('b) (Word.the_int w)\ + for w :: \'a::len word\ + by transfer simp + +lemma [code_abbrev, simp]: + \Word.cast = ucast\ + by (rule; transfer) simp + +lemma [code]: + \Word.the_int (Word.signed_cast w :: 'b::len word) = take_bit LENGTH('b) (Word.the_signed_int w)\ + for w :: \'a::len word\ + by transfer simp + +lemma [code_abbrev, simp]: + \Word.signed_cast = scast\ + by (rule; transfer) simp + +lemma [code]: + \unsigned w = of_nat (nat (Word.the_int w))\ + by transfer simp + +lemma [code]: + \signed w = of_int (Word.the_signed_int w)\ + by transfer simp + + +subsubsection \Basic ordering\ + +instantiation word :: (len) linorder +begin + +lift_definition less_eq_word :: "'a word \ 'a word \ bool" + is "\a b. take_bit LENGTH('a) a \ take_bit LENGTH('a) b" + by simp + +lift_definition less_word :: "'a word \ 'a word \ bool" + is "\a b. take_bit LENGTH('a) a < take_bit LENGTH('a) b" + by simp + +instance + by (standard; transfer) auto + +end + +interpretation word_order: ordering_top \(\)\ \(<)\ \- 1 :: 'a::len word\ + by (standard; transfer) (simp add: take_bit_eq_mod zmod_minus1) + +interpretation word_coorder: ordering_top \(\)\ \(>)\ \0 :: 'a::len word\ + by (standard; transfer) simp + +lemma word_of_nat_less_eq_iff: + \word_of_nat m \ (word_of_nat n :: 'a::len word) \ take_bit LENGTH('a) m \ take_bit LENGTH('a) n\ + by transfer (simp add: take_bit_of_nat) + +lemma word_of_int_less_eq_iff: + \word_of_int k \ (word_of_int l :: 'a::len word) \ take_bit LENGTH('a) k \ take_bit LENGTH('a) l\ + by transfer rule + +lemma word_of_nat_less_iff: + \word_of_nat m < (word_of_nat n :: 'a::len word) \ take_bit LENGTH('a) m < take_bit LENGTH('a) n\ + by transfer (simp add: take_bit_of_nat) + +lemma word_of_int_less_iff: + \word_of_int k < (word_of_int l :: 'a::len word) \ take_bit LENGTH('a) k < take_bit LENGTH('a) l\ + by transfer rule + +lemma word_le_def [code]: + "a \ b \ uint a \ uint b" + by transfer rule + +lemma word_less_def [code]: + "a < b \ uint a < uint b" + by transfer rule + +lemma word_greater_zero_iff: + \a > 0 \ a \ 0\ for a :: \'a::len word\ + by transfer (simp add: less_le) + +lemma of_nat_word_less_eq_iff: + \of_nat m \ (of_nat n :: 'a::len word) \ take_bit LENGTH('a) m \ take_bit LENGTH('a) n\ + by transfer (simp add: take_bit_of_nat) + +lemma of_nat_word_less_iff: + \of_nat m < (of_nat n :: 'a::len word) \ take_bit LENGTH('a) m < take_bit LENGTH('a) n\ + by transfer (simp add: take_bit_of_nat) + +lemma of_int_word_less_eq_iff: + \of_int k \ (of_int l :: 'a::len word) \ take_bit LENGTH('a) k \ take_bit LENGTH('a) l\ + by transfer rule + +lemma of_int_word_less_iff: + \of_int k < (of_int l :: 'a::len word) \ take_bit LENGTH('a) k < take_bit LENGTH('a) l\ + by transfer rule + + + +subsection \Enumeration\ + +lemma inj_on_word_of_nat: + \inj_on (word_of_nat :: nat \ 'a::len word) {0..<2 ^ LENGTH('a)}\ + by (rule inj_onI; transfer) (simp_all add: take_bit_int_eq_self) + +lemma UNIV_word_eq_word_of_nat: + \(UNIV :: 'a::len word set) = word_of_nat ` {0..<2 ^ LENGTH('a)}\ (is \_ = ?A\) +proof + show \word_of_nat ` {0..<2 ^ LENGTH('a)} \ UNIV\ + by simp + show \UNIV \ ?A\ + proof + fix w :: \'a word\ + show \w \ (word_of_nat ` {0..<2 ^ LENGTH('a)} :: 'a word set)\ + by (rule image_eqI [of _ _ \unat w\]; transfer) simp_all + qed +qed + +instantiation word :: (len) enum +begin + +definition enum_word :: \'a word list\ + where \enum_word = map word_of_nat [0..<2 ^ LENGTH('a)]\ + +definition enum_all_word :: \('a word \ bool) \ bool\ + where \enum_all_word = Ball UNIV\ + +definition enum_ex_word :: \('a word \ bool) \ bool\ + where \enum_ex_word = Bex UNIV\ + +lemma [code]: + \Enum.enum_all P \ Ball UNIV P\ + \Enum.enum_ex P \ Bex UNIV P\ for P :: \'a word \ bool\ + by (simp_all add: enum_all_word_def enum_ex_word_def) + +instance + by standard (simp_all add: UNIV_word_eq_word_of_nat inj_on_word_of_nat enum_word_def enum_all_word_def enum_ex_word_def distinct_map) + +end + + +subsection \Bit-wise operations\ + +instantiation word :: (len) semiring_modulo +begin + +lift_definition divide_word :: \'a word \ 'a word \ 'a word\ + is \\a b. take_bit LENGTH('a) a div take_bit LENGTH('a) b\ + by simp + +lift_definition modulo_word :: \'a word \ 'a word \ 'a word\ + is \\a b. take_bit LENGTH('a) a mod take_bit LENGTH('a) b\ + by simp + +instance proof + show "a div b * b + a mod b = a" for a b :: "'a word" + proof transfer + fix k l :: int + define r :: int where "r = 2 ^ LENGTH('a)" + then have r: "take_bit LENGTH('a) k = k mod r" for k + by (simp add: take_bit_eq_mod) + have "k mod r = ((k mod r) div (l mod r) * (l mod r) + + (k mod r) mod (l mod r)) mod r" + by (simp add: div_mult_mod_eq) + also have "... = (((k mod r) div (l mod r) * (l mod r)) mod r + + (k mod r) mod (l mod r)) mod r" + by (simp add: mod_add_left_eq) + also have "... = (((k mod r) div (l mod r) * l) mod r + + (k mod r) mod (l mod r)) mod r" + by (simp add: mod_mult_right_eq) + finally have "k mod r = ((k mod r) div (l mod r) * l + + (k mod r) mod (l mod r)) mod r" + by (simp add: mod_simps) + with r show "take_bit LENGTH('a) (take_bit LENGTH('a) k div take_bit LENGTH('a) l * l + + take_bit LENGTH('a) k mod take_bit LENGTH('a) l) = take_bit LENGTH('a) k" + by simp + qed +qed + +end + +instance word :: (len) semiring_parity +proof + show "\ 2 dvd (1::'a word)" + by transfer simp + show even_iff_mod_2_eq_0: "2 dvd a \ a mod 2 = 0" + for a :: "'a word" + by transfer (simp_all add: mod_2_eq_odd take_bit_Suc) + show "\ 2 dvd a \ a mod 2 = 1" + for a :: "'a word" + by transfer (simp_all add: mod_2_eq_odd take_bit_Suc) +qed + +lemma word_bit_induct [case_names zero even odd]: + \P a\ if word_zero: \P 0\ + and word_even: \\a. P a \ 0 < a \ a < 2 ^ (LENGTH('a) - Suc 0) \ P (2 * a)\ + and word_odd: \\a. P a \ a < 2 ^ (LENGTH('a) - Suc 0) \ P (1 + 2 * a)\ + for P and a :: \'a::len word\ +proof - + define m :: nat where \m = LENGTH('a) - Suc 0\ + then have l: \LENGTH('a) = Suc m\ + by simp + define n :: nat where \n = unat a\ + then have \n < 2 ^ LENGTH('a)\ + by transfer (simp add: take_bit_eq_mod) + then have \n < 2 * 2 ^ m\ + by (simp add: l) + then have \P (of_nat n)\ + proof (induction n rule: nat_bit_induct) + case zero + show ?case + by simp (rule word_zero) + next + case (even n) + then have \n < 2 ^ m\ + by simp + with even.IH have \P (of_nat n)\ + by simp + moreover from \n < 2 ^ m\ even.hyps have \0 < (of_nat n :: 'a word)\ + by (auto simp add: word_greater_zero_iff l) + moreover from \n < 2 ^ m\ have \(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - Suc 0)\ + using of_nat_word_less_iff [where ?'a = 'a, of n \2 ^ m\] + by (simp add: l take_bit_eq_mod) + ultimately have \P (2 * of_nat n)\ + by (rule word_even) + then show ?case + by simp + next + case (odd n) + then have \Suc n \ 2 ^ m\ + by simp + with odd.IH have \P (of_nat n)\ + by simp + moreover from \Suc n \ 2 ^ m\ have \(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - Suc 0)\ + using of_nat_word_less_iff [where ?'a = 'a, of n \2 ^ m\] + by (simp add: l take_bit_eq_mod) + ultimately have \P (1 + 2 * of_nat n)\ + by (rule word_odd) + then show ?case + by simp + qed + moreover have \of_nat (nat (uint a)) = a\ + by transfer simp + ultimately show ?thesis + by (simp add: n_def) +qed + +lemma bit_word_half_eq: + \(of_bool b + a * 2) div 2 = a\ + if \a < 2 ^ (LENGTH('a) - Suc 0)\ + for a :: \'a::len word\ +proof (cases \2 \ LENGTH('a::len)\) + case False + have \of_bool (odd k) < (1 :: int) \ even k\ for k :: int + by auto + with False that show ?thesis + by transfer (simp add: eq_iff) +next + case True + obtain n where length: \LENGTH('a) = Suc n\ + by (cases \LENGTH('a)\) simp_all + show ?thesis proof (cases b) + case False + moreover have \a * 2 div 2 = a\ + using that proof transfer + fix k :: int + from length have \k * 2 mod 2 ^ LENGTH('a) = (k mod 2 ^ n) * 2\ + by simp + moreover assume \take_bit LENGTH('a) k < take_bit LENGTH('a) (2 ^ (LENGTH('a) - Suc 0))\ + with \LENGTH('a) = Suc n\ + have \k mod 2 ^ LENGTH('a) = k mod 2 ^ n\ + by (simp add: take_bit_eq_mod divmod_digit_0) + ultimately have \take_bit LENGTH('a) (k * 2) = take_bit LENGTH('a) k * 2\ + by (simp add: take_bit_eq_mod) + with True show \take_bit LENGTH('a) (take_bit LENGTH('a) (k * 2) div take_bit LENGTH('a) 2) + = take_bit LENGTH('a) k\ + by simp + qed + ultimately show ?thesis + by simp + next + case True + moreover have \(1 + a * 2) div 2 = a\ + using that proof transfer + fix k :: int + from length have \(1 + k * 2) mod 2 ^ LENGTH('a) = 1 + (k mod 2 ^ n) * 2\ + using pos_zmod_mult_2 [of \2 ^ n\ k] by (simp add: ac_simps) + moreover assume \take_bit LENGTH('a) k < take_bit LENGTH('a) (2 ^ (LENGTH('a) - Suc 0))\ + with \LENGTH('a) = Suc n\ + have \k mod 2 ^ LENGTH('a) = k mod 2 ^ n\ + by (simp add: take_bit_eq_mod divmod_digit_0) + ultimately have \take_bit LENGTH('a) (1 + k * 2) = 1 + take_bit LENGTH('a) k * 2\ + by (simp add: take_bit_eq_mod) + with True show \take_bit LENGTH('a) (take_bit LENGTH('a) (1 + k * 2) div take_bit LENGTH('a) 2) + = take_bit LENGTH('a) k\ + by (auto simp add: take_bit_Suc) + qed + ultimately show ?thesis + by simp + qed +qed + +lemma even_mult_exp_div_word_iff: + \even (a * 2 ^ m div 2 ^ n) \ \ ( + m \ n \ + n < LENGTH('a) \ odd (a div 2 ^ (n - m)))\ for a :: \'a::len word\ + by transfer + (auto simp flip: drop_bit_eq_div simp add: even_drop_bit_iff_not_bit bit_take_bit_iff, + simp_all flip: push_bit_eq_mult add: bit_push_bit_iff_int) + +instantiation word :: (len) semiring_bits +begin + +lift_definition bit_word :: \'a word \ nat \ bool\ + is \\k n. n < LENGTH('a) \ bit k n\ +proof + fix k l :: int and n :: nat + assume *: \take_bit LENGTH('a) k = take_bit LENGTH('a) l\ + show \n < LENGTH('a) \ bit k n \ n < LENGTH('a) \ bit l n\ + proof (cases \n < LENGTH('a)\) + case True + from * have \bit (take_bit LENGTH('a) k) n \ bit (take_bit LENGTH('a) l) n\ + by simp + then show ?thesis + by (simp add: bit_take_bit_iff) + next + case False + then show ?thesis + by simp + qed +qed + +instance proof + show \P a\ if stable: \\a. a div 2 = a \ P a\ + and rec: \\a b. P a \ (of_bool b + 2 * a) div 2 = a \ P (of_bool b + 2 * a)\ + for P and a :: \'a word\ + proof (induction a rule: word_bit_induct) + case zero + have \0 div 2 = (0::'a word)\ + by transfer simp + with stable [of 0] show ?case + by simp + next + case (even a) + with rec [of a False] show ?case + using bit_word_half_eq [of a False] by (simp add: ac_simps) + next + case (odd a) + with rec [of a True] show ?case + using bit_word_half_eq [of a True] by (simp add: ac_simps) + qed + show \bit a n \ odd (a div 2 ^ n)\ for a :: \'a word\ and n + by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit bit_iff_odd_drop_bit) + show \0 div a = 0\ + for a :: \'a word\ + by transfer simp + show \a div 1 = a\ + for a :: \'a word\ + by transfer simp + show \a mod b div b = 0\ + for a b :: \'a word\ + apply transfer + apply (simp add: take_bit_eq_mod) + apply (subst (3) mod_pos_pos_trivial [of _ \2 ^ LENGTH('a)\]) + apply simp_all + apply (metis le_less mod_by_0 pos_mod_conj zero_less_numeral zero_less_power) + using pos_mod_bound [of \2 ^ LENGTH('a)\] apply simp + proof - + fix aa :: int and ba :: int + have f1: "\i n. (i::int) mod 2 ^ n = 0 \ 0 < i mod 2 ^ n" + by (metis le_less take_bit_eq_mod take_bit_nonnegative) + have "(0::int) < 2 ^ len_of (TYPE('a)::'a itself) \ ba mod 2 ^ len_of (TYPE('a)::'a itself) \ 0 \ aa mod 2 ^ len_of (TYPE('a)::'a itself) mod (ba mod 2 ^ len_of (TYPE('a)::'a itself)) < 2 ^ len_of (TYPE('a)::'a itself)" + by (metis (no_types) mod_by_0 unique_euclidean_semiring_numeral_class.pos_mod_bound zero_less_numeral zero_less_power) + then show "aa mod 2 ^ len_of (TYPE('a)::'a itself) mod (ba mod 2 ^ len_of (TYPE('a)::'a itself)) < 2 ^ len_of (TYPE('a)::'a itself)" + using f1 by (meson le_less less_le_trans unique_euclidean_semiring_numeral_class.pos_mod_bound) + qed + show \(1 + a) div 2 = a div 2\ + if \even a\ + for a :: \'a word\ + using that by transfer + (auto dest: le_Suc_ex simp add: mod_2_eq_odd take_bit_Suc elim!: evenE) + show \(2 :: 'a word) ^ m div 2 ^ n = of_bool ((2 :: 'a word) ^ m \ 0 \ n \ m) * 2 ^ (m - n)\ + for m n :: nat + by transfer (simp, simp add: exp_div_exp_eq) + show "a div 2 ^ m div 2 ^ n = a div 2 ^ (m + n)" + for a :: "'a word" and m n :: nat + apply transfer + apply (auto simp add: not_less take_bit_drop_bit ac_simps simp flip: drop_bit_eq_div) + apply (simp add: drop_bit_take_bit) + done + show "a mod 2 ^ m mod 2 ^ n = a mod 2 ^ min m n" + for a :: "'a word" and m n :: nat + by transfer (auto simp flip: take_bit_eq_mod simp add: ac_simps) + show \a * 2 ^ m mod 2 ^ n = a mod 2 ^ (n - m) * 2 ^ m\ + if \m \ n\ for a :: "'a word" and m n :: nat + using that apply transfer + apply (auto simp flip: take_bit_eq_mod) + apply (auto simp flip: push_bit_eq_mult simp add: push_bit_take_bit split: split_min_lin) + done + show \a div 2 ^ n mod 2 ^ m = a mod (2 ^ (n + m)) div 2 ^ n\ + for a :: "'a word" and m n :: nat + by transfer (auto simp add: not_less take_bit_drop_bit ac_simps simp flip: take_bit_eq_mod drop_bit_eq_div split: split_min_lin) + show \even ((2 ^ m - 1) div (2::'a word) ^ n) \ 2 ^ n = (0::'a word) \ m \ n\ + for m n :: nat + by transfer (auto simp add: take_bit_of_mask even_mask_div_iff) + show \even (a * 2 ^ m div 2 ^ n) \ n < m \ (2::'a word) ^ n = 0 \ m \ n \ even (a div 2 ^ (n - m))\ + for a :: \'a word\ and m n :: nat + proof transfer + show \even (take_bit LENGTH('a) (k * 2 ^ m) div take_bit LENGTH('a) (2 ^ n)) \ + n < m + \ take_bit LENGTH('a) ((2::int) ^ n) = take_bit LENGTH('a) 0 + \ (m \ n \ even (take_bit LENGTH('a) k div take_bit LENGTH('a) (2 ^ (n - m))))\ + for m n :: nat and k l :: int + by (auto simp flip: take_bit_eq_mod drop_bit_eq_div push_bit_eq_mult + simp add: div_push_bit_of_1_eq_drop_bit drop_bit_take_bit drop_bit_push_bit_int [of n m]) + qed +qed + +end + +lemma bit_word_eqI: + \a = b\ if \\n. n < LENGTH('a) \ bit a n \ bit b n\ + for a b :: \'a::len word\ + using that by transfer (auto simp add: nat_less_le bit_eq_iff bit_take_bit_iff) + +lemma bit_imp_le_length: + \n < LENGTH('a)\ if \bit w n\ + for w :: \'a::len word\ + using that by transfer simp + +lemma not_bit_length [simp]: + \\ bit w LENGTH('a)\ for w :: \'a::len word\ + by transfer simp + +instantiation word :: (len) semiring_bit_shifts +begin + +lift_definition push_bit_word :: \nat \ 'a word \ 'a word\ + is push_bit +proof - + show \take_bit LENGTH('a) (push_bit n k) = take_bit LENGTH('a) (push_bit n l)\ + if \take_bit LENGTH('a) k = take_bit LENGTH('a) l\ for k l :: int and n :: nat + proof - + from that + have \take_bit (LENGTH('a) - n) (take_bit LENGTH('a) k) + = take_bit (LENGTH('a) - n) (take_bit LENGTH('a) l)\ + by simp + moreover have \min (LENGTH('a) - n) LENGTH('a) = LENGTH('a) - n\ + by simp + ultimately show ?thesis + by (simp add: take_bit_push_bit) + qed +qed + +lift_definition drop_bit_word :: \nat \ 'a word \ 'a word\ + is \\n. drop_bit n \ take_bit LENGTH('a)\ + by (simp add: take_bit_eq_mod) + +lift_definition take_bit_word :: \nat \ 'a word \ 'a word\ + is \\n. take_bit (min LENGTH('a) n)\ + by (simp add: ac_simps) (simp only: flip: take_bit_take_bit) + +instance proof + show \push_bit n a = a * 2 ^ n\ for n :: nat and a :: \'a word\ + by transfer (simp add: push_bit_eq_mult) + show \drop_bit n a = a div 2 ^ n\ for n :: nat and a :: \'a word\ + by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit) + show \take_bit n a = a mod 2 ^ n\ for n :: nat and a :: \'a word\ + by transfer (auto simp flip: take_bit_eq_mod) +qed + +end + +lemma [code]: + \push_bit n w = w * 2 ^ n\ for w :: \'a::len word\ + by (fact push_bit_eq_mult) + +lemma [code]: + \Word.the_int (drop_bit n w) = drop_bit n (Word.the_int w)\ + by transfer (simp add: drop_bit_take_bit min_def le_less less_diff_conv) + +lemma [code]: + \Word.the_int (take_bit n w) = (if n < LENGTH('a::len) then take_bit n (Word.the_int w) else Word.the_int w)\ + for w :: \'a::len word\ + by transfer (simp add: not_le not_less ac_simps min_absorb2) + + +instantiation word :: (len) ring_bit_operations +begin + +lift_definition not_word :: \'a word \ 'a word\ + is not + by (simp add: take_bit_not_iff) + +lift_definition and_word :: \'a word \ 'a word \ 'a word\ + is \and\ + by simp + +lift_definition or_word :: \'a word \ 'a word \ 'a word\ + is or + by simp + +lift_definition xor_word :: \'a word \ 'a word \ 'a word\ + is xor + by simp + +lift_definition mask_word :: \nat \ 'a word\ + is mask + . + +instance by (standard; transfer) + (auto simp add: minus_eq_not_minus_1 mask_eq_exp_minus_1 + bit_not_iff bit_and_iff bit_or_iff bit_xor_iff) + +end + +lemma [code_abbrev]: + \push_bit n 1 = (2 :: 'a::len word) ^ n\ + by (fact push_bit_of_1) + +lemma [code]: + \NOT w = Word.of_int (NOT (Word.the_int w))\ + for w :: \'a::len word\ + by transfer (simp add: take_bit_not_take_bit) + +lemma [code]: + \Word.the_int (v AND w) = Word.the_int v AND Word.the_int w\ + by transfer simp + +lemma [code]: + \Word.the_int (v OR w) = Word.the_int v OR Word.the_int w\ + by transfer simp + +lemma [code]: + \Word.the_int (v XOR w) = Word.the_int v XOR Word.the_int w\ + by transfer simp + +lemma [code]: + \Word.the_int (mask n :: 'a::len word) = mask (min LENGTH('a) n)\ + by transfer simp + +context + includes lifting_syntax +begin + +lemma set_bit_word_transfer [transfer_rule]: + \((=) ===> pcr_word ===> pcr_word) set_bit set_bit\ + by (unfold set_bit_def) transfer_prover + +lemma unset_bit_word_transfer [transfer_rule]: + \((=) ===> pcr_word ===> pcr_word) unset_bit unset_bit\ + by (unfold unset_bit_def) transfer_prover + +lemma flip_bit_word_transfer [transfer_rule]: + \((=) ===> pcr_word ===> pcr_word) flip_bit flip_bit\ + by (unfold flip_bit_def) transfer_prover + +lemma signed_take_bit_word_transfer [transfer_rule]: + \((=) ===> pcr_word ===> pcr_word) + (\n k. signed_take_bit n (take_bit LENGTH('a::len) k)) + (signed_take_bit :: nat \ 'a word \ 'a word)\ +proof - + let ?K = \\n (k :: int). take_bit (min LENGTH('a) n) k OR of_bool (n < LENGTH('a) \ bit k n) * NOT (mask n)\ + let ?W = \\n (w :: 'a word). take_bit n w OR of_bool (bit w n) * NOT (mask n)\ + have \((=) ===> pcr_word ===> pcr_word) ?K ?W\ + by transfer_prover + also have \?K = (\n k. signed_take_bit n (take_bit LENGTH('a::len) k))\ + by (simp add: fun_eq_iff signed_take_bit_def bit_take_bit_iff ac_simps) + also have \?W = signed_take_bit\ + by (simp add: fun_eq_iff signed_take_bit_def) + finally show ?thesis . +qed + +end + + +subsection \Conversions including casts\ + +subsubsection \Generic unsigned conversion\ + +context semiring_bits +begin + +lemma bit_unsigned_iff: + \bit (unsigned w) n \ 2 ^ n \ 0 \ bit w n\ + for w :: \'b::len word\ + by (transfer fixing: bit) (simp add: bit_of_nat_iff bit_nat_iff bit_take_bit_iff) + +end + +context semiring_bit_shifts +begin + +lemma unsigned_push_bit_eq: + \unsigned (push_bit n w) = take_bit LENGTH('b) (push_bit n (unsigned w))\ + for w :: \'b::len word\ +proof (rule bit_eqI) + fix m + assume \2 ^ m \ 0\ + show \bit (unsigned (push_bit n w)) m = bit (take_bit LENGTH('b) (push_bit n (unsigned w))) m\ + proof (cases \n \ m\) + case True + with \2 ^ m \ 0\ have \2 ^ (m - n) \ 0\ + by (metis (full_types) diff_add exp_add_not_zero_imp) + with True show ?thesis + by (simp add: bit_unsigned_iff bit_push_bit_iff Parity.bit_push_bit_iff bit_take_bit_iff not_le exp_eq_zero_iff ac_simps) + next + case False + then show ?thesis + by (simp add: not_le bit_unsigned_iff bit_push_bit_iff Parity.bit_push_bit_iff bit_take_bit_iff) + qed +qed + +lemma unsigned_take_bit_eq: + \unsigned (take_bit n w) = take_bit n (unsigned w)\ + for w :: \'b::len word\ + by (rule bit_eqI) (simp add: bit_unsigned_iff bit_take_bit_iff Parity.bit_take_bit_iff) + +end + +context unique_euclidean_semiring_with_bit_shifts +begin + +lemma unsigned_drop_bit_eq: + \unsigned (drop_bit n w) = drop_bit n (take_bit LENGTH('b) (unsigned w))\ + for w :: \'b::len word\ + by (rule bit_eqI) (auto simp add: bit_unsigned_iff bit_take_bit_iff bit_drop_bit_eq Parity.bit_drop_bit_eq dest: bit_imp_le_length) + +end + +context semiring_bit_operations +begin + +lemma unsigned_and_eq: + \unsigned (v AND w) = unsigned v AND unsigned w\ + for v w :: \'b::len word\ + by (rule bit_eqI) (simp add: bit_unsigned_iff bit_and_iff Bit_Operations.bit_and_iff) + +lemma unsigned_or_eq: + \unsigned (v OR w) = unsigned v OR unsigned w\ + for v w :: \'b::len word\ + by (rule bit_eqI) (simp add: bit_unsigned_iff bit_or_iff Bit_Operations.bit_or_iff) + +lemma unsigned_xor_eq: + \unsigned (v XOR w) = unsigned v XOR unsigned w\ + for v w :: \'b::len word\ + by (rule bit_eqI) (simp add: bit_unsigned_iff bit_xor_iff Bit_Operations.bit_xor_iff) + +end + +context ring_bit_operations +begin + +lemma unsigned_not_eq: + \unsigned (NOT w) = take_bit LENGTH('b) (NOT (unsigned w))\ + for w :: \'b::len word\ + by (rule bit_eqI) + (simp add: bit_unsigned_iff bit_take_bit_iff bit_not_iff Bit_Operations.bit_not_iff exp_eq_zero_iff not_le) + +end + +context unique_euclidean_semiring_numeral +begin + +lemma unsigned_greater_eq [simp]: + \0 \ unsigned w\ for w :: \'b::len word\ + by (transfer fixing: less_eq) simp + +lemma unsigned_less [simp]: + \unsigned w < 2 ^ LENGTH('b)\ for w :: \'b::len word\ + by (transfer fixing: less) simp + +end + +context linordered_semidom +begin + +lemma word_less_eq_iff_unsigned: + "a \ b \ unsigned a \ unsigned b" + by (transfer fixing: less_eq) (simp add: nat_le_eq_zle) + +lemma word_less_iff_unsigned: + "a < b \ unsigned a < unsigned b" + by (transfer fixing: less) (auto dest: preorder_class.le_less_trans [OF take_bit_nonnegative]) + +end + + +subsubsection \Generic signed conversion\ + +context ring_bit_operations +begin + +lemma bit_signed_iff: + \bit (signed w) n \ 2 ^ n \ 0 \ bit w (min (LENGTH('b) - Suc 0) n)\ + for w :: \'b::len word\ + by (transfer fixing: bit) + (auto simp add: bit_of_int_iff Bit_Operations.bit_signed_take_bit_iff min_def) + +lemma signed_push_bit_eq: + \signed (push_bit n w) = signed_take_bit (LENGTH('b) - Suc 0) (push_bit n (signed w :: 'a))\ + for w :: \'b::len word\ +proof (rule bit_eqI) + fix m + assume \2 ^ m \ 0\ + define q where \q = LENGTH('b) - Suc 0\ + then have *: \LENGTH('b) = Suc q\ + by simp + show \bit (signed (push_bit n w)) m \ + bit (signed_take_bit (LENGTH('b) - Suc 0) (push_bit n (signed w :: 'a))) m\ + proof (cases \q \ m\) + case True + moreover define r where \r = m - q\ + ultimately have \m = q + r\ + by simp + moreover from \m = q + r\ \2 ^ m \ 0\ have \2 ^ q \ 0\ \2 ^ r \ 0\ + using exp_add_not_zero_imp_left [of q r] exp_add_not_zero_imp_right [of q r] + by simp_all + moreover from \2 ^ q \ 0\ have \2 ^ (q - n) \ 0\ + by (rule exp_not_zero_imp_exp_diff_not_zero) + ultimately show ?thesis + by (auto simp add: bit_signed_iff bit_signed_take_bit_iff bit_push_bit_iff Parity.bit_push_bit_iff + min_def * exp_eq_zero_iff le_diff_conv2) + next + case False + then show ?thesis + using exp_not_zero_imp_exp_diff_not_zero [of m n] + by (auto simp add: bit_signed_iff bit_signed_take_bit_iff bit_push_bit_iff Parity.bit_push_bit_iff + min_def not_le not_less * le_diff_conv2 less_diff_conv2 Parity.exp_eq_0_imp_not_bit exp_eq_0_imp_not_bit + exp_eq_zero_iff) + qed +qed + +lemma signed_take_bit_eq: + \signed (take_bit n w) = (if n < LENGTH('b) then take_bit n (signed w) else signed w)\ + for w :: \'b::len word\ + by (transfer fixing: take_bit; cases \LENGTH('b)\) + (auto simp add: Bit_Operations.signed_take_bit_take_bit Bit_Operations.take_bit_signed_take_bit take_bit_of_int min_def less_Suc_eq) + +lemma signed_not_eq: + \signed (NOT w) = signed_take_bit LENGTH('b) (NOT (signed w))\ + for w :: \'b::len word\ +proof (rule bit_eqI) + fix n + assume \2 ^ n \ 0\ + define q where \q = LENGTH('b) - Suc 0\ + then have *: \LENGTH('b) = Suc q\ + by simp + show \bit (signed (NOT w)) n \ + bit (signed_take_bit LENGTH('b) (NOT (signed w))) n\ + proof (cases \q < n\) + case True + moreover define r where \r = n - Suc q\ + ultimately have \n = r + Suc q\ + by simp + moreover from \2 ^ n \ 0\ \n = r + Suc q\ + have \2 ^ Suc q \ 0\ + using exp_add_not_zero_imp_right by blast + ultimately show ?thesis + by (simp add: * bit_signed_iff bit_not_iff bit_signed_take_bit_iff Bit_Operations.bit_not_iff min_def + exp_eq_zero_iff) + next + case False + then show ?thesis + by (auto simp add: * bit_signed_iff bit_not_iff bit_signed_take_bit_iff Bit_Operations.bit_not_iff min_def + exp_eq_zero_iff) + qed +qed + +lemma signed_and_eq: + \signed (v AND w) = signed v AND signed w\ + for v w :: \'b::len word\ + by (rule bit_eqI) (simp add: bit_signed_iff bit_and_iff Bit_Operations.bit_and_iff) + +lemma signed_or_eq: + \signed (v OR w) = signed v OR signed w\ + for v w :: \'b::len word\ + by (rule bit_eqI) (simp add: bit_signed_iff bit_or_iff Bit_Operations.bit_or_iff) + +lemma signed_xor_eq: + \signed (v XOR w) = signed v XOR signed w\ + for v w :: \'b::len word\ + by (rule bit_eqI) (simp add: bit_signed_iff bit_xor_iff Bit_Operations.bit_xor_iff) + +end + + +subsubsection \More\ + +lemma sint_greater_eq: + \- (2 ^ (LENGTH('a) - Suc 0)) \ sint w\ for w :: \'a::len word\ +proof (cases \bit w (LENGTH('a) - Suc 0)\) + case True + then show ?thesis + by transfer (simp add: signed_take_bit_eq_if_negative minus_exp_eq_not_mask or_greater_eq ac_simps) +next + have *: \- (2 ^ (LENGTH('a) - Suc 0)) \ (0::int)\ + by simp + case False + then show ?thesis + by transfer (auto simp add: signed_take_bit_eq intro: order_trans *) +qed + +lemma sint_less: + \sint w < 2 ^ (LENGTH('a) - Suc 0)\ for w :: \'a::len word\ + by (cases \bit w (LENGTH('a) - Suc 0)\; transfer) + (simp_all add: signed_take_bit_eq signed_take_bit_def not_eq_complement mask_eq_exp_minus_1 OR_upper) + +lemma unat_div_distrib: + \unat (v div w) = unat v div unat w\ +proof transfer + fix k l + have \nat (take_bit LENGTH('a) k) div nat (take_bit LENGTH('a) l) \ nat (take_bit LENGTH('a) k)\ + by (rule div_le_dividend) + also have \nat (take_bit LENGTH('a) k) < 2 ^ LENGTH('a)\ + by (simp add: nat_less_iff) + finally show \(nat \ take_bit LENGTH('a)) (take_bit LENGTH('a) k div take_bit LENGTH('a) l) = + (nat \ take_bit LENGTH('a)) k div (nat \ take_bit LENGTH('a)) l\ + by (simp add: nat_take_bit_eq div_int_pos_iff nat_div_distrib take_bit_nat_eq_self_iff) +qed + +lemma unat_mod_distrib: + \unat (v mod w) = unat v mod unat w\ +proof transfer + fix k l + have \nat (take_bit LENGTH('a) k) mod nat (take_bit LENGTH('a) l) \ nat (take_bit LENGTH('a) k)\ + by (rule mod_less_eq_dividend) + also have \nat (take_bit LENGTH('a) k) < 2 ^ LENGTH('a)\ + by (simp add: nat_less_iff) + finally show \(nat \ take_bit LENGTH('a)) (take_bit LENGTH('a) k mod take_bit LENGTH('a) l) = + (nat \ take_bit LENGTH('a)) k mod (nat \ take_bit LENGTH('a)) l\ + by (simp add: nat_take_bit_eq mod_int_pos_iff less_le nat_mod_distrib take_bit_nat_eq_self_iff) +qed + +lemma uint_div_distrib: + \uint (v div w) = uint v div uint w\ +proof - + have \int (unat (v div w)) = int (unat v div unat w)\ + by (simp add: unat_div_distrib) + then show ?thesis + by (simp add: of_nat_div) +qed + +lemma unat_drop_bit_eq: + \unat (drop_bit n w) = drop_bit n (unat w)\ + by (rule bit_eqI) (simp add: bit_unsigned_iff bit_drop_bit_eq) + +lemma uint_mod_distrib: + \uint (v mod w) = uint v mod uint w\ +proof - + have \int (unat (v mod w)) = int (unat v mod unat w)\ + by (simp add: unat_mod_distrib) + then show ?thesis + by (simp add: of_nat_mod) +qed + +context semiring_bit_shifts +begin + +lemma unsigned_ucast_eq: + \unsigned (ucast w :: 'c::len word) = take_bit LENGTH('c) (unsigned w)\ + for w :: \'b::len word\ + by (rule bit_eqI) (simp add: bit_unsigned_iff Word.bit_unsigned_iff bit_take_bit_iff exp_eq_zero_iff not_le) + +end + +context ring_bit_operations +begin + +lemma signed_ucast_eq: + \signed (ucast w :: 'c::len word) = signed_take_bit (LENGTH('c) - Suc 0) (unsigned w)\ + for w :: \'b::len word\ +proof (rule bit_eqI) + fix n + assume \2 ^ n \ 0\ + then have \2 ^ (min (LENGTH('c) - Suc 0) n) \ 0\ + by (simp add: min_def) + (metis (mono_tags) diff_diff_cancel exp_not_zero_imp_exp_diff_not_zero) + then show \bit (signed (ucast w :: 'c::len word)) n \ bit (signed_take_bit (LENGTH('c) - Suc 0) (unsigned w)) n\ + by (simp add: bit_signed_iff bit_unsigned_iff Word.bit_unsigned_iff bit_signed_take_bit_iff exp_eq_zero_iff not_le) +qed + +lemma signed_scast_eq: + \signed (scast w :: 'c::len word) = signed_take_bit (LENGTH('c) - Suc 0) (signed w)\ + for w :: \'b::len word\ +proof (rule bit_eqI) + fix n + assume \2 ^ n \ 0\ + then have \2 ^ (min (LENGTH('c) - Suc 0) n) \ 0\ + by (simp add: min_def) + (metis (mono_tags) diff_diff_cancel exp_not_zero_imp_exp_diff_not_zero) + then show \bit (signed (scast w :: 'c::len word)) n \ bit (signed_take_bit (LENGTH('c) - Suc 0) (signed w)) n\ + by (simp add: bit_signed_iff bit_unsigned_iff Word.bit_signed_iff bit_signed_take_bit_iff exp_eq_zero_iff not_le) +qed + +end + +lemma uint_nonnegative: "0 \ uint w" + by (fact unsigned_greater_eq) + +lemma uint_bounded: "uint w < 2 ^ LENGTH('a)" + for w :: "'a::len word" + by (fact unsigned_less) + +lemma uint_idem: "uint w mod 2 ^ LENGTH('a) = uint w" + for w :: "'a::len word" + by transfer (simp add: take_bit_eq_mod) + +lemma word_uint_eqI: "uint a = uint b \ a = b" + by (fact unsigned_word_eqI) + +lemma word_uint_eq_iff: "a = b \ uint a = uint b" + by (fact word_eq_iff_unsigned) + +lemma uint_word_of_int_eq: + \uint (word_of_int k :: 'a::len word) = take_bit LENGTH('a) k\ + by transfer rule + +lemma uint_word_of_int: "uint (word_of_int k :: 'a::len word) = k mod 2 ^ LENGTH('a)" + by (simp add: uint_word_of_int_eq take_bit_eq_mod) + +lemma word_of_int_uint: "word_of_int (uint w) = w" + by transfer simp + +lemma word_div_def [code]: + "a div b = word_of_int (uint a div uint b)" + by transfer rule + +lemma word_mod_def [code]: + "a mod b = word_of_int (uint a mod uint b)" + by transfer rule + +lemma split_word_all: "(\x::'a::len word. PROP P x) \ (\x. PROP P (word_of_int x))" +proof + fix x :: "'a word" + assume "\x. PROP P (word_of_int x)" + then have "PROP P (word_of_int (uint x))" . + then show "PROP P x" + by (simp only: word_of_int_uint) +qed + +lemma sint_uint: + \sint w = signed_take_bit (LENGTH('a) - Suc 0) (uint w)\ + for w :: \'a::len word\ + by (cases \LENGTH('a)\; transfer) (simp_all add: signed_take_bit_take_bit) + +lemma unat_eq_nat_uint: + \unat w = nat (uint w)\ + by simp + +lemma ucast_eq: + \ucast w = word_of_int (uint w)\ + by transfer simp + +lemma scast_eq: + \scast w = word_of_int (sint w)\ + by transfer simp + +lemma uint_0_eq: + \uint 0 = 0\ + by (fact unsigned_0) + +lemma uint_1_eq: + \uint 1 = 1\ + by (fact unsigned_1) + +lemma word_m1_wi: "- 1 = word_of_int (- 1)" + by simp + +lemma uint_0_iff: "uint x = 0 \ x = 0" + by (auto simp add: unsigned_word_eqI) + +lemma unat_0_iff: "unat x = 0 \ x = 0" + by (auto simp add: unsigned_word_eqI) + +lemma unat_0: "unat 0 = 0" + by (fact unsigned_0) + +lemma unat_gt_0: "0 < unat x \ x \ 0" + by (auto simp: unat_0_iff [symmetric]) + +lemma ucast_0: "ucast 0 = 0" + by (fact unsigned_0) + +lemma sint_0: "sint 0 = 0" + by (fact signed_0) + +lemma scast_0: "scast 0 = 0" + by (fact signed_0) + +lemma sint_n1: "sint (- 1) = - 1" + by (fact signed_minus_1) + +lemma scast_n1: "scast (- 1) = - 1" + by (fact signed_minus_1) + +lemma uint_1: "uint (1::'a::len word) = 1" + by (fact uint_1_eq) + +lemma unat_1: "unat (1::'a::len word) = 1" + by (fact unsigned_1) + +lemma ucast_1: "ucast (1::'a::len word) = 1" + by (fact unsigned_1) + +instantiation word :: (len) size +begin + +lift_definition size_word :: \'a word \ nat\ + is \\_. LENGTH('a)\ .. + +instance .. + +end + +lemma word_size [code]: + \size w = LENGTH('a)\ for w :: \'a::len word\ + by (fact size_word.rep_eq) + +lemma word_size_gt_0 [iff]: "0 < size w" + for w :: "'a::len word" + by (simp add: word_size) + +lemmas lens_gt_0 = word_size_gt_0 len_gt_0 + +lemma lens_not_0 [iff]: + \size w \ 0\ for w :: \'a::len word\ + by auto + +lift_definition source_size :: \('a::len word \ 'b) \ nat\ + is \\_. LENGTH('a)\ . + +lift_definition target_size :: \('a \ 'b::len word) \ nat\ + is \\_. LENGTH('b)\ .. + +lift_definition is_up :: \('a::len word \ 'b::len word) \ bool\ + is \\_. LENGTH('a) \ LENGTH('b)\ .. + +lift_definition is_down :: \('a::len word \ 'b::len word) \ bool\ + is \\_. LENGTH('a) \ LENGTH('b)\ .. + +lemma is_up_eq: + \is_up f \ source_size f \ target_size f\ + for f :: \'a::len word \ 'b::len word\ + by (simp add: source_size.rep_eq target_size.rep_eq is_up.rep_eq) + +lemma is_down_eq: + \is_down f \ target_size f \ source_size f\ + for f :: \'a::len word \ 'b::len word\ + by (simp add: source_size.rep_eq target_size.rep_eq is_down.rep_eq) + +lift_definition word_int_case :: \(int \ 'b) \ 'a::len word \ 'b\ + is \\f. f \ take_bit LENGTH('a)\ by simp + +lemma word_int_case_eq_uint [code]: + \word_int_case f w = f (uint w)\ + by transfer simp + +translations + "case x of XCONST of_int y \ b" \ "CONST word_int_case (\y. b) x" + "case x of (XCONST of_int :: 'a) y \ b" \ "CONST word_int_case (\y. b) x" + + +subsection \Arithmetic operations\ + +text \Legacy theorems:\ + +lemma word_add_def [code]: + "a + b = word_of_int (uint a + uint b)" + by transfer (simp add: take_bit_add) + +lemma word_sub_wi [code]: + "a - b = word_of_int (uint a - uint b)" + by transfer (simp add: take_bit_diff) + +lemma word_mult_def [code]: + "a * b = word_of_int (uint a * uint b)" + by transfer (simp add: take_bit_eq_mod mod_simps) + +lemma word_minus_def [code]: + "- a = word_of_int (- uint a)" + by transfer (simp add: take_bit_minus) + +lemma word_0_wi: + "0 = word_of_int 0" + by transfer simp + +lemma word_1_wi: + "1 = word_of_int 1" + by transfer simp + +lift_definition word_succ :: "'a::len word \ 'a word" is "\x. x + 1" + by (auto simp add: take_bit_eq_mod intro: mod_add_cong) + +lift_definition word_pred :: "'a::len word \ 'a word" is "\x. x - 1" + by (auto simp add: take_bit_eq_mod intro: mod_diff_cong) + +lemma word_succ_alt [code]: + "word_succ a = word_of_int (uint a + 1)" + by transfer (simp add: take_bit_eq_mod mod_simps) + +lemma word_pred_alt [code]: + "word_pred a = word_of_int (uint a - 1)" + by transfer (simp add: take_bit_eq_mod mod_simps) + +lemmas word_arith_wis = + word_add_def word_sub_wi word_mult_def + word_minus_def word_succ_alt word_pred_alt + word_0_wi word_1_wi + +lemma wi_homs: + shows wi_hom_add: "word_of_int a + word_of_int b = word_of_int (a + b)" + and wi_hom_sub: "word_of_int a - word_of_int b = word_of_int (a - b)" + and wi_hom_mult: "word_of_int a * word_of_int b = word_of_int (a * b)" + and wi_hom_neg: "- word_of_int a = word_of_int (- a)" + and wi_hom_succ: "word_succ (word_of_int a) = word_of_int (a + 1)" + and wi_hom_pred: "word_pred (word_of_int a) = word_of_int (a - 1)" + by (transfer, simp)+ + +lemmas wi_hom_syms = wi_homs [symmetric] + +lemmas word_of_int_homs = wi_homs word_0_wi word_1_wi + +lemmas word_of_int_hom_syms = word_of_int_homs [symmetric] + +lemma double_eq_zero_iff: + \2 * a = 0 \ a = 0 \ a = 2 ^ (LENGTH('a) - Suc 0)\ + for a :: \'a::len word\ +proof - + define n where \n = LENGTH('a) - Suc 0\ + then have *: \LENGTH('a) = Suc n\ + by simp + have \a = 0\ if \2 * a = 0\ and \a \ 2 ^ (LENGTH('a) - Suc 0)\ + using that by transfer + (auto simp add: take_bit_eq_0_iff take_bit_eq_mod *) + moreover have \2 ^ LENGTH('a) = (0 :: 'a word)\ + by transfer simp + then have \2 * 2 ^ (LENGTH('a) - Suc 0) = (0 :: 'a word)\ + by (simp add: *) + ultimately show ?thesis + by auto +qed + + +subsection \Ordering\ + +lift_definition word_sle :: \'a::len word \ 'a word \ bool\ + is \\k l. signed_take_bit (LENGTH('a) - Suc 0) k \ signed_take_bit (LENGTH('a) - Suc 0) l\ + by (simp flip: signed_take_bit_decr_length_iff) + +lift_definition word_sless :: \'a::len word \ 'a word \ bool\ + is \\k l. signed_take_bit (LENGTH('a) - Suc 0) k < signed_take_bit (LENGTH('a) - Suc 0) l\ + by (simp flip: signed_take_bit_decr_length_iff) + +notation + word_sle ("'(\s')") and + word_sle ("(_/ \s _)" [51, 51] 50) and + word_sless ("'(a <=s b \ sint a \ sint b\ + by transfer simp + +lemma [code]: + \a sint a < sint b\ + by transfer simp + +lemma signed_ordering: \ordering word_sle word_sless\ + apply (standard; transfer) + apply simp_all + using signed_take_bit_decr_length_iff apply force + using signed_take_bit_decr_length_iff apply force + done + +lemma signed_linorder: \class.linorder word_sle word_sless\ + by (standard; transfer) (auto simp add: signed_take_bit_decr_length_iff) + +interpretation signed: linorder word_sle word_sless + by (fact signed_linorder) + +lemma word_sless_eq: + \x x <=s y \ x \ y\ + by (fact signed.less_le) + +lemma word_less_alt: "a < b \ uint a < uint b" + by (fact word_less_def) + +lemma word_zero_le [simp]: "0 \ y" + for y :: "'a::len word" + by (fact word_coorder.extremum) + +lemma word_m1_ge [simp] : "word_pred 0 \ y" (* FIXME: delete *) + by transfer (simp add: take_bit_minus_one_eq_mask mask_eq_exp_minus_1 ) + +lemma word_n1_ge [simp]: "y \ -1" + for y :: "'a::len word" + by (fact word_order.extremum) + +lemmas word_not_simps [simp] = + word_zero_le [THEN leD] word_m1_ge [THEN leD] word_n1_ge [THEN leD] + +lemma word_gt_0: "0 < y \ 0 \ y" + for y :: "'a::len word" + by (simp add: less_le) + +lemmas word_gt_0_no [simp] = word_gt_0 [of "numeral y"] for y + +lemma word_sless_alt: "a sint a < sint b" + by transfer simp + +lemma word_le_nat_alt: "a \ b \ unat a \ unat b" + by transfer (simp add: nat_le_eq_zle) + +lemma word_less_nat_alt: "a < b \ unat a < unat b" + by transfer (auto simp add: less_le [of 0]) + +lemmas unat_mono = word_less_nat_alt [THEN iffD1] + +instance word :: (len) wellorder +proof + fix P :: "'a word \ bool" and a + assume *: "(\b. (\a. a < b \ P a) \ P b)" + have "wf (measure unat)" .. + moreover have "{(a, b :: ('a::len) word). a < b} \ measure unat" + by (auto simp add: word_less_nat_alt) + ultimately have "wf {(a, b :: ('a::len) word). a < b}" + by (rule wf_subset) + then show "P a" using * + by induction blast +qed + +lemma wi_less: + "(word_of_int n < (word_of_int m :: 'a::len word)) = + (n mod 2 ^ LENGTH('a) < m mod 2 ^ LENGTH('a))" + by transfer (simp add: take_bit_eq_mod) + +lemma wi_le: + "(word_of_int n \ (word_of_int m :: 'a::len word)) = + (n mod 2 ^ LENGTH('a) \ m mod 2 ^ LENGTH('a))" + by transfer (simp add: take_bit_eq_mod) + + +subsection \Bit-wise operations\ + +lemma uint_take_bit_eq: + \uint (take_bit n w) = take_bit n (uint w)\ + by transfer (simp add: ac_simps) + +lemma take_bit_word_eq_self: + \take_bit n w = w\ if \LENGTH('a) \ n\ for w :: \'a::len word\ + using that by transfer simp + +lemma take_bit_length_eq [simp]: + \take_bit LENGTH('a) w = w\ for w :: \'a::len word\ + by (rule take_bit_word_eq_self) simp + +lemma bit_word_of_int_iff: + \bit (word_of_int k :: 'a::len word) n \ n < LENGTH('a) \ bit k n\ + by transfer rule + +lemma bit_uint_iff: + \bit (uint w) n \ n < LENGTH('a) \ bit w n\ + for w :: \'a::len word\ + by transfer (simp add: bit_take_bit_iff) + +lemma bit_sint_iff: + \bit (sint w) n \ n \ LENGTH('a) \ bit w (LENGTH('a) - 1) \ bit w n\ + for w :: \'a::len word\ + by transfer (auto simp add: bit_signed_take_bit_iff min_def le_less not_less) + +lemma bit_word_ucast_iff: + \bit (ucast w :: 'b::len word) n \ n < LENGTH('a) \ n < LENGTH('b) \ bit w n\ + for w :: \'a::len word\ + by transfer (simp add: bit_take_bit_iff ac_simps) + +lemma bit_word_scast_iff: + \bit (scast w :: 'b::len word) n \ + n < LENGTH('b) \ (bit w n \ LENGTH('a) \ n \ bit w (LENGTH('a) - Suc 0))\ + for w :: \'a::len word\ + by transfer (auto simp add: bit_signed_take_bit_iff le_less min_def) + +lift_definition shiftl1 :: \'a::len word \ 'a word\ + is \(*) 2\ + by (auto simp add: take_bit_eq_mod intro: mod_mult_cong) + +lemma shiftl1_eq: + \shiftl1 w = word_of_int (2 * uint w)\ + by transfer (simp add: take_bit_eq_mod mod_simps) + +lemma shiftl1_eq_mult_2: + \shiftl1 = (*) 2\ + by (rule ext, transfer) simp + +lemma bit_shiftl1_iff: + \bit (shiftl1 w) n \ 0 < n \ n < LENGTH('a) \ bit w (n - 1)\ + for w :: \'a::len word\ + by (simp add: shiftl1_eq_mult_2 bit_double_iff exp_eq_zero_iff not_le) (simp add: ac_simps) + +lift_definition shiftr1 :: \'a::len word \ 'a word\ + \ \shift right as unsigned or as signed, ie logical or arithmetic\ + is \\k. take_bit LENGTH('a) k div 2\ + by simp + +lemma shiftr1_eq_div_2: + \shiftr1 w = w div 2\ + by transfer simp + +lemma bit_shiftr1_iff: + \bit (shiftr1 w) n \ bit w (Suc n)\ + by transfer (auto simp flip: bit_Suc simp add: bit_take_bit_iff) + +lemma shiftr1_eq: + \shiftr1 w = word_of_int (uint w div 2)\ + by transfer simp + +lemma bit_word_iff_drop_bit_and [code]: + \bit a n \ drop_bit n a AND 1 = 1\ for a :: \'a::len word\ + by (simp add: bit_iff_odd_drop_bit odd_iff_mod_2_eq_one and_one_eq) + +lemma + word_not_def: "NOT (a::'a::len word) = word_of_int (NOT (uint a))" + and word_and_def: "(a::'a word) AND b = word_of_int (uint a AND uint b)" + and word_or_def: "(a::'a word) OR b = word_of_int (uint a OR uint b)" + and word_xor_def: "(a::'a word) XOR b = word_of_int (uint a XOR uint b)" + by (transfer, simp add: take_bit_not_take_bit)+ + +lift_definition setBit :: \'a::len word \ nat \ 'a word\ + is \\k n. set_bit n k\ + by (simp add: take_bit_set_bit_eq) + +lemma set_Bit_eq: + \setBit w n = set_bit n w\ + by transfer simp + +lemma bit_setBit_iff: + \bit (setBit w m) n \ (m = n \ n < LENGTH('a) \ bit w n)\ + for w :: \'a::len word\ + by transfer (auto simp add: bit_set_bit_iff) + +lift_definition clearBit :: \'a::len word \ nat \ 'a word\ + is \\k n. unset_bit n k\ + by (simp add: take_bit_unset_bit_eq) + +lemma clear_Bit_eq: + \clearBit w n = unset_bit n w\ + by transfer simp + +lemma bit_clearBit_iff: + \bit (clearBit w m) n \ m \ n \ bit w n\ + for w :: \'a::len word\ + by transfer (auto simp add: bit_unset_bit_iff) + +definition even_word :: \'a::len word \ bool\ + where [code_abbrev]: \even_word = even\ + +lemma even_word_iff [code]: + \even_word a \ a AND 1 = 0\ + by (simp add: and_one_eq even_iff_mod_2_eq_zero even_word_def) + +lemma map_bit_range_eq_if_take_bit_eq: + \map (bit k) [0.. + if \take_bit n k = take_bit n l\ for k l :: int +using that proof (induction n arbitrary: k l) + case 0 + then show ?case + by simp +next + case (Suc n) + from Suc.prems have \take_bit n (k div 2) = take_bit n (l div 2)\ + by (simp add: take_bit_Suc) + then have \map (bit (k div 2)) [0.. + by (rule Suc.IH) + moreover have \bit (r div 2) = bit r \ Suc\ for r :: int + by (simp add: fun_eq_iff bit_Suc) + moreover from Suc.prems have \even k \ even l\ + by (auto simp add: take_bit_Suc elim!: evenE oddE) arith+ + ultimately show ?case + by (simp only: map_Suc_upt upt_conv_Cons flip: list.map_comp) simp +qed + +lemma + take_bit_word_Bit0_eq [simp]: \take_bit (numeral n) (numeral (num.Bit0 m) :: 'a::len word) + = 2 * take_bit (pred_numeral n) (numeral m)\ (is ?P) + and take_bit_word_Bit1_eq [simp]: \take_bit (numeral n) (numeral (num.Bit1 m) :: 'a::len word) + = 1 + 2 * take_bit (pred_numeral n) (numeral m)\ (is ?Q) + and take_bit_word_minus_Bit0_eq [simp]: \take_bit (numeral n) (- numeral (num.Bit0 m) :: 'a::len word) + = 2 * take_bit (pred_numeral n) (- numeral m)\ (is ?R) + and take_bit_word_minus_Bit1_eq [simp]: \take_bit (numeral n) (- numeral (num.Bit1 m) :: 'a::len word) + = 1 + 2 * take_bit (pred_numeral n) (- numeral (Num.inc m))\ (is ?S) +proof - + define w :: \'a::len word\ + where \w = numeral m\ + moreover define q :: nat + where \q = pred_numeral n\ + ultimately have num: + \numeral m = w\ + \numeral (num.Bit0 m) = 2 * w\ + \numeral (num.Bit1 m) = 1 + 2 * w\ + \numeral (Num.inc m) = 1 + w\ + \pred_numeral n = q\ + \numeral n = Suc q\ + by (simp_all only: w_def q_def numeral_Bit0 [of m] numeral_Bit1 [of m] ac_simps + numeral_inc numeral_eq_Suc flip: mult_2) + have even: \take_bit (Suc q) (2 * w) = 2 * take_bit q w\ for w :: \'a::len word\ + by (rule bit_word_eqI) + (auto simp add: bit_take_bit_iff bit_double_iff) + have odd: \take_bit (Suc q) (1 + 2 * w) = 1 + 2 * take_bit q w\ for w :: \'a::len word\ + by (rule bit_eqI) + (auto simp add: bit_take_bit_iff bit_double_iff even_bit_succ_iff) + show ?P + using even [of w] by (simp add: num) + show ?Q + using odd [of w] by (simp add: num) + show ?R + using even [of \- w\] by (simp add: num) + show ?S + using odd [of \- (1 + w)\] by (simp add: num) +qed + + +subsection \More shift operations\ + +lift_definition signed_drop_bit :: \nat \ 'a word \ 'a::len word\ + is \\n. drop_bit n \ signed_take_bit (LENGTH('a) - Suc 0)\ + using signed_take_bit_decr_length_iff + by (simp add: take_bit_drop_bit) force + +lemma bit_signed_drop_bit_iff: + \bit (signed_drop_bit m w) n \ bit w (if LENGTH('a) - m \ n \ n < LENGTH('a) then LENGTH('a) - 1 else m + n)\ + for w :: \'a::len word\ + apply transfer + apply (auto simp add: bit_drop_bit_eq bit_signed_take_bit_iff not_le min_def) + apply (metis add.commute le_antisym less_diff_conv less_eq_decr_length_iff) + apply (metis le_antisym less_eq_decr_length_iff) + done + +lemma [code]: + \Word.the_int (signed_drop_bit n w) = take_bit LENGTH('a) (drop_bit n (Word.the_signed_int w))\ + for w :: \'a::len word\ + by transfer simp + +lemma signed_drop_bit_signed_drop_bit [simp]: + \signed_drop_bit m (signed_drop_bit n w) = signed_drop_bit (m + n) w\ + for w :: \'a::len word\ + apply (cases \LENGTH('a)\) + apply simp_all + apply (rule bit_word_eqI) + apply (auto simp add: bit_signed_drop_bit_iff not_le less_diff_conv ac_simps) + done + +lemma signed_drop_bit_0 [simp]: + \signed_drop_bit 0 w = w\ + by transfer (simp add: take_bit_signed_take_bit) + +lemma sint_signed_drop_bit_eq: + \sint (signed_drop_bit n w) = drop_bit n (sint w)\ + apply (cases \LENGTH('a)\; cases n) + apply simp_all + apply (rule bit_eqI) + apply (auto simp add: bit_sint_iff bit_drop_bit_eq bit_signed_drop_bit_iff dest: bit_imp_le_length) + done + +lift_definition sshiftr1 :: \'a::len word \ 'a word\ + is \\k. take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - Suc 0) k div 2)\ + by (simp flip: signed_take_bit_decr_length_iff) + +lift_definition bshiftr1 :: \bool \ 'a::len word \ 'a word\ + is \\b k. take_bit LENGTH('a) k div 2 + of_bool b * 2 ^ (LENGTH('a) - Suc 0)\ + by (fact arg_cong) + +lemma sshiftr1_eq_signed_drop_bit_Suc_0: + \sshiftr1 = signed_drop_bit (Suc 0)\ + by (rule ext) (transfer, simp add: drop_bit_Suc) + +lemma sshiftr1_eq: + \sshiftr1 w = word_of_int (sint w div 2)\ + by transfer simp + + +subsection \Rotation\ + +lift_definition word_rotr :: \nat \ 'a::len word \ 'a::len word\ + is \\n k. concat_bit (LENGTH('a) - n mod LENGTH('a)) + (drop_bit (n mod LENGTH('a)) (take_bit LENGTH('a) k)) + (take_bit (n mod LENGTH('a)) k)\ + subgoal for n k l + apply (simp add: concat_bit_def nat_le_iff less_imp_le + take_bit_tightened [of \LENGTH('a)\ k l \n mod LENGTH('a::len)\]) + done + done + +lift_definition word_rotl :: \nat \ 'a::len word \ 'a::len word\ + is \\n k. concat_bit (n mod LENGTH('a)) + (drop_bit (LENGTH('a) - n mod LENGTH('a)) (take_bit LENGTH('a) k)) + (take_bit (LENGTH('a) - n mod LENGTH('a)) k)\ + subgoal for n k l + apply (simp add: concat_bit_def nat_le_iff less_imp_le + take_bit_tightened [of \LENGTH('a)\ k l \LENGTH('a) - n mod LENGTH('a::len)\]) + done + done + +lift_definition word_roti :: \int \ 'a::len word \ 'a::len word\ + is \\r k. concat_bit (LENGTH('a) - nat (r mod int LENGTH('a))) + (drop_bit (nat (r mod int LENGTH('a))) (take_bit LENGTH('a) k)) + (take_bit (nat (r mod int LENGTH('a))) k)\ + subgoal for r k l + apply (simp add: concat_bit_def nat_le_iff less_imp_le + take_bit_tightened [of \LENGTH('a)\ k l \nat (r mod int LENGTH('a::len))\]) + done + done + +lemma word_rotl_eq_word_rotr [code]: + \word_rotl n = (word_rotr (LENGTH('a) - n mod LENGTH('a)) :: 'a::len word \ 'a word)\ + by (rule ext, cases \n mod LENGTH('a) = 0\; transfer) simp_all + +lemma word_roti_eq_word_rotr_word_rotl [code]: + \word_roti i w = + (if i \ 0 then word_rotr (nat i) w else word_rotl (nat (- i)) w)\ +proof (cases \i \ 0\) + case True + moreover define n where \n = nat i\ + ultimately have \i = int n\ + by simp + moreover have \word_roti (int n) = (word_rotr n :: _ \ 'a word)\ + by (rule ext, transfer) (simp add: nat_mod_distrib) + ultimately show ?thesis + by simp +next + case False + moreover define n where \n = nat (- i)\ + ultimately have \i = - int n\ \n > 0\ + by simp_all + moreover have \word_roti (- int n) = (word_rotl n :: _ \ 'a word)\ + by (rule ext, transfer) + (simp add: zmod_zminus1_eq_if flip: of_nat_mod of_nat_diff) + ultimately show ?thesis + by simp +qed + +lemma bit_word_rotr_iff: + \bit (word_rotr m w) n \ + n < LENGTH('a) \ bit w ((n + m) mod LENGTH('a))\ + for w :: \'a::len word\ +proof transfer + fix k :: int and m n :: nat + define q where \q = m mod LENGTH('a)\ + have \q < LENGTH('a)\ + by (simp add: q_def) + then have \q \ LENGTH('a)\ + by simp + have \m mod LENGTH('a) = q\ + by (simp add: q_def) + moreover have \(n + m) mod LENGTH('a) = (n + q) mod LENGTH('a)\ + by (subst mod_add_right_eq [symmetric]) (simp add: \m mod LENGTH('a) = q\) + moreover have \n < LENGTH('a) \ + bit (concat_bit (LENGTH('a) - q) (drop_bit q (take_bit LENGTH('a) k)) (take_bit q k)) n \ + n < LENGTH('a) \ bit k ((n + q) mod LENGTH('a))\ + using \q < LENGTH('a)\ + by (cases \q + n \ LENGTH('a)\) + (auto simp add: bit_concat_bit_iff bit_drop_bit_eq + bit_take_bit_iff le_mod_geq ac_simps) + ultimately show \n < LENGTH('a) \ + bit (concat_bit (LENGTH('a) - m mod LENGTH('a)) + (drop_bit (m mod LENGTH('a)) (take_bit LENGTH('a) k)) + (take_bit (m mod LENGTH('a)) k)) n + \ n < LENGTH('a) \ + (n + m) mod LENGTH('a) < LENGTH('a) \ + bit k ((n + m) mod LENGTH('a))\ + by simp +qed + +lemma bit_word_rotl_iff: + \bit (word_rotl m w) n \ + n < LENGTH('a) \ bit w ((n + (LENGTH('a) - m mod LENGTH('a))) mod LENGTH('a))\ + for w :: \'a::len word\ + by (simp add: word_rotl_eq_word_rotr bit_word_rotr_iff) + +lemma bit_word_roti_iff: + \bit (word_roti k w) n \ + n < LENGTH('a) \ bit w (nat ((int n + k) mod int LENGTH('a)))\ + for w :: \'a::len word\ +proof transfer + fix k l :: int and n :: nat + define m where \m = nat (k mod int LENGTH('a))\ + have \m < LENGTH('a)\ + by (simp add: nat_less_iff m_def) + then have \m \ LENGTH('a)\ + by simp + have \k mod int LENGTH('a) = int m\ + by (simp add: nat_less_iff m_def) + moreover have \(int n + k) mod int LENGTH('a) = int ((n + m) mod LENGTH('a))\ + by (subst mod_add_right_eq [symmetric]) (simp add: of_nat_mod \k mod int LENGTH('a) = int m\) + moreover have \n < LENGTH('a) \ + bit (concat_bit (LENGTH('a) - m) (drop_bit m (take_bit LENGTH('a) l)) (take_bit m l)) n \ + n < LENGTH('a) \ bit l ((n + m) mod LENGTH('a))\ + using \m < LENGTH('a)\ + by (cases \m + n \ LENGTH('a)\) + (auto simp add: bit_concat_bit_iff bit_drop_bit_eq + bit_take_bit_iff nat_less_iff not_le not_less ac_simps + le_diff_conv le_mod_geq) + ultimately show \n < LENGTH('a) + \ bit (concat_bit (LENGTH('a) - nat (k mod int LENGTH('a))) + (drop_bit (nat (k mod int LENGTH('a))) (take_bit LENGTH('a) l)) + (take_bit (nat (k mod int LENGTH('a))) l)) n \ + n < LENGTH('a) + \ nat ((int n + k) mod int LENGTH('a)) < LENGTH('a) + \ bit l (nat ((int n + k) mod int LENGTH('a)))\ + by simp +qed + +lemma uint_word_rotr_eq: + \uint (word_rotr n w) = concat_bit (LENGTH('a) - n mod LENGTH('a)) + (drop_bit (n mod LENGTH('a)) (uint w)) + (uint (take_bit (n mod LENGTH('a)) w))\ + for w :: \'a::len word\ + apply transfer + apply (simp add: concat_bit_def take_bit_drop_bit push_bit_take_bit min_def) + using mod_less_divisor not_less apply blast + done + +lemma [code]: + \Word.the_int (word_rotr n w) = concat_bit (LENGTH('a) - n mod LENGTH('a)) + (drop_bit (n mod LENGTH('a)) (Word.the_int w)) + (Word.the_int (take_bit (n mod LENGTH('a)) w))\ + for w :: \'a::len word\ + using uint_word_rotr_eq [of n w] by simp + + +subsection \Split and cat operations\ + +lift_definition word_cat :: \'a::len word \ 'b::len word \ 'c::len word\ + is \\k l. concat_bit LENGTH('b) l (take_bit LENGTH('a) k)\ + by (simp add: bit_eq_iff bit_concat_bit_iff bit_take_bit_iff) + +lemma word_cat_eq: + \(word_cat v w :: 'c::len word) = push_bit LENGTH('b) (ucast v) + ucast w\ + for v :: \'a::len word\ and w :: \'b::len word\ + by transfer (simp add: concat_bit_eq ac_simps) + +lemma word_cat_eq' [code]: + \word_cat a b = word_of_int (concat_bit LENGTH('b) (uint b) (uint a))\ + for a :: \'a::len word\ and b :: \'b::len word\ + by transfer (simp add: concat_bit_take_bit_eq) + +lemma bit_word_cat_iff: + \bit (word_cat v w :: 'c::len word) n \ n < LENGTH('c) \ (if n < LENGTH('b) then bit w n else bit v (n - LENGTH('b)))\ + for v :: \'a::len word\ and w :: \'b::len word\ + by transfer (simp add: bit_concat_bit_iff bit_take_bit_iff) + +definition word_split :: \'a::len word \ 'b::len word \ 'c::len word\ + where \word_split w = + (ucast (drop_bit LENGTH('c) w) :: 'b::len word, ucast w :: 'c::len word)\ + +definition word_rcat :: \'a::len word list \ 'b::len word\ + where \word_rcat = word_of_int \ horner_sum uint (2 ^ LENGTH('a)) \ rev\ + +abbreviation (input) max_word :: \'a::len word\ + \ \Largest representable machine integer.\ + where "max_word \ - 1" + + +subsection \More on conversions\ + +lemma int_word_sint: + \sint (word_of_int x :: 'a::len word) = (x + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - 2 ^ (LENGTH('a) - 1)\ + by transfer (simp flip: take_bit_eq_mod add: signed_take_bit_eq_take_bit_shift) + +lemma sint_sbintrunc': "sint (word_of_int bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) bin" + by simp + +lemma uint_sint: "uint w = take_bit LENGTH('a) (sint w)" + for w :: "'a::len word" + by transfer (simp add: take_bit_signed_take_bit) + +lemma bintr_uint: "LENGTH('a) \ n \ take_bit n (uint w) = uint w" + for w :: "'a::len word" + by transfer (simp add: min_def) + +lemma wi_bintr: + "LENGTH('a::len) \ n \ + word_of_int (take_bit n w) = (word_of_int w :: 'a word)" + by transfer simp + +lemma word_numeral_alt: "numeral b = word_of_int (numeral b)" + by (induct b, simp_all only: numeral.simps word_of_int_homs) + +declare word_numeral_alt [symmetric, code_abbrev] + +lemma word_neg_numeral_alt: "- numeral b = word_of_int (- numeral b)" + by (simp only: word_numeral_alt wi_hom_neg) + +declare word_neg_numeral_alt [symmetric, code_abbrev] + +lemma uint_bintrunc [simp]: + "uint (numeral bin :: 'a word) = + take_bit (LENGTH('a::len)) (numeral bin)" + by transfer rule + +lemma uint_bintrunc_neg [simp]: + "uint (- numeral bin :: 'a word) = take_bit (LENGTH('a::len)) (- numeral bin)" + by transfer rule + +lemma sint_sbintrunc [simp]: + "sint (numeral bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) (numeral bin)" + by transfer simp + +lemma sint_sbintrunc_neg [simp]: + "sint (- numeral bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) (- numeral bin)" + by transfer simp + +lemma unat_bintrunc [simp]: + "unat (numeral bin :: 'a::len word) = nat (take_bit (LENGTH('a)) (numeral bin))" + by transfer simp + +lemma unat_bintrunc_neg [simp]: + "unat (- numeral bin :: 'a::len word) = nat (take_bit (LENGTH('a)) (- numeral bin))" + by transfer simp + +lemma size_0_eq: "size w = 0 \ v = w" + for v w :: "'a::len word" + by transfer simp + +lemma uint_ge_0 [iff]: "0 \ uint x" + by (fact unsigned_greater_eq) + +lemma uint_lt2p [iff]: "uint x < 2 ^ LENGTH('a)" + for x :: "'a::len word" + by (fact unsigned_less) + +lemma sint_ge: "- (2 ^ (LENGTH('a) - 1)) \ sint x" + for x :: "'a::len word" + using sint_greater_eq [of x] by simp + +lemma sint_lt: "sint x < 2 ^ (LENGTH('a) - 1)" + for x :: "'a::len word" + using sint_less [of x] by simp + +lemma uint_m2p_neg: "uint x - 2 ^ LENGTH('a) < 0" + for x :: "'a::len word" + by (simp only: diff_less_0_iff_less uint_lt2p) + +lemma uint_m2p_not_non_neg: "\ 0 \ uint x - 2 ^ LENGTH('a)" + for x :: "'a::len word" + by (simp only: not_le uint_m2p_neg) + +lemma lt2p_lem: "LENGTH('a) \ n \ uint w < 2 ^ n" + for w :: "'a::len word" + using uint_bounded [of w] by (rule less_le_trans) simp + +lemma uint_le_0_iff [simp]: "uint x \ 0 \ uint x = 0" + by (fact uint_ge_0 [THEN leD, THEN antisym_conv1]) + +lemma uint_nat: "uint w = int (unat w)" + by transfer simp + +lemma uint_numeral: "uint (numeral b :: 'a::len word) = numeral b mod 2 ^ LENGTH('a)" + by (simp flip: take_bit_eq_mod add: of_nat_take_bit) + +lemma uint_neg_numeral: "uint (- numeral b :: 'a::len word) = - numeral b mod 2 ^ LENGTH('a)" + by (simp flip: take_bit_eq_mod add: of_nat_take_bit) + +lemma unat_numeral: "unat (numeral b :: 'a::len word) = numeral b mod 2 ^ LENGTH('a)" + by transfer (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq) + +lemma sint_numeral: + "sint (numeral b :: 'a::len word) = + (numeral b + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - 2 ^ (LENGTH('a) - 1)" + apply (transfer fixing: b) + using int_word_sint [of \numeral b\] + apply simp + done + +lemma word_of_int_0 [simp, code_post]: "word_of_int 0 = 0" + by (fact of_int_0) + +lemma word_of_int_1 [simp, code_post]: "word_of_int 1 = 1" + by (fact of_int_1) + +lemma word_of_int_neg_1 [simp]: "word_of_int (- 1) = - 1" + by (simp add: wi_hom_syms) + +lemma word_of_int_numeral [simp] : "(word_of_int (numeral bin) :: 'a::len word) = numeral bin" + by (fact of_int_numeral) + +lemma word_of_int_neg_numeral [simp]: + "(word_of_int (- numeral bin) :: 'a::len word) = - numeral bin" + by (fact of_int_neg_numeral) + +lemma word_int_case_wi: + "word_int_case f (word_of_int i :: 'b word) = f (i mod 2 ^ LENGTH('b::len))" + by transfer (simp add: take_bit_eq_mod) + +lemma word_int_split: + "P (word_int_case f x) = + (\i. x = (word_of_int i :: 'b::len word) \ 0 \ i \ i < 2 ^ LENGTH('b) \ P (f i))" + by transfer (auto simp add: take_bit_eq_mod) + +lemma word_int_split_asm: + "P (word_int_case f x) = + (\n. x = (word_of_int n :: 'b::len word) \ 0 \ n \ n < 2 ^ LENGTH('b::len) \ \ P (f n))" + by transfer (auto simp add: take_bit_eq_mod) + +lemma uint_range_size: "0 \ uint w \ uint w < 2 ^ size w" + by transfer simp + +lemma sint_range_size: "- (2 ^ (size w - Suc 0)) \ sint w \ sint w < 2 ^ (size w - Suc 0)" + by (simp add: word_size sint_greater_eq sint_less) + +lemma sint_above_size: "2 ^ (size w - 1) \ x \ sint w < x" + for w :: "'a::len word" + unfolding word_size by (rule less_le_trans [OF sint_lt]) + +lemma sint_below_size: "x \ - (2 ^ (size w - 1)) \ x \ sint w" + for w :: "'a::len word" + unfolding word_size by (rule order_trans [OF _ sint_ge]) + + +subsection \Testing bits\ + +lemma bin_nth_uint_imp: "bit (uint w) n \ n < LENGTH('a)" + for w :: "'a::len word" + by transfer (simp add: bit_take_bit_iff) + +lemma bin_nth_sint: + "LENGTH('a) \ n \ + bit (sint w) n = bit (sint w) (LENGTH('a) - 1)" + for w :: "'a::len word" + by (transfer fixing: n) (simp add: bit_signed_take_bit_iff le_diff_conv min_def) + +lemma num_of_bintr': + "take_bit (LENGTH('a::len)) (numeral a :: int) = (numeral b) \ + numeral a = (numeral b :: 'a word)" +proof (transfer fixing: a b) + assume \take_bit LENGTH('a) (numeral a :: int) = numeral b\ + then have \take_bit LENGTH('a) (take_bit LENGTH('a) (numeral a :: int)) = take_bit LENGTH('a) (numeral b)\ + by simp + then show \take_bit LENGTH('a) (numeral a :: int) = take_bit LENGTH('a) (numeral b)\ + by simp +qed + +lemma num_of_sbintr': + "signed_take_bit (LENGTH('a::len) - 1) (numeral a :: int) = (numeral b) \ + numeral a = (numeral b :: 'a word)" +proof (transfer fixing: a b) + assume \signed_take_bit (LENGTH('a) - 1) (numeral a :: int) = numeral b\ + then have \take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - 1) (numeral a :: int)) = take_bit LENGTH('a) (numeral b)\ + by simp + then show \take_bit LENGTH('a) (numeral a :: int) = take_bit LENGTH('a) (numeral b)\ + by (simp add: take_bit_signed_take_bit) +qed + +lemma num_abs_bintr: + "(numeral x :: 'a word) = + word_of_int (take_bit (LENGTH('a::len)) (numeral x))" + by transfer simp + +lemma num_abs_sbintr: + "(numeral x :: 'a word) = + word_of_int (signed_take_bit (LENGTH('a::len) - 1) (numeral x))" + by transfer (simp add: take_bit_signed_take_bit) + +text \ + \cast\ -- note, no arg for new length, as it's determined by type of result, + thus in \cast w = w\, the type means cast to length of \w\! +\ + +lemma bit_ucast_iff: + \bit (ucast a :: 'a::len word) n \ n < LENGTH('a::len) \ Parity.bit a n\ + by transfer (simp add: bit_take_bit_iff) + +lemma ucast_id [simp]: "ucast w = w" + by transfer simp + +lemma scast_id [simp]: "scast w = w" + by transfer (simp add: take_bit_signed_take_bit) + +lemma ucast_mask_eq: + \ucast (mask n :: 'b word) = mask (min LENGTH('b::len) n)\ + by (simp add: bit_eq_iff) (auto simp add: bit_mask_iff bit_ucast_iff exp_eq_zero_iff) + +\ \literal u(s)cast\ +lemma ucast_bintr [simp]: + "ucast (numeral w :: 'a::len word) = + word_of_int (take_bit (LENGTH('a)) (numeral w))" + by transfer simp + +(* TODO: neg_numeral *) + +lemma scast_sbintr [simp]: + "scast (numeral w ::'a::len word) = + word_of_int (signed_take_bit (LENGTH('a) - Suc 0) (numeral w))" + by transfer simp + +lemma source_size: "source_size (c::'a::len word \ _) = LENGTH('a)" + by transfer simp + +lemma target_size: "target_size (c::_ \ 'b::len word) = LENGTH('b)" + by transfer simp + +lemma is_down: "is_down c \ LENGTH('b) \ LENGTH('a)" + for c :: "'a::len word \ 'b::len word" + by transfer simp + +lemma is_up: "is_up c \ LENGTH('a) \ LENGTH('b)" + for c :: "'a::len word \ 'b::len word" + by transfer simp + +lemma is_up_down: + \is_up c \ is_down d\ + for c :: \'a::len word \ 'b::len word\ + and d :: \'b::len word \ 'a::len word\ + by transfer simp + +context + fixes dummy_types :: \'a::len \ 'b::len\ +begin + +private abbreviation (input) UCAST :: \'a::len word \ 'b::len word\ + where \UCAST == ucast\ + +private abbreviation (input) SCAST :: \'a::len word \ 'b::len word\ + where \SCAST == scast\ + +lemma down_cast_same: + \UCAST = scast\ if \is_down UCAST\ + by (rule ext, use that in transfer) (simp add: take_bit_signed_take_bit) + +lemma sint_up_scast: + \sint (SCAST w) = sint w\ if \is_up SCAST\ + using that by transfer (simp add: min_def Suc_leI le_diff_iff) + +lemma uint_up_ucast: + \uint (UCAST w) = uint w\ if \is_up UCAST\ + using that by transfer (simp add: min_def) + +lemma ucast_up_ucast: + \ucast (UCAST w) = ucast w\ if \is_up UCAST\ + using that by transfer (simp add: ac_simps) + +lemma ucast_up_ucast_id: + \ucast (UCAST w) = w\ if \is_up UCAST\ + using that by (simp add: ucast_up_ucast) + +lemma scast_up_scast: + \scast (SCAST w) = scast w\ if \is_up SCAST\ + using that by transfer (simp add: ac_simps) + +lemma scast_up_scast_id: + \scast (SCAST w) = w\ if \is_up SCAST\ + using that by (simp add: scast_up_scast) + +lemma isduu: + \is_up UCAST\ if \is_down d\ + for d :: \'b word \ 'a word\ + using that is_up_down [of UCAST d] by simp + +lemma isdus: + \is_up SCAST\ if \is_down d\ + for d :: \'b word \ 'a word\ + using that is_up_down [of SCAST d] by simp + +lemmas ucast_down_ucast_id = isduu [THEN ucast_up_ucast_id] +lemmas scast_down_scast_id = isdus [THEN scast_up_scast_id] + +lemma up_ucast_surj: + \surj (ucast :: 'b word \ 'a word)\ if \is_up UCAST\ + by (rule surjI) (use that in \rule ucast_up_ucast_id\) + +lemma up_scast_surj: + \surj (scast :: 'b word \ 'a word)\ if \is_up SCAST\ + by (rule surjI) (use that in \rule scast_up_scast_id\) + +lemma down_ucast_inj: + \inj_on UCAST A\ if \is_down (ucast :: 'b word \ 'a word)\ + by (rule inj_on_inverseI) (use that in \rule ucast_down_ucast_id\) + +lemma down_scast_inj: + \inj_on SCAST A\ if \is_down (scast :: 'b word \ 'a word)\ + by (rule inj_on_inverseI) (use that in \rule scast_down_scast_id\) + +lemma ucast_down_wi: + \UCAST (word_of_int x) = word_of_int x\ if \is_down UCAST\ + using that by transfer simp + +lemma ucast_down_no: + \UCAST (numeral bin) = numeral bin\ if \is_down UCAST\ + using that by transfer simp + +end + +lemmas word_log_defs = word_and_def word_or_def word_xor_def word_not_def + +lemma bit_last_iff: + \bit w (LENGTH('a) - Suc 0) \ sint w < 0\ (is \?P \ ?Q\) + for w :: \'a::len word\ +proof - + have \?P \ bit (uint w) (LENGTH('a) - Suc 0)\ + by (simp add: bit_uint_iff) + also have \\ \ ?Q\ + by (simp add: sint_uint) + finally show ?thesis . +qed + +lemma drop_bit_eq_zero_iff_not_bit_last: + \drop_bit (LENGTH('a) - Suc 0) w = 0 \ \ bit w (LENGTH('a) - Suc 0)\ + for w :: "'a::len word" + apply (cases \LENGTH('a)\) + apply simp_all + apply (simp add: bit_iff_odd_drop_bit) + apply transfer + apply (simp add: take_bit_drop_bit) + apply (auto simp add: drop_bit_eq_div take_bit_eq_mod min_def) + apply (auto elim!: evenE) + apply (metis div_exp_eq mod_div_trivial mult.commute nonzero_mult_div_cancel_left power_Suc0_right power_add zero_neq_numeral) + done + + +subsection \Word Arithmetic\ + +lemmas word_div_no [simp] = word_div_def [of "numeral a" "numeral b"] for a b +lemmas word_mod_no [simp] = word_mod_def [of "numeral a" "numeral b"] for a b +lemmas word_less_no [simp] = word_less_def [of "numeral a" "numeral b"] for a b +lemmas word_le_no [simp] = word_le_def [of "numeral a" "numeral b"] for a b +lemmas word_sless_no [simp] = word_sless_eq [of "numeral a" "numeral b"] for a b +lemmas word_sle_no [simp] = word_sle_eq [of "numeral a" "numeral b"] for a b + +lemma size_0_same': "size w = 0 \ w = v" + for v w :: "'a::len word" + by (unfold word_size) simp + +lemmas size_0_same = size_0_same' [unfolded word_size] + +lemmas unat_eq_0 = unat_0_iff +lemmas unat_eq_zero = unat_0_iff + + +subsection \Transferring goals from words to ints\ + +lemma word_ths: + shows word_succ_p1: "word_succ a = a + 1" + and word_pred_m1: "word_pred a = a - 1" + and word_pred_succ: "word_pred (word_succ a) = a" + and word_succ_pred: "word_succ (word_pred a) = a" + and word_mult_succ: "word_succ a * b = b + a * b" + by (transfer, simp add: algebra_simps)+ + +lemma uint_cong: "x = y \ uint x = uint y" + by simp + +lemma uint_word_ariths: + fixes a b :: "'a::len word" + shows "uint (a + b) = (uint a + uint b) mod 2 ^ LENGTH('a::len)" + and "uint (a - b) = (uint a - uint b) mod 2 ^ LENGTH('a)" + and "uint (a * b) = uint a * uint b mod 2 ^ LENGTH('a)" + and "uint (- a) = - uint a mod 2 ^ LENGTH('a)" + and "uint (word_succ a) = (uint a + 1) mod 2 ^ LENGTH('a)" + and "uint (word_pred a) = (uint a - 1) mod 2 ^ LENGTH('a)" + and "uint (0 :: 'a word) = 0 mod 2 ^ LENGTH('a)" + and "uint (1 :: 'a word) = 1 mod 2 ^ LENGTH('a)" + by (simp_all only: word_arith_wis uint_word_of_int_eq flip: take_bit_eq_mod) + +lemma uint_word_arith_bintrs: + fixes a b :: "'a::len word" + shows "uint (a + b) = take_bit (LENGTH('a)) (uint a + uint b)" + and "uint (a - b) = take_bit (LENGTH('a)) (uint a - uint b)" + and "uint (a * b) = take_bit (LENGTH('a)) (uint a * uint b)" + and "uint (- a) = take_bit (LENGTH('a)) (- uint a)" + and "uint (word_succ a) = take_bit (LENGTH('a)) (uint a + 1)" + and "uint (word_pred a) = take_bit (LENGTH('a)) (uint a - 1)" + and "uint (0 :: 'a word) = take_bit (LENGTH('a)) 0" + and "uint (1 :: 'a word) = take_bit (LENGTH('a)) 1" + by (simp_all add: uint_word_ariths take_bit_eq_mod) + +lemma sint_word_ariths: + fixes a b :: "'a::len word" + shows "sint (a + b) = signed_take_bit (LENGTH('a) - 1) (sint a + sint b)" + and "sint (a - b) = signed_take_bit (LENGTH('a) - 1) (sint a - sint b)" + and "sint (a * b) = signed_take_bit (LENGTH('a) - 1) (sint a * sint b)" + and "sint (- a) = signed_take_bit (LENGTH('a) - 1) (- sint a)" + and "sint (word_succ a) = signed_take_bit (LENGTH('a) - 1) (sint a + 1)" + and "sint (word_pred a) = signed_take_bit (LENGTH('a) - 1) (sint a - 1)" + and "sint (0 :: 'a word) = signed_take_bit (LENGTH('a) - 1) 0" + and "sint (1 :: 'a word) = signed_take_bit (LENGTH('a) - 1) 1" + apply transfer apply (simp add: signed_take_bit_add) + apply transfer apply (simp add: signed_take_bit_diff) + apply transfer apply (simp add: signed_take_bit_mult) + apply transfer apply (simp add: signed_take_bit_minus) + apply (metis of_int_sint scast_id sint_sbintrunc' wi_hom_succ) + apply (metis of_int_sint scast_id sint_sbintrunc' wi_hom_pred) + apply (simp_all add: sint_uint) + done + +lemma word_pred_0_n1: "word_pred 0 = word_of_int (- 1)" + unfolding word_pred_m1 by simp + +lemma succ_pred_no [simp]: + "word_succ (numeral w) = numeral w + 1" + "word_pred (numeral w) = numeral w - 1" + "word_succ (- numeral w) = - numeral w + 1" + "word_pred (- numeral w) = - numeral w - 1" + by (simp_all add: word_succ_p1 word_pred_m1) + +lemma word_sp_01 [simp]: + "word_succ (- 1) = 0 \ word_succ 0 = 1 \ word_pred 0 = - 1 \ word_pred 1 = 0" + by (simp_all add: word_succ_p1 word_pred_m1) + +\ \alternative approach to lifting arithmetic equalities\ +lemma word_of_int_Ex: "\y. x = word_of_int y" + by (rule_tac x="uint x" in exI) simp + + +subsection \Order on fixed-length words\ + +lift_definition udvd :: \'a::len word \ 'a::len word \ bool\ (infixl \udvd\ 50) + is \\k l. take_bit LENGTH('a) k dvd take_bit LENGTH('a) l\ by simp + +lemma udvd_iff_dvd: + \x udvd y \ unat x dvd unat y\ + by transfer (simp add: nat_dvd_iff) + +lemma udvd_iff_dvd_int: + \v udvd w \ uint v dvd uint w\ + by transfer rule + +lemma udvdI [intro]: + \v udvd w\ if \unat w = unat v * unat u\ +proof - + from that have \unat v dvd unat w\ .. + then show ?thesis + by (simp add: udvd_iff_dvd) +qed + +lemma udvdE [elim]: + fixes v w :: \'a::len word\ + assumes \v udvd w\ + obtains u :: \'a word\ where \unat w = unat v * unat u\ +proof (cases \v = 0\) + case True + moreover from True \v udvd w\ have \w = 0\ + by transfer simp + ultimately show thesis + using that by simp +next + case False + then have \unat v > 0\ + by (simp add: unat_gt_0) + from \v udvd w\ have \unat v dvd unat w\ + by (simp add: udvd_iff_dvd) + then obtain n where \unat w = unat v * n\ .. + moreover have \n < 2 ^ LENGTH('a)\ + proof (rule ccontr) + assume \\ n < 2 ^ LENGTH('a)\ + then have \n \ 2 ^ LENGTH('a)\ + by (simp add: not_le) + then have \unat v * n \ 2 ^ LENGTH('a)\ + using \unat v > 0\ mult_le_mono [of 1 \unat v\ \2 ^ LENGTH('a)\ n] + by simp + with \unat w = unat v * n\ + have \unat w \ 2 ^ LENGTH('a)\ + by simp + with unsigned_less [of w, where ?'a = nat] show False + by linarith + qed + ultimately have \unat w = unat v * unat (word_of_nat n :: 'a word)\ + by (auto simp add: take_bit_nat_eq_self_iff intro: sym) + with that show thesis . +qed + +lemma udvd_imp_mod_eq_0: + \w mod v = 0\ if \v udvd w\ + using that by transfer simp + +lemma mod_eq_0_imp_udvd [intro?]: + \v udvd w\ if \w mod v = 0\ +proof - + from that have \unat (w mod v) = unat 0\ + by simp + then have \unat w mod unat v = 0\ + by (simp add: unat_mod_distrib) + then have \unat v dvd unat w\ .. + then show ?thesis + by (simp add: udvd_iff_dvd) +qed + +lemma udvd_imp_dvd: + \v dvd w\ if \v udvd w\ for v w :: \'a::len word\ +proof - + from that obtain u :: \'a word\ where \unat w = unat v * unat u\ .. + then have \(word_of_nat (unat w) :: 'a word) = word_of_nat (unat v * unat u)\ + by simp + then have \w = v * u\ + by simp + then show \v dvd w\ .. +qed + +lemma exp_dvd_iff_exp_udvd: + \2 ^ n dvd w \ 2 ^ n udvd w\ for v w :: \'a::len word\ +proof + assume \2 ^ n udvd w\ then show \2 ^ n dvd w\ + by (rule udvd_imp_dvd) +next + assume \2 ^ n dvd w\ + then obtain u :: \'a word\ where \w = 2 ^ n * u\ .. + then have \w = push_bit n u\ + by (simp add: push_bit_eq_mult) + then show \2 ^ n udvd w\ + by transfer (simp add: take_bit_push_bit dvd_eq_mod_eq_0 flip: take_bit_eq_mod) +qed + +lemma udvd_nat_alt: + \a udvd b \ (\n. unat b = n * unat a)\ + by (auto simp add: udvd_iff_dvd) + +lemma udvd_unfold_int: + \a udvd b \ (\n\0. uint b = n * uint a)\ + apply (auto elim!: dvdE simp add: udvd_iff_dvd) + apply (simp only: uint_nat) + apply auto + using of_nat_0_le_iff apply blast + apply (simp only: unat_eq_nat_uint) + apply (simp add: nat_mult_distrib) + done + +lemma unat_minus_one: + \unat (w - 1) = unat w - 1\ if \w \ 0\ +proof - + have "0 \ uint w" by (fact uint_nonnegative) + moreover from that have "0 \ uint w" + by (simp add: uint_0_iff) + ultimately have "1 \ uint w" + by arith + from uint_lt2p [of w] have "uint w - 1 < 2 ^ LENGTH('a)" + by arith + with \1 \ uint w\ have "(uint w - 1) mod 2 ^ LENGTH('a) = uint w - 1" + by (auto intro: mod_pos_pos_trivial) + with \1 \ uint w\ have "nat ((uint w - 1) mod 2 ^ LENGTH('a)) = nat (uint w) - 1" + by (auto simp del: nat_uint_eq) + then show ?thesis + by (simp only: unat_eq_nat_uint word_arith_wis mod_diff_right_eq) + (metis of_int_1 uint_word_of_int unsigned_1) +qed + +lemma measure_unat: "p \ 0 \ unat (p - 1) < unat p" + by (simp add: unat_minus_one) (simp add: unat_0_iff [symmetric]) + +lemmas uint_add_ge0 [simp] = add_nonneg_nonneg [OF uint_ge_0 uint_ge_0] +lemmas uint_mult_ge0 [simp] = mult_nonneg_nonneg [OF uint_ge_0 uint_ge_0] + +lemma uint_sub_lt2p [simp]: "uint x - uint y < 2 ^ LENGTH('a)" + for x :: "'a::len word" and y :: "'b::len word" + using uint_ge_0 [of y] uint_lt2p [of x] by arith + + +subsection \Conditions for the addition (etc) of two words to overflow\ + +lemma uint_add_lem: + "(uint x + uint y < 2 ^ LENGTH('a)) = + (uint (x + y) = uint x + uint y)" + for x y :: "'a::len word" + by (metis add.right_neutral add_mono_thms_linordered_semiring(1) mod_pos_pos_trivial of_nat_0_le_iff uint_lt2p uint_nat uint_word_ariths(1)) + +lemma uint_mult_lem: + "(uint x * uint y < 2 ^ LENGTH('a)) = + (uint (x * y) = uint x * uint y)" + for x y :: "'a::len word" + by (metis mod_pos_pos_trivial uint_lt2p uint_mult_ge0 uint_word_ariths(3)) + +lemma uint_sub_lem: "uint x \ uint y \ uint (x - y) = uint x - uint y" + by (metis diff_ge_0_iff_ge of_nat_0_le_iff uint_nat uint_sub_lt2p uint_word_of_int unique_euclidean_semiring_numeral_class.mod_less word_sub_wi) + +lemma uint_add_le: "uint (x + y) \ uint x + uint y" + unfolding uint_word_ariths by (simp add: zmod_le_nonneg_dividend) + +lemma uint_sub_ge: "uint (x - y) \ uint x - uint y" + unfolding uint_word_ariths + by (simp flip: take_bit_eq_mod add: take_bit_int_greater_eq_self_iff) + +lemma int_mod_ge: \a \ a mod n\ if \a < n\ \0 < n\ + for a n :: int +proof (cases \a < 0\) + case True + with \0 < n\ show ?thesis + by (metis less_trans not_less pos_mod_conj) + +next + case False + with \a < n\ show ?thesis + by simp +qed + +lemma mod_add_if_z: + "x < z \ y < z \ 0 \ y \ 0 \ x \ 0 \ z \ + (x + y) mod z = (if x + y < z then x + y else x + y - z)" + for x y z :: int + apply (auto simp add: not_less) + apply (rule antisym) + apply (metis diff_ge_0_iff_ge minus_mod_self2 zmod_le_nonneg_dividend) + apply (simp only: flip: minus_mod_self2 [of \x + y\ z]) + apply (metis add.commute add_less_cancel_left add_mono_thms_linordered_field(5) diff_add_cancel diff_ge_0_iff_ge mod_pos_pos_trivial order_refl) + done + +lemma uint_plus_if': + "uint (a + b) = + (if uint a + uint b < 2 ^ LENGTH('a) then uint a + uint b + else uint a + uint b - 2 ^ LENGTH('a))" + for a b :: "'a::len word" + using mod_add_if_z [of "uint a" _ "uint b"] by (simp add: uint_word_ariths) + +lemma mod_sub_if_z: + "x < z \ y < z \ 0 \ y \ 0 \ x \ 0 \ z \ + (x - y) mod z = (if y \ x then x - y else x - y + z)" + for x y z :: int + apply (auto simp add: not_le) + apply (rule antisym) + apply (simp only: flip: mod_add_self2 [of \x - y\ z]) + apply (rule zmod_le_nonneg_dividend) + apply simp + apply (metis add.commute add.right_neutral add_le_cancel_left diff_ge_0_iff_ge int_mod_ge le_less le_less_trans mod_add_self1 not_less) + done + +lemma uint_sub_if': + "uint (a - b) = + (if uint b \ uint a then uint a - uint b + else uint a - uint b + 2 ^ LENGTH('a))" + for a b :: "'a::len word" + using mod_sub_if_z [of "uint a" _ "uint b"] by (simp add: uint_word_ariths) + + +subsection \Definition of \uint_arith\\ + +lemma word_of_int_inverse: + "word_of_int r = a \ 0 \ r \ r < 2 ^ LENGTH('a) \ uint a = r" + for a :: "'a::len word" + apply transfer + apply (drule sym) + apply (simp add: take_bit_int_eq_self) + done + +lemma uint_split: + "P (uint x) = (\i. word_of_int i = x \ 0 \ i \ i < 2^LENGTH('a) \ P i)" + for x :: "'a::len word" + by transfer (auto simp add: take_bit_eq_mod) + +lemma uint_split_asm: + "P (uint x) = (\i. word_of_int i = x \ 0 \ i \ i < 2^LENGTH('a) \ \ P i)" + for x :: "'a::len word" + by auto (metis take_bit_int_eq_self_iff) + +lemmas uint_splits = uint_split uint_split_asm + +lemmas uint_arith_simps = + word_le_def word_less_alt + word_uint_eq_iff + uint_sub_if' uint_plus_if' + +\ \use this to stop, eg. \2 ^ LENGTH(32)\ being simplified\ +lemma power_False_cong: "False \ a ^ b = c ^ d" + by auto + +\ \\uint_arith_tac\: reduce to arithmetic on int, try to solve by arith\ +ML \ +val uint_arith_simpset = + @{context} + |> fold Simplifier.add_simp @{thms uint_arith_simps} + |> fold Splitter.add_split @{thms if_split_asm} + |> fold Simplifier.add_cong @{thms power_False_cong} + |> simpset_of; + +fun uint_arith_tacs ctxt = + let + fun arith_tac' n t = + Arith_Data.arith_tac ctxt n t + handle Cooper.COOPER _ => Seq.empty; + in + [ clarify_tac ctxt 1, + full_simp_tac (put_simpset uint_arith_simpset ctxt) 1, + ALLGOALS (full_simp_tac + (put_simpset HOL_ss ctxt + |> fold Splitter.add_split @{thms uint_splits} + |> fold Simplifier.add_cong @{thms power_False_cong})), + rewrite_goals_tac ctxt @{thms word_size}, + ALLGOALS (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN + REPEAT (eresolve_tac ctxt [conjE] n) THEN + REPEAT (dresolve_tac ctxt @{thms word_of_int_inverse} n + THEN assume_tac ctxt n + THEN assume_tac ctxt n)), + TRYALL arith_tac' ] + end + +fun uint_arith_tac ctxt = SELECT_GOAL (EVERY (uint_arith_tacs ctxt)) +\ + +method_setup uint_arith = + \Scan.succeed (SIMPLE_METHOD' o uint_arith_tac)\ + "solving word arithmetic via integers and arith" + + +subsection \More on overflows and monotonicity\ + +lemma no_plus_overflow_uint_size: "x \ x + y \ uint x + uint y < 2 ^ size x" + for x y :: "'a::len word" + unfolding word_size by uint_arith + +lemmas no_olen_add = no_plus_overflow_uint_size [unfolded word_size] + +lemma no_ulen_sub: "x \ x - y \ uint y \ uint x" + for x y :: "'a::len word" + by uint_arith + +lemma no_olen_add': "x \ y + x \ uint y + uint x < 2 ^ LENGTH('a)" + for x y :: "'a::len word" + by (simp add: ac_simps no_olen_add) + +lemmas olen_add_eqv = trans [OF no_olen_add no_olen_add' [symmetric]] + +lemmas uint_plus_simple_iff = trans [OF no_olen_add uint_add_lem] +lemmas uint_plus_simple = uint_plus_simple_iff [THEN iffD1] +lemmas uint_minus_simple_iff = trans [OF no_ulen_sub uint_sub_lem] +lemmas uint_minus_simple_alt = uint_sub_lem [folded word_le_def] +lemmas word_sub_le_iff = no_ulen_sub [folded word_le_def] +lemmas word_sub_le = word_sub_le_iff [THEN iffD2] + +lemma word_less_sub1: "x \ 0 \ 1 < x \ 0 < x - 1" + for x :: "'a::len word" + by uint_arith + +lemma word_le_sub1: "x \ 0 \ 1 \ x \ 0 \ x - 1" + for x :: "'a::len word" + by uint_arith + +lemma sub_wrap_lt: "x < x - z \ x < z" + for x z :: "'a::len word" + by uint_arith + +lemma sub_wrap: "x \ x - z \ z = 0 \ x < z" + for x z :: "'a::len word" + by uint_arith + +lemma plus_minus_not_NULL_ab: "x \ ab - c \ c \ ab \ c \ 0 \ x + c \ 0" + for x ab c :: "'a::len word" + by uint_arith + +lemma plus_minus_no_overflow_ab: "x \ ab - c \ c \ ab \ x \ x + c" + for x ab c :: "'a::len word" + by uint_arith + +lemma le_minus': "a + c \ b \ a \ a + c \ c \ b - a" + for a b c :: "'a::len word" + by uint_arith + +lemma le_plus': "a \ b \ c \ b - a \ a + c \ b" + for a b c :: "'a::len word" + by uint_arith + +lemmas le_plus = le_plus' [rotated] + +lemmas le_minus = leD [THEN thin_rl, THEN le_minus'] (* FIXME *) + +lemma word_plus_mono_right: "y \ z \ x \ x + z \ x + y \ x + z" + for x y z :: "'a::len word" + by uint_arith + +lemma word_less_minus_cancel: "y - x < z - x \ x \ z \ y < z" + for x y z :: "'a::len word" + by uint_arith + +lemma word_less_minus_mono_left: "y < z \ x \ y \ y - x < z - x" + for x y z :: "'a::len word" + by uint_arith + +lemma word_less_minus_mono: "a < c \ d < b \ a - b < a \ c - d < c \ a - b < c - d" + for a b c d :: "'a::len word" + by uint_arith + +lemma word_le_minus_cancel: "y - x \ z - x \ x \ z \ y \ z" + for x y z :: "'a::len word" + by uint_arith + +lemma word_le_minus_mono_left: "y \ z \ x \ y \ y - x \ z - x" + for x y z :: "'a::len word" + by uint_arith + +lemma word_le_minus_mono: + "a \ c \ d \ b \ a - b \ a \ c - d \ c \ a - b \ c - d" + for a b c d :: "'a::len word" + by uint_arith + +lemma plus_le_left_cancel_wrap: "x + y' < x \ x + y < x \ x + y' < x + y \ y' < y" + for x y y' :: "'a::len word" + by uint_arith + +lemma plus_le_left_cancel_nowrap: "x \ x + y' \ x \ x + y \ x + y' < x + y \ y' < y" + for x y y' :: "'a::len word" + by uint_arith + +lemma word_plus_mono_right2: "a \ a + b \ c \ b \ a \ a + c" + for a b c :: "'a::len word" + by uint_arith + +lemma word_less_add_right: "x < y - z \ z \ y \ x + z < y" + for x y z :: "'a::len word" + by uint_arith + +lemma word_less_sub_right: "x < y + z \ y \ x \ x - y < z" + for x y z :: "'a::len word" + by uint_arith + +lemma word_le_plus_either: "x \ y \ x \ z \ y \ y + z \ x \ y + z" + for x y z :: "'a::len word" + by uint_arith + +lemma word_less_nowrapI: "x < z - k \ k \ z \ 0 < k \ x < x + k" + for x z k :: "'a::len word" + by uint_arith + +lemma inc_le: "i < m \ i + 1 \ m" + for i m :: "'a::len word" + by uint_arith + +lemma inc_i: "1 \ i \ i < m \ 1 \ i + 1 \ i + 1 \ m" + for i m :: "'a::len word" + by uint_arith + +lemma udvd_incr_lem: + "up < uq \ up = ua + n * uint K \ + uq = ua + n' * uint K \ up + uint K \ uq" + by auto (metis int_distrib(1) linorder_not_less mult.left_neutral mult_right_mono uint_nonnegative zless_imp_add1_zle) + +lemma udvd_incr': + "p < q \ uint p = ua + n * uint K \ + uint q = ua + n' * uint K \ p + K \ q" + apply (unfold word_less_alt word_le_def) + apply (drule (2) udvd_incr_lem) + apply (erule uint_add_le [THEN order_trans]) + done + +lemma udvd_decr': + "p < q \ uint p = ua + n * uint K \ + uint q = ua + n' * uint K \ p \ q - K" + apply (unfold word_less_alt word_le_def) + apply (drule (2) udvd_incr_lem) + apply (drule le_diff_eq [THEN iffD2]) + apply (erule order_trans) + apply (rule uint_sub_ge) + done + +lemmas udvd_incr_lem0 = udvd_incr_lem [where ua=0, unfolded add_0_left] +lemmas udvd_incr0 = udvd_incr' [where ua=0, unfolded add_0_left] +lemmas udvd_decr0 = udvd_decr' [where ua=0, unfolded add_0_left] + +lemma udvd_minus_le': "xy < k \ z udvd xy \ z udvd k \ xy \ k - z" + apply (unfold udvd_unfold_int) + apply clarify + apply (erule (2) udvd_decr0) + done + +lemma udvd_incr2_K: + "p < a + s \ a \ a + s \ K udvd s \ K udvd p - a \ a \ p \ + 0 < K \ p \ p + K \ p + K \ a + s" + supply [[simproc del: linordered_ring_less_cancel_factor]] + apply (unfold udvd_unfold_int) + apply clarify + apply (simp add: uint_arith_simps split: if_split_asm) + prefer 2 + using uint_lt2p [of s] apply simp + apply (drule add.commute [THEN xtrans(1)]) + apply (simp flip: diff_less_eq) + apply (subst (asm) mult_less_cancel_right) + apply simp + apply (simp add: diff_eq_eq not_less) + apply (subst (asm) (3) zless_iff_Suc_zadd) + apply auto + apply (auto simp add: algebra_simps) + apply (drule less_le_trans [of _ \2 ^ LENGTH('a)\]) apply assumption + apply (simp add: mult_less_0_iff) + done + + +subsection \Arithmetic type class instantiations\ + +lemmas word_le_0_iff [simp] = + word_zero_le [THEN leD, THEN antisym_conv1] + +lemma word_of_int_nat: "0 \ x \ word_of_int x = of_nat (nat x)" + by simp + +text \ + note that \iszero_def\ is only for class \comm_semiring_1_cancel\, + which requires word length \\ 1\, ie \'a::len word\ +\ +lemma iszero_word_no [simp]: + "iszero (numeral bin :: 'a::len word) = + iszero (take_bit LENGTH('a) (numeral bin :: int))" + apply (simp add: iszero_def) + apply transfer + apply simp + done + +text \Use \iszero\ to simplify equalities between word numerals.\ + +lemmas word_eq_numeral_iff_iszero [simp] = + eq_numeral_iff_iszero [where 'a="'a::len word"] + + +subsection \Word and nat\ + +lemma word_nchotomy: "\w :: 'a::len word. \n. w = of_nat n \ n < 2 ^ LENGTH('a)" + apply (rule allI) + apply (rule exI [of _ \unat w\ for w :: \'a word\]) + apply simp + done + +lemma of_nat_eq: "of_nat n = w \ (\q. n = unat w + q * 2 ^ LENGTH('a))" + for w :: "'a::len word" + using mod_div_mult_eq [of n "2 ^ LENGTH('a)", symmetric] + by (auto simp flip: take_bit_eq_mod) + +lemma of_nat_eq_size: "of_nat n = w \ (\q. n = unat w + q * 2 ^ size w)" + unfolding word_size by (rule of_nat_eq) + +lemma of_nat_0: "of_nat m = (0::'a::len word) \ (\q. m = q * 2 ^ LENGTH('a))" + by (simp add: of_nat_eq) + +lemma of_nat_2p [simp]: "of_nat (2 ^ LENGTH('a)) = (0::'a::len word)" + by (fact mult_1 [symmetric, THEN iffD2 [OF of_nat_0 exI]]) + +lemma of_nat_gt_0: "of_nat k \ 0 \ 0 < k" + by (cases k) auto + +lemma of_nat_neq_0: "0 < k \ k < 2 ^ LENGTH('a::len) \ of_nat k \ (0 :: 'a word)" + by (auto simp add : of_nat_0) + +lemma Abs_fnat_hom_add: "of_nat a + of_nat b = of_nat (a + b)" + by simp + +lemma Abs_fnat_hom_mult: "of_nat a * of_nat b = (of_nat (a * b) :: 'a::len word)" + by (simp add: wi_hom_mult) + +lemma Abs_fnat_hom_Suc: "word_succ (of_nat a) = of_nat (Suc a)" + by transfer (simp add: ac_simps) + +lemma Abs_fnat_hom_0: "(0::'a::len word) = of_nat 0" + by simp + +lemma Abs_fnat_hom_1: "(1::'a::len word) = of_nat (Suc 0)" + by simp + +lemmas Abs_fnat_homs = + Abs_fnat_hom_add Abs_fnat_hom_mult Abs_fnat_hom_Suc + Abs_fnat_hom_0 Abs_fnat_hom_1 + +lemma word_arith_nat_add: "a + b = of_nat (unat a + unat b)" + by simp + +lemma word_arith_nat_mult: "a * b = of_nat (unat a * unat b)" + by simp + +lemma word_arith_nat_Suc: "word_succ a = of_nat (Suc (unat a))" + by (subst Abs_fnat_hom_Suc [symmetric]) simp + +lemma word_arith_nat_div: "a div b = of_nat (unat a div unat b)" + by (metis of_int_of_nat_eq of_nat_unat of_nat_div word_div_def) + +lemma word_arith_nat_mod: "a mod b = of_nat (unat a mod unat b)" + by (metis of_int_of_nat_eq of_nat_mod of_nat_unat word_mod_def) + +lemmas word_arith_nat_defs = + word_arith_nat_add word_arith_nat_mult + word_arith_nat_Suc Abs_fnat_hom_0 + Abs_fnat_hom_1 word_arith_nat_div + word_arith_nat_mod + +lemma unat_cong: "x = y \ unat x = unat y" + by (fact arg_cong) + +lemma unat_of_nat: + \unat (word_of_nat x :: 'a::len word) = x mod 2 ^ LENGTH('a)\ + by transfer (simp flip: take_bit_eq_mod add: nat_take_bit_eq) + +lemmas unat_word_ariths = word_arith_nat_defs + [THEN trans [OF unat_cong unat_of_nat]] + +lemmas word_sub_less_iff = word_sub_le_iff + [unfolded linorder_not_less [symmetric] Not_eq_iff] + +lemma unat_add_lem: + "unat x + unat y < 2 ^ LENGTH('a) \ unat (x + y) = unat x + unat y" + for x y :: "'a::len word" + apply (auto simp: unat_word_ariths) + apply (drule sym) + apply (metis unat_of_nat unsigned_less) + done + +lemma unat_mult_lem: + "unat x * unat y < 2 ^ LENGTH('a) \ unat (x * y) = unat x * unat y" + for x y :: "'a::len word" + apply (auto simp: unat_word_ariths) + apply (drule sym) + apply (metis unat_of_nat unsigned_less) + done + +lemma unat_plus_if': + \unat (a + b) = + (if unat a + unat b < 2 ^ LENGTH('a) + then unat a + unat b + else unat a + unat b - 2 ^ LENGTH('a))\ for a b :: \'a::len word\ + apply (auto simp: unat_word_ariths not_less) + apply (subst (asm) le_iff_add) + apply auto + apply (simp flip: take_bit_eq_mod add: take_bit_nat_eq_self_iff) + apply (metis add.commute add_less_cancel_right le_less_trans less_imp_le unsigned_less) + done + +lemma le_no_overflow: "x \ b \ a \ a + b \ x \ a + b" + for a b x :: "'a::len word" + apply (erule order_trans) + apply (erule olen_add_eqv [THEN iffD1]) + done + +lemmas un_ui_le = + trans [OF word_le_nat_alt [symmetric] word_le_def] + +lemma unat_sub_if_size: + "unat (x - y) = + (if unat y \ unat x + then unat x - unat y + else unat x + 2 ^ size x - unat y)" + supply nat_uint_eq [simp del] + apply (unfold word_size) + apply (simp add: un_ui_le) + apply (auto simp add: unat_eq_nat_uint uint_sub_if') + apply (rule nat_diff_distrib) + prefer 3 + apply (simp add: algebra_simps) + apply (rule nat_diff_distrib [THEN trans]) + prefer 3 + apply (subst nat_add_distrib) + prefer 3 + apply (simp add: nat_power_eq) + apply auto + apply uint_arith + done + +lemmas unat_sub_if' = unat_sub_if_size [unfolded word_size] + +lemma uint_div: + \uint (x div y) = uint x div uint y\ + by (fact uint_div_distrib) + +lemma unat_div: + \unat (x div y) = unat x div unat y\ + by (fact unat_div_distrib) + +lemma uint_mod: + \uint (x mod y) = uint x mod uint y\ + by (fact uint_mod_distrib) + +lemma unat_mod: + \unat (x mod y) = unat x mod unat y\ + by (fact unat_mod_distrib) + + +text \Definition of \unat_arith\ tactic\ + +lemma unat_split: "P (unat x) \ (\n. of_nat n = x \ n < 2^LENGTH('a) \ P n)" + for x :: "'a::len word" + by auto (metis take_bit_nat_eq_self_iff) + +lemma unat_split_asm: "P (unat x) \ (\n. of_nat n = x \ n < 2^LENGTH('a) \ \ P n)" + for x :: "'a::len word" + by auto (metis take_bit_nat_eq_self_iff) + +lemma of_nat_inverse: + \word_of_nat r = a \ r < 2 ^ LENGTH('a) \ unat a = r\ + for a :: \'a::len word\ + apply (drule sym) + apply transfer + apply (simp add: take_bit_int_eq_self) + done + +lemma word_unat_eq_iff: + \v = w \ unat v = unat w\ + for v w :: \'a::len word\ + by (fact word_eq_iff_unsigned) + +lemmas unat_splits = unat_split unat_split_asm + +lemmas unat_arith_simps = + word_le_nat_alt word_less_nat_alt + word_unat_eq_iff + unat_sub_if' unat_plus_if' unat_div unat_mod + +\ \\unat_arith_tac\: tactic to reduce word arithmetic to \nat\, try to solve via \arith\\ +ML \ +val unat_arith_simpset = + @{context} + |> fold Simplifier.add_simp @{thms unat_arith_simps} + |> fold Splitter.add_split @{thms if_split_asm} + |> fold Simplifier.add_cong @{thms power_False_cong} + |> simpset_of + +fun unat_arith_tacs ctxt = + let + fun arith_tac' n t = + Arith_Data.arith_tac ctxt n t + handle Cooper.COOPER _ => Seq.empty; + in + [ clarify_tac ctxt 1, + full_simp_tac (put_simpset unat_arith_simpset ctxt) 1, + ALLGOALS (full_simp_tac + (put_simpset HOL_ss ctxt + |> fold Splitter.add_split @{thms unat_splits} + |> fold Simplifier.add_cong @{thms power_False_cong})), + rewrite_goals_tac ctxt @{thms word_size}, + ALLGOALS (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN + REPEAT (eresolve_tac ctxt [conjE] n) THEN + REPEAT (dresolve_tac ctxt @{thms of_nat_inverse} n THEN assume_tac ctxt n)), + TRYALL arith_tac' ] + end + +fun unat_arith_tac ctxt = SELECT_GOAL (EVERY (unat_arith_tacs ctxt)) +\ + +method_setup unat_arith = + \Scan.succeed (SIMPLE_METHOD' o unat_arith_tac)\ + "solving word arithmetic via natural numbers and arith" + +lemma no_plus_overflow_unat_size: "x \ x + y \ unat x + unat y < 2 ^ size x" + for x y :: "'a::len word" + unfolding word_size by unat_arith + +lemmas no_olen_add_nat = + no_plus_overflow_unat_size [unfolded word_size] + +lemmas unat_plus_simple = + trans [OF no_olen_add_nat unat_add_lem] + +lemma word_div_mult: "0 < y \ unat x * unat y < 2 ^ LENGTH('a) \ x * y div y = x" + for x y :: "'a::len word" + apply unat_arith + apply clarsimp + apply (subst unat_mult_lem [THEN iffD1]) + apply auto + done + +lemma div_lt': "i \ k div x \ unat i * unat x < 2 ^ LENGTH('a)" + for i k x :: "'a::len word" + apply unat_arith + apply clarsimp + apply (drule mult_le_mono1) + apply (erule order_le_less_trans) + apply (metis add_lessD1 div_mult_mod_eq unsigned_less) + done + +lemmas div_lt'' = order_less_imp_le [THEN div_lt'] + +lemma div_lt_mult: "i < k div x \ 0 < x \ i * x < k" + for i k x :: "'a::len word" + apply (frule div_lt'' [THEN unat_mult_lem [THEN iffD1]]) + apply (simp add: unat_arith_simps) + apply (drule (1) mult_less_mono1) + apply (erule order_less_le_trans) + apply auto + done + +lemma div_le_mult: "i \ k div x \ 0 < x \ i * x \ k" + for i k x :: "'a::len word" + apply (frule div_lt' [THEN unat_mult_lem [THEN iffD1]]) + apply (simp add: unat_arith_simps) + apply (drule mult_le_mono1) + apply (erule order_trans) + apply auto + done + +lemma div_lt_uint': "i \ k div x \ uint i * uint x < 2 ^ LENGTH('a)" + for i k x :: "'a::len word" + apply (unfold uint_nat) + apply (drule div_lt') + apply (metis of_nat_less_iff of_nat_mult of_nat_numeral of_nat_power) + done + +lemmas div_lt_uint'' = order_less_imp_le [THEN div_lt_uint'] + +lemma word_le_exists': "x \ y \ \z. y = x + z \ uint x + uint z < 2 ^ LENGTH('a)" + for x y z :: "'a::len word" + by (metis add_diff_cancel_left' add_diff_eq uint_add_lem uint_plus_simple) + +lemmas plus_minus_not_NULL = order_less_imp_le [THEN plus_minus_not_NULL_ab] + +lemmas plus_minus_no_overflow = + order_less_imp_le [THEN plus_minus_no_overflow_ab] + +lemmas mcs = word_less_minus_cancel word_less_minus_mono_left + word_le_minus_cancel word_le_minus_mono_left + +lemmas word_l_diffs = mcs [where y = "w + x", unfolded add_diff_cancel] for w x +lemmas word_diff_ls = mcs [where z = "w + x", unfolded add_diff_cancel] for w x +lemmas word_plus_mcs = word_diff_ls [where y = "v + x", unfolded add_diff_cancel] for v x + +lemma le_unat_uoi: + \y \ unat z \ unat (word_of_nat y :: 'a word) = y\ + for z :: \'a::len word\ + by transfer (simp add: nat_take_bit_eq take_bit_nat_eq_self_iff le_less_trans) + +lemmas thd = times_div_less_eq_dividend + +lemmas uno_simps [THEN le_unat_uoi] = mod_le_divisor div_le_dividend + +lemma word_mod_div_equality: "(n div b) * b + (n mod b) = n" + for n b :: "'a::len word" + by (fact div_mult_mod_eq) + +lemma word_div_mult_le: "a div b * b \ a" + for a b :: "'a::len word" + by (metis div_le_mult mult_not_zero order.not_eq_order_implies_strict order_refl word_zero_le) + +lemma word_mod_less_divisor: "0 < n \ m mod n < n" + for m n :: "'a::len word" + by (simp add: unat_arith_simps) + +lemma word_of_int_power_hom: "word_of_int a ^ n = (word_of_int (a ^ n) :: 'a::len word)" + by (induct n) (simp_all add: wi_hom_mult [symmetric]) + +lemma word_arith_power_alt: "a ^ n = (word_of_int (uint a ^ n) :: 'a::len word)" + by (simp add : word_of_int_power_hom [symmetric]) + +lemma unatSuc: "1 + n \ 0 \ unat (1 + n) = Suc (unat n)" + for n :: "'a::len word" + by unat_arith + + +subsection \Cardinality, finiteness of set of words\ + +lemma inj_on_word_of_int: \inj_on (word_of_int :: int \ 'a word) {0..<2 ^ LENGTH('a::len)}\ + apply (rule inj_onI) + apply transfer + apply (simp add: take_bit_eq_mod) + done + +lemma inj_uint: \inj uint\ + by (fact inj_unsigned) + +lemma range_uint: \range (uint :: 'a word \ int) = {0..<2 ^ LENGTH('a::len)}\ + apply transfer + apply (auto simp add: image_iff) + apply (metis take_bit_int_eq_self_iff) + done + +lemma UNIV_eq: \(UNIV :: 'a word set) = word_of_int ` {0..<2 ^ LENGTH('a::len)}\ + by (auto simp add: image_iff) (metis atLeastLessThan_iff linorder_not_le uint_split) + +lemma card_word: "CARD('a word) = 2 ^ LENGTH('a::len)" + by (simp add: UNIV_eq card_image inj_on_word_of_int) + +lemma card_word_size: "CARD('a word) = 2 ^ size x" + for x :: "'a::len word" + unfolding word_size by (rule card_word) + +instance word :: (len) finite + by standard (simp add: UNIV_eq) + + +subsection \Bitwise Operations on Words\ + +lemma word_wi_log_defs: + "NOT (word_of_int a) = word_of_int (NOT a)" + "word_of_int a AND word_of_int b = word_of_int (a AND b)" + "word_of_int a OR word_of_int b = word_of_int (a OR b)" + "word_of_int a XOR word_of_int b = word_of_int (a XOR b)" + by (transfer, rule refl)+ + +lemma word_no_log_defs [simp]: + "NOT (numeral a) = word_of_int (NOT (numeral a))" + "NOT (- numeral a) = word_of_int (NOT (- numeral a))" + "numeral a AND numeral b = word_of_int (numeral a AND numeral b)" + "numeral a AND - numeral b = word_of_int (numeral a AND - numeral b)" + "- numeral a AND numeral b = word_of_int (- numeral a AND numeral b)" + "- numeral a AND - numeral b = word_of_int (- numeral a AND - numeral b)" + "numeral a OR numeral b = word_of_int (numeral a OR numeral b)" + "numeral a OR - numeral b = word_of_int (numeral a OR - numeral b)" + "- numeral a OR numeral b = word_of_int (- numeral a OR numeral b)" + "- numeral a OR - numeral b = word_of_int (- numeral a OR - numeral b)" + "numeral a XOR numeral b = word_of_int (numeral a XOR numeral b)" + "numeral a XOR - numeral b = word_of_int (numeral a XOR - numeral b)" + "- numeral a XOR numeral b = word_of_int (- numeral a XOR numeral b)" + "- numeral a XOR - numeral b = word_of_int (- numeral a XOR - numeral b)" + by (transfer, rule refl)+ + +text \Special cases for when one of the arguments equals 1.\ + +lemma word_bitwise_1_simps [simp]: + "NOT (1::'a::len word) = -2" + "1 AND numeral b = word_of_int (1 AND numeral b)" + "1 AND - numeral b = word_of_int (1 AND - numeral b)" + "numeral a AND 1 = word_of_int (numeral a AND 1)" + "- numeral a AND 1 = word_of_int (- numeral a AND 1)" + "1 OR numeral b = word_of_int (1 OR numeral b)" + "1 OR - numeral b = word_of_int (1 OR - numeral b)" + "numeral a OR 1 = word_of_int (numeral a OR 1)" + "- numeral a OR 1 = word_of_int (- numeral a OR 1)" + "1 XOR numeral b = word_of_int (1 XOR numeral b)" + "1 XOR - numeral b = word_of_int (1 XOR - numeral b)" + "numeral a XOR 1 = word_of_int (numeral a XOR 1)" + "- numeral a XOR 1 = word_of_int (- numeral a XOR 1)" + by (transfer, simp)+ + +text \Special cases for when one of the arguments equals -1.\ + +lemma word_bitwise_m1_simps [simp]: + "NOT (-1::'a::len word) = 0" + "(-1::'a::len word) AND x = x" + "x AND (-1::'a::len word) = x" + "(-1::'a::len word) OR x = -1" + "x OR (-1::'a::len word) = -1" + " (-1::'a::len word) XOR x = NOT x" + "x XOR (-1::'a::len word) = NOT x" + by (transfer, simp)+ + +lemma uint_and: + \uint (x AND y) = uint x AND uint y\ + by transfer simp + +lemma uint_or: + \uint (x OR y) = uint x OR uint y\ + by transfer simp + +lemma uint_xor: + \uint (x XOR y) = uint x XOR uint y\ + by transfer simp + +\ \get from commutativity, associativity etc of \int_and\ etc to same for \word_and etc\\ +lemmas bwsimps = + wi_hom_add + word_wi_log_defs + +lemma word_bw_assocs: + "(x AND y) AND z = x AND y AND z" + "(x OR y) OR z = x OR y OR z" + "(x XOR y) XOR z = x XOR y XOR z" + for x :: "'a::len word" + by (fact ac_simps)+ + +lemma word_bw_comms: + "x AND y = y AND x" + "x OR y = y OR x" + "x XOR y = y XOR x" + for x :: "'a::len word" + by (fact ac_simps)+ + +lemma word_bw_lcs: + "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 :: "'a::len word" + by (fact ac_simps)+ + +lemma word_log_esimps: + "x AND 0 = 0" + "x AND -1 = x" + "x OR 0 = x" + "x OR -1 = -1" + "x XOR 0 = x" + "x XOR -1 = NOT x" + "0 AND x = 0" + "-1 AND x = x" + "0 OR x = x" + "-1 OR x = -1" + "0 XOR x = x" + "-1 XOR x = NOT x" + for x :: "'a::len word" + by simp_all + +lemma word_not_dist: + "NOT (x OR y) = NOT x AND NOT y" + "NOT (x AND y) = NOT x OR NOT y" + for x :: "'a::len word" + by simp_all + +lemma word_bw_same: + "x AND x = x" + "x OR x = x" + "x XOR x = 0" + for x :: "'a::len word" + by simp_all + +lemma word_ao_absorbs [simp]: + "x AND (y OR x) = x" + "x OR y AND x = x" + "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 :: "'a::len word" + by (auto intro: bit_eqI simp add: bit_and_iff bit_or_iff) + +lemma word_not_not [simp]: "NOT (NOT x) = x" + for x :: "'a::len word" + by (fact bit.double_compl) + +lemma word_ao_dist: "(x OR y) AND z = x AND z OR y AND z" + for x :: "'a::len word" + by (fact bit.conj_disj_distrib2) + +lemma word_oa_dist: "x AND y OR z = (x OR z) AND (y OR z)" + for x :: "'a::len word" + by (fact bit.disj_conj_distrib2) + +lemma word_add_not [simp]: "x + NOT x = -1" + for x :: "'a::len word" + by (simp add: not_eq_complement) + +lemma word_plus_and_or [simp]: "(x AND y) + (x OR y) = x + y" + for x :: "'a::len word" + by transfer (simp add: plus_and_or) + +lemma leoa: "w = x OR y \ y = w AND y" + for x :: "'a::len word" + by auto + +lemma leao: "w' = x' AND y' \ x' = x' OR w'" + for x' :: "'a::len word" + by auto + +lemma word_ao_equiv: "w = w OR w' \ w' = w AND w'" + for w w' :: "'a::len word" + by (auto intro: leoa leao) + +lemma le_word_or2: "x \ x OR y" + for x y :: "'a::len word" + by (simp add: or_greater_eq uint_or word_le_def) + +lemmas le_word_or1 = xtrans(3) [OF word_bw_comms (2) le_word_or2] +lemmas word_and_le1 = xtrans(3) [OF word_ao_absorbs (4) [symmetric] le_word_or2] +lemmas word_and_le2 = xtrans(3) [OF word_ao_absorbs (8) [symmetric] le_word_or2] + +lemma bit_horner_sum_bit_word_iff: + \bit (horner_sum of_bool (2 :: 'a::len word) bs) n + \ n < min LENGTH('a) (length bs) \ bs ! n\ + by transfer (simp add: bit_horner_sum_bit_iff) + +definition word_reverse :: \'a::len word \ 'a word\ + where \word_reverse w = horner_sum of_bool 2 (rev (map (bit w) [0.. + +lemma bit_word_reverse_iff: + \bit (word_reverse w) n \ n < LENGTH('a) \ bit w (LENGTH('a) - Suc n)\ + for w :: \'a::len word\ + by (cases \n < LENGTH('a)\) + (simp_all add: word_reverse_def bit_horner_sum_bit_word_iff rev_nth) + +lemma word_rev_rev [simp] : "word_reverse (word_reverse w) = w" + by (rule bit_word_eqI) + (auto simp add: bit_word_reverse_iff bit_imp_le_length Suc_diff_Suc) + +lemma word_rev_gal: "word_reverse w = u \ word_reverse u = w" + by (metis word_rev_rev) + +lemma word_rev_gal': "u = word_reverse w \ w = word_reverse u" + by simp + +lemma uint_2p: "(0::'a::len word) < 2 ^ n \ uint (2 ^ n::'a::len word) = 2 ^ n" + apply (cases \n < LENGTH('a)\; transfer) + apply auto + done + +lemma word_of_int_2p: "(word_of_int (2 ^ n) :: 'a::len word) = 2 ^ n" + by (induct n) (simp_all add: wi_hom_syms) + + +subsection \Shifting, Rotating, and Splitting Words\ + +lemma shiftl1_wi [simp]: "shiftl1 (word_of_int w) = word_of_int (2 * w)" + by transfer simp + +lemma shiftl1_numeral [simp]: "shiftl1 (numeral w) = numeral (Num.Bit0 w)" + unfolding word_numeral_alt shiftl1_wi by simp + +lemma shiftl1_neg_numeral [simp]: "shiftl1 (- numeral w) = - numeral (Num.Bit0 w)" + unfolding word_neg_numeral_alt shiftl1_wi by simp + +lemma shiftl1_0 [simp] : "shiftl1 0 = 0" + by transfer simp + +lemma shiftl1_def_u: "shiftl1 w = word_of_int (2 * uint w)" + by (fact shiftl1_eq) + +lemma shiftl1_def_s: "shiftl1 w = word_of_int (2 * sint w)" + by (simp add: shiftl1_def_u wi_hom_syms) + +lemma shiftr1_0 [simp]: "shiftr1 0 = 0" + by transfer simp + +lemma sshiftr1_0 [simp]: "sshiftr1 0 = 0" + by transfer simp + +lemma sshiftr1_n1 [simp]: "sshiftr1 (- 1) = - 1" + by transfer simp + +text \ + see paper page 10, (1), (2), \shiftr1_def\ is of the form of (1), + where \f\ (ie \_ div 2\) takes normal arguments to normal results, + thus we get (2) from (1) +\ + +lemma uint_shiftr1: "uint (shiftr1 w) = uint w div 2" + using drop_bit_eq_div [of 1 \uint w\, symmetric] + apply simp + apply transfer + apply (simp add: drop_bit_take_bit min_def) + done + +lemma bit_sshiftr1_iff: + \bit (sshiftr1 w) n \ bit w (if n = LENGTH('a) - 1 then LENGTH('a) - 1 else Suc n)\ + for w :: \'a::len word\ + apply transfer + apply (auto simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def simp flip: bit_Suc) + using le_less_Suc_eq apply fastforce + using le_less_Suc_eq apply fastforce + done + +lemma shiftr1_div_2: "uint (shiftr1 w) = uint w div 2" + by (fact uint_shiftr1) + +lemma sshiftr1_div_2: "sint (sshiftr1 w) = sint w div 2" + using sint_signed_drop_bit_eq [of 1 w] + by (simp add: drop_bit_Suc sshiftr1_eq_signed_drop_bit_Suc_0) + +lemma bit_bshiftr1_iff: + \bit (bshiftr1 b w) n \ b \ n = LENGTH('a) - 1 \ bit w (Suc n)\ + for w :: \'a::len word\ + apply transfer + apply (simp add: bit_take_bit_iff flip: bit_Suc) + apply (subst disjunctive_add) + apply (auto simp add: bit_take_bit_iff bit_or_iff bit_exp_iff simp flip: bit_Suc) + done + + +subsubsection \shift functions in terms of lists of bools\ + +lemma shiftl1_rev: "shiftl1 w = word_reverse (shiftr1 (word_reverse w))" + apply (rule bit_word_eqI) + apply (auto simp add: bit_shiftl1_iff bit_word_reverse_iff bit_shiftr1_iff Suc_diff_Suc) + done + +\ \note -- the following results use \'a::len word < number_ring\\ + +lemma shiftl1_2t: "shiftl1 w = 2 * w" + for w :: "'a::len word" + by (simp add: shiftl1_eq wi_hom_mult [symmetric]) + +lemma shiftl1_p: "shiftl1 w = w + w" + for w :: "'a::len word" + by (simp add: shiftl1_2t) + +lemma shiftr1_bintr [simp]: + "(shiftr1 (numeral w) :: 'a::len word) = + word_of_int (take_bit LENGTH('a) (numeral w) div 2)" + by transfer simp + +lemma sshiftr1_sbintr [simp]: + "(sshiftr1 (numeral w) :: 'a::len word) = + word_of_int (signed_take_bit (LENGTH('a) - 1) (numeral w) div 2)" + by transfer simp + +text \TODO: rules for \<^term>\- (numeral n)\\ + +lemma drop_bit_word_numeral [simp]: + \drop_bit (numeral n) (numeral k) = + (word_of_int (drop_bit (numeral n) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\ + by transfer simp + +lemma zip_replicate: "n \ length ys \ zip (replicate n x) ys = map (\y. (x, y)) ys" + apply (induct ys arbitrary: n) + apply simp_all + apply (case_tac n) + apply simp_all + done + +lemma align_lem_or [rule_format] : + "\x m. length x = n + m \ length y = n + m \ + drop m x = replicate n False \ take m y = replicate m False \ + map2 (|) x y = take m x @ drop m y" + apply (induct y) + apply force + apply clarsimp + apply (case_tac x) + apply force + apply (case_tac m) + apply auto + apply (drule_tac t="length xs" for xs in sym) + apply (auto simp: zip_replicate o_def) + done + +lemma align_lem_and [rule_format] : + "\x m. length x = n + m \ length y = n + m \ + drop m x = replicate n False \ take m y = replicate m False \ + map2 (\) x y = replicate (n + m) False" + apply (induct y) + apply force + apply clarsimp + apply (case_tac x) + apply force + apply (case_tac m) + apply auto + apply (drule_tac t="length xs" for xs in sym) + apply (auto simp: zip_replicate o_def map_replicate_const) + done + + +subsubsection \Mask\ + +lemma minus_1_eq_mask: + \- 1 = (mask LENGTH('a) :: 'a::len word)\ + by (rule bit_eqI) (simp add: bit_exp_iff bit_mask_iff exp_eq_zero_iff) + +lemma mask_eq_decr_exp: + \mask n = 2 ^ n - (1 :: 'a::len word)\ + by (fact mask_eq_exp_minus_1) + +lemma mask_Suc_rec: + \mask (Suc n) = 2 * mask n + (1 :: 'a::len word)\ + by (simp add: mask_eq_exp_minus_1) + +context +begin + +qualified lemma bit_mask_iff: + \bit (mask m :: 'a::len word) n \ n < min LENGTH('a) m\ + by (simp add: bit_mask_iff exp_eq_zero_iff not_le) + +end + +lemma mask_bin: "mask n = word_of_int (take_bit n (- 1))" + by transfer (simp add: take_bit_minus_one_eq_mask) + +lemma and_mask_bintr: "w AND mask n = word_of_int (take_bit n (uint w))" + by transfer (simp add: ac_simps take_bit_eq_mask) + +lemma and_mask_wi: "word_of_int i AND mask n = word_of_int (take_bit n i)" + by (auto simp add: and_mask_bintr min_def not_le wi_bintr) + +lemma and_mask_wi': + "word_of_int i AND mask n = (word_of_int (take_bit (min LENGTH('a) n) i) :: 'a::len word)" + by (auto simp add: and_mask_wi min_def wi_bintr) + +lemma and_mask_no: "numeral i AND mask n = word_of_int (take_bit n (numeral i))" + unfolding word_numeral_alt by (rule and_mask_wi) + +lemma and_mask_mod_2p: "w AND mask n = word_of_int (uint w mod 2 ^ n)" + by (simp only: and_mask_bintr take_bit_eq_mod) + +lemma uint_mask_eq: + \uint (mask n :: 'a::len word) = mask (min LENGTH('a) n)\ + by transfer simp + +lemma and_mask_lt_2p: "uint (w AND mask n) < 2 ^ n" + apply (simp flip: take_bit_eq_mask) + apply transfer + apply (auto simp add: min_def) + using antisym_conv take_bit_int_eq_self_iff by fastforce + +lemma mask_eq_iff: "w AND mask n = w \ uint w < 2 ^ n" + apply (auto simp flip: take_bit_eq_mask) + apply (metis take_bit_int_eq_self_iff uint_take_bit_eq) + apply (simp add: take_bit_int_eq_self unsigned_take_bit_eq word_uint_eqI) + done + +lemma and_mask_dvd: "2 ^ n dvd uint w \ w AND mask n = 0" + by (simp flip: take_bit_eq_mask take_bit_eq_mod unsigned_take_bit_eq add: dvd_eq_mod_eq_0 uint_0_iff) + +lemma and_mask_dvd_nat: "2 ^ n dvd unat w \ w AND mask n = 0" + by (simp flip: take_bit_eq_mask take_bit_eq_mod unsigned_take_bit_eq add: dvd_eq_mod_eq_0 unat_0_iff uint_0_iff) + +lemma word_2p_lem: "n < size w \ w < 2 ^ n = (uint w < 2 ^ n)" + for w :: "'a::len word" + by transfer simp + +lemma less_mask_eq: "x < 2 ^ n \ x AND mask n = x" + for x :: "'a::len word" + apply (cases \n < LENGTH('a)\) + apply (simp_all add: not_less flip: take_bit_eq_mask exp_eq_zero_iff) + apply transfer + apply (simp add: min_def) + apply (metis min_def nat_less_le take_bit_int_eq_self_iff take_bit_take_bit) + done + +lemmas mask_eq_iff_w2p = trans [OF mask_eq_iff word_2p_lem [symmetric]] + +lemmas and_mask_less' = iffD2 [OF word_2p_lem and_mask_lt_2p, simplified word_size] + +lemma and_mask_less_size: "n < size x \ x AND mask n < 2 ^ n" + for x :: \'a::len word\ + unfolding word_size by (erule and_mask_less') + +lemma word_mod_2p_is_mask [OF refl]: "c = 2 ^ n \ c > 0 \ x mod c = x AND mask n" + for c x :: "'a::len word" + by (auto simp: word_mod_def uint_2p and_mask_mod_2p) + +lemma mask_eqs: + "(a AND mask n) + b AND mask n = a + b AND mask n" + "a + (b AND mask n) AND mask n = a + b AND mask n" + "(a AND mask n) - b AND mask n = a - b AND mask n" + "a - (b AND mask n) AND mask n = a - b AND mask n" + "a * (b AND mask n) AND mask n = a * b AND mask n" + "(b AND mask n) * a AND mask n = b * a AND mask n" + "(a AND mask n) + (b AND mask n) AND mask n = a + b AND mask n" + "(a AND mask n) - (b AND mask n) AND mask n = a - b AND mask n" + "(a AND mask n) * (b AND mask n) AND mask n = a * b AND mask n" + "- (a AND mask n) AND mask n = - a AND mask n" + "word_succ (a AND mask n) AND mask n = word_succ a AND mask n" + "word_pred (a AND mask n) AND mask n = word_pred a AND mask n" + using word_of_int_Ex [where x=a] word_of_int_Ex [where x=b] + apply (auto simp flip: take_bit_eq_mask) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + done + +lemma mask_power_eq: "(x AND mask n) ^ k AND mask n = x ^ k AND mask n" + for x :: \'a::len word\ + using word_of_int_Ex [where x=x] + apply (auto simp flip: take_bit_eq_mask) + apply transfer + apply (simp add: take_bit_eq_mod mod_simps) + done + +lemma mask_full [simp]: "mask LENGTH('a) = (- 1 :: 'a::len word)" + by transfer (simp add: take_bit_minus_one_eq_mask) + + +subsubsection \Slices\ + +definition slice1 :: \nat \ 'a::len word \ 'b::len word\ + where \slice1 n w = (if n < LENGTH('a) + then ucast (drop_bit (LENGTH('a) - n) w) + else push_bit (n - LENGTH('a)) (ucast w))\ + +lemma bit_slice1_iff: + \bit (slice1 m w :: 'b::len word) n \ m - LENGTH('a) \ n \ n < min LENGTH('b) m + \ bit w (n + (LENGTH('a) - m) - (m - LENGTH('a)))\ + for w :: \'a::len word\ + by (auto simp add: slice1_def bit_ucast_iff bit_drop_bit_eq bit_push_bit_iff exp_eq_zero_iff not_less not_le ac_simps + dest: bit_imp_le_length) + +definition slice :: \nat \ 'a::len word \ 'b::len word\ + where \slice n = slice1 (LENGTH('a) - n)\ + +lemma bit_slice_iff: + \bit (slice m w :: 'b::len word) n \ n < min LENGTH('b) (LENGTH('a) - m) \ bit w (n + LENGTH('a) - (LENGTH('a) - m))\ + for w :: \'a::len word\ + by (simp add: slice_def word_size bit_slice1_iff) + +lemma slice1_0 [simp] : "slice1 n 0 = 0" + unfolding slice1_def by simp + +lemma slice_0 [simp] : "slice n 0 = 0" + unfolding slice_def by auto + +lemma ucast_slice1: "ucast w = slice1 (size w) w" + apply (simp add: slice1_def) + apply transfer + apply simp + done + +lemma ucast_slice: "ucast w = slice 0 w" + by (simp add: slice_def slice1_def) + +lemma slice_id: "slice 0 t = t" + by (simp only: ucast_slice [symmetric] ucast_id) + +lemma rev_slice1: + \slice1 n (word_reverse w :: 'b::len word) = word_reverse (slice1 k w :: 'a::len word)\ + if \n + k = LENGTH('a) + LENGTH('b)\ +proof (rule bit_word_eqI) + fix m + assume *: \m < LENGTH('a)\ + from that have **: \LENGTH('b) = n + k - LENGTH('a)\ + by simp + show \bit (slice1 n (word_reverse w :: 'b word) :: 'a word) m \ bit (word_reverse (slice1 k w :: 'a word)) m\ + apply (simp add: bit_slice1_iff bit_word_reverse_iff) + using * ** + apply (cases \n \ LENGTH('a)\; cases \k \ LENGTH('a)\) + apply auto + done +qed + +lemma rev_slice: + "n + k + LENGTH('a::len) = LENGTH('b::len) \ + slice n (word_reverse (w::'b word)) = word_reverse (slice k w :: 'a word)" + apply (unfold slice_def word_size) + apply (rule rev_slice1) + apply arith + done + + +subsubsection \Revcast\ + +definition revcast :: \'a::len word \ 'b::len word\ + where \revcast = slice1 LENGTH('b)\ + +lemma bit_revcast_iff: + \bit (revcast w :: 'b::len word) n \ LENGTH('b) - LENGTH('a) \ n \ n < LENGTH('b) + \ bit w (n + (LENGTH('a) - LENGTH('b)) - (LENGTH('b) - LENGTH('a)))\ + for w :: \'a::len word\ + by (simp add: revcast_def bit_slice1_iff) + +lemma revcast_slice1 [OF refl]: "rc = revcast w \ slice1 (size rc) w = rc" + by (simp add: revcast_def word_size) + +lemma revcast_rev_ucast [OF refl refl refl]: + "cs = [rc, uc] \ rc = revcast (word_reverse w) \ uc = ucast w \ + rc = word_reverse uc" + apply auto + apply (rule bit_word_eqI) + apply (cases \LENGTH('a) \ LENGTH('b)\) + apply (simp_all add: bit_revcast_iff bit_word_reverse_iff bit_ucast_iff not_le + bit_imp_le_length) + using bit_imp_le_length apply fastforce + using bit_imp_le_length apply fastforce + done + +lemma revcast_ucast: "revcast w = word_reverse (ucast (word_reverse w))" + using revcast_rev_ucast [of "word_reverse w"] by simp + +lemma ucast_revcast: "ucast w = word_reverse (revcast (word_reverse w))" + by (fact revcast_rev_ucast [THEN word_rev_gal']) + +lemma ucast_rev_revcast: "ucast (word_reverse w) = word_reverse (revcast w)" + by (fact revcast_ucast [THEN word_rev_gal']) + + +text "linking revcast and cast via shift" + +lemmas wsst_TYs = source_size target_size word_size + +lemmas sym_notr = + not_iff [THEN iffD2, THEN not_sym, THEN not_iff [THEN iffD1]] + + +subsection \Split and cat\ + +lemmas word_split_bin' = word_split_def +lemmas word_cat_bin' = word_cat_eq + +\ \this odd result is analogous to \ucast_id\, + result to the length given by the result type\ + +lemma word_cat_id: "word_cat a b = b" + by transfer (simp add: take_bit_concat_bit_eq) + +lemma word_cat_split_alt: "size w \ size u + size v \ word_split w = (u, v) \ word_cat u v = w" + apply (rule bit_word_eqI) + apply (auto simp add: bit_word_cat_iff not_less word_size word_split_def bit_ucast_iff bit_drop_bit_eq) + done + +lemmas word_cat_split_size = sym [THEN [2] word_cat_split_alt [symmetric]] + + +subsubsection \Split and slice\ + +lemma split_slices: "word_split w = (u, v) \ u = slice (size v) w \ v = slice 0 w" + apply (auto simp add: word_split_def word_size) + apply (rule bit_word_eqI) + apply (simp add: bit_slice_iff bit_ucast_iff bit_drop_bit_eq) + apply (cases \LENGTH('c) \ LENGTH('b)\) + apply (auto simp add: ac_simps dest: bit_imp_le_length) + apply (rule bit_word_eqI) + apply (auto simp add: bit_slice_iff bit_ucast_iff dest: bit_imp_le_length) + done + +lemma slice_cat1 [OF refl]: + "wc = word_cat a b \ size wc >= size a + size b \ slice (size b) wc = a" + apply (rule bit_word_eqI) + apply (auto simp add: bit_slice_iff bit_word_cat_iff word_size) + done + +lemmas slice_cat2 = trans [OF slice_id word_cat_id] + +lemma cat_slices: + "a = slice n c \ b = slice 0 c \ n = size b \ + size a + size b >= size c \ word_cat a b = c" + apply (rule bit_word_eqI) + apply (auto simp add: bit_slice_iff bit_word_cat_iff word_size) + done + +lemma word_split_cat_alt: + "w = word_cat u v \ size u + size v \ size w \ word_split w = (u, v)" + apply (auto simp add: word_split_def word_size) + apply (rule bit_eqI) + apply (auto simp add: bit_ucast_iff bit_drop_bit_eq bit_word_cat_iff dest: bit_imp_le_length) + apply (rule bit_eqI) + apply (auto simp add: bit_ucast_iff bit_drop_bit_eq bit_word_cat_iff dest: bit_imp_le_length) + done + +lemma horner_sum_uint_exp_Cons_eq: + \horner_sum uint (2 ^ LENGTH('a)) (w # ws) = + concat_bit LENGTH('a) (uint w) (horner_sum uint (2 ^ LENGTH('a)) ws)\ + for ws :: \'a::len word list\ + apply (simp add: concat_bit_eq push_bit_eq_mult) + apply transfer + apply simp + done + +lemma bit_horner_sum_uint_exp_iff: + \bit (horner_sum uint (2 ^ LENGTH('a)) ws) n \ + n div LENGTH('a) < length ws \ bit (ws ! (n div LENGTH('a))) (n mod LENGTH('a))\ + for ws :: \'a::len word list\ +proof (induction ws arbitrary: n) + case Nil + then show ?case + by simp +next + case (Cons w ws) + then show ?case + by (cases \n \ LENGTH('a)\) + (simp_all only: horner_sum_uint_exp_Cons_eq, simp_all add: bit_concat_bit_iff le_div_geq le_mod_geq bit_uint_iff Cons) +qed + + +subsection \Rotation\ + +lemma word_rotr_word_rotr_eq: + \word_rotr m (word_rotr n w) = word_rotr (m + n) w\ + by (rule bit_word_eqI) (simp add: bit_word_rotr_iff ac_simps mod_add_right_eq) + +lemma word_rot_rl [simp]: + \word_rotl k (word_rotr k v) = v\ + apply (rule bit_word_eqI) + apply (simp add: word_rotl_eq_word_rotr word_rotr_word_rotr_eq bit_word_rotr_iff algebra_simps) + apply (auto dest: bit_imp_le_length) + apply (metis (no_types, lifting) add.right_neutral add_diff_cancel_right' div_mult_mod_eq mod_add_right_eq mod_add_self2 mod_if mod_mult_self2_is_0) + apply (metis (no_types, lifting) add.right_neutral add_diff_cancel_right' div_mult_mod_eq mod_add_right_eq mod_add_self2 mod_less mod_mult_self2_is_0) + done + +lemma word_rot_lr [simp]: + \word_rotr k (word_rotl k v) = v\ + apply (rule bit_word_eqI) + apply (simp add: word_rotl_eq_word_rotr word_rotr_word_rotr_eq bit_word_rotr_iff algebra_simps) + apply (auto dest: bit_imp_le_length) + apply (metis (no_types, lifting) add.right_neutral add_diff_cancel_right' div_mult_mod_eq mod_add_right_eq mod_add_self2 mod_if mod_mult_self2_is_0) + apply (metis (no_types, lifting) add.right_neutral add_diff_cancel_right' div_mult_mod_eq mod_add_right_eq mod_add_self2 mod_less mod_mult_self2_is_0) + done + +lemma word_rot_gal: + \word_rotr n v = w \ word_rotl n w = v\ + by auto + +lemma word_rot_gal': + \w = word_rotr n v \ v = word_rotl n w\ + by auto + +lemma word_rotr_rev: + \word_rotr n w = word_reverse (word_rotl n (word_reverse w))\ +proof (rule bit_word_eqI) + fix m + assume \m < LENGTH('a)\ + moreover have \1 + + ((int m + int n mod int LENGTH('a)) mod int LENGTH('a) + + ((int LENGTH('a) * 2) mod int LENGTH('a) - (1 + (int m + int n mod int LENGTH('a)))) mod int LENGTH('a)) = + int LENGTH('a)\ + apply (cases \(1 + (int m + int n mod int LENGTH('a))) mod + int LENGTH('a) = 0\) + using zmod_zminus1_eq_if [of \1 + (int m + int n mod int LENGTH('a))\ \int LENGTH('a)\] + apply simp_all + apply (auto simp add: algebra_simps) + apply (simp add: minus_equation_iff [of \int m\]) + apply (drule sym [of _ \int m\]) + apply simp + apply (metis add.commute add_minus_cancel diff_minus_eq_add len_gt_0 less_imp_of_nat_less less_nat_zero_code mod_mult_self3 of_nat_gt_0 zmod_minus1) + apply (metis (no_types, hide_lams) Abs_fnat_hom_add less_not_refl mod_Suc of_nat_Suc of_nat_gt_0 of_nat_mod) + done + then have \int ((m + n) mod LENGTH('a)) = + int (LENGTH('a) - Suc ((LENGTH('a) - Suc m + LENGTH('a) - n mod LENGTH('a)) mod LENGTH('a)))\ + using \m < LENGTH('a)\ + by (simp only: of_nat_mod mod_simps) + (simp add: of_nat_diff of_nat_mod Suc_le_eq add_less_mono algebra_simps mod_simps) + then have \(m + n) mod LENGTH('a) = + LENGTH('a) - Suc ((LENGTH('a) - Suc m + LENGTH('a) - n mod LENGTH('a)) mod LENGTH('a))\ + by simp + ultimately show \bit (word_rotr n w) m \ bit (word_reverse (word_rotl n (word_reverse w))) m\ + by (simp add: word_rotl_eq_word_rotr bit_word_rotr_iff bit_word_reverse_iff) +qed + +lemma word_roti_0 [simp]: "word_roti 0 w = w" + by transfer simp + +lemma word_roti_add: "word_roti (m + n) w = word_roti m (word_roti n w)" + by (rule bit_word_eqI) + (simp add: bit_word_roti_iff nat_less_iff mod_simps ac_simps) + +lemma word_roti_conv_mod': + "word_roti n w = word_roti (n mod int (size w)) w" + by transfer simp + +lemmas word_roti_conv_mod = word_roti_conv_mod' [unfolded word_size] + + +subsubsection \"Word rotation commutes with bit-wise operations\ + +\ \using locale to not pollute lemma namespace\ +locale word_rotate +begin + +lemma word_rot_logs: + "word_rotl n (NOT v) = NOT (word_rotl n v)" + "word_rotr n (NOT v) = NOT (word_rotr n v)" + "word_rotl n (x AND y) = word_rotl n x AND word_rotl n y" + "word_rotr n (x AND y) = word_rotr n x AND word_rotr n y" + "word_rotl n (x OR y) = word_rotl n x OR word_rotl n y" + "word_rotr n (x OR y) = word_rotr n x OR word_rotr n y" + "word_rotl n (x XOR y) = word_rotl n x XOR word_rotl n y" + "word_rotr n (x XOR y) = word_rotr n x XOR word_rotr n y" + apply (rule bit_word_eqI) + apply (auto simp add: bit_word_rotl_iff bit_not_iff algebra_simps exp_eq_zero_iff not_le) + apply (rule bit_word_eqI) + apply (auto simp add: bit_word_rotr_iff bit_not_iff algebra_simps exp_eq_zero_iff not_le) + apply (rule bit_word_eqI) + apply (auto simp add: bit_word_rotl_iff bit_and_iff algebra_simps exp_eq_zero_iff not_le) + apply (rule bit_word_eqI) + apply (auto simp add: bit_word_rotr_iff bit_and_iff algebra_simps exp_eq_zero_iff not_le) + apply (rule bit_word_eqI) + apply (auto simp add: bit_word_rotl_iff bit_or_iff algebra_simps exp_eq_zero_iff not_le) + apply (rule bit_word_eqI) + apply (auto simp add: bit_word_rotr_iff bit_or_iff algebra_simps exp_eq_zero_iff not_le) + apply (rule bit_word_eqI) + apply (auto simp add: bit_word_rotl_iff bit_xor_iff algebra_simps exp_eq_zero_iff not_le) + apply (rule bit_word_eqI) + apply (auto simp add: bit_word_rotr_iff bit_xor_iff algebra_simps exp_eq_zero_iff not_le) + done + +end + +lemmas word_rot_logs = word_rotate.word_rot_logs + +lemma word_rotx_0 [simp] : "word_rotr i 0 = 0 \ word_rotl i 0 = 0" + by transfer simp_all + +lemma word_roti_0' [simp] : "word_roti n 0 = 0" + by transfer simp + +declare word_roti_eq_word_rotr_word_rotl [simp] + + +subsection \Maximum machine word\ + +lemma word_int_cases: + fixes x :: "'a::len word" + obtains n where "x = word_of_int n" and "0 \ n" and "n < 2^LENGTH('a)" + by (rule that [of \uint x\]) simp_all + +lemma word_nat_cases [cases type: word]: + fixes x :: "'a::len word" + obtains n where "x = of_nat n" and "n < 2^LENGTH('a)" + by (rule that [of \unat x\]) simp_all + +lemma max_word_max [intro!]: "n \ max_word" + by (fact word_order.extremum) + +lemma word_of_int_2p_len: "word_of_int (2 ^ LENGTH('a)) = (0::'a::len word)" + by simp + +lemma word_pow_0: "(2::'a::len word) ^ LENGTH('a) = 0" + by (fact word_exp_length_eq_0) + +lemma max_word_wrap: "x + 1 = 0 \ x = max_word" + by (simp add: eq_neg_iff_add_eq_0) + +lemma word_and_max: "x AND max_word = x" + by (fact word_log_esimps) + +lemma word_or_max: "x OR max_word = max_word" + by (fact word_log_esimps) + +lemma word_ao_dist2: "x AND (y OR z) = x AND y OR x AND z" + for x y z :: "'a::len word" + by (fact bit.conj_disj_distrib) + +lemma word_oa_dist2: "x OR y AND z = (x OR y) AND (x OR z)" + for x y z :: "'a::len word" + by (fact bit.disj_conj_distrib) + +lemma word_and_not [simp]: "x AND NOT x = 0" + for x :: "'a::len word" + by (fact bit.conj_cancel_right) + +lemma word_or_not [simp]: "x OR NOT x = max_word" + by (fact bit.disj_cancel_right) + +lemma word_xor_and_or: "x XOR y = x AND NOT y OR NOT x AND y" + for x y :: "'a::len word" + by (fact bit.xor_def) + +lemma uint_lt_0 [simp]: "uint x < 0 = False" + by (simp add: linorder_not_less) + +lemma shiftr1_1 [simp]: "shiftr1 (1::'a::len word) = 0" + by transfer simp + +lemma word_less_1 [simp]: "x < 1 \ x = 0" + for x :: "'a::len word" + by (simp add: word_less_nat_alt unat_0_iff) + +lemma uint_plus_if_size: + "uint (x + y) = + (if uint x + uint y < 2^size x + then uint x + uint y + else uint x + uint y - 2^size x)" + apply (simp only: word_arith_wis word_size uint_word_of_int_eq) + apply (auto simp add: not_less take_bit_int_eq_self_iff) + apply (metis not_less take_bit_eq_mod uint_plus_if' uint_word_ariths(1)) + done + +lemma unat_plus_if_size: + "unat (x + y) = + (if unat x + unat y < 2^size x + then unat x + unat y + else unat x + unat y - 2^size x)" + for x y :: "'a::len word" + apply (subst word_arith_nat_defs) + apply (subst unat_of_nat) + apply (auto simp add: not_less word_size) + apply (metis not_le unat_plus_if' unat_word_ariths(1)) + done + +lemma word_neq_0_conv: "w \ 0 \ 0 < w" + for w :: "'a::len word" + by (fact word_coorder.not_eq_extremum) + +lemma max_lt: "unat (max a b div c) = unat (max a b) div unat c" + for c :: "'a::len word" + by (fact unat_div) + +lemma uint_sub_if_size: + "uint (x - y) = + (if uint y \ uint x + then uint x - uint y + else uint x - uint y + 2^size x)" + apply (simp only: word_arith_wis word_size uint_word_of_int_eq) + apply (auto simp add: take_bit_int_eq_self_iff not_le) + apply (metis not_less uint_sub_if' uint_word_arith_bintrs(2)) + done + +lemma unat_sub: + \unat (a - b) = unat a - unat b\ + if \b \ a\ +proof - + from that have \unat b \ unat a\ + by transfer simp + with that show ?thesis + apply transfer + apply simp + apply (subst take_bit_diff [symmetric]) + apply (subst nat_take_bit_eq) + apply (simp add: nat_le_eq_zle) + apply (simp add: nat_diff_distrib take_bit_nat_eq_self_iff less_imp_diff_less) + done +qed + +lemmas word_less_sub1_numberof [simp] = word_less_sub1 [of "numeral w"] for w +lemmas word_le_sub1_numberof [simp] = word_le_sub1 [of "numeral w"] for w + +lemma word_of_int_minus: "word_of_int (2^LENGTH('a) - i) = (word_of_int (-i)::'a::len word)" + apply transfer + apply (subst take_bit_diff [symmetric]) + apply (simp add: take_bit_minus) + done + +lemma word_of_int_inj: + \(word_of_int x :: 'a::len word) = word_of_int y \ x = y\ + if \0 \ x \ x < 2 ^ LENGTH('a)\ \0 \ y \ y < 2 ^ LENGTH('a)\ + using that by (transfer fixing: x y) (simp add: take_bit_int_eq_self) + +lemma word_le_less_eq: "x \ y \ x = y \ x < y" + for x y :: "'z::len word" + by (auto simp add: order_class.le_less) + +lemma mod_plus_cong: + fixes b b' :: int + assumes 1: "b = b'" + and 2: "x mod b' = x' mod b'" + and 3: "y mod b' = y' mod b'" + and 4: "x' + y' = z'" + shows "(x + y) mod b = z' mod b'" +proof - + from 1 2[symmetric] 3[symmetric] have "(x + y) mod b = (x' mod b' + y' mod b') mod b'" + by (simp add: mod_add_eq) + also have "\ = (x' + y') mod b'" + by (simp add: mod_add_eq) + finally show ?thesis + by (simp add: 4) +qed + +lemma mod_minus_cong: + fixes b b' :: int + assumes "b = b'" + and "x mod b' = x' mod b'" + and "y mod b' = y' mod b'" + and "x' - y' = z'" + shows "(x - y) mod b = z' mod b'" + using assms [symmetric] by (auto intro: mod_diff_cong) + +lemma word_induct_less: + \P m\ if zero: \P 0\ and less: \\n. n < m \ P n \ P (1 + n)\ + for m :: \'a::len word\ +proof - + define q where \q = unat m\ + with less have \\n. n < word_of_nat q \ P n \ P (1 + n)\ + by simp + then have \P (word_of_nat q :: 'a word)\ + proof (induction q) + case 0 + show ?case + by (simp add: zero) + next + case (Suc q) + show ?case + proof (cases \1 + word_of_nat q = (0 :: 'a word)\) + case True + then show ?thesis + by (simp add: zero) + next + case False + then have *: \word_of_nat q < (word_of_nat (Suc q) :: 'a word)\ + by (simp add: unatSuc word_less_nat_alt) + then have **: \n < (1 + word_of_nat q :: 'a word) \ n \ (word_of_nat q :: 'a word)\ for n + by (metis (no_types, lifting) add.commute inc_le le_less_trans not_less of_nat_Suc) + have \P (word_of_nat q)\ + apply (rule Suc.IH) + apply (rule Suc.prems) + apply (erule less_trans) + apply (rule *) + apply assumption + done + with * have \P (1 + word_of_nat q)\ + by (rule Suc.prems) + then show ?thesis + by simp + qed + qed + with \q = unat m\ show ?thesis + by simp +qed + +lemma word_induct: "P 0 \ (\n. P n \ P (1 + n)) \ P m" + for P :: "'a::len word \ bool" + by (rule word_induct_less) + +lemma word_induct2 [induct type]: "P 0 \ (\n. 1 + n \ 0 \ P n \ P (1 + n)) \ P n" + for P :: "'b::len word \ bool" + apply (rule word_induct_less) + apply simp_all + apply (case_tac "1 + na = 0") + apply auto + done + + +subsection \Recursion combinator for words\ + +definition word_rec :: "'a \ ('b::len word \ 'a \ 'a) \ 'b word \ 'a" + where "word_rec forZero forSuc n = rec_nat forZero (forSuc \ of_nat) (unat n)" + +lemma word_rec_0: "word_rec z s 0 = z" + by (simp add: word_rec_def) + +lemma word_rec_Suc: "1 + n \ 0 \ word_rec z s (1 + n) = s n (word_rec z s n)" + for n :: "'a::len word" + apply (auto simp add: word_rec_def unat_word_ariths) + apply (metis (mono_tags, lifting) Abs_fnat_hom_add add_diff_cancel_left' o_def of_nat_1 old.nat.simps(7) plus_1_eq_Suc unatSuc unat_word_ariths(1) unsigned_1 word_arith_nat_add) + done + +lemma word_rec_Pred: "n \ 0 \ word_rec z s n = s (n - 1) (word_rec z s (n - 1))" + apply (rule subst[where t="n" and s="1 + (n - 1)"]) + apply simp + apply (subst word_rec_Suc) + apply simp + apply simp + done + +lemma word_rec_in: "f (word_rec z (\_. f) n) = word_rec (f z) (\_. f) n" + by (induct n) (simp_all add: word_rec_0 word_rec_Suc) + +lemma word_rec_in2: "f n (word_rec z f n) = word_rec (f 0 z) (f \ (+) 1) n" + by (induct n) (simp_all add: word_rec_0 word_rec_Suc) + +lemma word_rec_twice: + "m \ n \ word_rec z f n = word_rec (word_rec z f (n - m)) (f \ (+) (n - m)) m" + apply (erule rev_mp) + apply (rule_tac x=z in spec) + apply (rule_tac x=f in spec) + apply (induct n) + apply (simp add: word_rec_0) + apply clarsimp + apply (rule_tac t="1 + n - m" and s="1 + (n - m)" in subst) + apply simp + apply (case_tac "1 + (n - m) = 0") + apply (simp add: word_rec_0) + apply (rule_tac f = "word_rec a b" for a b in arg_cong) + apply (rule_tac t="m" and s="m + (1 + (n - m))" in subst) + apply simp + apply (simp (no_asm_use)) + apply (simp add: word_rec_Suc word_rec_in2) + apply (erule impE) + apply uint_arith + apply (drule_tac x="x \ (+) 1" in spec) + apply (drule_tac x="x 0 xa" in spec) + apply simp + apply (rule_tac t="\a. x (1 + (n - m + a))" and s="\a. x (1 + (n - m) + a)" in subst) + apply (clarsimp simp add: fun_eq_iff) + apply (rule_tac t="(1 + (n - m + xb))" and s="1 + (n - m) + xb" in subst) + apply simp + apply (rule refl) + apply (rule refl) + done + +lemma word_rec_id: "word_rec z (\_. id) n = z" + by (induct n) (auto simp add: word_rec_0 word_rec_Suc) + +lemma word_rec_id_eq: "\m < n. f m = id \ word_rec z f n = z" + apply (erule rev_mp) + apply (induct n) + apply (auto simp add: word_rec_0 word_rec_Suc) + apply (drule spec, erule mp) + apply uint_arith + apply (drule_tac x=n in spec, erule impE) + apply uint_arith + apply simp + done + +lemma word_rec_max: + "\m\n. m \ - 1 \ f m = id \ word_rec z f (- 1) = word_rec z f n" + apply (subst word_rec_twice[where n="-1" and m="-1 - n"]) + apply simp + apply simp + apply (rule word_rec_id_eq) + apply clarsimp + apply (drule spec, rule mp, erule mp) + apply (rule word_plus_mono_right2[OF _ order_less_imp_le]) + prefer 2 + apply assumption + apply simp + apply (erule contrapos_pn) + apply simp + apply (drule arg_cong[where f="\x. x - n"]) + apply simp + done + + +subsection \More\ + +lemma mask_1: "mask 1 = 1" + by simp + +lemma mask_Suc_0: "mask (Suc 0) = 1" + by simp + +lemma bin_last_bintrunc: "odd (take_bit l n) \ l > 0 \ odd n" + by simp + + +lemma push_bit_word_beyond [simp]: + \push_bit n w = 0\ if \LENGTH('a) \ n\ for w :: \'a::len word\ + using that by (transfer fixing: n) (simp add: take_bit_push_bit) + +lemma drop_bit_word_beyond [simp]: + \drop_bit n w = 0\ if \LENGTH('a) \ n\ for w :: \'a::len word\ + using that by (transfer fixing: n) (simp add: drop_bit_take_bit) + +lemma signed_drop_bit_beyond: + \signed_drop_bit n w = (if bit w (LENGTH('a) - Suc 0) then - 1 else 0)\ + if \LENGTH('a) \ n\ for w :: \'a::len word\ + by (rule bit_word_eqI) (simp add: bit_signed_drop_bit_iff that) + + +subsection \SMT support\ + +ML_file \Tools/smt_word.ML\ + +end diff -r d8661799afb2 -r c7038c397ae3 src/HOL/ROOT --- a/src/HOL/ROOT Thu Oct 29 09:59:40 2020 +0000 +++ b/src/HOL/ROOT Thu Oct 29 10:03:03 2020 +0000 @@ -324,7 +324,6 @@ "HOL-Number_Theory" "HOL-Data_Structures" "HOL-Examples" - "HOL-Word" theories Generate Generate_Binary_Nat @@ -615,8 +614,6 @@ description " Miscellaneous examples for Higher-Order Logic. " - sessions - "HOL-Word" theories Antiquote Argo_Examples @@ -892,15 +889,6 @@ "Tests/TLList_Friends" "Tests/Type_Class" -session "HOL-Word" (main timing) in Word = HOL + - sessions - "HOL-Library" - theories - Word - More_Word - Word_Examples - document_files "root.bib" "root.tex" - session "HOL-Statespace" in Statespace = HOL + theories [skip_proofs = false] StateSpaceEx @@ -925,8 +913,10 @@ options [timeout = 60] theories Ex -session "HOL-Word-SMT_Examples" (timing) in SMT_Examples = "HOL-Word" + +session "HOL-SMT_Examples" (timing) in SMT_Examples = HOL + options [quick_and_dirty] + sessions + "HOL-Library" theories Boogie SMT_Examples @@ -935,7 +925,9 @@ SMT_Tests_Verit SMT_Examples_Verit -session "HOL-SPARK" in "SPARK" = "HOL-Word" + +session "HOL-SPARK" in "SPARK" = HOL + + sessions + "HOL-Library" theories SPARK @@ -959,6 +951,7 @@ session "HOL-SPARK-Manual" in "SPARK/Manual" = "HOL-SPARK" + options [show_question_marks = false] sessions + "HOL-Library" "HOL-SPARK-Examples" theories Example_Verification diff -r d8661799afb2 -r c7038c397ae3 src/HOL/SMT_Examples/SMT_Word_Examples.thy --- a/src/HOL/SMT_Examples/SMT_Word_Examples.thy Thu Oct 29 09:59:40 2020 +0000 +++ b/src/HOL/SMT_Examples/SMT_Word_Examples.thy Thu Oct 29 10:03:03 2020 +0000 @@ -5,7 +5,7 @@ section \Word examples for for SMT binding\ theory SMT_Word_Examples -imports "HOL-Word.Word" +imports "HOL-Library.Word" begin declare [[smt_oracle = true]] diff -r d8661799afb2 -r c7038c397ae3 src/HOL/SPARK/Examples/RIPEMD-160/RMD.thy --- a/src/HOL/SPARK/Examples/RIPEMD-160/RMD.thy Thu Oct 29 09:59:40 2020 +0000 +++ b/src/HOL/SPARK/Examples/RIPEMD-160/RMD.thy Thu Oct 29 10:03:03 2020 +0000 @@ -5,7 +5,7 @@ *) theory RMD -imports "HOL-Word.Word" +imports "HOL-Library.Word" begin diff -r d8661799afb2 -r c7038c397ae3 src/HOL/SPARK/SPARK_Setup.thy --- a/src/HOL/SPARK/SPARK_Setup.thy Thu Oct 29 09:59:40 2020 +0000 +++ b/src/HOL/SPARK/SPARK_Setup.thy Thu Oct 29 10:03:03 2020 +0000 @@ -7,7 +7,7 @@ theory SPARK_Setup imports - "HOL-Word.Word" + "HOL-Library.Word" keywords "spark_open_vcg" :: thy_load ("vcg", "fdl", "rls") and "spark_open" :: thy_load ("siv", "fdl", "rls") and diff -r d8661799afb2 -r c7038c397ae3 src/HOL/Word/Ancient_Numeral.thy --- a/src/HOL/Word/Ancient_Numeral.thy Thu Oct 29 09:59:40 2020 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,229 +0,0 @@ -theory Ancient_Numeral - imports Main Bits_Int Misc_lsb Misc_msb -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 -r d8661799afb2 -r c7038c397ae3 src/HOL/Word/Bit_Comprehension.thy --- a/src/HOL/Word/Bit_Comprehension.thy Thu Oct 29 09:59:40 2020 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,248 +0,0 @@ -(* Title: HOL/Word/Bit_Comprehension.thy - Author: Brian Huffman, PSU; Jeremy Dawson and Gerwin Klein, NICTA -*) - -section \Comprehension syntax for bit expressions\ - -theory Bit_Comprehension - imports Word -begin - -class bit_comprehension = ring_bit_operations + - fixes set_bits :: \(nat \ bool) \ 'a\ (binder \BITS \ 10) - assumes set_bits_bit_eq: \set_bits (bit a) = a\ -begin - -lemma set_bits_False_eq [simp]: - \(BITS _. False) = 0\ - using set_bits_bit_eq [of 0] by (simp add: bot_fun_def) - -end - -instantiation int :: bit_comprehension -begin - -definition - \set_bits f = ( - if \n. \m\n. f m = f n then - let n = LEAST n. \m\n. f m = f n - in signed_take_bit n (horner_sum of_bool 2 (map f [0.. - -instance proof - fix k :: int - from int_bit_bound [of k] - obtain n where *: \\m. n \ m \ bit k m \ bit k n\ - and **: \n > 0 \ bit k (n - 1) \ bit k n\ - by blast - then have ***: \\n. \n'\n. bit k n' \ bit k n\ - by meson - have l: \(LEAST q. \m\q. bit k m \ bit k q) = n\ - apply (rule Least_equality) - using * apply blast - apply (metis "**" One_nat_def Suc_pred le_cases le0 neq0_conv not_less_eq_eq) - done - show \set_bits (bit k) = k\ - apply (simp only: *** set_bits_int_def horner_sum_bit_eq_take_bit l) - apply simp - apply (rule bit_eqI) - apply (simp add: bit_signed_take_bit_iff min_def) - apply (auto simp add: not_le bit_take_bit_iff dest: *) - done -qed - -end - -lemma int_set_bits_K_False [simp]: "(BITS _. False) = (0 :: int)" - by (simp add: set_bits_int_def) - -lemma int_set_bits_K_True [simp]: "(BITS _. True) = (-1 :: int)" - by (simp add: set_bits_int_def) - -instantiation word :: (len) bit_comprehension -begin - -definition word_set_bits_def: - \(BITS n. P n) = (horner_sum of_bool 2 (map P [0.. - -instance by standard - (simp add: word_set_bits_def horner_sum_bit_eq_take_bit) - -end - -lemma bit_set_bits_word_iff: - \bit (set_bits P :: 'a::len word) n \ n < LENGTH('a) \ P n\ - by (auto simp add: word_set_bits_def bit_horner_sum_bit_word_iff) - -lemma set_bits_K_False [simp]: - \set_bits (\_. False) = (0 :: 'a :: len word)\ - by (rule bit_word_eqI) (simp add: bit_set_bits_word_iff) - -lemma set_bits_int_unfold': - \set_bits f = - (if \n. \n'\n. \ f n' then - let n = LEAST n. \n'\n. \ f n' - in horner_sum of_bool 2 (map f [0..n. \n'\n. f n' then - let n = LEAST n. \n'\n. f n' - in signed_take_bit n (horner_sum of_bool 2 (map f [0.. -proof (cases \\n. \m\n. f m \ f n\) - case True - then obtain q where q: \\m\q. f m \ f q\ - by blast - define n where \n = (LEAST n. \m\n. f m \ f n)\ - have \\m\n. f m \ f n\ - unfolding n_def - using q by (rule LeastI [of _ q]) - then have n: \\m. n \ m \ f m \ f n\ - by blast - from n_def have n_eq: \(LEAST q. \m\q. f m \ f n) = n\ - by (smt Least_equality Least_le \\m\n. f m = f n\ dual_order.refl le_refl n order_refl) - show ?thesis - proof (cases \f n\) - case False - with n have *: \\n. \n'\n. \ f n'\ - by blast - have **: \(LEAST n. \n'\n. \ f n') = n\ - using False n_eq by simp - from * False show ?thesis - apply (simp add: set_bits_int_def n_def [symmetric] ** del: upt.upt_Suc) - apply (auto simp add: take_bit_horner_sum_bit_eq - bit_horner_sum_bit_iff take_map - signed_take_bit_def set_bits_int_def - horner_sum_bit_eq_take_bit simp del: upt.upt_Suc) - done - next - case True - with n have *: \\n. \n'\n. f n'\ - by blast - have ***: \\ (\n. \n'\n. \ f n')\ - apply (rule ccontr) - using * nat_le_linear by auto - have **: \(LEAST n. \n'\n. f n') = n\ - using True n_eq by simp - from * *** True show ?thesis - apply (simp add: set_bits_int_def n_def [symmetric] ** del: upt.upt_Suc) - apply (auto simp add: take_bit_horner_sum_bit_eq - bit_horner_sum_bit_iff take_map - signed_take_bit_def set_bits_int_def - horner_sum_bit_eq_take_bit nth_append simp del: upt.upt_Suc) - done - qed -next - case False - then show ?thesis - by (auto simp add: set_bits_int_def) -qed - -inductive wf_set_bits_int :: "(nat \ bool) \ bool" - for f :: "nat \ bool" -where - zeros: "\n' \ n. \ f n' \ wf_set_bits_int f" -| ones: "\n' \ n. f n' \ wf_set_bits_int f" - -lemma wf_set_bits_int_simps: "wf_set_bits_int f \ (\n. (\n'\n. \ f n') \ (\n'\n. f n'))" -by(auto simp add: wf_set_bits_int.simps) - -lemma wf_set_bits_int_const [simp]: "wf_set_bits_int (\_. b)" -by(cases b)(auto intro: wf_set_bits_int.intros) - -lemma wf_set_bits_int_fun_upd [simp]: - "wf_set_bits_int (f(n := b)) \ wf_set_bits_int f" (is "?lhs \ ?rhs") -proof - assume ?lhs - then obtain n' - where "(\n''\n'. \ (f(n := b)) n'') \ (\n''\n'. (f(n := b)) n'')" - by(auto simp add: wf_set_bits_int_simps) - hence "(\n''\max (Suc n) n'. \ f n'') \ (\n''\max (Suc n) n'. f n'')" by auto - thus ?rhs by(auto simp only: wf_set_bits_int_simps) -next - assume ?rhs - then obtain n' where "(\n''\n'. \ f n'') \ (\n''\n'. f n'')" (is "?wf f n'") - by(auto simp add: wf_set_bits_int_simps) - hence "?wf (f(n := b)) (max (Suc n) n')" by auto - thus ?lhs by(auto simp only: wf_set_bits_int_simps) -qed - -lemma wf_set_bits_int_Suc [simp]: - "wf_set_bits_int (\n. f (Suc n)) \ wf_set_bits_int f" (is "?lhs \ ?rhs") -by(auto simp add: wf_set_bits_int_simps intro: le_SucI dest: Suc_le_D) - -context - fixes f - assumes wff: "wf_set_bits_int f" -begin - -lemma int_set_bits_unfold_BIT: - "set_bits f = of_bool (f 0) + (2 :: int) * set_bits (f \ Suc)" -using wff proof cases - case (zeros n) - show ?thesis - proof(cases "\n. \ f n") - case True - hence "f = (\_. False)" by auto - thus ?thesis using True by(simp add: o_def) - next - case False - then obtain n' where "f n'" by blast - with zeros have "(LEAST n. \n'\n. \ f n') = Suc (LEAST n. \n'\Suc n. \ f n')" - by(auto intro: Least_Suc) - also have "(\n. \n'\Suc n. \ f n') = (\n. \n'\n. \ f (Suc n'))" by(auto dest: Suc_le_D) - also from zeros have "\n'\n. \ f (Suc n')" by auto - ultimately show ?thesis using zeros - apply (simp (no_asm_simp) add: set_bits_int_unfold' exI - del: upt.upt_Suc flip: map_map split del: if_split) - apply (simp only: map_Suc_upt upt_conv_Cons) - apply simp - done - qed -next - case (ones n) - show ?thesis - proof(cases "\n. f n") - case True - hence "f = (\_. True)" by auto - thus ?thesis using True by(simp add: o_def) - next - case False - then obtain n' where "\ f n'" by blast - with ones have "(LEAST n. \n'\n. f n') = Suc (LEAST n. \n'\Suc n. f n')" - by(auto intro: Least_Suc) - also have "(\n. \n'\Suc n. f n') = (\n. \n'\n. f (Suc n'))" by(auto dest: Suc_le_D) - also from ones have "\n'\n. f (Suc n')" by auto - moreover from ones have "(\n. \n'\n. \ f n') = False" - by(auto intro!: exI[where x="max n m" for n m] simp add: max_def split: if_split_asm) - moreover hence "(\n. \n'\n. \ f (Suc n')) = False" - by(auto elim: allE[where x="Suc n" for n] dest: Suc_le_D) - ultimately show ?thesis using ones - apply (simp (no_asm_simp) add: set_bits_int_unfold' exI split del: if_split) - apply (auto simp add: Let_def hd_map map_tl[symmetric] map_map[symmetric] map_Suc_upt upt_conv_Cons signed_take_bit_Suc - not_le simp del: map_map) - done - qed -qed - -lemma bin_last_set_bits [simp]: - "odd (set_bits f :: int) = f 0" - by (subst int_set_bits_unfold_BIT) simp_all - -lemma bin_rest_set_bits [simp]: - "set_bits f div (2 :: int) = set_bits (f \ Suc)" - by (subst int_set_bits_unfold_BIT) simp_all - -lemma bin_nth_set_bits [simp]: - "bit (set_bits f :: int) m \ f m" -using wff proof (induction m arbitrary: f) - case 0 - then show ?case - by (simp add: Bit_Comprehension.bin_last_set_bits) -next - case Suc - from Suc.IH [of "f \ Suc"] Suc.prems show ?case - by (simp add: Bit_Comprehension.bin_rest_set_bits comp_def bit_Suc) -qed - -end - -end diff -r d8661799afb2 -r c7038c397ae3 src/HOL/Word/Misc_Arithmetic.thy --- a/src/HOL/Word/Misc_Arithmetic.thy Thu Oct 29 09:59:40 2020 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,462 +0,0 @@ -(* Title: HOL/Word/Misc_Arithmetic.thy *) - -section \Miscellaneous lemmas, mostly for arithmetic\ - -theory Misc_Arithmetic - imports Main Bits_Int -begin - -lemma int_mod_lem: "0 < n \ 0 \ b \ b < n \ b mod n = b" - for b n :: int - apply safe - apply (erule (1) mod_pos_pos_trivial) - apply (erule_tac [!] subst) - apply auto - done - -lemma int_mod_ge': "b < 0 \ 0 < n \ b + n \ b mod n" - for b n :: int - by (metis add_less_same_cancel2 int_mod_ge mod_add_self2) - -lemma int_mod_le': "0 \ b - n \ b mod n \ b - n" - for b n :: int - by (metis minus_mod_self2 zmod_le_nonneg_dividend) - -lemma emep1: "even n \ even d \ 0 \ d \ (n + 1) mod d = (n mod d) + 1" - for n d :: int - by (auto simp add: pos_zmod_mult_2 add.commute dvd_def) - -lemma m1mod2k: "- 1 mod 2 ^ n = (2 ^ n - 1 :: int)" - by (rule zmod_minus1) simp - -lemma sb_inc_lem: "a + 2^k < 0 \ a + 2^k + 2^(Suc k) \ (a + 2^k) mod 2^(Suc k)" - for a :: int - using int_mod_ge' [where n = "2 ^ (Suc k)" and b = "a + 2 ^ k"] - by simp - -lemma sb_inc_lem': "a < - (2^k) \ a + 2^k + 2^(Suc k) \ (a + 2^k) mod 2^(Suc k)" - for a :: int - by (rule sb_inc_lem) simp - -lemma sb_dec_lem: "0 \ - (2 ^ k) + a \ (a + 2 ^ k) mod (2 * 2 ^ k) \ - (2 ^ k) + a" - for a :: int - using int_mod_le'[where n = "2 ^ (Suc k)" and b = "a + 2 ^ k"] by simp - -lemma sb_dec_lem': "2 ^ k \ a \ (a + 2 ^ k) mod (2 * 2 ^ k) \ - (2 ^ k) + a" - for a :: int - by (rule sb_dec_lem) simp - -lemma one_mod_exp_eq_one [simp]: - "1 mod (2 * 2 ^ n) = (1::int)" - using power_gt1 [of 2 n] by (auto intro: mod_pos_pos_trivial) - -lemma mod_2_neq_1_eq_eq_0: "k mod 2 \ 1 \ k mod 2 = 0" - for k :: int - by (fact not_mod_2_eq_1_eq_0) - -lemma z1pmod2: "(2 * b + 1) mod 2 = (1::int)" - for b :: int - by arith - -lemma diff_le_eq': "a - b \ c \ a \ b + c" - for a b c :: int - by arith - -lemma zless2: "0 < (2 :: int)" - by (fact zero_less_numeral) - -lemma zless2p: "0 < (2 ^ n :: int)" - by arith - -lemma zle2p: "0 \ (2 ^ n :: int)" - by arith - -lemma p1mod22k': "(1 + 2 * b) mod (2 * 2 ^ n) = 1 + 2 * (b mod 2 ^ n)" - for b :: int - using zle2p by (rule pos_zmod_mult_2) - -lemma p1mod22k: "(2 * b + 1) mod (2 * 2 ^ n) = 2 * (b mod 2 ^ n) + 1" - for b :: int - by (simp add: p1mod22k' add.commute) - -lemma ex_eq_or: "(\m. n = Suc m \ (m = k \ P m)) \ n = Suc k \ (\m. n = Suc m \ P m)" - by auto - -lemma power_minus_simp: "0 < n \ a ^ n = a * a ^ (n - 1)" - by (auto dest: gr0_implies_Suc) - -lemma funpow_minus_simp: "0 < n \ f ^^ n = f \ f ^^ (n - 1)" - by (auto dest: gr0_implies_Suc) - -lemma power_numeral: "a ^ numeral k = a * a ^ (pred_numeral k)" - by (simp add: numeral_eq_Suc) - -lemma funpow_numeral [simp]: "f ^^ numeral k = f \ f ^^ (pred_numeral k)" - by (simp add: numeral_eq_Suc) - -lemma replicate_numeral [simp]: "replicate (numeral k) x = x # replicate (pred_numeral k) x" - by (simp add: numeral_eq_Suc) - -lemma rco_alt: "(f \ g) ^^ n \ f = f \ (g \ f) ^^ n" - apply (rule ext) - apply (induct n) - apply (simp_all add: o_def) - done - -lemma list_exhaust_size_gt0: - assumes "\a list. y = a # list \ P" - shows "0 < length y \ P" - apply (cases y) - apply simp - apply (rule assms) - apply fastforce - done - -lemma list_exhaust_size_eq0: - assumes "y = [] \ P" - shows "length y = 0 \ P" - apply (cases y) - apply (rule assms) - apply simp - apply simp - done - -lemma size_Cons_lem_eq: "y = xa # list \ size y = Suc k \ size list = k" - by auto - -lemmas ls_splits = prod.split prod.split_asm if_split_asm - -\ \simplifications for specific word lengths\ -lemmas n2s_ths [THEN eq_reflection] = add_2_eq_Suc add_2_eq_Suc' - -lemmas s2n_ths = n2s_ths [symmetric] - -lemma and_len: "xs = ys \ xs = ys \ length xs = length ys" - by auto - -lemma size_if: "size (if p then xs else ys) = (if p then size xs else size ys)" - by auto - -lemma tl_if: "tl (if p then xs else ys) = (if p then tl xs else tl ys)" - by auto - -lemma hd_if: "hd (if p then xs else ys) = (if p then hd xs else hd ys)" - by auto - -lemma if_Not_x: "(if p then \ x else x) = (p = (\ x))" - by auto - -lemma if_x_Not: "(if p then x else \ x) = (p = x)" - by auto - -lemma if_same_and: "(If p x y \ If p u v) = (if p then x \ u else y \ v)" - by auto - -lemma if_same_eq: "(If p x y = (If p u v)) = (if p then x = u else y = v)" - by auto - -lemma if_same_eq_not: "(If p x y = (\ If p u v)) = (if p then x = (\ u) else y = (\ v))" - by auto - -\ \note -- \if_Cons\ can cause blowup in the size, if \p\ is complex, so make a simproc\ -lemma if_Cons: "(if p then x # xs else y # ys) = If p x y # If p xs ys" - by auto - -lemma if_single: "(if xc then [xab] else [an]) = [if xc then xab else an]" - by auto - -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 - -lemmas if_simps = - if_x_Not if_Not_x if_cancel if_True if_False if_bool_simps - -lemmas seqr = eq_reflection [where x = "size w"] for w (* FIXME: delete *) - -lemma the_elemI: "y = {x} \ the_elem y = x" - by simp - -lemma nonemptyE: "S \ {} \ (\x. x \ S \ R) \ R" - by auto - -lemma gt_or_eq_0: "0 < y \ 0 = y" - for y :: nat - by arith - -lemmas xtr1 = xtrans(1) -lemmas xtr2 = xtrans(2) -lemmas xtr3 = xtrans(3) -lemmas xtr4 = xtrans(4) -lemmas xtr5 = xtrans(5) -lemmas xtr6 = xtrans(6) -lemmas xtr7 = xtrans(7) -lemmas xtr8 = xtrans(8) - -lemmas nat_simps = diff_add_inverse2 diff_add_inverse -lemmas nat_iffs = le_add1 le_add2 - -lemma sum_imp_diff: "j = k + i \ j - i = k" - for k :: nat - by arith - -lemmas pos_mod_sign2 = zless2 [THEN pos_mod_sign [where b = "2::int"]] -lemmas pos_mod_bound2 = zless2 [THEN pos_mod_bound [where b = "2::int"]] - -lemma nmod2: "n mod 2 = 0 \ n mod 2 = 1" - for n :: int - by arith - -lemma eme1p: - "even n \ even d \ 0 \ d \ (1 + n) mod d = 1 + n mod d" for n d :: int - using emep1 [of n d] by (simp add: ac_simps) - -lemma le_diff_eq': "a \ c - b \ b + a \ c" - for a b c :: int - by arith - -lemma less_diff_eq': "a < c - b \ b + a < c" - for a b c :: int - by arith - -lemma diff_less_eq': "a - b < c \ a < b + c" - for a b c :: int - by arith - -lemmas m1mod22k = mult_pos_pos [OF zless2 zless2p, THEN zmod_minus1] - -lemma z1pdiv2: "(2 * b + 1) div 2 = b" - for b :: int - by arith - -lemmas zdiv_le_dividend = xtr3 [OF div_by_1 [symmetric] zdiv_mono2, - simplified int_one_le_iff_zero_less, simplified] - -lemma axxbyy: "a + m + m = b + n + n \ a = 0 \ a = 1 \ b = 0 \ b = 1 \ a = b \ m = n" - for a b m n :: int - by arith - -lemma axxmod2: "(1 + x + x) mod 2 = 1 \ (0 + x + x) mod 2 = 0" - for x :: int - by arith - -lemma axxdiv2: "(1 + x + x) div 2 = x \ (0 + x + x) div 2 = x" - for x :: int - by arith - -lemmas iszero_minus = - trans [THEN trans, OF iszero_def neg_equal_0_iff_equal iszero_def [symmetric]] - -lemmas zadd_diff_inverse = - trans [OF diff_add_cancel [symmetric] add.commute] - -lemmas add_diff_cancel2 = - add.commute [THEN diff_eq_eq [THEN iffD2]] - -lemmas rdmods [symmetric] = mod_minus_eq - mod_diff_left_eq mod_diff_right_eq mod_add_left_eq - mod_add_right_eq mod_mult_right_eq mod_mult_left_eq - -lemma mod_plus_right: "(a + x) mod m = (b + x) mod m \ a mod m = b mod m" - for a b m x :: nat - by (induct x) (simp_all add: mod_Suc, arith) - -lemma nat_minus_mod: "(n - n mod m) mod m = 0" - for m n :: nat - by (induct n) (simp_all add: mod_Suc) - -lemmas nat_minus_mod_plus_right = - trans [OF nat_minus_mod mod_0 [symmetric], - THEN mod_plus_right [THEN iffD2], simplified] - -lemmas push_mods' = mod_add_eq - mod_mult_eq mod_diff_eq - mod_minus_eq - -lemmas push_mods = push_mods' [THEN eq_reflection] -lemmas pull_mods = push_mods [symmetric] rdmods [THEN eq_reflection] - -lemma nat_mod_eq: "b < n \ a mod n = b mod n \ a mod n = b" - for a b n :: nat - by (induct a) auto - -lemmas nat_mod_eq' = refl [THEN [2] nat_mod_eq] - -lemma nat_mod_lem: "0 < n \ b < n \ b mod n = b" - for b n :: nat - apply safe - apply (erule nat_mod_eq') - apply (erule subst) - apply (erule mod_less_divisor) - done - -lemma mod_nat_add: "x < z \ y < z \ (x + y) mod z = (if x + y < z then x + y else x + y - z)" - for x y z :: nat - apply (rule nat_mod_eq) - apply auto - apply (rule trans) - apply (rule le_mod_geq) - apply simp - apply (rule nat_mod_eq') - apply arith - done - -lemma mod_nat_sub: "x < z \ (x - y) mod z = x - y" - for x y :: nat - by (rule nat_mod_eq') arith - -lemma int_mod_eq: "0 \ b \ b < n \ a mod n = b mod n \ a mod n = b" - for a b n :: int - by (metis mod_pos_pos_trivial) - -lemmas int_mod_eq' = mod_pos_pos_trivial (* FIXME delete *) - -lemmas int_mod_le = zmod_le_nonneg_dividend (* FIXME: delete *) - -lemma mod_add_if_z: - "x < z \ y < z \ 0 \ y \ 0 \ x \ 0 \ z \ - (x + y) mod z = (if x + y < z then x + y else x + y - z)" - for x y z :: int - by (auto intro: int_mod_eq) - -lemma mod_sub_if_z: - "x < z \ y < z \ 0 \ y \ 0 \ x \ 0 \ z \ - (x - y) mod z = (if y \ x then x - y else x - y + z)" - for x y z :: int - by (auto intro: int_mod_eq) - -lemmas zmde = mult_div_mod_eq [symmetric, THEN diff_eq_eq [THEN iffD2], symmetric] -lemmas mcl = mult_cancel_left [THEN iffD1, THEN make_pos_rule] - -(* already have this for naturals, div_mult_self1/2, but not for ints *) -lemma zdiv_mult_self: "m \ 0 \ (a + m * n) div m = a div m + n" - for a m n :: int - apply (rule mcl) - prefer 2 - apply (erule asm_rl) - apply (simp add: zmde ring_distribs) - done - -lemma mod_power_lem: "a > 1 \ a ^ n mod a ^ m = (if m \ n then 0 else a ^ n)" - for a :: int - by (simp add: mod_eq_0_iff_dvd le_imp_power_dvd) - -lemma pl_pl_rels: "a + b = c + d \ a \ c \ b \ d \ a \ c \ b \ d" - for a b c d :: nat - by arith - -lemmas pl_pl_rels' = add.commute [THEN [2] trans, THEN pl_pl_rels] - -lemma minus_eq: "m - k = m \ k = 0 \ m = 0" - for k m :: nat - by arith - -lemma pl_pl_mm: "a + b = c + d \ a - c = d - b" - for a b c d :: nat - by arith - -lemmas pl_pl_mm' = add.commute [THEN [2] trans, THEN pl_pl_mm] - -lemmas dme = div_mult_mod_eq -lemmas dtle = div_times_less_eq_dividend -lemmas th2 = order_trans [OF order_refl [THEN [2] mult_le_mono] div_times_less_eq_dividend] - -lemma td_gal: "0 < c \ a \ b * c \ a div c \ b" - for a b c :: nat - apply safe - apply (erule (1) xtr4 [OF div_le_mono div_mult_self_is_m]) - apply (erule th2) - done - -lemmas td_gal_lt = td_gal [simplified not_less [symmetric], simplified] - -lemmas div_mult_le = div_times_less_eq_dividend - -lemmas sdl = div_nat_eqI - -lemma given_quot: "f > 0 \ (f * l + (f - 1)) div f = l" - for f l :: nat - by (rule div_nat_eqI) (simp_all) - -lemma given_quot_alt: "f > 0 \ (l * f + f - Suc 0) div f = l" - for f l :: nat - apply (frule given_quot) - apply (rule trans) - prefer 2 - apply (erule asm_rl) - apply (rule_tac f="\n. n div f" in arg_cong) - apply (simp add : ac_simps) - done - -lemma diff_mod_le: "a < d \ b dvd d \ a - a mod b \ d - b" - for a b d :: nat - apply (unfold dvd_def) - apply clarify - apply (case_tac k) - apply clarsimp - apply clarify - apply (cases "b > 0") - apply (drule mult.commute [THEN xtr1]) - apply (frule (1) td_gal_lt [THEN iffD1]) - apply (clarsimp simp: le_simps) - apply (rule minus_mod_eq_mult_div [symmetric, THEN [2] xtr4]) - apply (rule mult_mono) - apply auto - done - -lemma less_le_mult': "w * c < b * c \ 0 \ c \ (w + 1) * c \ b * c" - for b c w :: int - apply (rule mult_right_mono) - apply (rule zless_imp_add1_zle) - apply (erule (1) mult_right_less_imp_less) - apply assumption - done - -lemma less_le_mult: "w * c < b * c \ 0 \ c \ w * c + c \ b * c" - for b c w :: int - using less_le_mult' [of w c b] by (simp add: algebra_simps) - -lemmas less_le_mult_minus = iffD2 [OF le_diff_eq less_le_mult, - simplified left_diff_distrib] - -lemma gen_minus: "0 < n \ f n = f (Suc (n - 1))" - by auto - -lemma mpl_lem: "j \ i \ k < j \ i - j + k < i" - for i j k :: nat - by arith - -lemma nonneg_mod_div: "0 \ a \ 0 \ b \ 0 \ (a mod b) \ 0 \ a div b" - for a b :: int - by (cases "b = 0") (auto intro: pos_imp_zdiv_nonneg_iff [THEN iffD2]) - -declare iszero_0 [intro] - -lemma min_pm [simp]: "min a b + (a - b) = a" - for a b :: nat - by arith - -lemma min_pm1 [simp]: "a - b + min a b = a" - for a b :: nat - by arith - -lemma rev_min_pm [simp]: "min b a + (a - b) = a" - for a b :: nat - by arith - -lemma rev_min_pm1 [simp]: "a - b + min b a = a" - for a b :: nat - by arith - -lemma min_minus [simp]: "min m (m - k) = m - k" - for m k :: nat - by arith - -lemma min_minus' [simp]: "min (m - k) m = m - k" - for m k :: nat - by arith - -lemmas m2pths = pos_mod_sign pos_mod_bound [OF zless2p] - -end diff -r d8661799afb2 -r c7038c397ae3 src/HOL/Word/Misc_Auxiliary.thy --- a/src/HOL/Word/Misc_Auxiliary.thy Thu Oct 29 09:59:40 2020 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,25 +0,0 @@ -(* Title: HOL/Word/Misc_Auxiliary.thy - Author: Jeremy Dawson, NICTA -*) - -section \Generic auxiliary\ - -theory Misc_Auxiliary - imports Main -begin - -subsection \Lemmas on list operations\ - -lemma butlast_power: "(butlast ^^ n) bl = take (length bl - n) bl" - by (induct n) (auto simp: butlast_take) - -lemma nth_rev: "n < length xs \ rev xs ! n = xs ! (length xs - 1 - n)" - using rev_nth by simp - -lemma nth_rev_alt: "n < length ys \ ys ! n = rev ys ! (length ys - Suc n)" - by (simp add: nth_rev) - -lemma hd_butlast: "length xs > 1 \ hd (butlast xs) = hd xs" - by (cases xs) auto - -end diff -r d8661799afb2 -r c7038c397ae3 src/HOL/Word/Misc_Typedef.thy --- a/src/HOL/Word/Misc_Typedef.thy Thu Oct 29 09:59:40 2020 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,356 +0,0 @@ -(* - Author: Jeremy Dawson and Gerwin Klein, NICTA - - Consequences of type definition theorems, and of extended type definition. -*) - -section \Type Definition Theorems\ - -theory Misc_Typedef - imports Main Word Bit_Comprehension Bits_Int -begin - -subsection "More lemmas about normal type definitions" - -lemma tdD1: "type_definition Rep Abs A \ \x. Rep x \ A" - and tdD2: "type_definition Rep Abs A \ \x. Abs (Rep x) = x" - and tdD3: "type_definition Rep Abs A \ \y. y \ A \ Rep (Abs y) = y" - by (auto simp: type_definition_def) - -lemma td_nat_int: "type_definition int nat (Collect ((\) 0))" - unfolding type_definition_def by auto - -context type_definition -begin - -declare Rep [iff] Rep_inverse [simp] Rep_inject [simp] - -lemma Abs_eqD: "Abs x = Abs y \ x \ A \ y \ A \ x = y" - by (simp add: Abs_inject) - -lemma Abs_inverse': "r \ A \ Abs r = a \ Rep a = r" - by (safe elim!: Abs_inverse) - -lemma Rep_comp_inverse: "Rep \ f = g \ Abs \ g = f" - using Rep_inverse by auto - -lemma Rep_eqD [elim!]: "Rep x = Rep y \ x = y" - by simp - -lemma Rep_inverse': "Rep a = r \ Abs r = a" - by (safe intro!: Rep_inverse) - -lemma comp_Abs_inverse: "f \ Abs = g \ g \ Rep = f" - using Rep_inverse by auto - -lemma set_Rep: "A = range Rep" -proof (rule set_eqI) - show "x \ A \ x \ range Rep" for x - by (auto dest: Abs_inverse [of x, symmetric]) -qed - -lemma set_Rep_Abs: "A = range (Rep \ Abs)" -proof (rule set_eqI) - show "x \ A \ x \ range (Rep \ Abs)" for x - by (auto dest: Abs_inverse [of x, symmetric]) -qed - -lemma Abs_inj_on: "inj_on Abs A" - unfolding inj_on_def - by (auto dest: Abs_inject [THEN iffD1]) - -lemma image: "Abs ` A = UNIV" - by (fact Abs_image) - -lemmas td_thm = type_definition_axioms - -lemma fns1: "Rep \ fa = fr \ Rep \ fa \ Abs = Abs \ fr \ Abs \ fr \ Rep = fa" - by (auto dest: Rep_comp_inverse elim: comp_Abs_inverse simp: o_assoc) - -lemmas fns1a = disjI1 [THEN fns1] -lemmas fns1b = disjI2 [THEN fns1] - -lemma fns4: "Rep \ fa \ Abs = fr \ Rep \ fa = fr \ Rep \ fa \ Abs = Abs \ fr" - by auto - -end - -interpretation nat_int: type_definition int nat "Collect ((\) 0)" - by (rule td_nat_int) - -declare - nat_int.Rep_cases [cases del] - nat_int.Abs_cases [cases del] - nat_int.Rep_induct [induct del] - nat_int.Abs_induct [induct del] - - -subsection "Extended form of type definition predicate" - -lemma td_conds: - "norm \ norm = norm \ - fr \ norm = norm \ fr \ norm \ fr \ norm = fr \ norm \ norm \ fr \ norm = norm \ fr" - apply safe - apply (simp_all add: comp_assoc) - apply (simp_all add: o_assoc) - done - -lemma fn_comm_power: "fa \ tr = tr \ fr \ fa ^^ n \ tr = tr \ fr ^^ n" - apply (rule ext) - apply (induct n) - apply (auto dest: fun_cong) - done - -lemmas fn_comm_power' = - ext [THEN fn_comm_power, THEN fun_cong, unfolded o_def] - - -locale td_ext = type_definition + - fixes norm - assumes eq_norm: "\x. Rep (Abs x) = norm x" -begin - -lemma Abs_norm [simp]: "Abs (norm x) = Abs x" - using eq_norm [of x] by (auto elim: Rep_inverse') - -lemma td_th: "g \ Abs = f \ f (Rep x) = g x" - by (drule comp_Abs_inverse [symmetric]) simp - -lemma eq_norm': "Rep \ Abs = norm" - by (auto simp: eq_norm) - -lemma norm_Rep [simp]: "norm (Rep x) = Rep x" - by (auto simp: eq_norm' intro: td_th) - -lemmas td = td_thm - -lemma set_iff_norm: "w \ A \ w = norm w" - by (auto simp: set_Rep_Abs eq_norm' eq_norm [symmetric]) - -lemma inverse_norm: "Abs n = w \ Rep w = norm n" - apply (rule iffI) - apply (clarsimp simp add: eq_norm) - apply (simp add: eq_norm' [symmetric]) - done - -lemma norm_eq_iff: "norm x = norm y \ Abs x = Abs y" - by (simp add: eq_norm' [symmetric]) - -lemma norm_comps: - "Abs \ norm = Abs" - "norm \ Rep = Rep" - "norm \ norm = norm" - by (auto simp: eq_norm' [symmetric] o_def) - -lemmas norm_norm [simp] = norm_comps - -lemma fns5: "Rep \ fa \ Abs = fr \ fr \ norm = fr \ norm \ fr = fr" - by (fold eq_norm') auto - -text \ - following give conditions for converses to \td_fns1\ - \<^item> the condition \norm \ fr \ norm = fr \ norm\ says that - \fr\ takes normalised arguments to normalised results - \<^item> \norm \ fr \ norm = norm \ fr\ says that \fr\ - takes norm-equivalent arguments to norm-equivalent results - \<^item> \fr \ norm = fr\ says that \fr\ - takes norm-equivalent arguments to the same result - \<^item> \norm \ fr = fr\ says that \fr\ takes any argument to a normalised result -\ -lemma fns2: "Abs \ fr \ Rep = fa \ norm \ fr \ norm = fr \ norm \ Rep \ fa = fr \ Rep" - apply (fold eq_norm') - apply safe - prefer 2 - apply (simp add: o_assoc) - apply (rule ext) - apply (drule_tac x="Rep x" in fun_cong) - apply auto - done - -lemma fns3: "Abs \ fr \ Rep = fa \ norm \ fr \ norm = norm \ fr \ fa \ Abs = Abs \ fr" - apply (fold eq_norm') - apply safe - prefer 2 - apply (simp add: comp_assoc) - apply (rule ext) - apply (drule_tac f="a \ b" for a b in fun_cong) - apply simp - done - -lemma fns: "fr \ norm = norm \ fr \ fa \ Abs = Abs \ fr \ Rep \ fa = fr \ Rep" - apply safe - apply (frule fns1b) - prefer 2 - apply (frule fns1a) - apply (rule fns3 [THEN iffD1]) - prefer 3 - apply (rule fns2 [THEN iffD1]) - apply (simp_all add: comp_assoc) - apply (simp_all add: o_assoc) - done - -lemma range_norm: "range (Rep \ Abs) = A" - by (simp add: set_Rep_Abs) - -end - -lemmas td_ext_def' = - td_ext_def [unfolded type_definition_def td_ext_axioms_def] - - -subsection \Type-definition locale instantiations\ - -definition uints :: "nat \ int set" - \ \the sets of integers representing the words\ - where "uints n = range (take_bit n)" - -definition sints :: "nat \ int set" - where "sints n = range (signed_take_bit (n - 1))" - -lemma uints_num: "uints n = {i. 0 \ i \ i < 2 ^ n}" - by (simp add: uints_def range_bintrunc) - -lemma sints_num: "sints n = {i. - (2 ^ (n - 1)) \ i \ i < 2 ^ (n - 1)}" - by (simp add: sints_def range_sbintrunc) - -definition unats :: "nat \ nat set" - where "unats n = {i. i < 2 ^ n}" - -\ \naturals\ -lemma uints_unats: "uints n = int ` unats n" - apply (unfold unats_def uints_num) - apply safe - apply (rule_tac image_eqI) - apply (erule_tac nat_0_le [symmetric]) - by auto - -lemma unats_uints: "unats n = nat ` uints n" - by (auto simp: uints_unats image_iff) - -lemma td_ext_uint: - "td_ext (uint :: 'a word \ int) word_of_int (uints (LENGTH('a::len))) - (\w::int. w mod 2 ^ LENGTH('a))" - apply (unfold td_ext_def') - apply transfer - apply (simp add: uints_num take_bit_eq_mod) - done - -interpretation word_uint: - td_ext - "uint::'a::len word \ int" - word_of_int - "uints (LENGTH('a::len))" - "\w. w mod 2 ^ LENGTH('a::len)" - by (fact td_ext_uint) - -lemmas td_uint = word_uint.td_thm -lemmas int_word_uint = word_uint.eq_norm - -lemma td_ext_ubin: - "td_ext (uint :: 'a word \ int) word_of_int (uints (LENGTH('a::len))) - (take_bit (LENGTH('a)))" - apply standard - apply transfer - apply simp - done - -interpretation word_ubin: - td_ext - "uint::'a::len word \ int" - word_of_int - "uints (LENGTH('a::len))" - "take_bit (LENGTH('a::len))" - by (fact td_ext_ubin) - -lemma td_ext_unat [OF refl]: - "n = LENGTH('a::len) \ - td_ext (unat :: 'a word \ nat) of_nat (unats n) (\i. i mod 2 ^ n)" - apply (standard; transfer) - apply (simp_all add: unats_def take_bit_of_nat take_bit_nat_eq_self_iff - flip: take_bit_eq_mod) - done - -lemmas unat_of_nat = td_ext_unat [THEN td_ext.eq_norm] - -interpretation word_unat: - td_ext - "unat::'a::len word \ nat" - of_nat - "unats (LENGTH('a::len))" - "\i. i mod 2 ^ LENGTH('a::len)" - by (rule td_ext_unat) - -lemmas td_unat = word_unat.td_thm - -lemmas unat_lt2p [iff] = word_unat.Rep [unfolded unats_def mem_Collect_eq] - -lemma unat_le: "y \ unat z \ y \ unats (LENGTH('a))" - for z :: "'a::len word" - apply (unfold unats_def) - apply clarsimp - apply (rule xtrans, rule unat_lt2p, assumption) - done - -lemma td_ext_sbin: - "td_ext (sint :: 'a word \ int) word_of_int (sints (LENGTH('a::len))) - (signed_take_bit (LENGTH('a) - 1))" - by (standard; transfer) (auto simp add: sints_def) - -lemma td_ext_sint: - "td_ext (sint :: 'a word \ int) word_of_int (sints (LENGTH('a::len))) - (\w. (w + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - - 2 ^ (LENGTH('a) - 1))" - using td_ext_sbin [where ?'a = 'a] by (simp add: no_sbintr_alt2) - -text \ - We do \sint\ before \sbin\, before \sint\ is the user version - and interpretations do not produce thm duplicates. I.e. - we get the name \word_sint.Rep_eqD\, but not \word_sbin.Req_eqD\, - because the latter is the same thm as the former. -\ -interpretation word_sint: - td_ext - "sint ::'a::len word \ int" - word_of_int - "sints (LENGTH('a::len))" - "\w. (w + 2^(LENGTH('a::len) - 1)) mod 2^LENGTH('a::len) - - 2 ^ (LENGTH('a::len) - 1)" - by (rule td_ext_sint) - -interpretation word_sbin: - td_ext - "sint ::'a::len word \ int" - word_of_int - "sints (LENGTH('a::len))" - "signed_take_bit (LENGTH('a::len) - 1)" - by (rule td_ext_sbin) - -lemmas int_word_sint = td_ext_sint [THEN td_ext.eq_norm] - -lemmas td_sint = word_sint.td - -lemma uints_mod: "uints n = range (\w. w mod 2 ^ n)" - by (fact uints_def [unfolded no_bintr_alt1]) - -lemmas uint_range' = word_uint.Rep [unfolded uints_num mem_Collect_eq] -lemmas sint_range' = word_sint.Rep [unfolded One_nat_def sints_num mem_Collect_eq] - -lemmas bintr_num = - word_ubin.norm_eq_iff [of "numeral a" "numeral b", symmetric, folded word_numeral_alt] for a b -lemmas sbintr_num = - word_sbin.norm_eq_iff [of "numeral a" "numeral b", symmetric, folded word_numeral_alt] for a b - -lemmas uint_div_alt = word_div_def [THEN trans [OF uint_cong int_word_uint]] -lemmas uint_mod_alt = word_mod_def [THEN trans [OF uint_cong int_word_uint]] - -interpretation test_bit: - td_ext - "(!!) :: 'a::len word \ nat \ bool" - set_bits - "{f. \i. f i \ i < LENGTH('a::len)}" - "(\h i. h i \ i < LENGTH('a::len))" - by standard (auto simp add: test_bit_word_eq bit_imp_le_length bit_set_bits_word_iff set_bits_bit_eq) - -lemmas td_nth = test_bit.td_thm - -end diff -r d8661799afb2 -r c7038c397ae3 src/HOL/Word/Misc_lsb.thy --- a/src/HOL/Word/Misc_lsb.thy Thu Oct 29 09:59:40 2020 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,89 +0,0 @@ -(* Author: Jeremy Dawson, NICTA -*) - -section \Operation variant for the least significant bit\ - -theory Misc_lsb - imports - Word - Reversed_Bit_Lists -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" - 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_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 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)" - 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)" - by (simp add: word_lsb_alt) - -end diff -r d8661799afb2 -r c7038c397ae3 src/HOL/Word/Misc_msb.thy --- a/src/HOL/Word/Misc_msb.thy Thu Oct 29 09:59:40 2020 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,137 +0,0 @@ -(* Author: Jeremy Dawson, NICTA -*) - -section \Operation variant for the most significant bit\ - -theory Misc_msb - imports - Word - Reversed_Bit_Lists -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 (bin_rest x) = msb x" -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\ - -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)" - by (simp add: word_msb_def word_sbin.eq_norm bin_sign_lem) - -lemma word_msb_numeral [simp]: - "msb (numeral w::'a::len word) = bin_nth (numeral w) (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)" - 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)" - for w :: "'a::len word" - by (simp add: word_msb_def sint_uint bin_sign_lem) - -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 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 exp_eq_zero_iff 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]] - -end diff -r d8661799afb2 -r c7038c397ae3 src/HOL/Word/Misc_set_bit.thy --- a/src/HOL/Word/Misc_set_bit.thy Thu Oct 29 09:59:40 2020 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,167 +0,0 @@ -(* Author: Jeremy Dawson, NICTA -*) - -section \Operation variant for setting and unsetting bits\ - -theory Misc_set_bit - imports Word Misc_msb -begin - -class set_bit = ring_bit_operations + - fixes set_bit :: \'a \ nat \ bool \ 'a\ - assumes set_bit_eq: \set_bit a n b = (if b then Bit_Operations.set_bit else unset_bit) n a\ - -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 add: set_bit_int_def bin_sc_eq) - -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)" - 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)" - 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 - apply standard - apply (simp add: word_set_bit_def bin_sc_eq Bit_Operations.set_bit_def unset_bit_def) - apply transfer - apply simp - done - -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 (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: set_bit_unfold bit_unset_bit_iff bit_set_bit_iff exp_eq_zero_iff not_le 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 word_ubin.eq_norm 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: word_ubin.eq_norm 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 word_ubin.eq_norm 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 (unfold word_set_bit_def word_le_def word_ubin.eq_norm) - 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 (unfold word_set_bit_def word_le_def word_ubin.eq_norm) - 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) - -end diff -r d8661799afb2 -r c7038c397ae3 src/HOL/Word/More_Word.thy --- a/src/HOL/Word/More_Word.thy Thu Oct 29 09:59:40 2020 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,58 +0,0 @@ -(* Title: HOL/Word/More_thy -*) - -section \Ancient comprehensive Word Library\ - -theory More_Word -imports - Word - Ancient_Numeral - Reversed_Bit_Lists - Bits_Int - Bit_Comprehension - Misc_Auxiliary - Misc_Arithmetic - Misc_set_bit - Misc_lsb - Misc_Typedef - Word_rsplit -begin - -declare signed_take_bit_Suc [simp] - -lemmas bshiftr1_def = bshiftr1_eq -lemmas is_down_def = is_down_eq -lemmas is_up_def = is_up_eq -lemmas mask_def = mask_eq_decr_exp -lemmas scast_def = scast_eq -lemmas shiftl1_def = shiftl1_eq -lemmas shiftr1_def = shiftr1_eq -lemmas sshiftr1_def = sshiftr1_eq -lemmas sshiftr_def = sshiftr_eq_funpow_sshiftr1 -lemmas to_bl_def = to_bl_eq -lemmas ucast_def = ucast_eq -lemmas unat_def = unat_eq_nat_uint -lemmas word_cat_def = word_cat_eq -lemmas word_reverse_def = word_reverse_eq_of_bl_rev_to_bl -lemmas word_roti_def = word_roti_eq_word_rotr_word_rotl -lemmas word_rotl_def = word_rotl_eq -lemmas word_rotr_def = word_rotr_eq -lemmas word_sle_def = word_sle_eq -lemmas word_sless_def = word_sless_eq - -lemmas uint_0 = uint_nonnegative -lemmas uint_lt = uint_bounded -lemmas uint_mod_same = uint_idem -lemmas of_nth_def = word_set_bits_def - -lemmas of_nat_word_eq_iff = word_of_nat_eq_iff -lemmas of_nat_word_eq_0_iff = word_of_nat_eq_0_iff -lemmas of_int_word_eq_iff = word_of_int_eq_iff -lemmas of_int_word_eq_0_iff = word_of_int_eq_0_iff - -lemma shiftl_transfer [transfer_rule]: - includes lifting_syntax - shows "(pcr_word ===> (=) ===> pcr_word) (<<) (<<)" - by (unfold shiftl_eq_push_bit) transfer_prover - -end diff -r d8661799afb2 -r c7038c397ae3 src/HOL/Word/Reversed_Bit_Lists.thy --- a/src/HOL/Word/Reversed_Bit_Lists.thy Thu Oct 29 09:59:40 2020 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1870 +0,0 @@ -(* Title: HOL/Word/Reversed_Bit_Lists.thy - Author: Jeremy Dawson, NICTA -*) - -section \Bit values as reversed lists of bools\ - -theory Reversed_Bit_Lists - imports Word Misc_Typedef -begin - -context comm_semiring_1 -begin - -lemma horner_sum_append: - \horner_sum f a (xs @ ys) = horner_sum f a xs + a ^ length xs * horner_sum f a ys\ - using sum.atLeastLessThan_shift_bounds [of _ 0 \length xs\ \length ys\] - atLeastLessThan_add_Un [of 0 \length xs\ \length ys\] - by (simp add: horner_sum_eq_sum sum_distrib_left sum.union_disjoint ac_simps nth_append power_add) - -end - -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)" - -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)" - 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" - 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" - by (auto intro: bl_to_bin_inj) - -lemma bin_to_bl_aux_bintr: - "bin_to_bl_aux n (bintrunc 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" - 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)" - 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)" - 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))" - 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)" - 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" - 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)" - 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)" - 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))" - 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)" - by (induct bl arbitrary: w) auto - -lemma butlast_rest_bl2bin: "bl_to_bin (butlast bl) = bin_rest (bl_to_bin bl)" - 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)" -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)" - 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" - by (simp add: trunc_bl2bin) - -lemma bl2bin_drop: "bl_to_bin (drop k bl) = bintrunc (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)" - 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)" - by (induct xs arbitrary: w) auto - -lemma last_bin_last: "size xs > 0 \ last xs \ bin_last (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)" - 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)" - 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 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)" - 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_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)" - 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" - 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)" - 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 (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 - apply (simp 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 - apply (simp add: bin_to_bl_aux_alt) - apply (simp 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 (induct xs) auto - -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)" - 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 (simp add: bin_cat_assoc_sym min.absorb2) - 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" -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)" -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 (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" - 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]: - \uint (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 (simp add: to_bl_unfold rev_map) - -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 (simp add: word_size word_ops_nth_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 (max_word::'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] - -end diff -r d8661799afb2 -r c7038c397ae3 src/HOL/Word/Tools/smt_word.ML --- a/src/HOL/Word/Tools/smt_word.ML Thu Oct 29 09:59:40 2020 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,158 +0,0 @@ -(* Title: HOL/Word/Tools/smt_word.ML - Author: Sascha Boehme, TU Muenchen - -SMT setup for words. -*) - -signature SMT_WORD = -sig - val add_word_shift': term * string -> Context.generic -> Context.generic -end - -structure SMT_Word : SMT_WORD = -struct - -open Word_Lib - - -(* SMT-LIB logic *) - -(* "QF_AUFBV" is too restrictive for Isabelle's problems, which contain aritmetic and quantifiers. - Better set the logic to "" and make at least Z3 happy. *) -fun smtlib_logic ts = - if exists (Term.exists_type (Term.exists_subtype is_wordT)) ts then SOME "" else NONE - - -(* SMT-LIB builtins *) - -local - val smtlibC = SMTLIB_Interface.smtlibC - - val wordT = \<^typ>\'a::len word\ - - fun index1 s i = "(_ " ^ s ^ " " ^ string_of_int i ^ ")" - fun index2 s i j = "(_ " ^ s ^ " " ^ string_of_int i ^ " " ^ string_of_int j ^ ")" - - fun word_typ (Type (\<^type_name>\word\, [T])) = - Option.map (rpair [] o index1 "BitVec") (try dest_binT T) - | word_typ _ = NONE - - fun word_num (Type (\<^type_name>\word\, [T])) k = - Option.map (index1 ("bv" ^ string_of_int k)) (try dest_binT T) - | word_num _ _ = NONE - - fun if_fixed pred m n T ts = - let val (Us, U) = Term.strip_type T - in - if pred (U, Us) then SOME (n, length Us, ts, Term.list_comb o pair (Const (m, T))) else NONE - end - - fun if_fixed_all m = if_fixed (forall (can dest_wordT) o (op ::)) m - fun if_fixed_args m = if_fixed (forall (can dest_wordT) o snd) m - - fun add_word_fun f (t, n) = - let val (m, _) = Term.dest_Const t - in SMT_Builtin.add_builtin_fun smtlibC (Term.dest_Const t, K (f m n)) end - - val mk_nat = HOLogic.mk_number \<^typ>\nat\ - - fun mk_shift c [u, t] = Const c $ mk_nat (snd (HOLogic.dest_number u)) $ t - | mk_shift c ts = raise TERM ("bad arguments", Const c :: ts) - - fun shift m n T ts = - let val U = Term.domain_type (Term.range_type T) - in - (case (can dest_wordT U, try (snd o HOLogic.dest_number o hd) ts) of - (true, SOME i) => - SOME (n, 2, [hd (tl ts), HOLogic.mk_number U i], mk_shift (m, T)) - | _ => NONE) (* FIXME: also support non-numerical shifts *) - end - - fun mk_shift' c [t, u] = Const c $ t $ mk_nat (snd (HOLogic.dest_number u)) - | mk_shift' c ts = raise TERM ("bad arguments", Const c :: ts) - - fun shift' m n T ts = - let val U = Term.domain_type T - in - (case (can dest_wordT U, try (snd o HOLogic.dest_number o hd o tl) ts) of - (true, SOME i) => - SOME (n, 2, [hd ts, HOLogic.mk_number U i], mk_shift' (m, T)) - | _ => NONE) (* FIXME: also support non-numerical shifts *) - end - - fun mk_extract c i ts = Term.list_comb (Const c, mk_nat i :: ts) - - fun extract m n T ts = - let val U = Term.range_type (Term.range_type T) - in - (case (try (snd o HOLogic.dest_number o hd) ts, try dest_wordT U) of - (SOME lb, SOME i) => - SOME (index2 n (i + lb - 1) lb, 1, tl ts, mk_extract (m, T) lb) - | _ => NONE) - end - - fun mk_extend c ts = Term.list_comb (Const c, ts) - - fun extend m n T ts = - let val (U1, U2) = Term.dest_funT T - in - (case (try dest_wordT U1, try dest_wordT U2) of - (SOME i, SOME j) => - if j-i >= 0 then SOME (index1 n (j-i), 1, ts, mk_extend (m, T)) - else NONE - | _ => NONE) - end - - fun mk_rotate c i ts = Term.list_comb (Const c, mk_nat i :: ts) - - fun rotate m n T ts = - let val U = Term.domain_type (Term.range_type T) - in - (case (can dest_wordT U, try (snd o HOLogic.dest_number o hd) ts) of - (true, SOME i) => SOME (index1 n i, 1, tl ts, mk_rotate (m, T) i) - | _ => NONE) - end -in - -val setup_builtins = - SMT_Builtin.add_builtin_typ smtlibC (wordT, word_typ, word_num) #> - fold (add_word_fun if_fixed_all) [ - (\<^term>\uminus :: 'a::len word \ _\, "bvneg"), - (\<^term>\plus :: 'a::len word \ _\, "bvadd"), - (\<^term>\minus :: 'a::len word \ _\, "bvsub"), - (\<^term>\times :: 'a::len word \ _\, "bvmul"), - (\<^term>\NOT :: 'a::len word \ _\, "bvnot"), - (\<^term>\(AND) :: 'a::len word \ _\, "bvand"), - (\<^term>\(OR) :: 'a::len word \ _\, "bvor"), - (\<^term>\(XOR) :: 'a::len word \ _\, "bvxor"), - (\<^term>\word_cat :: 'a::len word \ _\, "concat") ] #> - fold (add_word_fun shift) [ - (\<^term>\push_bit :: nat \ 'a::len word \ _ \, "bvshl"), - (\<^term>\drop_bit :: nat \ 'a::len word \ _\, "bvlshr"), - (\<^term>\signed_drop_bit :: nat \ 'a::len word \ _\, "bvashr") ] #> - add_word_fun extract - (\<^term>\slice :: _ \ 'a::len word \ _\, "extract") #> - fold (add_word_fun extend) [ - (\<^term>\ucast :: 'a::len word \ _\, "zero_extend"), - (\<^term>\scast :: 'a::len word \ _\, "sign_extend") ] #> - fold (add_word_fun rotate) [ - (\<^term>\word_rotl\, "rotate_left"), - (\<^term>\word_rotr\, "rotate_right") ] #> - fold (add_word_fun if_fixed_args) [ - (\<^term>\less :: 'a::len word \ _\, "bvult"), - (\<^term>\less_eq :: 'a::len word \ _\, "bvule"), - (\<^term>\word_sless\, "bvslt"), - (\<^term>\word_sle\, "bvsle") ] - -val add_word_shift' = add_word_fun shift' - -end - - -(* setup *) - -val _ = Theory.setup (Context.theory_map ( - SMTLIB_Interface.add_logic (20, smtlib_logic) #> - setup_builtins)) - -end; diff -r d8661799afb2 -r c7038c397ae3 src/HOL/Word/Tools/word_lib.ML --- a/src/HOL/Word/Tools/word_lib.ML Thu Oct 29 09:59:40 2020 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,47 +0,0 @@ -(* Title: HOL/Word/Tools/word_lib.ML - Author: Sascha Boehme, TU Muenchen and Thomas Sewell, NICTA - -Helper routines for operating on the word type. -*) - -signature WORD_LIB = -sig - val dest_binT: typ -> int - val is_wordT: typ -> bool - val dest_wordT: typ -> int - val mk_wordT: int -> typ -end - -structure Word_Lib: WORD_LIB = -struct - -fun dest_binT T = - (case T of - Type (\<^type_name>\Numeral_Type.num0\, _) => 0 - | Type (\<^type_name>\Numeral_Type.num1\, _) => 1 - | Type (\<^type_name>\Numeral_Type.bit0\, [T]) => 2 * dest_binT T - | Type (\<^type_name>\Numeral_Type.bit1\, [T]) => 1 + 2 * dest_binT T - | _ => raise TYPE ("dest_binT", [T], [])) - -fun is_wordT (Type (\<^type_name>\word\, _)) = true - | is_wordT _ = false - -fun dest_wordT (Type (\<^type_name>\word\, [T])) = dest_binT T - | dest_wordT T = raise TYPE ("dest_wordT", [T], []) - - -fun mk_bitT i T = - if i = 0 - then Type (\<^type_name>\Numeral_Type.bit0\, [T]) - else Type (\<^type_name>\Numeral_Type.bit1\, [T]) - -fun mk_binT size = - if size = 0 then \<^typ>\Numeral_Type.num0\ - else if size = 1 then \<^typ>\Numeral_Type.num1\ - else let val (q, r) = Integer.div_mod size 2 in mk_bitT r (mk_binT q) end - -fun mk_wordT size = - if size >= 0 then Type (\<^type_name>\word\, [mk_binT size]) - else raise TYPE ("mk_wordT: " ^ string_of_int size, [], []) - -end diff -r d8661799afb2 -r c7038c397ae3 src/HOL/Word/Word.thy --- a/src/HOL/Word/Word.thy Thu Oct 29 09:59:40 2020 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,4588 +0,0 @@ -(* Title: HOL/Word/Word.thy - Author: Jeremy Dawson and Gerwin Klein, NICTA -*) - -section \A type of finite bit strings\ - -theory Word -imports - "HOL-Library.Type_Length" - "HOL-Library.Boolean_Algebra" - "HOL-Library.Bit_Operations" -begin - -subsection \Preliminaries\ - -lemma signed_take_bit_decr_length_iff: - \signed_take_bit (LENGTH('a::len) - Suc 0) k = signed_take_bit (LENGTH('a) - Suc 0) l - \ take_bit LENGTH('a) k = take_bit LENGTH('a) l\ - by (cases \LENGTH('a)\) - (simp_all add: signed_take_bit_eq_iff_take_bit_eq) - - -subsection \Fundamentals\ - -subsubsection \Type definition\ - -quotient_type (overloaded) 'a word = int / \\k l. take_bit LENGTH('a) k = take_bit LENGTH('a::len) l\ - morphisms rep Word by (auto intro!: equivpI reflpI sympI transpI) - -hide_const (open) rep \ \only for foundational purpose\ -hide_const (open) Word \ \only for code generation\ - - -subsubsection \Basic arithmetic\ - -instantiation word :: (len) comm_ring_1 -begin - -lift_definition zero_word :: \'a word\ - is 0 . - -lift_definition one_word :: \'a word\ - is 1 . - -lift_definition plus_word :: \'a word \ 'a word \ 'a word\ - is \(+)\ - by (auto simp add: take_bit_eq_mod intro: mod_add_cong) - -lift_definition minus_word :: \'a word \ 'a word \ 'a word\ - is \(-)\ - by (auto simp add: take_bit_eq_mod intro: mod_diff_cong) - -lift_definition uminus_word :: \'a word \ 'a word\ - is uminus - by (auto simp add: take_bit_eq_mod intro: mod_minus_cong) - -lift_definition times_word :: \'a word \ 'a word \ 'a word\ - is \(*)\ - by (auto simp add: take_bit_eq_mod intro: mod_mult_cong) - -instance - by (standard; transfer) (simp_all add: algebra_simps) - -end - -context - includes lifting_syntax - notes - power_transfer [transfer_rule] - transfer_rule_of_bool [transfer_rule] - transfer_rule_numeral [transfer_rule] - transfer_rule_of_nat [transfer_rule] - transfer_rule_of_int [transfer_rule] -begin - -lemma power_transfer_word [transfer_rule]: - \(pcr_word ===> (=) ===> pcr_word) (^) (^)\ - by transfer_prover - -lemma [transfer_rule]: - \((=) ===> pcr_word) of_bool of_bool\ - by transfer_prover - -lemma [transfer_rule]: - \((=) ===> pcr_word) numeral numeral\ - by transfer_prover - -lemma [transfer_rule]: - \((=) ===> pcr_word) int of_nat\ - by transfer_prover - -lemma [transfer_rule]: - \((=) ===> pcr_word) (\k. k) of_int\ -proof - - have \((=) ===> pcr_word) of_int of_int\ - by transfer_prover - then show ?thesis by (simp add: id_def) -qed - -lemma [transfer_rule]: - \(pcr_word ===> (\)) even ((dvd) 2 :: 'a::len word \ bool)\ -proof - - have even_word_unfold: "even k \ (\l. take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l))" (is "?P \ ?Q") - for k :: int - proof - assume ?P - then show ?Q - by auto - next - assume ?Q - then obtain l where "take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l)" .. - then have "even (take_bit LENGTH('a) k)" - by simp - then show ?P - by simp - qed - show ?thesis by (simp only: even_word_unfold [abs_def] dvd_def [where ?'a = "'a word", abs_def]) - transfer_prover -qed - -end - -lemma exp_eq_zero_iff [simp]: - \2 ^ n = (0 :: 'a::len word) \ n \ LENGTH('a)\ - by transfer simp - -lemma word_exp_length_eq_0 [simp]: - \(2 :: 'a::len word) ^ LENGTH('a) = 0\ - by simp - - -subsubsection \Basic tool setup\ - -ML_file \Tools/word_lib.ML\ - - -subsubsection \Basic code generation setup\ - -context -begin - -qualified lift_definition the_int :: \'a::len word \ int\ - is \take_bit LENGTH('a)\ . - -end - -lemma [code abstype]: - \Word.Word (Word.the_int w) = w\ - by transfer simp - -lemma Word_eq_word_of_int [code_post, simp]: - \Word.Word = of_int\ - by (rule; transfer) simp - -quickcheck_generator word - constructors: - \0 :: 'a::len word\, - \numeral :: num \ 'a::len word\ - -instantiation word :: (len) equal -begin - -lift_definition equal_word :: \'a word \ 'a word \ bool\ - is \\k l. take_bit LENGTH('a) k = take_bit LENGTH('a) l\ - by simp - -instance - by (standard; transfer) rule - -end - -lemma [code]: - \HOL.equal v w \ HOL.equal (Word.the_int v) (Word.the_int w)\ - by transfer (simp add: equal) - -lemma [code]: - \Word.the_int 0 = 0\ - by transfer simp - -lemma [code]: - \Word.the_int 1 = 1\ - by transfer simp - -lemma [code]: - \Word.the_int (v + w) = take_bit LENGTH('a) (Word.the_int v + Word.the_int w)\ - for v w :: \'a::len word\ - by transfer (simp add: take_bit_add) - -lemma [code]: - \Word.the_int (- w) = (let k = Word.the_int w in if w = 0 then 0 else 2 ^ LENGTH('a) - k)\ - for w :: \'a::len word\ - by transfer (auto simp add: take_bit_eq_mod zmod_zminus1_eq_if) - -lemma [code]: - \Word.the_int (v - w) = take_bit LENGTH('a) (Word.the_int v - Word.the_int w)\ - for v w :: \'a::len word\ - by transfer (simp add: take_bit_diff) - -lemma [code]: - \Word.the_int (v * w) = take_bit LENGTH('a) (Word.the_int v * Word.the_int w)\ - for v w :: \'a::len word\ - by transfer (simp add: take_bit_mult) - - -subsubsection \Basic conversions\ - -abbreviation word_of_nat :: \nat \ 'a::len word\ - where \word_of_nat \ of_nat\ - -abbreviation word_of_int :: \int \ 'a::len word\ - where \word_of_int \ of_int\ - -lemma word_of_nat_eq_iff: - \word_of_nat m = (word_of_nat n :: 'a::len word) \ take_bit LENGTH('a) m = take_bit LENGTH('a) n\ - by transfer (simp add: take_bit_of_nat) - -lemma word_of_int_eq_iff: - \word_of_int k = (word_of_int l :: 'a::len word) \ take_bit LENGTH('a) k = take_bit LENGTH('a) l\ - by transfer rule - -lemma word_of_nat_eq_0_iff [simp]: - \word_of_nat n = (0 :: 'a::len word) \ 2 ^ LENGTH('a) dvd n\ - using word_of_nat_eq_iff [where ?'a = 'a, of n 0] by (simp add: take_bit_eq_0_iff) - -lemma word_of_int_eq_0_iff [simp]: - \word_of_int k = (0 :: 'a::len word) \ 2 ^ LENGTH('a) dvd k\ - using word_of_int_eq_iff [where ?'a = 'a, of k 0] by (simp add: take_bit_eq_0_iff) - -context semiring_1 -begin - -lift_definition unsigned :: \'b::len word \ 'a\ - is \of_nat \ nat \ take_bit LENGTH('b)\ - by simp - -lemma unsigned_0 [simp]: - \unsigned 0 = 0\ - by transfer simp - -lemma unsigned_1 [simp]: - \unsigned 1 = 1\ - by transfer simp - -lemma unsigned_numeral [simp]: - \unsigned (numeral n :: 'b::len word) = of_nat (take_bit LENGTH('b) (numeral n))\ - by transfer (simp add: nat_take_bit_eq) - -lemma unsigned_neg_numeral [simp]: - \unsigned (- numeral n :: 'b::len word) = of_nat (nat (take_bit LENGTH('b) (- numeral n)))\ - by transfer simp - -end - -context semiring_1 -begin - -lemma unsigned_of_nat [simp]: - \unsigned (word_of_nat n :: 'b::len word) = of_nat (take_bit LENGTH('b) n)\ - by transfer (simp add: nat_eq_iff take_bit_of_nat) - -lemma unsigned_of_int [simp]: - \unsigned (word_of_int k :: 'b::len word) = of_nat (nat (take_bit LENGTH('b) k))\ - by transfer simp - -end - -context semiring_char_0 -begin - -lemma unsigned_word_eqI: - \v = w\ if \unsigned v = unsigned w\ - using that by transfer (simp add: eq_nat_nat_iff) - -lemma word_eq_iff_unsigned: - \v = w \ unsigned v = unsigned w\ - by (auto intro: unsigned_word_eqI) - -lemma inj_unsigned [simp]: - \inj unsigned\ - by (rule injI) (simp add: unsigned_word_eqI) - -lemma unsigned_eq_0_iff: - \unsigned w = 0 \ w = 0\ - using word_eq_iff_unsigned [of w 0] by simp - -end - -context ring_1 -begin - -lift_definition signed :: \'b::len word \ 'a\ - is \of_int \ signed_take_bit (LENGTH('b) - Suc 0)\ - by (simp flip: signed_take_bit_decr_length_iff) - -lemma signed_0 [simp]: - \signed 0 = 0\ - by transfer simp - -lemma signed_1 [simp]: - \signed (1 :: 'b::len word) = (if LENGTH('b) = 1 then - 1 else 1)\ - by (transfer fixing: uminus; cases \LENGTH('b)\) (auto dest: gr0_implies_Suc) - -lemma signed_minus_1 [simp]: - \signed (- 1 :: 'b::len word) = - 1\ - by (transfer fixing: uminus) simp - -lemma signed_numeral [simp]: - \signed (numeral n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - 1) (numeral n))\ - by transfer simp - -lemma signed_neg_numeral [simp]: - \signed (- numeral n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - 1) (- numeral n))\ - by transfer simp - -lemma signed_of_nat [simp]: - \signed (word_of_nat n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - Suc 0) (int n))\ - by transfer simp - -lemma signed_of_int [simp]: - \signed (word_of_int n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - Suc 0) n)\ - by transfer simp - -end - -context ring_char_0 -begin - -lemma signed_word_eqI: - \v = w\ if \signed v = signed w\ - using that by transfer (simp flip: signed_take_bit_decr_length_iff) - -lemma word_eq_iff_signed: - \v = w \ signed v = signed w\ - by (auto intro: signed_word_eqI) - -lemma inj_signed [simp]: - \inj signed\ - by (rule injI) (simp add: signed_word_eqI) - -lemma signed_eq_0_iff: - \signed w = 0 \ w = 0\ - using word_eq_iff_signed [of w 0] by simp - -end - -abbreviation unat :: \'a::len word \ nat\ - where \unat \ unsigned\ - -abbreviation uint :: \'a::len word \ int\ - where \uint \ unsigned\ - -abbreviation sint :: \'a::len word \ int\ - where \sint \ signed\ - -abbreviation ucast :: \'a::len word \ 'b::len word\ - where \ucast \ unsigned\ - -abbreviation scast :: \'a::len word \ 'b::len word\ - where \scast \ signed\ - -context - includes lifting_syntax -begin - -lemma [transfer_rule]: - \(pcr_word ===> (=)) (nat \ take_bit LENGTH('a)) (unat :: 'a::len word \ nat)\ - using unsigned.transfer [where ?'a = nat] by simp - -lemma [transfer_rule]: - \(pcr_word ===> (=)) (take_bit LENGTH('a)) (uint :: 'a::len word \ int)\ - using unsigned.transfer [where ?'a = int] by (simp add: comp_def) - -lemma [transfer_rule]: - \(pcr_word ===> (=)) (signed_take_bit (LENGTH('a) - Suc 0)) (sint :: 'a::len word \ int)\ - using signed.transfer [where ?'a = int] by simp - -lemma [transfer_rule]: - \(pcr_word ===> pcr_word) (take_bit LENGTH('a)) (ucast :: 'a::len word \ 'b::len word)\ -proof (rule rel_funI) - fix k :: int and w :: \'a word\ - assume \pcr_word k w\ - then have \w = word_of_int k\ - by (simp add: pcr_word_def cr_word_def relcompp_apply) - moreover have \pcr_word (take_bit LENGTH('a) k) (ucast (word_of_int k :: 'a word))\ - by transfer (simp add: pcr_word_def cr_word_def relcompp_apply) - ultimately show \pcr_word (take_bit LENGTH('a) k) (ucast w)\ - by simp -qed - -lemma [transfer_rule]: - \(pcr_word ===> pcr_word) (signed_take_bit (LENGTH('a) - Suc 0)) (scast :: 'a::len word \ 'b::len word)\ -proof (rule rel_funI) - fix k :: int and w :: \'a word\ - assume \pcr_word k w\ - then have \w = word_of_int k\ - by (simp add: pcr_word_def cr_word_def relcompp_apply) - moreover have \pcr_word (signed_take_bit (LENGTH('a) - Suc 0) k) (scast (word_of_int k :: 'a word))\ - by transfer (simp add: pcr_word_def cr_word_def relcompp_apply) - ultimately show \pcr_word (signed_take_bit (LENGTH('a) - Suc 0) k) (scast w)\ - by simp -qed - -end - -lemma of_nat_unat [simp]: - \of_nat (unat w) = unsigned w\ - by transfer simp - -lemma of_int_uint [simp]: - \of_int (uint w) = unsigned w\ - by transfer simp - -lemma of_int_sint [simp]: - \of_int (sint a) = signed a\ - by transfer (simp_all add: take_bit_signed_take_bit) - -lemma nat_uint_eq [simp]: - \nat (uint w) = unat w\ - by transfer simp - -lemma sgn_uint_eq [simp]: - \sgn (uint w) = of_bool (w \ 0)\ - by transfer (simp add: less_le) - -text \Aliasses only for code generation\ - -context -begin - -qualified lift_definition of_int :: \int \ 'a::len word\ - is \take_bit LENGTH('a)\ . - -qualified lift_definition of_nat :: \nat \ 'a::len word\ - is \int \ take_bit LENGTH('a)\ . - -qualified lift_definition the_nat :: \'a::len word \ nat\ - is \nat \ take_bit LENGTH('a)\ by simp - -qualified lift_definition the_signed_int :: \'a::len word \ int\ - is \signed_take_bit (LENGTH('a) - Suc 0)\ by (simp add: signed_take_bit_decr_length_iff) - -qualified lift_definition cast :: \'a::len word \ 'b::len word\ - is \take_bit LENGTH('a)\ by simp - -qualified lift_definition signed_cast :: \'a::len word \ 'b::len word\ - is \signed_take_bit (LENGTH('a) - Suc 0)\ by (metis signed_take_bit_decr_length_iff) - -end - -lemma [code_abbrev, simp]: - \Word.the_int = uint\ - by transfer rule - -lemma [code]: - \Word.the_int (Word.of_int k :: 'a::len word) = take_bit LENGTH('a) k\ - by transfer simp - -lemma [code_abbrev, simp]: - \Word.of_int = word_of_int\ - by (rule; transfer) simp - -lemma [code]: - \Word.the_int (Word.of_nat n :: 'a::len word) = take_bit LENGTH('a) (int n)\ - by transfer (simp add: take_bit_of_nat) - -lemma [code_abbrev, simp]: - \Word.of_nat = word_of_nat\ - by (rule; transfer) (simp add: take_bit_of_nat) - -lemma [code]: - \Word.the_nat w = nat (Word.the_int w)\ - by transfer simp - -lemma [code_abbrev, simp]: - \Word.the_nat = unat\ - by (rule; transfer) simp - -lemma [code]: - \Word.the_signed_int w = signed_take_bit (LENGTH('a) - Suc 0) (Word.the_int w)\ - for w :: \'a::len word\ - by transfer (simp add: signed_take_bit_take_bit) - -lemma [code_abbrev, simp]: - \Word.the_signed_int = sint\ - by (rule; transfer) simp - -lemma [code]: - \Word.the_int (Word.cast w :: 'b::len word) = take_bit LENGTH('b) (Word.the_int w)\ - for w :: \'a::len word\ - by transfer simp - -lemma [code_abbrev, simp]: - \Word.cast = ucast\ - by (rule; transfer) simp - -lemma [code]: - \Word.the_int (Word.signed_cast w :: 'b::len word) = take_bit LENGTH('b) (Word.the_signed_int w)\ - for w :: \'a::len word\ - by transfer simp - -lemma [code_abbrev, simp]: - \Word.signed_cast = scast\ - by (rule; transfer) simp - -lemma [code]: - \unsigned w = of_nat (nat (Word.the_int w))\ - by transfer simp - -lemma [code]: - \signed w = of_int (Word.the_signed_int w)\ - by transfer simp - - -subsubsection \Basic ordering\ - -instantiation word :: (len) linorder -begin - -lift_definition less_eq_word :: "'a word \ 'a word \ bool" - is "\a b. take_bit LENGTH('a) a \ take_bit LENGTH('a) b" - by simp - -lift_definition less_word :: "'a word \ 'a word \ bool" - is "\a b. take_bit LENGTH('a) a < take_bit LENGTH('a) b" - by simp - -instance - by (standard; transfer) auto - -end - -interpretation word_order: ordering_top \(\)\ \(<)\ \- 1 :: 'a::len word\ - by (standard; transfer) (simp add: take_bit_eq_mod zmod_minus1) - -interpretation word_coorder: ordering_top \(\)\ \(>)\ \0 :: 'a::len word\ - by (standard; transfer) simp - -lemma word_of_nat_less_eq_iff: - \word_of_nat m \ (word_of_nat n :: 'a::len word) \ take_bit LENGTH('a) m \ take_bit LENGTH('a) n\ - by transfer (simp add: take_bit_of_nat) - -lemma word_of_int_less_eq_iff: - \word_of_int k \ (word_of_int l :: 'a::len word) \ take_bit LENGTH('a) k \ take_bit LENGTH('a) l\ - by transfer rule - -lemma word_of_nat_less_iff: - \word_of_nat m < (word_of_nat n :: 'a::len word) \ take_bit LENGTH('a) m < take_bit LENGTH('a) n\ - by transfer (simp add: take_bit_of_nat) - -lemma word_of_int_less_iff: - \word_of_int k < (word_of_int l :: 'a::len word) \ take_bit LENGTH('a) k < take_bit LENGTH('a) l\ - by transfer rule - -lemma word_le_def [code]: - "a \ b \ uint a \ uint b" - by transfer rule - -lemma word_less_def [code]: - "a < b \ uint a < uint b" - by transfer rule - -lemma word_greater_zero_iff: - \a > 0 \ a \ 0\ for a :: \'a::len word\ - by transfer (simp add: less_le) - -lemma of_nat_word_less_eq_iff: - \of_nat m \ (of_nat n :: 'a::len word) \ take_bit LENGTH('a) m \ take_bit LENGTH('a) n\ - by transfer (simp add: take_bit_of_nat) - -lemma of_nat_word_less_iff: - \of_nat m < (of_nat n :: 'a::len word) \ take_bit LENGTH('a) m < take_bit LENGTH('a) n\ - by transfer (simp add: take_bit_of_nat) - -lemma of_int_word_less_eq_iff: - \of_int k \ (of_int l :: 'a::len word) \ take_bit LENGTH('a) k \ take_bit LENGTH('a) l\ - by transfer rule - -lemma of_int_word_less_iff: - \of_int k < (of_int l :: 'a::len word) \ take_bit LENGTH('a) k < take_bit LENGTH('a) l\ - by transfer rule - - - -subsection \Enumeration\ - -lemma inj_on_word_of_nat: - \inj_on (word_of_nat :: nat \ 'a::len word) {0..<2 ^ LENGTH('a)}\ - by (rule inj_onI; transfer) (simp_all add: take_bit_int_eq_self) - -lemma UNIV_word_eq_word_of_nat: - \(UNIV :: 'a::len word set) = word_of_nat ` {0..<2 ^ LENGTH('a)}\ (is \_ = ?A\) -proof - show \word_of_nat ` {0..<2 ^ LENGTH('a)} \ UNIV\ - by simp - show \UNIV \ ?A\ - proof - fix w :: \'a word\ - show \w \ (word_of_nat ` {0..<2 ^ LENGTH('a)} :: 'a word set)\ - by (rule image_eqI [of _ _ \unat w\]; transfer) simp_all - qed -qed - -instantiation word :: (len) enum -begin - -definition enum_word :: \'a word list\ - where \enum_word = map word_of_nat [0..<2 ^ LENGTH('a)]\ - -definition enum_all_word :: \('a word \ bool) \ bool\ - where \enum_all_word = Ball UNIV\ - -definition enum_ex_word :: \('a word \ bool) \ bool\ - where \enum_ex_word = Bex UNIV\ - -lemma [code]: - \Enum.enum_all P \ Ball UNIV P\ - \Enum.enum_ex P \ Bex UNIV P\ for P :: \'a word \ bool\ - by (simp_all add: enum_all_word_def enum_ex_word_def) - -instance - by standard (simp_all add: UNIV_word_eq_word_of_nat inj_on_word_of_nat enum_word_def enum_all_word_def enum_ex_word_def distinct_map) - -end - - -subsection \Bit-wise operations\ - -instantiation word :: (len) semiring_modulo -begin - -lift_definition divide_word :: \'a word \ 'a word \ 'a word\ - is \\a b. take_bit LENGTH('a) a div take_bit LENGTH('a) b\ - by simp - -lift_definition modulo_word :: \'a word \ 'a word \ 'a word\ - is \\a b. take_bit LENGTH('a) a mod take_bit LENGTH('a) b\ - by simp - -instance proof - show "a div b * b + a mod b = a" for a b :: "'a word" - proof transfer - fix k l :: int - define r :: int where "r = 2 ^ LENGTH('a)" - then have r: "take_bit LENGTH('a) k = k mod r" for k - by (simp add: take_bit_eq_mod) - have "k mod r = ((k mod r) div (l mod r) * (l mod r) - + (k mod r) mod (l mod r)) mod r" - by (simp add: div_mult_mod_eq) - also have "... = (((k mod r) div (l mod r) * (l mod r)) mod r - + (k mod r) mod (l mod r)) mod r" - by (simp add: mod_add_left_eq) - also have "... = (((k mod r) div (l mod r) * l) mod r - + (k mod r) mod (l mod r)) mod r" - by (simp add: mod_mult_right_eq) - finally have "k mod r = ((k mod r) div (l mod r) * l - + (k mod r) mod (l mod r)) mod r" - by (simp add: mod_simps) - with r show "take_bit LENGTH('a) (take_bit LENGTH('a) k div take_bit LENGTH('a) l * l - + take_bit LENGTH('a) k mod take_bit LENGTH('a) l) = take_bit LENGTH('a) k" - by simp - qed -qed - -end - -instance word :: (len) semiring_parity -proof - show "\ 2 dvd (1::'a word)" - by transfer simp - show even_iff_mod_2_eq_0: "2 dvd a \ a mod 2 = 0" - for a :: "'a word" - by transfer (simp_all add: mod_2_eq_odd take_bit_Suc) - show "\ 2 dvd a \ a mod 2 = 1" - for a :: "'a word" - by transfer (simp_all add: mod_2_eq_odd take_bit_Suc) -qed - -lemma word_bit_induct [case_names zero even odd]: - \P a\ if word_zero: \P 0\ - and word_even: \\a. P a \ 0 < a \ a < 2 ^ (LENGTH('a) - Suc 0) \ P (2 * a)\ - and word_odd: \\a. P a \ a < 2 ^ (LENGTH('a) - Suc 0) \ P (1 + 2 * a)\ - for P and a :: \'a::len word\ -proof - - define m :: nat where \m = LENGTH('a) - Suc 0\ - then have l: \LENGTH('a) = Suc m\ - by simp - define n :: nat where \n = unat a\ - then have \n < 2 ^ LENGTH('a)\ - by transfer (simp add: take_bit_eq_mod) - then have \n < 2 * 2 ^ m\ - by (simp add: l) - then have \P (of_nat n)\ - proof (induction n rule: nat_bit_induct) - case zero - show ?case - by simp (rule word_zero) - next - case (even n) - then have \n < 2 ^ m\ - by simp - with even.IH have \P (of_nat n)\ - by simp - moreover from \n < 2 ^ m\ even.hyps have \0 < (of_nat n :: 'a word)\ - by (auto simp add: word_greater_zero_iff l) - moreover from \n < 2 ^ m\ have \(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - Suc 0)\ - using of_nat_word_less_iff [where ?'a = 'a, of n \2 ^ m\] - by (simp add: l take_bit_eq_mod) - ultimately have \P (2 * of_nat n)\ - by (rule word_even) - then show ?case - by simp - next - case (odd n) - then have \Suc n \ 2 ^ m\ - by simp - with odd.IH have \P (of_nat n)\ - by simp - moreover from \Suc n \ 2 ^ m\ have \(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - Suc 0)\ - using of_nat_word_less_iff [where ?'a = 'a, of n \2 ^ m\] - by (simp add: l take_bit_eq_mod) - ultimately have \P (1 + 2 * of_nat n)\ - by (rule word_odd) - then show ?case - by simp - qed - moreover have \of_nat (nat (uint a)) = a\ - by transfer simp - ultimately show ?thesis - by (simp add: n_def) -qed - -lemma bit_word_half_eq: - \(of_bool b + a * 2) div 2 = a\ - if \a < 2 ^ (LENGTH('a) - Suc 0)\ - for a :: \'a::len word\ -proof (cases \2 \ LENGTH('a::len)\) - case False - have \of_bool (odd k) < (1 :: int) \ even k\ for k :: int - by auto - with False that show ?thesis - by transfer (simp add: eq_iff) -next - case True - obtain n where length: \LENGTH('a) = Suc n\ - by (cases \LENGTH('a)\) simp_all - show ?thesis proof (cases b) - case False - moreover have \a * 2 div 2 = a\ - using that proof transfer - fix k :: int - from length have \k * 2 mod 2 ^ LENGTH('a) = (k mod 2 ^ n) * 2\ - by simp - moreover assume \take_bit LENGTH('a) k < take_bit LENGTH('a) (2 ^ (LENGTH('a) - Suc 0))\ - with \LENGTH('a) = Suc n\ - have \k mod 2 ^ LENGTH('a) = k mod 2 ^ n\ - by (simp add: take_bit_eq_mod divmod_digit_0) - ultimately have \take_bit LENGTH('a) (k * 2) = take_bit LENGTH('a) k * 2\ - by (simp add: take_bit_eq_mod) - with True show \take_bit LENGTH('a) (take_bit LENGTH('a) (k * 2) div take_bit LENGTH('a) 2) - = take_bit LENGTH('a) k\ - by simp - qed - ultimately show ?thesis - by simp - next - case True - moreover have \(1 + a * 2) div 2 = a\ - using that proof transfer - fix k :: int - from length have \(1 + k * 2) mod 2 ^ LENGTH('a) = 1 + (k mod 2 ^ n) * 2\ - using pos_zmod_mult_2 [of \2 ^ n\ k] by (simp add: ac_simps) - moreover assume \take_bit LENGTH('a) k < take_bit LENGTH('a) (2 ^ (LENGTH('a) - Suc 0))\ - with \LENGTH('a) = Suc n\ - have \k mod 2 ^ LENGTH('a) = k mod 2 ^ n\ - by (simp add: take_bit_eq_mod divmod_digit_0) - ultimately have \take_bit LENGTH('a) (1 + k * 2) = 1 + take_bit LENGTH('a) k * 2\ - by (simp add: take_bit_eq_mod) - with True show \take_bit LENGTH('a) (take_bit LENGTH('a) (1 + k * 2) div take_bit LENGTH('a) 2) - = take_bit LENGTH('a) k\ - by (auto simp add: take_bit_Suc) - qed - ultimately show ?thesis - by simp - qed -qed - -lemma even_mult_exp_div_word_iff: - \even (a * 2 ^ m div 2 ^ n) \ \ ( - m \ n \ - n < LENGTH('a) \ odd (a div 2 ^ (n - m)))\ for a :: \'a::len word\ - by transfer - (auto simp flip: drop_bit_eq_div simp add: even_drop_bit_iff_not_bit bit_take_bit_iff, - simp_all flip: push_bit_eq_mult add: bit_push_bit_iff_int) - -instantiation word :: (len) semiring_bits -begin - -lift_definition bit_word :: \'a word \ nat \ bool\ - is \\k n. n < LENGTH('a) \ bit k n\ -proof - fix k l :: int and n :: nat - assume *: \take_bit LENGTH('a) k = take_bit LENGTH('a) l\ - show \n < LENGTH('a) \ bit k n \ n < LENGTH('a) \ bit l n\ - proof (cases \n < LENGTH('a)\) - case True - from * have \bit (take_bit LENGTH('a) k) n \ bit (take_bit LENGTH('a) l) n\ - by simp - then show ?thesis - by (simp add: bit_take_bit_iff) - next - case False - then show ?thesis - by simp - qed -qed - -instance proof - show \P a\ if stable: \\a. a div 2 = a \ P a\ - and rec: \\a b. P a \ (of_bool b + 2 * a) div 2 = a \ P (of_bool b + 2 * a)\ - for P and a :: \'a word\ - proof (induction a rule: word_bit_induct) - case zero - have \0 div 2 = (0::'a word)\ - by transfer simp - with stable [of 0] show ?case - by simp - next - case (even a) - with rec [of a False] show ?case - using bit_word_half_eq [of a False] by (simp add: ac_simps) - next - case (odd a) - with rec [of a True] show ?case - using bit_word_half_eq [of a True] by (simp add: ac_simps) - qed - show \bit a n \ odd (a div 2 ^ n)\ for a :: \'a word\ and n - by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit bit_iff_odd_drop_bit) - show \0 div a = 0\ - for a :: \'a word\ - by transfer simp - show \a div 1 = a\ - for a :: \'a word\ - by transfer simp - show \a mod b div b = 0\ - for a b :: \'a word\ - apply transfer - apply (simp add: take_bit_eq_mod) - apply (subst (3) mod_pos_pos_trivial [of _ \2 ^ LENGTH('a)\]) - apply simp_all - apply (metis le_less mod_by_0 pos_mod_conj zero_less_numeral zero_less_power) - using pos_mod_bound [of \2 ^ LENGTH('a)\] apply simp - proof - - fix aa :: int and ba :: int - have f1: "\i n. (i::int) mod 2 ^ n = 0 \ 0 < i mod 2 ^ n" - by (metis le_less take_bit_eq_mod take_bit_nonnegative) - have "(0::int) < 2 ^ len_of (TYPE('a)::'a itself) \ ba mod 2 ^ len_of (TYPE('a)::'a itself) \ 0 \ aa mod 2 ^ len_of (TYPE('a)::'a itself) mod (ba mod 2 ^ len_of (TYPE('a)::'a itself)) < 2 ^ len_of (TYPE('a)::'a itself)" - by (metis (no_types) mod_by_0 unique_euclidean_semiring_numeral_class.pos_mod_bound zero_less_numeral zero_less_power) - then show "aa mod 2 ^ len_of (TYPE('a)::'a itself) mod (ba mod 2 ^ len_of (TYPE('a)::'a itself)) < 2 ^ len_of (TYPE('a)::'a itself)" - using f1 by (meson le_less less_le_trans unique_euclidean_semiring_numeral_class.pos_mod_bound) - qed - show \(1 + a) div 2 = a div 2\ - if \even a\ - for a :: \'a word\ - using that by transfer - (auto dest: le_Suc_ex simp add: mod_2_eq_odd take_bit_Suc elim!: evenE) - show \(2 :: 'a word) ^ m div 2 ^ n = of_bool ((2 :: 'a word) ^ m \ 0 \ n \ m) * 2 ^ (m - n)\ - for m n :: nat - by transfer (simp, simp add: exp_div_exp_eq) - show "a div 2 ^ m div 2 ^ n = a div 2 ^ (m + n)" - for a :: "'a word" and m n :: nat - apply transfer - apply (auto simp add: not_less take_bit_drop_bit ac_simps simp flip: drop_bit_eq_div) - apply (simp add: drop_bit_take_bit) - done - show "a mod 2 ^ m mod 2 ^ n = a mod 2 ^ min m n" - for a :: "'a word" and m n :: nat - by transfer (auto simp flip: take_bit_eq_mod simp add: ac_simps) - show \a * 2 ^ m mod 2 ^ n = a mod 2 ^ (n - m) * 2 ^ m\ - if \m \ n\ for a :: "'a word" and m n :: nat - using that apply transfer - apply (auto simp flip: take_bit_eq_mod) - apply (auto simp flip: push_bit_eq_mult simp add: push_bit_take_bit split: split_min_lin) - done - show \a div 2 ^ n mod 2 ^ m = a mod (2 ^ (n + m)) div 2 ^ n\ - for a :: "'a word" and m n :: nat - by transfer (auto simp add: not_less take_bit_drop_bit ac_simps simp flip: take_bit_eq_mod drop_bit_eq_div split: split_min_lin) - show \even ((2 ^ m - 1) div (2::'a word) ^ n) \ 2 ^ n = (0::'a word) \ m \ n\ - for m n :: nat - by transfer (auto simp add: take_bit_of_mask even_mask_div_iff) - show \even (a * 2 ^ m div 2 ^ n) \ n < m \ (2::'a word) ^ n = 0 \ m \ n \ even (a div 2 ^ (n - m))\ - for a :: \'a word\ and m n :: nat - proof transfer - show \even (take_bit LENGTH('a) (k * 2 ^ m) div take_bit LENGTH('a) (2 ^ n)) \ - n < m - \ take_bit LENGTH('a) ((2::int) ^ n) = take_bit LENGTH('a) 0 - \ (m \ n \ even (take_bit LENGTH('a) k div take_bit LENGTH('a) (2 ^ (n - m))))\ - for m n :: nat and k l :: int - by (auto simp flip: take_bit_eq_mod drop_bit_eq_div push_bit_eq_mult - simp add: div_push_bit_of_1_eq_drop_bit drop_bit_take_bit drop_bit_push_bit_int [of n m]) - qed -qed - -end - -lemma bit_word_eqI: - \a = b\ if \\n. n < LENGTH('a) \ bit a n \ bit b n\ - for a b :: \'a::len word\ - using that by transfer (auto simp add: nat_less_le bit_eq_iff bit_take_bit_iff) - -lemma bit_imp_le_length: - \n < LENGTH('a)\ if \bit w n\ - for w :: \'a::len word\ - using that by transfer simp - -lemma not_bit_length [simp]: - \\ bit w LENGTH('a)\ for w :: \'a::len word\ - by transfer simp - -instantiation word :: (len) semiring_bit_shifts -begin - -lift_definition push_bit_word :: \nat \ 'a word \ 'a word\ - is push_bit -proof - - show \take_bit LENGTH('a) (push_bit n k) = take_bit LENGTH('a) (push_bit n l)\ - if \take_bit LENGTH('a) k = take_bit LENGTH('a) l\ for k l :: int and n :: nat - proof - - from that - have \take_bit (LENGTH('a) - n) (take_bit LENGTH('a) k) - = take_bit (LENGTH('a) - n) (take_bit LENGTH('a) l)\ - by simp - moreover have \min (LENGTH('a) - n) LENGTH('a) = LENGTH('a) - n\ - by simp - ultimately show ?thesis - by (simp add: take_bit_push_bit) - qed -qed - -lift_definition drop_bit_word :: \nat \ 'a word \ 'a word\ - is \\n. drop_bit n \ take_bit LENGTH('a)\ - by (simp add: take_bit_eq_mod) - -lift_definition take_bit_word :: \nat \ 'a word \ 'a word\ - is \\n. take_bit (min LENGTH('a) n)\ - by (simp add: ac_simps) (simp only: flip: take_bit_take_bit) - -instance proof - show \push_bit n a = a * 2 ^ n\ for n :: nat and a :: \'a word\ - by transfer (simp add: push_bit_eq_mult) - show \drop_bit n a = a div 2 ^ n\ for n :: nat and a :: \'a word\ - by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit) - show \take_bit n a = a mod 2 ^ n\ for n :: nat and a :: \'a word\ - by transfer (auto simp flip: take_bit_eq_mod) -qed - -end - -lemma [code]: - \push_bit n w = w * 2 ^ n\ for w :: \'a::len word\ - by (fact push_bit_eq_mult) - -lemma [code]: - \Word.the_int (drop_bit n w) = drop_bit n (Word.the_int w)\ - by transfer (simp add: drop_bit_take_bit min_def le_less less_diff_conv) - -lemma [code]: - \Word.the_int (take_bit n w) = (if n < LENGTH('a::len) then take_bit n (Word.the_int w) else Word.the_int w)\ - for w :: \'a::len word\ - by transfer (simp add: not_le not_less ac_simps min_absorb2) - - -instantiation word :: (len) ring_bit_operations -begin - -lift_definition not_word :: \'a word \ 'a word\ - is not - by (simp add: take_bit_not_iff) - -lift_definition and_word :: \'a word \ 'a word \ 'a word\ - is \and\ - by simp - -lift_definition or_word :: \'a word \ 'a word \ 'a word\ - is or - by simp - -lift_definition xor_word :: \'a word \ 'a word \ 'a word\ - is xor - by simp - -lift_definition mask_word :: \nat \ 'a word\ - is mask - . - -instance by (standard; transfer) - (auto simp add: minus_eq_not_minus_1 mask_eq_exp_minus_1 - bit_not_iff bit_and_iff bit_or_iff bit_xor_iff) - -end - -lemma [code_abbrev]: - \push_bit n 1 = (2 :: 'a::len word) ^ n\ - by (fact push_bit_of_1) - -lemma [code]: - \NOT w = Word.of_int (NOT (Word.the_int w))\ - for w :: \'a::len word\ - by transfer (simp add: take_bit_not_take_bit) - -lemma [code]: - \Word.the_int (v AND w) = Word.the_int v AND Word.the_int w\ - by transfer simp - -lemma [code]: - \Word.the_int (v OR w) = Word.the_int v OR Word.the_int w\ - by transfer simp - -lemma [code]: - \Word.the_int (v XOR w) = Word.the_int v XOR Word.the_int w\ - by transfer simp - -lemma [code]: - \Word.the_int (mask n :: 'a::len word) = mask (min LENGTH('a) n)\ - by transfer simp - -context - includes lifting_syntax -begin - -lemma set_bit_word_transfer [transfer_rule]: - \((=) ===> pcr_word ===> pcr_word) set_bit set_bit\ - by (unfold set_bit_def) transfer_prover - -lemma unset_bit_word_transfer [transfer_rule]: - \((=) ===> pcr_word ===> pcr_word) unset_bit unset_bit\ - by (unfold unset_bit_def) transfer_prover - -lemma flip_bit_word_transfer [transfer_rule]: - \((=) ===> pcr_word ===> pcr_word) flip_bit flip_bit\ - by (unfold flip_bit_def) transfer_prover - -lemma signed_take_bit_word_transfer [transfer_rule]: - \((=) ===> pcr_word ===> pcr_word) - (\n k. signed_take_bit n (take_bit LENGTH('a::len) k)) - (signed_take_bit :: nat \ 'a word \ 'a word)\ -proof - - let ?K = \\n (k :: int). take_bit (min LENGTH('a) n) k OR of_bool (n < LENGTH('a) \ bit k n) * NOT (mask n)\ - let ?W = \\n (w :: 'a word). take_bit n w OR of_bool (bit w n) * NOT (mask n)\ - have \((=) ===> pcr_word ===> pcr_word) ?K ?W\ - by transfer_prover - also have \?K = (\n k. signed_take_bit n (take_bit LENGTH('a::len) k))\ - by (simp add: fun_eq_iff signed_take_bit_def bit_take_bit_iff ac_simps) - also have \?W = signed_take_bit\ - by (simp add: fun_eq_iff signed_take_bit_def) - finally show ?thesis . -qed - -end - - -subsection \Conversions including casts\ - -subsubsection \Generic unsigned conversion\ - -context semiring_bits -begin - -lemma bit_unsigned_iff: - \bit (unsigned w) n \ 2 ^ n \ 0 \ bit w n\ - for w :: \'b::len word\ - by (transfer fixing: bit) (simp add: bit_of_nat_iff bit_nat_iff bit_take_bit_iff) - -end - -context semiring_bit_shifts -begin - -lemma unsigned_push_bit_eq: - \unsigned (push_bit n w) = take_bit LENGTH('b) (push_bit n (unsigned w))\ - for w :: \'b::len word\ -proof (rule bit_eqI) - fix m - assume \2 ^ m \ 0\ - show \bit (unsigned (push_bit n w)) m = bit (take_bit LENGTH('b) (push_bit n (unsigned w))) m\ - proof (cases \n \ m\) - case True - with \2 ^ m \ 0\ have \2 ^ (m - n) \ 0\ - by (metis (full_types) diff_add exp_add_not_zero_imp) - with True show ?thesis - by (simp add: bit_unsigned_iff bit_push_bit_iff Parity.bit_push_bit_iff bit_take_bit_iff not_le exp_eq_zero_iff ac_simps) - next - case False - then show ?thesis - by (simp add: not_le bit_unsigned_iff bit_push_bit_iff Parity.bit_push_bit_iff bit_take_bit_iff) - qed -qed - -lemma unsigned_take_bit_eq: - \unsigned (take_bit n w) = take_bit n (unsigned w)\ - for w :: \'b::len word\ - by (rule bit_eqI) (simp add: bit_unsigned_iff bit_take_bit_iff Parity.bit_take_bit_iff) - -end - -context unique_euclidean_semiring_with_bit_shifts -begin - -lemma unsigned_drop_bit_eq: - \unsigned (drop_bit n w) = drop_bit n (take_bit LENGTH('b) (unsigned w))\ - for w :: \'b::len word\ - by (rule bit_eqI) (auto simp add: bit_unsigned_iff bit_take_bit_iff bit_drop_bit_eq Parity.bit_drop_bit_eq dest: bit_imp_le_length) - -end - -context semiring_bit_operations -begin - -lemma unsigned_and_eq: - \unsigned (v AND w) = unsigned v AND unsigned w\ - for v w :: \'b::len word\ - by (rule bit_eqI) (simp add: bit_unsigned_iff bit_and_iff Bit_Operations.bit_and_iff) - -lemma unsigned_or_eq: - \unsigned (v OR w) = unsigned v OR unsigned w\ - for v w :: \'b::len word\ - by (rule bit_eqI) (simp add: bit_unsigned_iff bit_or_iff Bit_Operations.bit_or_iff) - -lemma unsigned_xor_eq: - \unsigned (v XOR w) = unsigned v XOR unsigned w\ - for v w :: \'b::len word\ - by (rule bit_eqI) (simp add: bit_unsigned_iff bit_xor_iff Bit_Operations.bit_xor_iff) - -end - -context ring_bit_operations -begin - -lemma unsigned_not_eq: - \unsigned (NOT w) = take_bit LENGTH('b) (NOT (unsigned w))\ - for w :: \'b::len word\ - by (rule bit_eqI) - (simp add: bit_unsigned_iff bit_take_bit_iff bit_not_iff Bit_Operations.bit_not_iff exp_eq_zero_iff not_le) - -end - -context unique_euclidean_semiring_numeral -begin - -lemma unsigned_greater_eq [simp]: - \0 \ unsigned w\ for w :: \'b::len word\ - by (transfer fixing: less_eq) simp - -lemma unsigned_less [simp]: - \unsigned w < 2 ^ LENGTH('b)\ for w :: \'b::len word\ - by (transfer fixing: less) simp - -end - -context linordered_semidom -begin - -lemma word_less_eq_iff_unsigned: - "a \ b \ unsigned a \ unsigned b" - by (transfer fixing: less_eq) (simp add: nat_le_eq_zle) - -lemma word_less_iff_unsigned: - "a < b \ unsigned a < unsigned b" - by (transfer fixing: less) (auto dest: preorder_class.le_less_trans [OF take_bit_nonnegative]) - -end - - -subsubsection \Generic signed conversion\ - -context ring_bit_operations -begin - -lemma bit_signed_iff: - \bit (signed w) n \ 2 ^ n \ 0 \ bit w (min (LENGTH('b) - Suc 0) n)\ - for w :: \'b::len word\ - by (transfer fixing: bit) - (auto simp add: bit_of_int_iff Bit_Operations.bit_signed_take_bit_iff min_def) - -lemma signed_push_bit_eq: - \signed (push_bit n w) = signed_take_bit (LENGTH('b) - Suc 0) (push_bit n (signed w :: 'a))\ - for w :: \'b::len word\ -proof (rule bit_eqI) - fix m - assume \2 ^ m \ 0\ - define q where \q = LENGTH('b) - Suc 0\ - then have *: \LENGTH('b) = Suc q\ - by simp - show \bit (signed (push_bit n w)) m \ - bit (signed_take_bit (LENGTH('b) - Suc 0) (push_bit n (signed w :: 'a))) m\ - proof (cases \q \ m\) - case True - moreover define r where \r = m - q\ - ultimately have \m = q + r\ - by simp - moreover from \m = q + r\ \2 ^ m \ 0\ have \2 ^ q \ 0\ \2 ^ r \ 0\ - using exp_add_not_zero_imp_left [of q r] exp_add_not_zero_imp_right [of q r] - by simp_all - moreover from \2 ^ q \ 0\ have \2 ^ (q - n) \ 0\ - by (rule exp_not_zero_imp_exp_diff_not_zero) - ultimately show ?thesis - by (auto simp add: bit_signed_iff bit_signed_take_bit_iff bit_push_bit_iff Parity.bit_push_bit_iff - min_def * exp_eq_zero_iff le_diff_conv2) - next - case False - then show ?thesis - using exp_not_zero_imp_exp_diff_not_zero [of m n] - by (auto simp add: bit_signed_iff bit_signed_take_bit_iff bit_push_bit_iff Parity.bit_push_bit_iff - min_def not_le not_less * le_diff_conv2 less_diff_conv2 Parity.exp_eq_0_imp_not_bit exp_eq_0_imp_not_bit - exp_eq_zero_iff) - qed -qed - -lemma signed_take_bit_eq: - \signed (take_bit n w) = (if n < LENGTH('b) then take_bit n (signed w) else signed w)\ - for w :: \'b::len word\ - by (transfer fixing: take_bit; cases \LENGTH('b)\) - (auto simp add: Bit_Operations.signed_take_bit_take_bit Bit_Operations.take_bit_signed_take_bit take_bit_of_int min_def less_Suc_eq) - -lemma signed_not_eq: - \signed (NOT w) = signed_take_bit LENGTH('b) (NOT (signed w))\ - for w :: \'b::len word\ -proof (rule bit_eqI) - fix n - assume \2 ^ n \ 0\ - define q where \q = LENGTH('b) - Suc 0\ - then have *: \LENGTH('b) = Suc q\ - by simp - show \bit (signed (NOT w)) n \ - bit (signed_take_bit LENGTH('b) (NOT (signed w))) n\ - proof (cases \q < n\) - case True - moreover define r where \r = n - Suc q\ - ultimately have \n = r + Suc q\ - by simp - moreover from \2 ^ n \ 0\ \n = r + Suc q\ - have \2 ^ Suc q \ 0\ - using exp_add_not_zero_imp_right by blast - ultimately show ?thesis - by (simp add: * bit_signed_iff bit_not_iff bit_signed_take_bit_iff Bit_Operations.bit_not_iff min_def - exp_eq_zero_iff) - next - case False - then show ?thesis - by (auto simp add: * bit_signed_iff bit_not_iff bit_signed_take_bit_iff Bit_Operations.bit_not_iff min_def - exp_eq_zero_iff) - qed -qed - -lemma signed_and_eq: - \signed (v AND w) = signed v AND signed w\ - for v w :: \'b::len word\ - by (rule bit_eqI) (simp add: bit_signed_iff bit_and_iff Bit_Operations.bit_and_iff) - -lemma signed_or_eq: - \signed (v OR w) = signed v OR signed w\ - for v w :: \'b::len word\ - by (rule bit_eqI) (simp add: bit_signed_iff bit_or_iff Bit_Operations.bit_or_iff) - -lemma signed_xor_eq: - \signed (v XOR w) = signed v XOR signed w\ - for v w :: \'b::len word\ - by (rule bit_eqI) (simp add: bit_signed_iff bit_xor_iff Bit_Operations.bit_xor_iff) - -end - - -subsubsection \More\ - -lemma sint_greater_eq: - \- (2 ^ (LENGTH('a) - Suc 0)) \ sint w\ for w :: \'a::len word\ -proof (cases \bit w (LENGTH('a) - Suc 0)\) - case True - then show ?thesis - by transfer (simp add: signed_take_bit_eq_if_negative minus_exp_eq_not_mask or_greater_eq ac_simps) -next - have *: \- (2 ^ (LENGTH('a) - Suc 0)) \ (0::int)\ - by simp - case False - then show ?thesis - by transfer (auto simp add: signed_take_bit_eq intro: order_trans *) -qed - -lemma sint_less: - \sint w < 2 ^ (LENGTH('a) - Suc 0)\ for w :: \'a::len word\ - by (cases \bit w (LENGTH('a) - Suc 0)\; transfer) - (simp_all add: signed_take_bit_eq signed_take_bit_def not_eq_complement mask_eq_exp_minus_1 OR_upper) - -lemma unat_div_distrib: - \unat (v div w) = unat v div unat w\ -proof transfer - fix k l - have \nat (take_bit LENGTH('a) k) div nat (take_bit LENGTH('a) l) \ nat (take_bit LENGTH('a) k)\ - by (rule div_le_dividend) - also have \nat (take_bit LENGTH('a) k) < 2 ^ LENGTH('a)\ - by (simp add: nat_less_iff) - finally show \(nat \ take_bit LENGTH('a)) (take_bit LENGTH('a) k div take_bit LENGTH('a) l) = - (nat \ take_bit LENGTH('a)) k div (nat \ take_bit LENGTH('a)) l\ - by (simp add: nat_take_bit_eq div_int_pos_iff nat_div_distrib take_bit_nat_eq_self_iff) -qed - -lemma unat_mod_distrib: - \unat (v mod w) = unat v mod unat w\ -proof transfer - fix k l - have \nat (take_bit LENGTH('a) k) mod nat (take_bit LENGTH('a) l) \ nat (take_bit LENGTH('a) k)\ - by (rule mod_less_eq_dividend) - also have \nat (take_bit LENGTH('a) k) < 2 ^ LENGTH('a)\ - by (simp add: nat_less_iff) - finally show \(nat \ take_bit LENGTH('a)) (take_bit LENGTH('a) k mod take_bit LENGTH('a) l) = - (nat \ take_bit LENGTH('a)) k mod (nat \ take_bit LENGTH('a)) l\ - by (simp add: nat_take_bit_eq mod_int_pos_iff less_le nat_mod_distrib take_bit_nat_eq_self_iff) -qed - -lemma uint_div_distrib: - \uint (v div w) = uint v div uint w\ -proof - - have \int (unat (v div w)) = int (unat v div unat w)\ - by (simp add: unat_div_distrib) - then show ?thesis - by (simp add: of_nat_div) -qed - -lemma unat_drop_bit_eq: - \unat (drop_bit n w) = drop_bit n (unat w)\ - by (rule bit_eqI) (simp add: bit_unsigned_iff bit_drop_bit_eq) - -lemma uint_mod_distrib: - \uint (v mod w) = uint v mod uint w\ -proof - - have \int (unat (v mod w)) = int (unat v mod unat w)\ - by (simp add: unat_mod_distrib) - then show ?thesis - by (simp add: of_nat_mod) -qed - -context semiring_bit_shifts -begin - -lemma unsigned_ucast_eq: - \unsigned (ucast w :: 'c::len word) = take_bit LENGTH('c) (unsigned w)\ - for w :: \'b::len word\ - by (rule bit_eqI) (simp add: bit_unsigned_iff Word.bit_unsigned_iff bit_take_bit_iff exp_eq_zero_iff not_le) - -end - -context ring_bit_operations -begin - -lemma signed_ucast_eq: - \signed (ucast w :: 'c::len word) = signed_take_bit (LENGTH('c) - Suc 0) (unsigned w)\ - for w :: \'b::len word\ -proof (rule bit_eqI) - fix n - assume \2 ^ n \ 0\ - then have \2 ^ (min (LENGTH('c) - Suc 0) n) \ 0\ - by (simp add: min_def) - (metis (mono_tags) diff_diff_cancel exp_not_zero_imp_exp_diff_not_zero) - then show \bit (signed (ucast w :: 'c::len word)) n \ bit (signed_take_bit (LENGTH('c) - Suc 0) (unsigned w)) n\ - by (simp add: bit_signed_iff bit_unsigned_iff Word.bit_unsigned_iff bit_signed_take_bit_iff exp_eq_zero_iff not_le) -qed - -lemma signed_scast_eq: - \signed (scast w :: 'c::len word) = signed_take_bit (LENGTH('c) - Suc 0) (signed w)\ - for w :: \'b::len word\ -proof (rule bit_eqI) - fix n - assume \2 ^ n \ 0\ - then have \2 ^ (min (LENGTH('c) - Suc 0) n) \ 0\ - by (simp add: min_def) - (metis (mono_tags) diff_diff_cancel exp_not_zero_imp_exp_diff_not_zero) - then show \bit (signed (scast w :: 'c::len word)) n \ bit (signed_take_bit (LENGTH('c) - Suc 0) (signed w)) n\ - by (simp add: bit_signed_iff bit_unsigned_iff Word.bit_signed_iff bit_signed_take_bit_iff exp_eq_zero_iff not_le) -qed - -end - -lemma uint_nonnegative: "0 \ uint w" - by (fact unsigned_greater_eq) - -lemma uint_bounded: "uint w < 2 ^ LENGTH('a)" - for w :: "'a::len word" - by (fact unsigned_less) - -lemma uint_idem: "uint w mod 2 ^ LENGTH('a) = uint w" - for w :: "'a::len word" - by transfer (simp add: take_bit_eq_mod) - -lemma word_uint_eqI: "uint a = uint b \ a = b" - by (fact unsigned_word_eqI) - -lemma word_uint_eq_iff: "a = b \ uint a = uint b" - by (fact word_eq_iff_unsigned) - -lemma uint_word_of_int_eq: - \uint (word_of_int k :: 'a::len word) = take_bit LENGTH('a) k\ - by transfer rule - -lemma uint_word_of_int: "uint (word_of_int k :: 'a::len word) = k mod 2 ^ LENGTH('a)" - by (simp add: uint_word_of_int_eq take_bit_eq_mod) - -lemma word_of_int_uint: "word_of_int (uint w) = w" - by transfer simp - -lemma word_div_def [code]: - "a div b = word_of_int (uint a div uint b)" - by transfer rule - -lemma word_mod_def [code]: - "a mod b = word_of_int (uint a mod uint b)" - by transfer rule - -lemma split_word_all: "(\x::'a::len word. PROP P x) \ (\x. PROP P (word_of_int x))" -proof - fix x :: "'a word" - assume "\x. PROP P (word_of_int x)" - then have "PROP P (word_of_int (uint x))" . - then show "PROP P x" - by (simp only: word_of_int_uint) -qed - -lemma sint_uint: - \sint w = signed_take_bit (LENGTH('a) - Suc 0) (uint w)\ - for w :: \'a::len word\ - by (cases \LENGTH('a)\; transfer) (simp_all add: signed_take_bit_take_bit) - -lemma unat_eq_nat_uint: - \unat w = nat (uint w)\ - by simp - -lemma ucast_eq: - \ucast w = word_of_int (uint w)\ - by transfer simp - -lemma scast_eq: - \scast w = word_of_int (sint w)\ - by transfer simp - -lemma uint_0_eq: - \uint 0 = 0\ - by (fact unsigned_0) - -lemma uint_1_eq: - \uint 1 = 1\ - by (fact unsigned_1) - -lemma word_m1_wi: "- 1 = word_of_int (- 1)" - by simp - -lemma uint_0_iff: "uint x = 0 \ x = 0" - by (auto simp add: unsigned_word_eqI) - -lemma unat_0_iff: "unat x = 0 \ x = 0" - by (auto simp add: unsigned_word_eqI) - -lemma unat_0: "unat 0 = 0" - by (fact unsigned_0) - -lemma unat_gt_0: "0 < unat x \ x \ 0" - by (auto simp: unat_0_iff [symmetric]) - -lemma ucast_0: "ucast 0 = 0" - by (fact unsigned_0) - -lemma sint_0: "sint 0 = 0" - by (fact signed_0) - -lemma scast_0: "scast 0 = 0" - by (fact signed_0) - -lemma sint_n1: "sint (- 1) = - 1" - by (fact signed_minus_1) - -lemma scast_n1: "scast (- 1) = - 1" - by (fact signed_minus_1) - -lemma uint_1: "uint (1::'a::len word) = 1" - by (fact uint_1_eq) - -lemma unat_1: "unat (1::'a::len word) = 1" - by (fact unsigned_1) - -lemma ucast_1: "ucast (1::'a::len word) = 1" - by (fact unsigned_1) - -instantiation word :: (len) size -begin - -lift_definition size_word :: \'a word \ nat\ - is \\_. LENGTH('a)\ .. - -instance .. - -end - -lemma word_size [code]: - \size w = LENGTH('a)\ for w :: \'a::len word\ - by (fact size_word.rep_eq) - -lemma word_size_gt_0 [iff]: "0 < size w" - for w :: "'a::len word" - by (simp add: word_size) - -lemmas lens_gt_0 = word_size_gt_0 len_gt_0 - -lemma lens_not_0 [iff]: - \size w \ 0\ for w :: \'a::len word\ - by auto - -lift_definition source_size :: \('a::len word \ 'b) \ nat\ - is \\_. LENGTH('a)\ . - -lift_definition target_size :: \('a \ 'b::len word) \ nat\ - is \\_. LENGTH('b)\ .. - -lift_definition is_up :: \('a::len word \ 'b::len word) \ bool\ - is \\_. LENGTH('a) \ LENGTH('b)\ .. - -lift_definition is_down :: \('a::len word \ 'b::len word) \ bool\ - is \\_. LENGTH('a) \ LENGTH('b)\ .. - -lemma is_up_eq: - \is_up f \ source_size f \ target_size f\ - for f :: \'a::len word \ 'b::len word\ - by (simp add: source_size.rep_eq target_size.rep_eq is_up.rep_eq) - -lemma is_down_eq: - \is_down f \ target_size f \ source_size f\ - for f :: \'a::len word \ 'b::len word\ - by (simp add: source_size.rep_eq target_size.rep_eq is_down.rep_eq) - -lift_definition word_int_case :: \(int \ 'b) \ 'a::len word \ 'b\ - is \\f. f \ take_bit LENGTH('a)\ by simp - -lemma word_int_case_eq_uint [code]: - \word_int_case f w = f (uint w)\ - by transfer simp - -translations - "case x of XCONST of_int y \ b" \ "CONST word_int_case (\y. b) x" - "case x of (XCONST of_int :: 'a) y \ b" \ "CONST word_int_case (\y. b) x" - - -subsection \Arithmetic operations\ - -text \Legacy theorems:\ - -lemma word_add_def [code]: - "a + b = word_of_int (uint a + uint b)" - by transfer (simp add: take_bit_add) - -lemma word_sub_wi [code]: - "a - b = word_of_int (uint a - uint b)" - by transfer (simp add: take_bit_diff) - -lemma word_mult_def [code]: - "a * b = word_of_int (uint a * uint b)" - by transfer (simp add: take_bit_eq_mod mod_simps) - -lemma word_minus_def [code]: - "- a = word_of_int (- uint a)" - by transfer (simp add: take_bit_minus) - -lemma word_0_wi: - "0 = word_of_int 0" - by transfer simp - -lemma word_1_wi: - "1 = word_of_int 1" - by transfer simp - -lift_definition word_succ :: "'a::len word \ 'a word" is "\x. x + 1" - by (auto simp add: take_bit_eq_mod intro: mod_add_cong) - -lift_definition word_pred :: "'a::len word \ 'a word" is "\x. x - 1" - by (auto simp add: take_bit_eq_mod intro: mod_diff_cong) - -lemma word_succ_alt [code]: - "word_succ a = word_of_int (uint a + 1)" - by transfer (simp add: take_bit_eq_mod mod_simps) - -lemma word_pred_alt [code]: - "word_pred a = word_of_int (uint a - 1)" - by transfer (simp add: take_bit_eq_mod mod_simps) - -lemmas word_arith_wis = - word_add_def word_sub_wi word_mult_def - word_minus_def word_succ_alt word_pred_alt - word_0_wi word_1_wi - -lemma wi_homs: - shows wi_hom_add: "word_of_int a + word_of_int b = word_of_int (a + b)" - and wi_hom_sub: "word_of_int a - word_of_int b = word_of_int (a - b)" - and wi_hom_mult: "word_of_int a * word_of_int b = word_of_int (a * b)" - and wi_hom_neg: "- word_of_int a = word_of_int (- a)" - and wi_hom_succ: "word_succ (word_of_int a) = word_of_int (a + 1)" - and wi_hom_pred: "word_pred (word_of_int a) = word_of_int (a - 1)" - by (transfer, simp)+ - -lemmas wi_hom_syms = wi_homs [symmetric] - -lemmas word_of_int_homs = wi_homs word_0_wi word_1_wi - -lemmas word_of_int_hom_syms = word_of_int_homs [symmetric] - -lemma double_eq_zero_iff: - \2 * a = 0 \ a = 0 \ a = 2 ^ (LENGTH('a) - Suc 0)\ - for a :: \'a::len word\ -proof - - define n where \n = LENGTH('a) - Suc 0\ - then have *: \LENGTH('a) = Suc n\ - by simp - have \a = 0\ if \2 * a = 0\ and \a \ 2 ^ (LENGTH('a) - Suc 0)\ - using that by transfer - (auto simp add: take_bit_eq_0_iff take_bit_eq_mod *) - moreover have \2 ^ LENGTH('a) = (0 :: 'a word)\ - by transfer simp - then have \2 * 2 ^ (LENGTH('a) - Suc 0) = (0 :: 'a word)\ - by (simp add: *) - ultimately show ?thesis - by auto -qed - - -subsection \Ordering\ - -lift_definition word_sle :: \'a::len word \ 'a word \ bool\ - is \\k l. signed_take_bit (LENGTH('a) - Suc 0) k \ signed_take_bit (LENGTH('a) - Suc 0) l\ - by (simp flip: signed_take_bit_decr_length_iff) - -lift_definition word_sless :: \'a::len word \ 'a word \ bool\ - is \\k l. signed_take_bit (LENGTH('a) - Suc 0) k < signed_take_bit (LENGTH('a) - Suc 0) l\ - by (simp flip: signed_take_bit_decr_length_iff) - -notation - word_sle ("'(\s')") and - word_sle ("(_/ \s _)" [51, 51] 50) and - word_sless ("'(a <=s b \ sint a \ sint b\ - by transfer simp - -lemma [code]: - \a sint a < sint b\ - by transfer simp - -lemma signed_ordering: \ordering word_sle word_sless\ - apply (standard; transfer) - apply simp_all - using signed_take_bit_decr_length_iff apply force - using signed_take_bit_decr_length_iff apply force - done - -lemma signed_linorder: \class.linorder word_sle word_sless\ - by (standard; transfer) (auto simp add: signed_take_bit_decr_length_iff) - -interpretation signed: linorder word_sle word_sless - by (fact signed_linorder) - -lemma word_sless_eq: - \x x <=s y \ x \ y\ - by (fact signed.less_le) - -lemma word_less_alt: "a < b \ uint a < uint b" - by (fact word_less_def) - -lemma word_zero_le [simp]: "0 \ y" - for y :: "'a::len word" - by (fact word_coorder.extremum) - -lemma word_m1_ge [simp] : "word_pred 0 \ y" (* FIXME: delete *) - by transfer (simp add: take_bit_minus_one_eq_mask mask_eq_exp_minus_1 ) - -lemma word_n1_ge [simp]: "y \ -1" - for y :: "'a::len word" - by (fact word_order.extremum) - -lemmas word_not_simps [simp] = - word_zero_le [THEN leD] word_m1_ge [THEN leD] word_n1_ge [THEN leD] - -lemma word_gt_0: "0 < y \ 0 \ y" - for y :: "'a::len word" - by (simp add: less_le) - -lemmas word_gt_0_no [simp] = word_gt_0 [of "numeral y"] for y - -lemma word_sless_alt: "a sint a < sint b" - by transfer simp - -lemma word_le_nat_alt: "a \ b \ unat a \ unat b" - by transfer (simp add: nat_le_eq_zle) - -lemma word_less_nat_alt: "a < b \ unat a < unat b" - by transfer (auto simp add: less_le [of 0]) - -lemmas unat_mono = word_less_nat_alt [THEN iffD1] - -instance word :: (len) wellorder -proof - fix P :: "'a word \ bool" and a - assume *: "(\b. (\a. a < b \ P a) \ P b)" - have "wf (measure unat)" .. - moreover have "{(a, b :: ('a::len) word). a < b} \ measure unat" - by (auto simp add: word_less_nat_alt) - ultimately have "wf {(a, b :: ('a::len) word). a < b}" - by (rule wf_subset) - then show "P a" using * - by induction blast -qed - -lemma wi_less: - "(word_of_int n < (word_of_int m :: 'a::len word)) = - (n mod 2 ^ LENGTH('a) < m mod 2 ^ LENGTH('a))" - by transfer (simp add: take_bit_eq_mod) - -lemma wi_le: - "(word_of_int n \ (word_of_int m :: 'a::len word)) = - (n mod 2 ^ LENGTH('a) \ m mod 2 ^ LENGTH('a))" - by transfer (simp add: take_bit_eq_mod) - - -subsection \Bit-wise operations\ - -lemma uint_take_bit_eq: - \uint (take_bit n w) = take_bit n (uint w)\ - by transfer (simp add: ac_simps) - -lemma take_bit_word_eq_self: - \take_bit n w = w\ if \LENGTH('a) \ n\ for w :: \'a::len word\ - using that by transfer simp - -lemma take_bit_length_eq [simp]: - \take_bit LENGTH('a) w = w\ for w :: \'a::len word\ - by (rule take_bit_word_eq_self) simp - -lemma bit_word_of_int_iff: - \bit (word_of_int k :: 'a::len word) n \ n < LENGTH('a) \ bit k n\ - by transfer rule - -lemma bit_uint_iff: - \bit (uint w) n \ n < LENGTH('a) \ bit w n\ - for w :: \'a::len word\ - by transfer (simp add: bit_take_bit_iff) - -lemma bit_sint_iff: - \bit (sint w) n \ n \ LENGTH('a) \ bit w (LENGTH('a) - 1) \ bit w n\ - for w :: \'a::len word\ - by transfer (auto simp add: bit_signed_take_bit_iff min_def le_less not_less) - -lemma bit_word_ucast_iff: - \bit (ucast w :: 'b::len word) n \ n < LENGTH('a) \ n < LENGTH('b) \ bit w n\ - for w :: \'a::len word\ - by transfer (simp add: bit_take_bit_iff ac_simps) - -lemma bit_word_scast_iff: - \bit (scast w :: 'b::len word) n \ - n < LENGTH('b) \ (bit w n \ LENGTH('a) \ n \ bit w (LENGTH('a) - Suc 0))\ - for w :: \'a::len word\ - by transfer (auto simp add: bit_signed_take_bit_iff le_less min_def) - -lift_definition shiftl1 :: \'a::len word \ 'a word\ - is \(*) 2\ - by (auto simp add: take_bit_eq_mod intro: mod_mult_cong) - -lemma shiftl1_eq: - \shiftl1 w = word_of_int (2 * uint w)\ - by transfer (simp add: take_bit_eq_mod mod_simps) - -lemma shiftl1_eq_mult_2: - \shiftl1 = (*) 2\ - by (rule ext, transfer) simp - -lemma bit_shiftl1_iff: - \bit (shiftl1 w) n \ 0 < n \ n < LENGTH('a) \ bit w (n - 1)\ - for w :: \'a::len word\ - by (simp add: shiftl1_eq_mult_2 bit_double_iff exp_eq_zero_iff not_le) (simp add: ac_simps) - -lift_definition shiftr1 :: \'a::len word \ 'a word\ - \ \shift right as unsigned or as signed, ie logical or arithmetic\ - is \\k. take_bit LENGTH('a) k div 2\ - by simp - -lemma shiftr1_eq_div_2: - \shiftr1 w = w div 2\ - by transfer simp - -lemma bit_shiftr1_iff: - \bit (shiftr1 w) n \ bit w (Suc n)\ - by transfer (auto simp flip: bit_Suc simp add: bit_take_bit_iff) - -lemma shiftr1_eq: - \shiftr1 w = word_of_int (uint w div 2)\ - by transfer simp - -lemma bit_word_iff_drop_bit_and [code]: - \bit a n \ drop_bit n a AND 1 = 1\ for a :: \'a::len word\ - by (simp add: bit_iff_odd_drop_bit odd_iff_mod_2_eq_one and_one_eq) - -lemma - word_not_def: "NOT (a::'a::len word) = word_of_int (NOT (uint a))" - and word_and_def: "(a::'a word) AND b = word_of_int (uint a AND uint b)" - and word_or_def: "(a::'a word) OR b = word_of_int (uint a OR uint b)" - and word_xor_def: "(a::'a word) XOR b = word_of_int (uint a XOR uint b)" - by (transfer, simp add: take_bit_not_take_bit)+ - -lift_definition setBit :: \'a::len word \ nat \ 'a word\ - is \\k n. set_bit n k\ - by (simp add: take_bit_set_bit_eq) - -lemma set_Bit_eq: - \setBit w n = set_bit n w\ - by transfer simp - -lemma bit_setBit_iff: - \bit (setBit w m) n \ (m = n \ n < LENGTH('a) \ bit w n)\ - for w :: \'a::len word\ - by transfer (auto simp add: bit_set_bit_iff) - -lift_definition clearBit :: \'a::len word \ nat \ 'a word\ - is \\k n. unset_bit n k\ - by (simp add: take_bit_unset_bit_eq) - -lemma clear_Bit_eq: - \clearBit w n = unset_bit n w\ - by transfer simp - -lemma bit_clearBit_iff: - \bit (clearBit w m) n \ m \ n \ bit w n\ - for w :: \'a::len word\ - by transfer (auto simp add: bit_unset_bit_iff) - -definition even_word :: \'a::len word \ bool\ - where [code_abbrev]: \even_word = even\ - -lemma even_word_iff [code]: - \even_word a \ a AND 1 = 0\ - by (simp add: and_one_eq even_iff_mod_2_eq_zero even_word_def) - -lemma map_bit_range_eq_if_take_bit_eq: - \map (bit k) [0.. - if \take_bit n k = take_bit n l\ for k l :: int -using that proof (induction n arbitrary: k l) - case 0 - then show ?case - by simp -next - case (Suc n) - from Suc.prems have \take_bit n (k div 2) = take_bit n (l div 2)\ - by (simp add: take_bit_Suc) - then have \map (bit (k div 2)) [0.. - by (rule Suc.IH) - moreover have \bit (r div 2) = bit r \ Suc\ for r :: int - by (simp add: fun_eq_iff bit_Suc) - moreover from Suc.prems have \even k \ even l\ - by (auto simp add: take_bit_Suc elim!: evenE oddE) arith+ - ultimately show ?case - by (simp only: map_Suc_upt upt_conv_Cons flip: list.map_comp) simp -qed - -lemma - take_bit_word_Bit0_eq [simp]: \take_bit (numeral n) (numeral (num.Bit0 m) :: 'a::len word) - = 2 * take_bit (pred_numeral n) (numeral m)\ (is ?P) - and take_bit_word_Bit1_eq [simp]: \take_bit (numeral n) (numeral (num.Bit1 m) :: 'a::len word) - = 1 + 2 * take_bit (pred_numeral n) (numeral m)\ (is ?Q) - and take_bit_word_minus_Bit0_eq [simp]: \take_bit (numeral n) (- numeral (num.Bit0 m) :: 'a::len word) - = 2 * take_bit (pred_numeral n) (- numeral m)\ (is ?R) - and take_bit_word_minus_Bit1_eq [simp]: \take_bit (numeral n) (- numeral (num.Bit1 m) :: 'a::len word) - = 1 + 2 * take_bit (pred_numeral n) (- numeral (Num.inc m))\ (is ?S) -proof - - define w :: \'a::len word\ - where \w = numeral m\ - moreover define q :: nat - where \q = pred_numeral n\ - ultimately have num: - \numeral m = w\ - \numeral (num.Bit0 m) = 2 * w\ - \numeral (num.Bit1 m) = 1 + 2 * w\ - \numeral (Num.inc m) = 1 + w\ - \pred_numeral n = q\ - \numeral n = Suc q\ - by (simp_all only: w_def q_def numeral_Bit0 [of m] numeral_Bit1 [of m] ac_simps - numeral_inc numeral_eq_Suc flip: mult_2) - have even: \take_bit (Suc q) (2 * w) = 2 * take_bit q w\ for w :: \'a::len word\ - by (rule bit_word_eqI) - (auto simp add: bit_take_bit_iff bit_double_iff) - have odd: \take_bit (Suc q) (1 + 2 * w) = 1 + 2 * take_bit q w\ for w :: \'a::len word\ - by (rule bit_eqI) - (auto simp add: bit_take_bit_iff bit_double_iff even_bit_succ_iff) - show ?P - using even [of w] by (simp add: num) - show ?Q - using odd [of w] by (simp add: num) - show ?R - using even [of \- w\] by (simp add: num) - show ?S - using odd [of \- (1 + w)\] by (simp add: num) -qed - - -subsection \More shift operations\ - -lift_definition signed_drop_bit :: \nat \ 'a word \ 'a::len word\ - is \\n. drop_bit n \ signed_take_bit (LENGTH('a) - Suc 0)\ - using signed_take_bit_decr_length_iff - by (simp add: take_bit_drop_bit) force - -lemma bit_signed_drop_bit_iff: - \bit (signed_drop_bit m w) n \ bit w (if LENGTH('a) - m \ n \ n < LENGTH('a) then LENGTH('a) - 1 else m + n)\ - for w :: \'a::len word\ - apply transfer - apply (auto simp add: bit_drop_bit_eq bit_signed_take_bit_iff not_le min_def) - apply (metis add.commute le_antisym less_diff_conv less_eq_decr_length_iff) - apply (metis le_antisym less_eq_decr_length_iff) - done - -lemma [code]: - \Word.the_int (signed_drop_bit n w) = take_bit LENGTH('a) (drop_bit n (Word.the_signed_int w))\ - for w :: \'a::len word\ - by transfer simp - -lemma signed_drop_bit_signed_drop_bit [simp]: - \signed_drop_bit m (signed_drop_bit n w) = signed_drop_bit (m + n) w\ - for w :: \'a::len word\ - apply (cases \LENGTH('a)\) - apply simp_all - apply (rule bit_word_eqI) - apply (auto simp add: bit_signed_drop_bit_iff not_le less_diff_conv ac_simps) - done - -lemma signed_drop_bit_0 [simp]: - \signed_drop_bit 0 w = w\ - by transfer (simp add: take_bit_signed_take_bit) - -lemma sint_signed_drop_bit_eq: - \sint (signed_drop_bit n w) = drop_bit n (sint w)\ - apply (cases \LENGTH('a)\; cases n) - apply simp_all - apply (rule bit_eqI) - apply (auto simp add: bit_sint_iff bit_drop_bit_eq bit_signed_drop_bit_iff dest: bit_imp_le_length) - done - -lift_definition sshiftr1 :: \'a::len word \ 'a word\ - is \\k. take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - Suc 0) k div 2)\ - by (simp flip: signed_take_bit_decr_length_iff) - -lift_definition bshiftr1 :: \bool \ 'a::len word \ 'a word\ - is \\b k. take_bit LENGTH('a) k div 2 + of_bool b * 2 ^ (LENGTH('a) - Suc 0)\ - by (fact arg_cong) - -lemma sshiftr1_eq_signed_drop_bit_Suc_0: - \sshiftr1 = signed_drop_bit (Suc 0)\ - by (rule ext) (transfer, simp add: drop_bit_Suc) - -lemma sshiftr1_eq: - \sshiftr1 w = word_of_int (sint w div 2)\ - by transfer simp - - -subsection \Rotation\ - -lift_definition word_rotr :: \nat \ 'a::len word \ 'a::len word\ - is \\n k. concat_bit (LENGTH('a) - n mod LENGTH('a)) - (drop_bit (n mod LENGTH('a)) (take_bit LENGTH('a) k)) - (take_bit (n mod LENGTH('a)) k)\ - subgoal for n k l - apply (simp add: concat_bit_def nat_le_iff less_imp_le - take_bit_tightened [of \LENGTH('a)\ k l \n mod LENGTH('a::len)\]) - done - done - -lift_definition word_rotl :: \nat \ 'a::len word \ 'a::len word\ - is \\n k. concat_bit (n mod LENGTH('a)) - (drop_bit (LENGTH('a) - n mod LENGTH('a)) (take_bit LENGTH('a) k)) - (take_bit (LENGTH('a) - n mod LENGTH('a)) k)\ - subgoal for n k l - apply (simp add: concat_bit_def nat_le_iff less_imp_le - take_bit_tightened [of \LENGTH('a)\ k l \LENGTH('a) - n mod LENGTH('a::len)\]) - done - done - -lift_definition word_roti :: \int \ 'a::len word \ 'a::len word\ - is \\r k. concat_bit (LENGTH('a) - nat (r mod int LENGTH('a))) - (drop_bit (nat (r mod int LENGTH('a))) (take_bit LENGTH('a) k)) - (take_bit (nat (r mod int LENGTH('a))) k)\ - subgoal for r k l - apply (simp add: concat_bit_def nat_le_iff less_imp_le - take_bit_tightened [of \LENGTH('a)\ k l \nat (r mod int LENGTH('a::len))\]) - done - done - -lemma word_rotl_eq_word_rotr [code]: - \word_rotl n = (word_rotr (LENGTH('a) - n mod LENGTH('a)) :: 'a::len word \ 'a word)\ - by (rule ext, cases \n mod LENGTH('a) = 0\; transfer) simp_all - -lemma word_roti_eq_word_rotr_word_rotl [code]: - \word_roti i w = - (if i \ 0 then word_rotr (nat i) w else word_rotl (nat (- i)) w)\ -proof (cases \i \ 0\) - case True - moreover define n where \n = nat i\ - ultimately have \i = int n\ - by simp - moreover have \word_roti (int n) = (word_rotr n :: _ \ 'a word)\ - by (rule ext, transfer) (simp add: nat_mod_distrib) - ultimately show ?thesis - by simp -next - case False - moreover define n where \n = nat (- i)\ - ultimately have \i = - int n\ \n > 0\ - by simp_all - moreover have \word_roti (- int n) = (word_rotl n :: _ \ 'a word)\ - by (rule ext, transfer) - (simp add: zmod_zminus1_eq_if flip: of_nat_mod of_nat_diff) - ultimately show ?thesis - by simp -qed - -lemma bit_word_rotr_iff: - \bit (word_rotr m w) n \ - n < LENGTH('a) \ bit w ((n + m) mod LENGTH('a))\ - for w :: \'a::len word\ -proof transfer - fix k :: int and m n :: nat - define q where \q = m mod LENGTH('a)\ - have \q < LENGTH('a)\ - by (simp add: q_def) - then have \q \ LENGTH('a)\ - by simp - have \m mod LENGTH('a) = q\ - by (simp add: q_def) - moreover have \(n + m) mod LENGTH('a) = (n + q) mod LENGTH('a)\ - by (subst mod_add_right_eq [symmetric]) (simp add: \m mod LENGTH('a) = q\) - moreover have \n < LENGTH('a) \ - bit (concat_bit (LENGTH('a) - q) (drop_bit q (take_bit LENGTH('a) k)) (take_bit q k)) n \ - n < LENGTH('a) \ bit k ((n + q) mod LENGTH('a))\ - using \q < LENGTH('a)\ - by (cases \q + n \ LENGTH('a)\) - (auto simp add: bit_concat_bit_iff bit_drop_bit_eq - bit_take_bit_iff le_mod_geq ac_simps) - ultimately show \n < LENGTH('a) \ - bit (concat_bit (LENGTH('a) - m mod LENGTH('a)) - (drop_bit (m mod LENGTH('a)) (take_bit LENGTH('a) k)) - (take_bit (m mod LENGTH('a)) k)) n - \ n < LENGTH('a) \ - (n + m) mod LENGTH('a) < LENGTH('a) \ - bit k ((n + m) mod LENGTH('a))\ - by simp -qed - -lemma bit_word_rotl_iff: - \bit (word_rotl m w) n \ - n < LENGTH('a) \ bit w ((n + (LENGTH('a) - m mod LENGTH('a))) mod LENGTH('a))\ - for w :: \'a::len word\ - by (simp add: word_rotl_eq_word_rotr bit_word_rotr_iff) - -lemma bit_word_roti_iff: - \bit (word_roti k w) n \ - n < LENGTH('a) \ bit w (nat ((int n + k) mod int LENGTH('a)))\ - for w :: \'a::len word\ -proof transfer - fix k l :: int and n :: nat - define m where \m = nat (k mod int LENGTH('a))\ - have \m < LENGTH('a)\ - by (simp add: nat_less_iff m_def) - then have \m \ LENGTH('a)\ - by simp - have \k mod int LENGTH('a) = int m\ - by (simp add: nat_less_iff m_def) - moreover have \(int n + k) mod int LENGTH('a) = int ((n + m) mod LENGTH('a))\ - by (subst mod_add_right_eq [symmetric]) (simp add: of_nat_mod \k mod int LENGTH('a) = int m\) - moreover have \n < LENGTH('a) \ - bit (concat_bit (LENGTH('a) - m) (drop_bit m (take_bit LENGTH('a) l)) (take_bit m l)) n \ - n < LENGTH('a) \ bit l ((n + m) mod LENGTH('a))\ - using \m < LENGTH('a)\ - by (cases \m + n \ LENGTH('a)\) - (auto simp add: bit_concat_bit_iff bit_drop_bit_eq - bit_take_bit_iff nat_less_iff not_le not_less ac_simps - le_diff_conv le_mod_geq) - ultimately show \n < LENGTH('a) - \ bit (concat_bit (LENGTH('a) - nat (k mod int LENGTH('a))) - (drop_bit (nat (k mod int LENGTH('a))) (take_bit LENGTH('a) l)) - (take_bit (nat (k mod int LENGTH('a))) l)) n \ - n < LENGTH('a) - \ nat ((int n + k) mod int LENGTH('a)) < LENGTH('a) - \ bit l (nat ((int n + k) mod int LENGTH('a)))\ - by simp -qed - -lemma uint_word_rotr_eq: - \uint (word_rotr n w) = concat_bit (LENGTH('a) - n mod LENGTH('a)) - (drop_bit (n mod LENGTH('a)) (uint w)) - (uint (take_bit (n mod LENGTH('a)) w))\ - for w :: \'a::len word\ - apply transfer - apply (simp add: concat_bit_def take_bit_drop_bit push_bit_take_bit min_def) - using mod_less_divisor not_less apply blast - done - -lemma [code]: - \Word.the_int (word_rotr n w) = concat_bit (LENGTH('a) - n mod LENGTH('a)) - (drop_bit (n mod LENGTH('a)) (Word.the_int w)) - (Word.the_int (take_bit (n mod LENGTH('a)) w))\ - for w :: \'a::len word\ - using uint_word_rotr_eq [of n w] by simp - - -subsection \Split and cat operations\ - -lift_definition word_cat :: \'a::len word \ 'b::len word \ 'c::len word\ - is \\k l. concat_bit LENGTH('b) l (take_bit LENGTH('a) k)\ - by (simp add: bit_eq_iff bit_concat_bit_iff bit_take_bit_iff) - -lemma word_cat_eq: - \(word_cat v w :: 'c::len word) = push_bit LENGTH('b) (ucast v) + ucast w\ - for v :: \'a::len word\ and w :: \'b::len word\ - by transfer (simp add: concat_bit_eq ac_simps) - -lemma word_cat_eq' [code]: - \word_cat a b = word_of_int (concat_bit LENGTH('b) (uint b) (uint a))\ - for a :: \'a::len word\ and b :: \'b::len word\ - by transfer (simp add: concat_bit_take_bit_eq) - -lemma bit_word_cat_iff: - \bit (word_cat v w :: 'c::len word) n \ n < LENGTH('c) \ (if n < LENGTH('b) then bit w n else bit v (n - LENGTH('b)))\ - for v :: \'a::len word\ and w :: \'b::len word\ - by transfer (simp add: bit_concat_bit_iff bit_take_bit_iff) - -definition word_split :: \'a::len word \ 'b::len word \ 'c::len word\ - where \word_split w = - (ucast (drop_bit LENGTH('c) w) :: 'b::len word, ucast w :: 'c::len word)\ - -definition word_rcat :: \'a::len word list \ 'b::len word\ - where \word_rcat = word_of_int \ horner_sum uint (2 ^ LENGTH('a)) \ rev\ - -abbreviation (input) max_word :: \'a::len word\ - \ \Largest representable machine integer.\ - where "max_word \ - 1" - - -subsection \More on conversions\ - -lemma int_word_sint: - \sint (word_of_int x :: 'a::len word) = (x + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - 2 ^ (LENGTH('a) - 1)\ - by transfer (simp flip: take_bit_eq_mod add: signed_take_bit_eq_take_bit_shift) - -lemma sint_sbintrunc': "sint (word_of_int bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) bin" - by simp - -lemma uint_sint: "uint w = take_bit LENGTH('a) (sint w)" - for w :: "'a::len word" - by transfer (simp add: take_bit_signed_take_bit) - -lemma bintr_uint: "LENGTH('a) \ n \ take_bit n (uint w) = uint w" - for w :: "'a::len word" - by transfer (simp add: min_def) - -lemma wi_bintr: - "LENGTH('a::len) \ n \ - word_of_int (take_bit n w) = (word_of_int w :: 'a word)" - by transfer simp - -lemma word_numeral_alt: "numeral b = word_of_int (numeral b)" - by (induct b, simp_all only: numeral.simps word_of_int_homs) - -declare word_numeral_alt [symmetric, code_abbrev] - -lemma word_neg_numeral_alt: "- numeral b = word_of_int (- numeral b)" - by (simp only: word_numeral_alt wi_hom_neg) - -declare word_neg_numeral_alt [symmetric, code_abbrev] - -lemma uint_bintrunc [simp]: - "uint (numeral bin :: 'a word) = - take_bit (LENGTH('a::len)) (numeral bin)" - by transfer rule - -lemma uint_bintrunc_neg [simp]: - "uint (- numeral bin :: 'a word) = take_bit (LENGTH('a::len)) (- numeral bin)" - by transfer rule - -lemma sint_sbintrunc [simp]: - "sint (numeral bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) (numeral bin)" - by transfer simp - -lemma sint_sbintrunc_neg [simp]: - "sint (- numeral bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) (- numeral bin)" - by transfer simp - -lemma unat_bintrunc [simp]: - "unat (numeral bin :: 'a::len word) = nat (take_bit (LENGTH('a)) (numeral bin))" - by transfer simp - -lemma unat_bintrunc_neg [simp]: - "unat (- numeral bin :: 'a::len word) = nat (take_bit (LENGTH('a)) (- numeral bin))" - by transfer simp - -lemma size_0_eq: "size w = 0 \ v = w" - for v w :: "'a::len word" - by transfer simp - -lemma uint_ge_0 [iff]: "0 \ uint x" - by (fact unsigned_greater_eq) - -lemma uint_lt2p [iff]: "uint x < 2 ^ LENGTH('a)" - for x :: "'a::len word" - by (fact unsigned_less) - -lemma sint_ge: "- (2 ^ (LENGTH('a) - 1)) \ sint x" - for x :: "'a::len word" - using sint_greater_eq [of x] by simp - -lemma sint_lt: "sint x < 2 ^ (LENGTH('a) - 1)" - for x :: "'a::len word" - using sint_less [of x] by simp - -lemma uint_m2p_neg: "uint x - 2 ^ LENGTH('a) < 0" - for x :: "'a::len word" - by (simp only: diff_less_0_iff_less uint_lt2p) - -lemma uint_m2p_not_non_neg: "\ 0 \ uint x - 2 ^ LENGTH('a)" - for x :: "'a::len word" - by (simp only: not_le uint_m2p_neg) - -lemma lt2p_lem: "LENGTH('a) \ n \ uint w < 2 ^ n" - for w :: "'a::len word" - using uint_bounded [of w] by (rule less_le_trans) simp - -lemma uint_le_0_iff [simp]: "uint x \ 0 \ uint x = 0" - by (fact uint_ge_0 [THEN leD, THEN antisym_conv1]) - -lemma uint_nat: "uint w = int (unat w)" - by transfer simp - -lemma uint_numeral: "uint (numeral b :: 'a::len word) = numeral b mod 2 ^ LENGTH('a)" - by (simp flip: take_bit_eq_mod add: of_nat_take_bit) - -lemma uint_neg_numeral: "uint (- numeral b :: 'a::len word) = - numeral b mod 2 ^ LENGTH('a)" - by (simp flip: take_bit_eq_mod add: of_nat_take_bit) - -lemma unat_numeral: "unat (numeral b :: 'a::len word) = numeral b mod 2 ^ LENGTH('a)" - by transfer (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq) - -lemma sint_numeral: - "sint (numeral b :: 'a::len word) = - (numeral b + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - 2 ^ (LENGTH('a) - 1)" - apply (transfer fixing: b) - using int_word_sint [of \numeral b\] - apply simp - done - -lemma word_of_int_0 [simp, code_post]: "word_of_int 0 = 0" - by (fact of_int_0) - -lemma word_of_int_1 [simp, code_post]: "word_of_int 1 = 1" - by (fact of_int_1) - -lemma word_of_int_neg_1 [simp]: "word_of_int (- 1) = - 1" - by (simp add: wi_hom_syms) - -lemma word_of_int_numeral [simp] : "(word_of_int (numeral bin) :: 'a::len word) = numeral bin" - by (fact of_int_numeral) - -lemma word_of_int_neg_numeral [simp]: - "(word_of_int (- numeral bin) :: 'a::len word) = - numeral bin" - by (fact of_int_neg_numeral) - -lemma word_int_case_wi: - "word_int_case f (word_of_int i :: 'b word) = f (i mod 2 ^ LENGTH('b::len))" - by transfer (simp add: take_bit_eq_mod) - -lemma word_int_split: - "P (word_int_case f x) = - (\i. x = (word_of_int i :: 'b::len word) \ 0 \ i \ i < 2 ^ LENGTH('b) \ P (f i))" - by transfer (auto simp add: take_bit_eq_mod) - -lemma word_int_split_asm: - "P (word_int_case f x) = - (\n. x = (word_of_int n :: 'b::len word) \ 0 \ n \ n < 2 ^ LENGTH('b::len) \ \ P (f n))" - by transfer (auto simp add: take_bit_eq_mod) - -lemma uint_range_size: "0 \ uint w \ uint w < 2 ^ size w" - by transfer simp - -lemma sint_range_size: "- (2 ^ (size w - Suc 0)) \ sint w \ sint w < 2 ^ (size w - Suc 0)" - by (simp add: word_size sint_greater_eq sint_less) - -lemma sint_above_size: "2 ^ (size w - 1) \ x \ sint w < x" - for w :: "'a::len word" - unfolding word_size by (rule less_le_trans [OF sint_lt]) - -lemma sint_below_size: "x \ - (2 ^ (size w - 1)) \ x \ sint w" - for w :: "'a::len word" - unfolding word_size by (rule order_trans [OF _ sint_ge]) - - -subsection \Testing bits\ - -lemma bin_nth_uint_imp: "bit (uint w) n \ n < LENGTH('a)" - for w :: "'a::len word" - by transfer (simp add: bit_take_bit_iff) - -lemma bin_nth_sint: - "LENGTH('a) \ n \ - bit (sint w) n = bit (sint w) (LENGTH('a) - 1)" - for w :: "'a::len word" - by (transfer fixing: n) (simp add: bit_signed_take_bit_iff le_diff_conv min_def) - -lemma num_of_bintr': - "take_bit (LENGTH('a::len)) (numeral a :: int) = (numeral b) \ - numeral a = (numeral b :: 'a word)" -proof (transfer fixing: a b) - assume \take_bit LENGTH('a) (numeral a :: int) = numeral b\ - then have \take_bit LENGTH('a) (take_bit LENGTH('a) (numeral a :: int)) = take_bit LENGTH('a) (numeral b)\ - by simp - then show \take_bit LENGTH('a) (numeral a :: int) = take_bit LENGTH('a) (numeral b)\ - by simp -qed - -lemma num_of_sbintr': - "signed_take_bit (LENGTH('a::len) - 1) (numeral a :: int) = (numeral b) \ - numeral a = (numeral b :: 'a word)" -proof (transfer fixing: a b) - assume \signed_take_bit (LENGTH('a) - 1) (numeral a :: int) = numeral b\ - then have \take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - 1) (numeral a :: int)) = take_bit LENGTH('a) (numeral b)\ - by simp - then show \take_bit LENGTH('a) (numeral a :: int) = take_bit LENGTH('a) (numeral b)\ - by (simp add: take_bit_signed_take_bit) -qed - -lemma num_abs_bintr: - "(numeral x :: 'a word) = - word_of_int (take_bit (LENGTH('a::len)) (numeral x))" - by transfer simp - -lemma num_abs_sbintr: - "(numeral x :: 'a word) = - word_of_int (signed_take_bit (LENGTH('a::len) - 1) (numeral x))" - by transfer (simp add: take_bit_signed_take_bit) - -text \ - \cast\ -- note, no arg for new length, as it's determined by type of result, - thus in \cast w = w\, the type means cast to length of \w\! -\ - -lemma bit_ucast_iff: - \bit (ucast a :: 'a::len word) n \ n < LENGTH('a::len) \ Parity.bit a n\ - by transfer (simp add: bit_take_bit_iff) - -lemma ucast_id [simp]: "ucast w = w" - by transfer simp - -lemma scast_id [simp]: "scast w = w" - by transfer (simp add: take_bit_signed_take_bit) - -lemma ucast_mask_eq: - \ucast (mask n :: 'b word) = mask (min LENGTH('b::len) n)\ - by (simp add: bit_eq_iff) (auto simp add: bit_mask_iff bit_ucast_iff exp_eq_zero_iff) - -\ \literal u(s)cast\ -lemma ucast_bintr [simp]: - "ucast (numeral w :: 'a::len word) = - word_of_int (take_bit (LENGTH('a)) (numeral w))" - by transfer simp - -(* TODO: neg_numeral *) - -lemma scast_sbintr [simp]: - "scast (numeral w ::'a::len word) = - word_of_int (signed_take_bit (LENGTH('a) - Suc 0) (numeral w))" - by transfer simp - -lemma source_size: "source_size (c::'a::len word \ _) = LENGTH('a)" - by transfer simp - -lemma target_size: "target_size (c::_ \ 'b::len word) = LENGTH('b)" - by transfer simp - -lemma is_down: "is_down c \ LENGTH('b) \ LENGTH('a)" - for c :: "'a::len word \ 'b::len word" - by transfer simp - -lemma is_up: "is_up c \ LENGTH('a) \ LENGTH('b)" - for c :: "'a::len word \ 'b::len word" - by transfer simp - -lemma is_up_down: - \is_up c \ is_down d\ - for c :: \'a::len word \ 'b::len word\ - and d :: \'b::len word \ 'a::len word\ - by transfer simp - -context - fixes dummy_types :: \'a::len \ 'b::len\ -begin - -private abbreviation (input) UCAST :: \'a::len word \ 'b::len word\ - where \UCAST == ucast\ - -private abbreviation (input) SCAST :: \'a::len word \ 'b::len word\ - where \SCAST == scast\ - -lemma down_cast_same: - \UCAST = scast\ if \is_down UCAST\ - by (rule ext, use that in transfer) (simp add: take_bit_signed_take_bit) - -lemma sint_up_scast: - \sint (SCAST w) = sint w\ if \is_up SCAST\ - using that by transfer (simp add: min_def Suc_leI le_diff_iff) - -lemma uint_up_ucast: - \uint (UCAST w) = uint w\ if \is_up UCAST\ - using that by transfer (simp add: min_def) - -lemma ucast_up_ucast: - \ucast (UCAST w) = ucast w\ if \is_up UCAST\ - using that by transfer (simp add: ac_simps) - -lemma ucast_up_ucast_id: - \ucast (UCAST w) = w\ if \is_up UCAST\ - using that by (simp add: ucast_up_ucast) - -lemma scast_up_scast: - \scast (SCAST w) = scast w\ if \is_up SCAST\ - using that by transfer (simp add: ac_simps) - -lemma scast_up_scast_id: - \scast (SCAST w) = w\ if \is_up SCAST\ - using that by (simp add: scast_up_scast) - -lemma isduu: - \is_up UCAST\ if \is_down d\ - for d :: \'b word \ 'a word\ - using that is_up_down [of UCAST d] by simp - -lemma isdus: - \is_up SCAST\ if \is_down d\ - for d :: \'b word \ 'a word\ - using that is_up_down [of SCAST d] by simp - -lemmas ucast_down_ucast_id = isduu [THEN ucast_up_ucast_id] -lemmas scast_down_scast_id = isdus [THEN scast_up_scast_id] - -lemma up_ucast_surj: - \surj (ucast :: 'b word \ 'a word)\ if \is_up UCAST\ - by (rule surjI) (use that in \rule ucast_up_ucast_id\) - -lemma up_scast_surj: - \surj (scast :: 'b word \ 'a word)\ if \is_up SCAST\ - by (rule surjI) (use that in \rule scast_up_scast_id\) - -lemma down_ucast_inj: - \inj_on UCAST A\ if \is_down (ucast :: 'b word \ 'a word)\ - by (rule inj_on_inverseI) (use that in \rule ucast_down_ucast_id\) - -lemma down_scast_inj: - \inj_on SCAST A\ if \is_down (scast :: 'b word \ 'a word)\ - by (rule inj_on_inverseI) (use that in \rule scast_down_scast_id\) - -lemma ucast_down_wi: - \UCAST (word_of_int x) = word_of_int x\ if \is_down UCAST\ - using that by transfer simp - -lemma ucast_down_no: - \UCAST (numeral bin) = numeral bin\ if \is_down UCAST\ - using that by transfer simp - -end - -lemmas word_log_defs = word_and_def word_or_def word_xor_def word_not_def - -lemma bit_last_iff: - \bit w (LENGTH('a) - Suc 0) \ sint w < 0\ (is \?P \ ?Q\) - for w :: \'a::len word\ -proof - - have \?P \ bit (uint w) (LENGTH('a) - Suc 0)\ - by (simp add: bit_uint_iff) - also have \\ \ ?Q\ - by (simp add: sint_uint) - finally show ?thesis . -qed - -lemma drop_bit_eq_zero_iff_not_bit_last: - \drop_bit (LENGTH('a) - Suc 0) w = 0 \ \ bit w (LENGTH('a) - Suc 0)\ - for w :: "'a::len word" - apply (cases \LENGTH('a)\) - apply simp_all - apply (simp add: bit_iff_odd_drop_bit) - apply transfer - apply (simp add: take_bit_drop_bit) - apply (auto simp add: drop_bit_eq_div take_bit_eq_mod min_def) - apply (auto elim!: evenE) - apply (metis div_exp_eq mod_div_trivial mult.commute nonzero_mult_div_cancel_left power_Suc0_right power_add zero_neq_numeral) - done - - -subsection \Word Arithmetic\ - -lemmas word_div_no [simp] = word_div_def [of "numeral a" "numeral b"] for a b -lemmas word_mod_no [simp] = word_mod_def [of "numeral a" "numeral b"] for a b -lemmas word_less_no [simp] = word_less_def [of "numeral a" "numeral b"] for a b -lemmas word_le_no [simp] = word_le_def [of "numeral a" "numeral b"] for a b -lemmas word_sless_no [simp] = word_sless_eq [of "numeral a" "numeral b"] for a b -lemmas word_sle_no [simp] = word_sle_eq [of "numeral a" "numeral b"] for a b - -lemma size_0_same': "size w = 0 \ w = v" - for v w :: "'a::len word" - by (unfold word_size) simp - -lemmas size_0_same = size_0_same' [unfolded word_size] - -lemmas unat_eq_0 = unat_0_iff -lemmas unat_eq_zero = unat_0_iff - - -subsection \Transferring goals from words to ints\ - -lemma word_ths: - shows word_succ_p1: "word_succ a = a + 1" - and word_pred_m1: "word_pred a = a - 1" - and word_pred_succ: "word_pred (word_succ a) = a" - and word_succ_pred: "word_succ (word_pred a) = a" - and word_mult_succ: "word_succ a * b = b + a * b" - by (transfer, simp add: algebra_simps)+ - -lemma uint_cong: "x = y \ uint x = uint y" - by simp - -lemma uint_word_ariths: - fixes a b :: "'a::len word" - shows "uint (a + b) = (uint a + uint b) mod 2 ^ LENGTH('a::len)" - and "uint (a - b) = (uint a - uint b) mod 2 ^ LENGTH('a)" - and "uint (a * b) = uint a * uint b mod 2 ^ LENGTH('a)" - and "uint (- a) = - uint a mod 2 ^ LENGTH('a)" - and "uint (word_succ a) = (uint a + 1) mod 2 ^ LENGTH('a)" - and "uint (word_pred a) = (uint a - 1) mod 2 ^ LENGTH('a)" - and "uint (0 :: 'a word) = 0 mod 2 ^ LENGTH('a)" - and "uint (1 :: 'a word) = 1 mod 2 ^ LENGTH('a)" - by (simp_all only: word_arith_wis uint_word_of_int_eq flip: take_bit_eq_mod) - -lemma uint_word_arith_bintrs: - fixes a b :: "'a::len word" - shows "uint (a + b) = take_bit (LENGTH('a)) (uint a + uint b)" - and "uint (a - b) = take_bit (LENGTH('a)) (uint a - uint b)" - and "uint (a * b) = take_bit (LENGTH('a)) (uint a * uint b)" - and "uint (- a) = take_bit (LENGTH('a)) (- uint a)" - and "uint (word_succ a) = take_bit (LENGTH('a)) (uint a + 1)" - and "uint (word_pred a) = take_bit (LENGTH('a)) (uint a - 1)" - and "uint (0 :: 'a word) = take_bit (LENGTH('a)) 0" - and "uint (1 :: 'a word) = take_bit (LENGTH('a)) 1" - by (simp_all add: uint_word_ariths take_bit_eq_mod) - -lemma sint_word_ariths: - fixes a b :: "'a::len word" - shows "sint (a + b) = signed_take_bit (LENGTH('a) - 1) (sint a + sint b)" - and "sint (a - b) = signed_take_bit (LENGTH('a) - 1) (sint a - sint b)" - and "sint (a * b) = signed_take_bit (LENGTH('a) - 1) (sint a * sint b)" - and "sint (- a) = signed_take_bit (LENGTH('a) - 1) (- sint a)" - and "sint (word_succ a) = signed_take_bit (LENGTH('a) - 1) (sint a + 1)" - and "sint (word_pred a) = signed_take_bit (LENGTH('a) - 1) (sint a - 1)" - and "sint (0 :: 'a word) = signed_take_bit (LENGTH('a) - 1) 0" - and "sint (1 :: 'a word) = signed_take_bit (LENGTH('a) - 1) 1" - apply transfer apply (simp add: signed_take_bit_add) - apply transfer apply (simp add: signed_take_bit_diff) - apply transfer apply (simp add: signed_take_bit_mult) - apply transfer apply (simp add: signed_take_bit_minus) - apply (metis of_int_sint scast_id sint_sbintrunc' wi_hom_succ) - apply (metis of_int_sint scast_id sint_sbintrunc' wi_hom_pred) - apply (simp_all add: sint_uint) - done - -lemma word_pred_0_n1: "word_pred 0 = word_of_int (- 1)" - unfolding word_pred_m1 by simp - -lemma succ_pred_no [simp]: - "word_succ (numeral w) = numeral w + 1" - "word_pred (numeral w) = numeral w - 1" - "word_succ (- numeral w) = - numeral w + 1" - "word_pred (- numeral w) = - numeral w - 1" - by (simp_all add: word_succ_p1 word_pred_m1) - -lemma word_sp_01 [simp]: - "word_succ (- 1) = 0 \ word_succ 0 = 1 \ word_pred 0 = - 1 \ word_pred 1 = 0" - by (simp_all add: word_succ_p1 word_pred_m1) - -\ \alternative approach to lifting arithmetic equalities\ -lemma word_of_int_Ex: "\y. x = word_of_int y" - by (rule_tac x="uint x" in exI) simp - - -subsection \Order on fixed-length words\ - -lift_definition udvd :: \'a::len word \ 'a::len word \ bool\ (infixl \udvd\ 50) - is \\k l. take_bit LENGTH('a) k dvd take_bit LENGTH('a) l\ by simp - -lemma udvd_iff_dvd: - \x udvd y \ unat x dvd unat y\ - by transfer (simp add: nat_dvd_iff) - -lemma udvd_iff_dvd_int: - \v udvd w \ uint v dvd uint w\ - by transfer rule - -lemma udvdI [intro]: - \v udvd w\ if \unat w = unat v * unat u\ -proof - - from that have \unat v dvd unat w\ .. - then show ?thesis - by (simp add: udvd_iff_dvd) -qed - -lemma udvdE [elim]: - fixes v w :: \'a::len word\ - assumes \v udvd w\ - obtains u :: \'a word\ where \unat w = unat v * unat u\ -proof (cases \v = 0\) - case True - moreover from True \v udvd w\ have \w = 0\ - by transfer simp - ultimately show thesis - using that by simp -next - case False - then have \unat v > 0\ - by (simp add: unat_gt_0) - from \v udvd w\ have \unat v dvd unat w\ - by (simp add: udvd_iff_dvd) - then obtain n where \unat w = unat v * n\ .. - moreover have \n < 2 ^ LENGTH('a)\ - proof (rule ccontr) - assume \\ n < 2 ^ LENGTH('a)\ - then have \n \ 2 ^ LENGTH('a)\ - by (simp add: not_le) - then have \unat v * n \ 2 ^ LENGTH('a)\ - using \unat v > 0\ mult_le_mono [of 1 \unat v\ \2 ^ LENGTH('a)\ n] - by simp - with \unat w = unat v * n\ - have \unat w \ 2 ^ LENGTH('a)\ - by simp - with unsigned_less [of w, where ?'a = nat] show False - by linarith - qed - ultimately have \unat w = unat v * unat (word_of_nat n :: 'a word)\ - by (auto simp add: take_bit_nat_eq_self_iff intro: sym) - with that show thesis . -qed - -lemma udvd_imp_mod_eq_0: - \w mod v = 0\ if \v udvd w\ - using that by transfer simp - -lemma mod_eq_0_imp_udvd [intro?]: - \v udvd w\ if \w mod v = 0\ -proof - - from that have \unat (w mod v) = unat 0\ - by simp - then have \unat w mod unat v = 0\ - by (simp add: unat_mod_distrib) - then have \unat v dvd unat w\ .. - then show ?thesis - by (simp add: udvd_iff_dvd) -qed - -lemma udvd_imp_dvd: - \v dvd w\ if \v udvd w\ for v w :: \'a::len word\ -proof - - from that obtain u :: \'a word\ where \unat w = unat v * unat u\ .. - then have \(word_of_nat (unat w) :: 'a word) = word_of_nat (unat v * unat u)\ - by simp - then have \w = v * u\ - by simp - then show \v dvd w\ .. -qed - -lemma exp_dvd_iff_exp_udvd: - \2 ^ n dvd w \ 2 ^ n udvd w\ for v w :: \'a::len word\ -proof - assume \2 ^ n udvd w\ then show \2 ^ n dvd w\ - by (rule udvd_imp_dvd) -next - assume \2 ^ n dvd w\ - then obtain u :: \'a word\ where \w = 2 ^ n * u\ .. - then have \w = push_bit n u\ - by (simp add: push_bit_eq_mult) - then show \2 ^ n udvd w\ - by transfer (simp add: take_bit_push_bit dvd_eq_mod_eq_0 flip: take_bit_eq_mod) -qed - -lemma udvd_nat_alt: - \a udvd b \ (\n. unat b = n * unat a)\ - by (auto simp add: udvd_iff_dvd) - -lemma udvd_unfold_int: - \a udvd b \ (\n\0. uint b = n * uint a)\ - apply (auto elim!: dvdE simp add: udvd_iff_dvd) - apply (simp only: uint_nat) - apply auto - using of_nat_0_le_iff apply blast - apply (simp only: unat_eq_nat_uint) - apply (simp add: nat_mult_distrib) - done - -lemma unat_minus_one: - \unat (w - 1) = unat w - 1\ if \w \ 0\ -proof - - have "0 \ uint w" by (fact uint_nonnegative) - moreover from that have "0 \ uint w" - by (simp add: uint_0_iff) - ultimately have "1 \ uint w" - by arith - from uint_lt2p [of w] have "uint w - 1 < 2 ^ LENGTH('a)" - by arith - with \1 \ uint w\ have "(uint w - 1) mod 2 ^ LENGTH('a) = uint w - 1" - by (auto intro: mod_pos_pos_trivial) - with \1 \ uint w\ have "nat ((uint w - 1) mod 2 ^ LENGTH('a)) = nat (uint w) - 1" - by (auto simp del: nat_uint_eq) - then show ?thesis - by (simp only: unat_eq_nat_uint word_arith_wis mod_diff_right_eq) - (metis of_int_1 uint_word_of_int unsigned_1) -qed - -lemma measure_unat: "p \ 0 \ unat (p - 1) < unat p" - by (simp add: unat_minus_one) (simp add: unat_0_iff [symmetric]) - -lemmas uint_add_ge0 [simp] = add_nonneg_nonneg [OF uint_ge_0 uint_ge_0] -lemmas uint_mult_ge0 [simp] = mult_nonneg_nonneg [OF uint_ge_0 uint_ge_0] - -lemma uint_sub_lt2p [simp]: "uint x - uint y < 2 ^ LENGTH('a)" - for x :: "'a::len word" and y :: "'b::len word" - using uint_ge_0 [of y] uint_lt2p [of x] by arith - - -subsection \Conditions for the addition (etc) of two words to overflow\ - -lemma uint_add_lem: - "(uint x + uint y < 2 ^ LENGTH('a)) = - (uint (x + y) = uint x + uint y)" - for x y :: "'a::len word" - by (metis add.right_neutral add_mono_thms_linordered_semiring(1) mod_pos_pos_trivial of_nat_0_le_iff uint_lt2p uint_nat uint_word_ariths(1)) - -lemma uint_mult_lem: - "(uint x * uint y < 2 ^ LENGTH('a)) = - (uint (x * y) = uint x * uint y)" - for x y :: "'a::len word" - by (metis mod_pos_pos_trivial uint_lt2p uint_mult_ge0 uint_word_ariths(3)) - -lemma uint_sub_lem: "uint x \ uint y \ uint (x - y) = uint x - uint y" - by (metis diff_ge_0_iff_ge of_nat_0_le_iff uint_nat uint_sub_lt2p uint_word_of_int unique_euclidean_semiring_numeral_class.mod_less word_sub_wi) - -lemma uint_add_le: "uint (x + y) \ uint x + uint y" - unfolding uint_word_ariths by (simp add: zmod_le_nonneg_dividend) - -lemma uint_sub_ge: "uint (x - y) \ uint x - uint y" - unfolding uint_word_ariths - by (simp flip: take_bit_eq_mod add: take_bit_int_greater_eq_self_iff) - -lemma int_mod_ge: \a \ a mod n\ if \a < n\ \0 < n\ - for a n :: int -proof (cases \a < 0\) - case True - with \0 < n\ show ?thesis - by (metis less_trans not_less pos_mod_conj) - -next - case False - with \a < n\ show ?thesis - by simp -qed - -lemma mod_add_if_z: - "x < z \ y < z \ 0 \ y \ 0 \ x \ 0 \ z \ - (x + y) mod z = (if x + y < z then x + y else x + y - z)" - for x y z :: int - apply (auto simp add: not_less) - apply (rule antisym) - apply (metis diff_ge_0_iff_ge minus_mod_self2 zmod_le_nonneg_dividend) - apply (simp only: flip: minus_mod_self2 [of \x + y\ z]) - apply (metis add.commute add_less_cancel_left add_mono_thms_linordered_field(5) diff_add_cancel diff_ge_0_iff_ge mod_pos_pos_trivial order_refl) - done - -lemma uint_plus_if': - "uint (a + b) = - (if uint a + uint b < 2 ^ LENGTH('a) then uint a + uint b - else uint a + uint b - 2 ^ LENGTH('a))" - for a b :: "'a::len word" - using mod_add_if_z [of "uint a" _ "uint b"] by (simp add: uint_word_ariths) - -lemma mod_sub_if_z: - "x < z \ y < z \ 0 \ y \ 0 \ x \ 0 \ z \ - (x - y) mod z = (if y \ x then x - y else x - y + z)" - for x y z :: int - apply (auto simp add: not_le) - apply (rule antisym) - apply (simp only: flip: mod_add_self2 [of \x - y\ z]) - apply (rule zmod_le_nonneg_dividend) - apply simp - apply (metis add.commute add.right_neutral add_le_cancel_left diff_ge_0_iff_ge int_mod_ge le_less le_less_trans mod_add_self1 not_less) - done - -lemma uint_sub_if': - "uint (a - b) = - (if uint b \ uint a then uint a - uint b - else uint a - uint b + 2 ^ LENGTH('a))" - for a b :: "'a::len word" - using mod_sub_if_z [of "uint a" _ "uint b"] by (simp add: uint_word_ariths) - - -subsection \Definition of \uint_arith\\ - -lemma word_of_int_inverse: - "word_of_int r = a \ 0 \ r \ r < 2 ^ LENGTH('a) \ uint a = r" - for a :: "'a::len word" - apply transfer - apply (drule sym) - apply (simp add: take_bit_int_eq_self) - done - -lemma uint_split: - "P (uint x) = (\i. word_of_int i = x \ 0 \ i \ i < 2^LENGTH('a) \ P i)" - for x :: "'a::len word" - by transfer (auto simp add: take_bit_eq_mod) - -lemma uint_split_asm: - "P (uint x) = (\i. word_of_int i = x \ 0 \ i \ i < 2^LENGTH('a) \ \ P i)" - for x :: "'a::len word" - by auto (metis take_bit_int_eq_self_iff) - -lemmas uint_splits = uint_split uint_split_asm - -lemmas uint_arith_simps = - word_le_def word_less_alt - word_uint_eq_iff - uint_sub_if' uint_plus_if' - -\ \use this to stop, eg. \2 ^ LENGTH(32)\ being simplified\ -lemma power_False_cong: "False \ a ^ b = c ^ d" - by auto - -\ \\uint_arith_tac\: reduce to arithmetic on int, try to solve by arith\ -ML \ -val uint_arith_simpset = - @{context} - |> fold Simplifier.add_simp @{thms uint_arith_simps} - |> fold Splitter.add_split @{thms if_split_asm} - |> fold Simplifier.add_cong @{thms power_False_cong} - |> simpset_of; - -fun uint_arith_tacs ctxt = - let - fun arith_tac' n t = - Arith_Data.arith_tac ctxt n t - handle Cooper.COOPER _ => Seq.empty; - in - [ clarify_tac ctxt 1, - full_simp_tac (put_simpset uint_arith_simpset ctxt) 1, - ALLGOALS (full_simp_tac - (put_simpset HOL_ss ctxt - |> fold Splitter.add_split @{thms uint_splits} - |> fold Simplifier.add_cong @{thms power_False_cong})), - rewrite_goals_tac ctxt @{thms word_size}, - ALLGOALS (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN - REPEAT (eresolve_tac ctxt [conjE] n) THEN - REPEAT (dresolve_tac ctxt @{thms word_of_int_inverse} n - THEN assume_tac ctxt n - THEN assume_tac ctxt n)), - TRYALL arith_tac' ] - end - -fun uint_arith_tac ctxt = SELECT_GOAL (EVERY (uint_arith_tacs ctxt)) -\ - -method_setup uint_arith = - \Scan.succeed (SIMPLE_METHOD' o uint_arith_tac)\ - "solving word arithmetic via integers and arith" - - -subsection \More on overflows and monotonicity\ - -lemma no_plus_overflow_uint_size: "x \ x + y \ uint x + uint y < 2 ^ size x" - for x y :: "'a::len word" - unfolding word_size by uint_arith - -lemmas no_olen_add = no_plus_overflow_uint_size [unfolded word_size] - -lemma no_ulen_sub: "x \ x - y \ uint y \ uint x" - for x y :: "'a::len word" - by uint_arith - -lemma no_olen_add': "x \ y + x \ uint y + uint x < 2 ^ LENGTH('a)" - for x y :: "'a::len word" - by (simp add: ac_simps no_olen_add) - -lemmas olen_add_eqv = trans [OF no_olen_add no_olen_add' [symmetric]] - -lemmas uint_plus_simple_iff = trans [OF no_olen_add uint_add_lem] -lemmas uint_plus_simple = uint_plus_simple_iff [THEN iffD1] -lemmas uint_minus_simple_iff = trans [OF no_ulen_sub uint_sub_lem] -lemmas uint_minus_simple_alt = uint_sub_lem [folded word_le_def] -lemmas word_sub_le_iff = no_ulen_sub [folded word_le_def] -lemmas word_sub_le = word_sub_le_iff [THEN iffD2] - -lemma word_less_sub1: "x \ 0 \ 1 < x \ 0 < x - 1" - for x :: "'a::len word" - by uint_arith - -lemma word_le_sub1: "x \ 0 \ 1 \ x \ 0 \ x - 1" - for x :: "'a::len word" - by uint_arith - -lemma sub_wrap_lt: "x < x - z \ x < z" - for x z :: "'a::len word" - by uint_arith - -lemma sub_wrap: "x \ x - z \ z = 0 \ x < z" - for x z :: "'a::len word" - by uint_arith - -lemma plus_minus_not_NULL_ab: "x \ ab - c \ c \ ab \ c \ 0 \ x + c \ 0" - for x ab c :: "'a::len word" - by uint_arith - -lemma plus_minus_no_overflow_ab: "x \ ab - c \ c \ ab \ x \ x + c" - for x ab c :: "'a::len word" - by uint_arith - -lemma le_minus': "a + c \ b \ a \ a + c \ c \ b - a" - for a b c :: "'a::len word" - by uint_arith - -lemma le_plus': "a \ b \ c \ b - a \ a + c \ b" - for a b c :: "'a::len word" - by uint_arith - -lemmas le_plus = le_plus' [rotated] - -lemmas le_minus = leD [THEN thin_rl, THEN le_minus'] (* FIXME *) - -lemma word_plus_mono_right: "y \ z \ x \ x + z \ x + y \ x + z" - for x y z :: "'a::len word" - by uint_arith - -lemma word_less_minus_cancel: "y - x < z - x \ x \ z \ y < z" - for x y z :: "'a::len word" - by uint_arith - -lemma word_less_minus_mono_left: "y < z \ x \ y \ y - x < z - x" - for x y z :: "'a::len word" - by uint_arith - -lemma word_less_minus_mono: "a < c \ d < b \ a - b < a \ c - d < c \ a - b < c - d" - for a b c d :: "'a::len word" - by uint_arith - -lemma word_le_minus_cancel: "y - x \ z - x \ x \ z \ y \ z" - for x y z :: "'a::len word" - by uint_arith - -lemma word_le_minus_mono_left: "y \ z \ x \ y \ y - x \ z - x" - for x y z :: "'a::len word" - by uint_arith - -lemma word_le_minus_mono: - "a \ c \ d \ b \ a - b \ a \ c - d \ c \ a - b \ c - d" - for a b c d :: "'a::len word" - by uint_arith - -lemma plus_le_left_cancel_wrap: "x + y' < x \ x + y < x \ x + y' < x + y \ y' < y" - for x y y' :: "'a::len word" - by uint_arith - -lemma plus_le_left_cancel_nowrap: "x \ x + y' \ x \ x + y \ x + y' < x + y \ y' < y" - for x y y' :: "'a::len word" - by uint_arith - -lemma word_plus_mono_right2: "a \ a + b \ c \ b \ a \ a + c" - for a b c :: "'a::len word" - by uint_arith - -lemma word_less_add_right: "x < y - z \ z \ y \ x + z < y" - for x y z :: "'a::len word" - by uint_arith - -lemma word_less_sub_right: "x < y + z \ y \ x \ x - y < z" - for x y z :: "'a::len word" - by uint_arith - -lemma word_le_plus_either: "x \ y \ x \ z \ y \ y + z \ x \ y + z" - for x y z :: "'a::len word" - by uint_arith - -lemma word_less_nowrapI: "x < z - k \ k \ z \ 0 < k \ x < x + k" - for x z k :: "'a::len word" - by uint_arith - -lemma inc_le: "i < m \ i + 1 \ m" - for i m :: "'a::len word" - by uint_arith - -lemma inc_i: "1 \ i \ i < m \ 1 \ i + 1 \ i + 1 \ m" - for i m :: "'a::len word" - by uint_arith - -lemma udvd_incr_lem: - "up < uq \ up = ua + n * uint K \ - uq = ua + n' * uint K \ up + uint K \ uq" - by auto (metis int_distrib(1) linorder_not_less mult.left_neutral mult_right_mono uint_nonnegative zless_imp_add1_zle) - -lemma udvd_incr': - "p < q \ uint p = ua + n * uint K \ - uint q = ua + n' * uint K \ p + K \ q" - apply (unfold word_less_alt word_le_def) - apply (drule (2) udvd_incr_lem) - apply (erule uint_add_le [THEN order_trans]) - done - -lemma udvd_decr': - "p < q \ uint p = ua + n * uint K \ - uint q = ua + n' * uint K \ p \ q - K" - apply (unfold word_less_alt word_le_def) - apply (drule (2) udvd_incr_lem) - apply (drule le_diff_eq [THEN iffD2]) - apply (erule order_trans) - apply (rule uint_sub_ge) - done - -lemmas udvd_incr_lem0 = udvd_incr_lem [where ua=0, unfolded add_0_left] -lemmas udvd_incr0 = udvd_incr' [where ua=0, unfolded add_0_left] -lemmas udvd_decr0 = udvd_decr' [where ua=0, unfolded add_0_left] - -lemma udvd_minus_le': "xy < k \ z udvd xy \ z udvd k \ xy \ k - z" - apply (unfold udvd_unfold_int) - apply clarify - apply (erule (2) udvd_decr0) - done - -lemma udvd_incr2_K: - "p < a + s \ a \ a + s \ K udvd s \ K udvd p - a \ a \ p \ - 0 < K \ p \ p + K \ p + K \ a + s" - supply [[simproc del: linordered_ring_less_cancel_factor]] - apply (unfold udvd_unfold_int) - apply clarify - apply (simp add: uint_arith_simps split: if_split_asm) - prefer 2 - using uint_lt2p [of s] apply simp - apply (drule add.commute [THEN xtrans(1)]) - apply (simp flip: diff_less_eq) - apply (subst (asm) mult_less_cancel_right) - apply simp - apply (simp add: diff_eq_eq not_less) - apply (subst (asm) (3) zless_iff_Suc_zadd) - apply auto - apply (auto simp add: algebra_simps) - apply (drule less_le_trans [of _ \2 ^ LENGTH('a)\]) apply assumption - apply (simp add: mult_less_0_iff) - done - - -subsection \Arithmetic type class instantiations\ - -lemmas word_le_0_iff [simp] = - word_zero_le [THEN leD, THEN antisym_conv1] - -lemma word_of_int_nat: "0 \ x \ word_of_int x = of_nat (nat x)" - by simp - -text \ - note that \iszero_def\ is only for class \comm_semiring_1_cancel\, - which requires word length \\ 1\, ie \'a::len word\ -\ -lemma iszero_word_no [simp]: - "iszero (numeral bin :: 'a::len word) = - iszero (take_bit LENGTH('a) (numeral bin :: int))" - apply (simp add: iszero_def) - apply transfer - apply simp - done - -text \Use \iszero\ to simplify equalities between word numerals.\ - -lemmas word_eq_numeral_iff_iszero [simp] = - eq_numeral_iff_iszero [where 'a="'a::len word"] - - -subsection \Word and nat\ - -lemma word_nchotomy: "\w :: 'a::len word. \n. w = of_nat n \ n < 2 ^ LENGTH('a)" - apply (rule allI) - apply (rule exI [of _ \unat w\ for w :: \'a word\]) - apply simp - done - -lemma of_nat_eq: "of_nat n = w \ (\q. n = unat w + q * 2 ^ LENGTH('a))" - for w :: "'a::len word" - using mod_div_mult_eq [of n "2 ^ LENGTH('a)", symmetric] - by (auto simp flip: take_bit_eq_mod) - -lemma of_nat_eq_size: "of_nat n = w \ (\q. n = unat w + q * 2 ^ size w)" - unfolding word_size by (rule of_nat_eq) - -lemma of_nat_0: "of_nat m = (0::'a::len word) \ (\q. m = q * 2 ^ LENGTH('a))" - by (simp add: of_nat_eq) - -lemma of_nat_2p [simp]: "of_nat (2 ^ LENGTH('a)) = (0::'a::len word)" - by (fact mult_1 [symmetric, THEN iffD2 [OF of_nat_0 exI]]) - -lemma of_nat_gt_0: "of_nat k \ 0 \ 0 < k" - by (cases k) auto - -lemma of_nat_neq_0: "0 < k \ k < 2 ^ LENGTH('a::len) \ of_nat k \ (0 :: 'a word)" - by (auto simp add : of_nat_0) - -lemma Abs_fnat_hom_add: "of_nat a + of_nat b = of_nat (a + b)" - by simp - -lemma Abs_fnat_hom_mult: "of_nat a * of_nat b = (of_nat (a * b) :: 'a::len word)" - by (simp add: wi_hom_mult) - -lemma Abs_fnat_hom_Suc: "word_succ (of_nat a) = of_nat (Suc a)" - by transfer (simp add: ac_simps) - -lemma Abs_fnat_hom_0: "(0::'a::len word) = of_nat 0" - by simp - -lemma Abs_fnat_hom_1: "(1::'a::len word) = of_nat (Suc 0)" - by simp - -lemmas Abs_fnat_homs = - Abs_fnat_hom_add Abs_fnat_hom_mult Abs_fnat_hom_Suc - Abs_fnat_hom_0 Abs_fnat_hom_1 - -lemma word_arith_nat_add: "a + b = of_nat (unat a + unat b)" - by simp - -lemma word_arith_nat_mult: "a * b = of_nat (unat a * unat b)" - by simp - -lemma word_arith_nat_Suc: "word_succ a = of_nat (Suc (unat a))" - by (subst Abs_fnat_hom_Suc [symmetric]) simp - -lemma word_arith_nat_div: "a div b = of_nat (unat a div unat b)" - by (metis of_int_of_nat_eq of_nat_unat of_nat_div word_div_def) - -lemma word_arith_nat_mod: "a mod b = of_nat (unat a mod unat b)" - by (metis of_int_of_nat_eq of_nat_mod of_nat_unat word_mod_def) - -lemmas word_arith_nat_defs = - word_arith_nat_add word_arith_nat_mult - word_arith_nat_Suc Abs_fnat_hom_0 - Abs_fnat_hom_1 word_arith_nat_div - word_arith_nat_mod - -lemma unat_cong: "x = y \ unat x = unat y" - by (fact arg_cong) - -lemma unat_of_nat: - \unat (word_of_nat x :: 'a::len word) = x mod 2 ^ LENGTH('a)\ - by transfer (simp flip: take_bit_eq_mod add: nat_take_bit_eq) - -lemmas unat_word_ariths = word_arith_nat_defs - [THEN trans [OF unat_cong unat_of_nat]] - -lemmas word_sub_less_iff = word_sub_le_iff - [unfolded linorder_not_less [symmetric] Not_eq_iff] - -lemma unat_add_lem: - "unat x + unat y < 2 ^ LENGTH('a) \ unat (x + y) = unat x + unat y" - for x y :: "'a::len word" - apply (auto simp: unat_word_ariths) - apply (drule sym) - apply (metis unat_of_nat unsigned_less) - done - -lemma unat_mult_lem: - "unat x * unat y < 2 ^ LENGTH('a) \ unat (x * y) = unat x * unat y" - for x y :: "'a::len word" - apply (auto simp: unat_word_ariths) - apply (drule sym) - apply (metis unat_of_nat unsigned_less) - done - -lemma unat_plus_if': - \unat (a + b) = - (if unat a + unat b < 2 ^ LENGTH('a) - then unat a + unat b - else unat a + unat b - 2 ^ LENGTH('a))\ for a b :: \'a::len word\ - apply (auto simp: unat_word_ariths not_less) - apply (subst (asm) le_iff_add) - apply auto - apply (simp flip: take_bit_eq_mod add: take_bit_nat_eq_self_iff) - apply (metis add.commute add_less_cancel_right le_less_trans less_imp_le unsigned_less) - done - -lemma le_no_overflow: "x \ b \ a \ a + b \ x \ a + b" - for a b x :: "'a::len word" - apply (erule order_trans) - apply (erule olen_add_eqv [THEN iffD1]) - done - -lemmas un_ui_le = - trans [OF word_le_nat_alt [symmetric] word_le_def] - -lemma unat_sub_if_size: - "unat (x - y) = - (if unat y \ unat x - then unat x - unat y - else unat x + 2 ^ size x - unat y)" - supply nat_uint_eq [simp del] - apply (unfold word_size) - apply (simp add: un_ui_le) - apply (auto simp add: unat_eq_nat_uint uint_sub_if') - apply (rule nat_diff_distrib) - prefer 3 - apply (simp add: algebra_simps) - apply (rule nat_diff_distrib [THEN trans]) - prefer 3 - apply (subst nat_add_distrib) - prefer 3 - apply (simp add: nat_power_eq) - apply auto - apply uint_arith - done - -lemmas unat_sub_if' = unat_sub_if_size [unfolded word_size] - -lemma uint_div: - \uint (x div y) = uint x div uint y\ - by (fact uint_div_distrib) - -lemma unat_div: - \unat (x div y) = unat x div unat y\ - by (fact unat_div_distrib) - -lemma uint_mod: - \uint (x mod y) = uint x mod uint y\ - by (fact uint_mod_distrib) - -lemma unat_mod: - \unat (x mod y) = unat x mod unat y\ - by (fact unat_mod_distrib) - - -text \Definition of \unat_arith\ tactic\ - -lemma unat_split: "P (unat x) \ (\n. of_nat n = x \ n < 2^LENGTH('a) \ P n)" - for x :: "'a::len word" - by auto (metis take_bit_nat_eq_self_iff) - -lemma unat_split_asm: "P (unat x) \ (\n. of_nat n = x \ n < 2^LENGTH('a) \ \ P n)" - for x :: "'a::len word" - by auto (metis take_bit_nat_eq_self_iff) - -lemma of_nat_inverse: - \word_of_nat r = a \ r < 2 ^ LENGTH('a) \ unat a = r\ - for a :: \'a::len word\ - apply (drule sym) - apply transfer - apply (simp add: take_bit_int_eq_self) - done - -lemma word_unat_eq_iff: - \v = w \ unat v = unat w\ - for v w :: \'a::len word\ - by (fact word_eq_iff_unsigned) - -lemmas unat_splits = unat_split unat_split_asm - -lemmas unat_arith_simps = - word_le_nat_alt word_less_nat_alt - word_unat_eq_iff - unat_sub_if' unat_plus_if' unat_div unat_mod - -\ \\unat_arith_tac\: tactic to reduce word arithmetic to \nat\, try to solve via \arith\\ -ML \ -val unat_arith_simpset = - @{context} - |> fold Simplifier.add_simp @{thms unat_arith_simps} - |> fold Splitter.add_split @{thms if_split_asm} - |> fold Simplifier.add_cong @{thms power_False_cong} - |> simpset_of - -fun unat_arith_tacs ctxt = - let - fun arith_tac' n t = - Arith_Data.arith_tac ctxt n t - handle Cooper.COOPER _ => Seq.empty; - in - [ clarify_tac ctxt 1, - full_simp_tac (put_simpset unat_arith_simpset ctxt) 1, - ALLGOALS (full_simp_tac - (put_simpset HOL_ss ctxt - |> fold Splitter.add_split @{thms unat_splits} - |> fold Simplifier.add_cong @{thms power_False_cong})), - rewrite_goals_tac ctxt @{thms word_size}, - ALLGOALS (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN - REPEAT (eresolve_tac ctxt [conjE] n) THEN - REPEAT (dresolve_tac ctxt @{thms of_nat_inverse} n THEN assume_tac ctxt n)), - TRYALL arith_tac' ] - end - -fun unat_arith_tac ctxt = SELECT_GOAL (EVERY (unat_arith_tacs ctxt)) -\ - -method_setup unat_arith = - \Scan.succeed (SIMPLE_METHOD' o unat_arith_tac)\ - "solving word arithmetic via natural numbers and arith" - -lemma no_plus_overflow_unat_size: "x \ x + y \ unat x + unat y < 2 ^ size x" - for x y :: "'a::len word" - unfolding word_size by unat_arith - -lemmas no_olen_add_nat = - no_plus_overflow_unat_size [unfolded word_size] - -lemmas unat_plus_simple = - trans [OF no_olen_add_nat unat_add_lem] - -lemma word_div_mult: "0 < y \ unat x * unat y < 2 ^ LENGTH('a) \ x * y div y = x" - for x y :: "'a::len word" - apply unat_arith - apply clarsimp - apply (subst unat_mult_lem [THEN iffD1]) - apply auto - done - -lemma div_lt': "i \ k div x \ unat i * unat x < 2 ^ LENGTH('a)" - for i k x :: "'a::len word" - apply unat_arith - apply clarsimp - apply (drule mult_le_mono1) - apply (erule order_le_less_trans) - apply (metis add_lessD1 div_mult_mod_eq unsigned_less) - done - -lemmas div_lt'' = order_less_imp_le [THEN div_lt'] - -lemma div_lt_mult: "i < k div x \ 0 < x \ i * x < k" - for i k x :: "'a::len word" - apply (frule div_lt'' [THEN unat_mult_lem [THEN iffD1]]) - apply (simp add: unat_arith_simps) - apply (drule (1) mult_less_mono1) - apply (erule order_less_le_trans) - apply auto - done - -lemma div_le_mult: "i \ k div x \ 0 < x \ i * x \ k" - for i k x :: "'a::len word" - apply (frule div_lt' [THEN unat_mult_lem [THEN iffD1]]) - apply (simp add: unat_arith_simps) - apply (drule mult_le_mono1) - apply (erule order_trans) - apply auto - done - -lemma div_lt_uint': "i \ k div x \ uint i * uint x < 2 ^ LENGTH('a)" - for i k x :: "'a::len word" - apply (unfold uint_nat) - apply (drule div_lt') - apply (metis of_nat_less_iff of_nat_mult of_nat_numeral of_nat_power) - done - -lemmas div_lt_uint'' = order_less_imp_le [THEN div_lt_uint'] - -lemma word_le_exists': "x \ y \ \z. y = x + z \ uint x + uint z < 2 ^ LENGTH('a)" - for x y z :: "'a::len word" - by (metis add_diff_cancel_left' add_diff_eq uint_add_lem uint_plus_simple) - -lemmas plus_minus_not_NULL = order_less_imp_le [THEN plus_minus_not_NULL_ab] - -lemmas plus_minus_no_overflow = - order_less_imp_le [THEN plus_minus_no_overflow_ab] - -lemmas mcs = word_less_minus_cancel word_less_minus_mono_left - word_le_minus_cancel word_le_minus_mono_left - -lemmas word_l_diffs = mcs [where y = "w + x", unfolded add_diff_cancel] for w x -lemmas word_diff_ls = mcs [where z = "w + x", unfolded add_diff_cancel] for w x -lemmas word_plus_mcs = word_diff_ls [where y = "v + x", unfolded add_diff_cancel] for v x - -lemma le_unat_uoi: - \y \ unat z \ unat (word_of_nat y :: 'a word) = y\ - for z :: \'a::len word\ - by transfer (simp add: nat_take_bit_eq take_bit_nat_eq_self_iff le_less_trans) - -lemmas thd = times_div_less_eq_dividend - -lemmas uno_simps [THEN le_unat_uoi] = mod_le_divisor div_le_dividend - -lemma word_mod_div_equality: "(n div b) * b + (n mod b) = n" - for n b :: "'a::len word" - by (fact div_mult_mod_eq) - -lemma word_div_mult_le: "a div b * b \ a" - for a b :: "'a::len word" - by (metis div_le_mult mult_not_zero order.not_eq_order_implies_strict order_refl word_zero_le) - -lemma word_mod_less_divisor: "0 < n \ m mod n < n" - for m n :: "'a::len word" - by (simp add: unat_arith_simps) - -lemma word_of_int_power_hom: "word_of_int a ^ n = (word_of_int (a ^ n) :: 'a::len word)" - by (induct n) (simp_all add: wi_hom_mult [symmetric]) - -lemma word_arith_power_alt: "a ^ n = (word_of_int (uint a ^ n) :: 'a::len word)" - by (simp add : word_of_int_power_hom [symmetric]) - -lemma unatSuc: "1 + n \ 0 \ unat (1 + n) = Suc (unat n)" - for n :: "'a::len word" - by unat_arith - - -subsection \Cardinality, finiteness of set of words\ - -lemma inj_on_word_of_int: \inj_on (word_of_int :: int \ 'a word) {0..<2 ^ LENGTH('a::len)}\ - apply (rule inj_onI) - apply transfer - apply (simp add: take_bit_eq_mod) - done - -lemma inj_uint: \inj uint\ - by (fact inj_unsigned) - -lemma range_uint: \range (uint :: 'a word \ int) = {0..<2 ^ LENGTH('a::len)}\ - apply transfer - apply (auto simp add: image_iff) - apply (metis take_bit_int_eq_self_iff) - done - -lemma UNIV_eq: \(UNIV :: 'a word set) = word_of_int ` {0..<2 ^ LENGTH('a::len)}\ - by (auto simp add: image_iff) (metis atLeastLessThan_iff linorder_not_le uint_split) - -lemma card_word: "CARD('a word) = 2 ^ LENGTH('a::len)" - by (simp add: UNIV_eq card_image inj_on_word_of_int) - -lemma card_word_size: "CARD('a word) = 2 ^ size x" - for x :: "'a::len word" - unfolding word_size by (rule card_word) - -instance word :: (len) finite - by standard (simp add: UNIV_eq) - - -subsection \Bitwise Operations on Words\ - -lemma word_wi_log_defs: - "NOT (word_of_int a) = word_of_int (NOT a)" - "word_of_int a AND word_of_int b = word_of_int (a AND b)" - "word_of_int a OR word_of_int b = word_of_int (a OR b)" - "word_of_int a XOR word_of_int b = word_of_int (a XOR b)" - by (transfer, rule refl)+ - -lemma word_no_log_defs [simp]: - "NOT (numeral a) = word_of_int (NOT (numeral a))" - "NOT (- numeral a) = word_of_int (NOT (- numeral a))" - "numeral a AND numeral b = word_of_int (numeral a AND numeral b)" - "numeral a AND - numeral b = word_of_int (numeral a AND - numeral b)" - "- numeral a AND numeral b = word_of_int (- numeral a AND numeral b)" - "- numeral a AND - numeral b = word_of_int (- numeral a AND - numeral b)" - "numeral a OR numeral b = word_of_int (numeral a OR numeral b)" - "numeral a OR - numeral b = word_of_int (numeral a OR - numeral b)" - "- numeral a OR numeral b = word_of_int (- numeral a OR numeral b)" - "- numeral a OR - numeral b = word_of_int (- numeral a OR - numeral b)" - "numeral a XOR numeral b = word_of_int (numeral a XOR numeral b)" - "numeral a XOR - numeral b = word_of_int (numeral a XOR - numeral b)" - "- numeral a XOR numeral b = word_of_int (- numeral a XOR numeral b)" - "- numeral a XOR - numeral b = word_of_int (- numeral a XOR - numeral b)" - by (transfer, rule refl)+ - -text \Special cases for when one of the arguments equals 1.\ - -lemma word_bitwise_1_simps [simp]: - "NOT (1::'a::len word) = -2" - "1 AND numeral b = word_of_int (1 AND numeral b)" - "1 AND - numeral b = word_of_int (1 AND - numeral b)" - "numeral a AND 1 = word_of_int (numeral a AND 1)" - "- numeral a AND 1 = word_of_int (- numeral a AND 1)" - "1 OR numeral b = word_of_int (1 OR numeral b)" - "1 OR - numeral b = word_of_int (1 OR - numeral b)" - "numeral a OR 1 = word_of_int (numeral a OR 1)" - "- numeral a OR 1 = word_of_int (- numeral a OR 1)" - "1 XOR numeral b = word_of_int (1 XOR numeral b)" - "1 XOR - numeral b = word_of_int (1 XOR - numeral b)" - "numeral a XOR 1 = word_of_int (numeral a XOR 1)" - "- numeral a XOR 1 = word_of_int (- numeral a XOR 1)" - by (transfer, simp)+ - -text \Special cases for when one of the arguments equals -1.\ - -lemma word_bitwise_m1_simps [simp]: - "NOT (-1::'a::len word) = 0" - "(-1::'a::len word) AND x = x" - "x AND (-1::'a::len word) = x" - "(-1::'a::len word) OR x = -1" - "x OR (-1::'a::len word) = -1" - " (-1::'a::len word) XOR x = NOT x" - "x XOR (-1::'a::len word) = NOT x" - by (transfer, simp)+ - -lemma uint_and: - \uint (x AND y) = uint x AND uint y\ - by transfer simp - -lemma uint_or: - \uint (x OR y) = uint x OR uint y\ - by transfer simp - -lemma uint_xor: - \uint (x XOR y) = uint x XOR uint y\ - by transfer simp - -\ \get from commutativity, associativity etc of \int_and\ etc to same for \word_and etc\\ -lemmas bwsimps = - wi_hom_add - word_wi_log_defs - -lemma word_bw_assocs: - "(x AND y) AND z = x AND y AND z" - "(x OR y) OR z = x OR y OR z" - "(x XOR y) XOR z = x XOR y XOR z" - for x :: "'a::len word" - by (fact ac_simps)+ - -lemma word_bw_comms: - "x AND y = y AND x" - "x OR y = y OR x" - "x XOR y = y XOR x" - for x :: "'a::len word" - by (fact ac_simps)+ - -lemma word_bw_lcs: - "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 :: "'a::len word" - by (fact ac_simps)+ - -lemma word_log_esimps: - "x AND 0 = 0" - "x AND -1 = x" - "x OR 0 = x" - "x OR -1 = -1" - "x XOR 0 = x" - "x XOR -1 = NOT x" - "0 AND x = 0" - "-1 AND x = x" - "0 OR x = x" - "-1 OR x = -1" - "0 XOR x = x" - "-1 XOR x = NOT x" - for x :: "'a::len word" - by simp_all - -lemma word_not_dist: - "NOT (x OR y) = NOT x AND NOT y" - "NOT (x AND y) = NOT x OR NOT y" - for x :: "'a::len word" - by simp_all - -lemma word_bw_same: - "x AND x = x" - "x OR x = x" - "x XOR x = 0" - for x :: "'a::len word" - by simp_all - -lemma word_ao_absorbs [simp]: - "x AND (y OR x) = x" - "x OR y AND x = x" - "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 :: "'a::len word" - by (auto intro: bit_eqI simp add: bit_and_iff bit_or_iff) - -lemma word_not_not [simp]: "NOT (NOT x) = x" - for x :: "'a::len word" - by (fact bit.double_compl) - -lemma word_ao_dist: "(x OR y) AND z = x AND z OR y AND z" - for x :: "'a::len word" - by (fact bit.conj_disj_distrib2) - -lemma word_oa_dist: "x AND y OR z = (x OR z) AND (y OR z)" - for x :: "'a::len word" - by (fact bit.disj_conj_distrib2) - -lemma word_add_not [simp]: "x + NOT x = -1" - for x :: "'a::len word" - by (simp add: not_eq_complement) - -lemma word_plus_and_or [simp]: "(x AND y) + (x OR y) = x + y" - for x :: "'a::len word" - by transfer (simp add: plus_and_or) - -lemma leoa: "w = x OR y \ y = w AND y" - for x :: "'a::len word" - by auto - -lemma leao: "w' = x' AND y' \ x' = x' OR w'" - for x' :: "'a::len word" - by auto - -lemma word_ao_equiv: "w = w OR w' \ w' = w AND w'" - for w w' :: "'a::len word" - by (auto intro: leoa leao) - -lemma le_word_or2: "x \ x OR y" - for x y :: "'a::len word" - by (simp add: or_greater_eq uint_or word_le_def) - -lemmas le_word_or1 = xtrans(3) [OF word_bw_comms (2) le_word_or2] -lemmas word_and_le1 = xtrans(3) [OF word_ao_absorbs (4) [symmetric] le_word_or2] -lemmas word_and_le2 = xtrans(3) [OF word_ao_absorbs (8) [symmetric] le_word_or2] - -lemma bit_horner_sum_bit_word_iff: - \bit (horner_sum of_bool (2 :: 'a::len word) bs) n - \ n < min LENGTH('a) (length bs) \ bs ! n\ - by transfer (simp add: bit_horner_sum_bit_iff) - -definition word_reverse :: \'a::len word \ 'a word\ - where \word_reverse w = horner_sum of_bool 2 (rev (map (bit w) [0.. - -lemma bit_word_reverse_iff: - \bit (word_reverse w) n \ n < LENGTH('a) \ bit w (LENGTH('a) - Suc n)\ - for w :: \'a::len word\ - by (cases \n < LENGTH('a)\) - (simp_all add: word_reverse_def bit_horner_sum_bit_word_iff rev_nth) - -lemma word_rev_rev [simp] : "word_reverse (word_reverse w) = w" - by (rule bit_word_eqI) - (auto simp add: bit_word_reverse_iff bit_imp_le_length Suc_diff_Suc) - -lemma word_rev_gal: "word_reverse w = u \ word_reverse u = w" - by (metis word_rev_rev) - -lemma word_rev_gal': "u = word_reverse w \ w = word_reverse u" - by simp - -lemma uint_2p: "(0::'a::len word) < 2 ^ n \ uint (2 ^ n::'a::len word) = 2 ^ n" - apply (cases \n < LENGTH('a)\; transfer) - apply auto - done - -lemma word_of_int_2p: "(word_of_int (2 ^ n) :: 'a::len word) = 2 ^ n" - by (induct n) (simp_all add: wi_hom_syms) - - -subsection \Shifting, Rotating, and Splitting Words\ - -lemma shiftl1_wi [simp]: "shiftl1 (word_of_int w) = word_of_int (2 * w)" - by transfer simp - -lemma shiftl1_numeral [simp]: "shiftl1 (numeral w) = numeral (Num.Bit0 w)" - unfolding word_numeral_alt shiftl1_wi by simp - -lemma shiftl1_neg_numeral [simp]: "shiftl1 (- numeral w) = - numeral (Num.Bit0 w)" - unfolding word_neg_numeral_alt shiftl1_wi by simp - -lemma shiftl1_0 [simp] : "shiftl1 0 = 0" - by transfer simp - -lemma shiftl1_def_u: "shiftl1 w = word_of_int (2 * uint w)" - by (fact shiftl1_eq) - -lemma shiftl1_def_s: "shiftl1 w = word_of_int (2 * sint w)" - by (simp add: shiftl1_def_u wi_hom_syms) - -lemma shiftr1_0 [simp]: "shiftr1 0 = 0" - by transfer simp - -lemma sshiftr1_0 [simp]: "sshiftr1 0 = 0" - by transfer simp - -lemma sshiftr1_n1 [simp]: "sshiftr1 (- 1) = - 1" - by transfer simp - -text \ - see paper page 10, (1), (2), \shiftr1_def\ is of the form of (1), - where \f\ (ie \_ div 2\) takes normal arguments to normal results, - thus we get (2) from (1) -\ - -lemma uint_shiftr1: "uint (shiftr1 w) = uint w div 2" - using drop_bit_eq_div [of 1 \uint w\, symmetric] - apply simp - apply transfer - apply (simp add: drop_bit_take_bit min_def) - done - -lemma bit_sshiftr1_iff: - \bit (sshiftr1 w) n \ bit w (if n = LENGTH('a) - 1 then LENGTH('a) - 1 else Suc n)\ - for w :: \'a::len word\ - apply transfer - apply (auto simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def simp flip: bit_Suc) - using le_less_Suc_eq apply fastforce - using le_less_Suc_eq apply fastforce - done - -lemma shiftr1_div_2: "uint (shiftr1 w) = uint w div 2" - by (fact uint_shiftr1) - -lemma sshiftr1_div_2: "sint (sshiftr1 w) = sint w div 2" - using sint_signed_drop_bit_eq [of 1 w] - by (simp add: drop_bit_Suc sshiftr1_eq_signed_drop_bit_Suc_0) - -lemma bit_bshiftr1_iff: - \bit (bshiftr1 b w) n \ b \ n = LENGTH('a) - 1 \ bit w (Suc n)\ - for w :: \'a::len word\ - apply transfer - apply (simp add: bit_take_bit_iff flip: bit_Suc) - apply (subst disjunctive_add) - apply (auto simp add: bit_take_bit_iff bit_or_iff bit_exp_iff simp flip: bit_Suc) - done - - -subsubsection \shift functions in terms of lists of bools\ - -lemma shiftl1_rev: "shiftl1 w = word_reverse (shiftr1 (word_reverse w))" - apply (rule bit_word_eqI) - apply (auto simp add: bit_shiftl1_iff bit_word_reverse_iff bit_shiftr1_iff Suc_diff_Suc) - done - -\ \note -- the following results use \'a::len word < number_ring\\ - -lemma shiftl1_2t: "shiftl1 w = 2 * w" - for w :: "'a::len word" - by (simp add: shiftl1_eq wi_hom_mult [symmetric]) - -lemma shiftl1_p: "shiftl1 w = w + w" - for w :: "'a::len word" - by (simp add: shiftl1_2t) - -lemma shiftr1_bintr [simp]: - "(shiftr1 (numeral w) :: 'a::len word) = - word_of_int (take_bit LENGTH('a) (numeral w) div 2)" - by transfer simp - -lemma sshiftr1_sbintr [simp]: - "(sshiftr1 (numeral w) :: 'a::len word) = - word_of_int (signed_take_bit (LENGTH('a) - 1) (numeral w) div 2)" - by transfer simp - -text \TODO: rules for \<^term>\- (numeral n)\\ - -lemma drop_bit_word_numeral [simp]: - \drop_bit (numeral n) (numeral k) = - (word_of_int (drop_bit (numeral n) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\ - by transfer simp - -lemma zip_replicate: "n \ length ys \ zip (replicate n x) ys = map (\y. (x, y)) ys" - apply (induct ys arbitrary: n) - apply simp_all - apply (case_tac n) - apply simp_all - done - -lemma align_lem_or [rule_format] : - "\x m. length x = n + m \ length y = n + m \ - drop m x = replicate n False \ take m y = replicate m False \ - map2 (|) x y = take m x @ drop m y" - apply (induct y) - apply force - apply clarsimp - apply (case_tac x) - apply force - apply (case_tac m) - apply auto - apply (drule_tac t="length xs" for xs in sym) - apply (auto simp: zip_replicate o_def) - done - -lemma align_lem_and [rule_format] : - "\x m. length x = n + m \ length y = n + m \ - drop m x = replicate n False \ take m y = replicate m False \ - map2 (\) x y = replicate (n + m) False" - apply (induct y) - apply force - apply clarsimp - apply (case_tac x) - apply force - apply (case_tac m) - apply auto - apply (drule_tac t="length xs" for xs in sym) - apply (auto simp: zip_replicate o_def map_replicate_const) - done - - -subsubsection \Mask\ - -lemma minus_1_eq_mask: - \- 1 = (mask LENGTH('a) :: 'a::len word)\ - by (rule bit_eqI) (simp add: bit_exp_iff bit_mask_iff exp_eq_zero_iff) - -lemma mask_eq_decr_exp: - \mask n = 2 ^ n - (1 :: 'a::len word)\ - by (fact mask_eq_exp_minus_1) - -lemma mask_Suc_rec: - \mask (Suc n) = 2 * mask n + (1 :: 'a::len word)\ - by (simp add: mask_eq_exp_minus_1) - -context -begin - -qualified lemma bit_mask_iff: - \bit (mask m :: 'a::len word) n \ n < min LENGTH('a) m\ - by (simp add: bit_mask_iff exp_eq_zero_iff not_le) - -end - -lemma mask_bin: "mask n = word_of_int (take_bit n (- 1))" - by transfer (simp add: take_bit_minus_one_eq_mask) - -lemma and_mask_bintr: "w AND mask n = word_of_int (take_bit n (uint w))" - by transfer (simp add: ac_simps take_bit_eq_mask) - -lemma and_mask_wi: "word_of_int i AND mask n = word_of_int (take_bit n i)" - by (auto simp add: and_mask_bintr min_def not_le wi_bintr) - -lemma and_mask_wi': - "word_of_int i AND mask n = (word_of_int (take_bit (min LENGTH('a) n) i) :: 'a::len word)" - by (auto simp add: and_mask_wi min_def wi_bintr) - -lemma and_mask_no: "numeral i AND mask n = word_of_int (take_bit n (numeral i))" - unfolding word_numeral_alt by (rule and_mask_wi) - -lemma and_mask_mod_2p: "w AND mask n = word_of_int (uint w mod 2 ^ n)" - by (simp only: and_mask_bintr take_bit_eq_mod) - -lemma uint_mask_eq: - \uint (mask n :: 'a::len word) = mask (min LENGTH('a) n)\ - by transfer simp - -lemma and_mask_lt_2p: "uint (w AND mask n) < 2 ^ n" - apply (simp flip: take_bit_eq_mask) - apply transfer - apply (auto simp add: min_def) - using antisym_conv take_bit_int_eq_self_iff by fastforce - -lemma mask_eq_iff: "w AND mask n = w \ uint w < 2 ^ n" - apply (auto simp flip: take_bit_eq_mask) - apply (metis take_bit_int_eq_self_iff uint_take_bit_eq) - apply (simp add: take_bit_int_eq_self unsigned_take_bit_eq word_uint_eqI) - done - -lemma and_mask_dvd: "2 ^ n dvd uint w \ w AND mask n = 0" - by (simp flip: take_bit_eq_mask take_bit_eq_mod unsigned_take_bit_eq add: dvd_eq_mod_eq_0 uint_0_iff) - -lemma and_mask_dvd_nat: "2 ^ n dvd unat w \ w AND mask n = 0" - by (simp flip: take_bit_eq_mask take_bit_eq_mod unsigned_take_bit_eq add: dvd_eq_mod_eq_0 unat_0_iff uint_0_iff) - -lemma word_2p_lem: "n < size w \ w < 2 ^ n = (uint w < 2 ^ n)" - for w :: "'a::len word" - by transfer simp - -lemma less_mask_eq: "x < 2 ^ n \ x AND mask n = x" - for x :: "'a::len word" - apply (cases \n < LENGTH('a)\) - apply (simp_all add: not_less flip: take_bit_eq_mask exp_eq_zero_iff) - apply transfer - apply (simp add: min_def) - apply (metis min_def nat_less_le take_bit_int_eq_self_iff take_bit_take_bit) - done - -lemmas mask_eq_iff_w2p = trans [OF mask_eq_iff word_2p_lem [symmetric]] - -lemmas and_mask_less' = iffD2 [OF word_2p_lem and_mask_lt_2p, simplified word_size] - -lemma and_mask_less_size: "n < size x \ x AND mask n < 2 ^ n" - for x :: \'a::len word\ - unfolding word_size by (erule and_mask_less') - -lemma word_mod_2p_is_mask [OF refl]: "c = 2 ^ n \ c > 0 \ x mod c = x AND mask n" - for c x :: "'a::len word" - by (auto simp: word_mod_def uint_2p and_mask_mod_2p) - -lemma mask_eqs: - "(a AND mask n) + b AND mask n = a + b AND mask n" - "a + (b AND mask n) AND mask n = a + b AND mask n" - "(a AND mask n) - b AND mask n = a - b AND mask n" - "a - (b AND mask n) AND mask n = a - b AND mask n" - "a * (b AND mask n) AND mask n = a * b AND mask n" - "(b AND mask n) * a AND mask n = b * a AND mask n" - "(a AND mask n) + (b AND mask n) AND mask n = a + b AND mask n" - "(a AND mask n) - (b AND mask n) AND mask n = a - b AND mask n" - "(a AND mask n) * (b AND mask n) AND mask n = a * b AND mask n" - "- (a AND mask n) AND mask n = - a AND mask n" - "word_succ (a AND mask n) AND mask n = word_succ a AND mask n" - "word_pred (a AND mask n) AND mask n = word_pred a AND mask n" - using word_of_int_Ex [where x=a] word_of_int_Ex [where x=b] - apply (auto simp flip: take_bit_eq_mask) - apply transfer - apply (simp add: take_bit_eq_mod mod_simps) - apply transfer - apply (simp add: take_bit_eq_mod mod_simps) - apply transfer - apply (simp add: take_bit_eq_mod mod_simps) - apply transfer - apply (simp add: take_bit_eq_mod mod_simps) - apply transfer - apply (simp add: take_bit_eq_mod mod_simps) - apply transfer - apply (simp add: take_bit_eq_mod mod_simps) - apply transfer - apply (simp add: take_bit_eq_mod mod_simps) - apply transfer - apply (simp add: take_bit_eq_mod mod_simps) - apply transfer - apply (simp add: take_bit_eq_mod mod_simps) - apply transfer - apply (simp add: take_bit_eq_mod mod_simps) - apply transfer - apply (simp add: take_bit_eq_mod mod_simps) - apply transfer - apply (simp add: take_bit_eq_mod mod_simps) - done - -lemma mask_power_eq: "(x AND mask n) ^ k AND mask n = x ^ k AND mask n" - for x :: \'a::len word\ - using word_of_int_Ex [where x=x] - apply (auto simp flip: take_bit_eq_mask) - apply transfer - apply (simp add: take_bit_eq_mod mod_simps) - done - -lemma mask_full [simp]: "mask LENGTH('a) = (- 1 :: 'a::len word)" - by transfer (simp add: take_bit_minus_one_eq_mask) - - -subsubsection \Slices\ - -definition slice1 :: \nat \ 'a::len word \ 'b::len word\ - where \slice1 n w = (if n < LENGTH('a) - then ucast (drop_bit (LENGTH('a) - n) w) - else push_bit (n - LENGTH('a)) (ucast w))\ - -lemma bit_slice1_iff: - \bit (slice1 m w :: 'b::len word) n \ m - LENGTH('a) \ n \ n < min LENGTH('b) m - \ bit w (n + (LENGTH('a) - m) - (m - LENGTH('a)))\ - for w :: \'a::len word\ - by (auto simp add: slice1_def bit_ucast_iff bit_drop_bit_eq bit_push_bit_iff exp_eq_zero_iff not_less not_le ac_simps - dest: bit_imp_le_length) - -definition slice :: \nat \ 'a::len word \ 'b::len word\ - where \slice n = slice1 (LENGTH('a) - n)\ - -lemma bit_slice_iff: - \bit (slice m w :: 'b::len word) n \ n < min LENGTH('b) (LENGTH('a) - m) \ bit w (n + LENGTH('a) - (LENGTH('a) - m))\ - for w :: \'a::len word\ - by (simp add: slice_def word_size bit_slice1_iff) - -lemma slice1_0 [simp] : "slice1 n 0 = 0" - unfolding slice1_def by simp - -lemma slice_0 [simp] : "slice n 0 = 0" - unfolding slice_def by auto - -lemma ucast_slice1: "ucast w = slice1 (size w) w" - apply (simp add: slice1_def) - apply transfer - apply simp - done - -lemma ucast_slice: "ucast w = slice 0 w" - by (simp add: slice_def slice1_def) - -lemma slice_id: "slice 0 t = t" - by (simp only: ucast_slice [symmetric] ucast_id) - -lemma rev_slice1: - \slice1 n (word_reverse w :: 'b::len word) = word_reverse (slice1 k w :: 'a::len word)\ - if \n + k = LENGTH('a) + LENGTH('b)\ -proof (rule bit_word_eqI) - fix m - assume *: \m < LENGTH('a)\ - from that have **: \LENGTH('b) = n + k - LENGTH('a)\ - by simp - show \bit (slice1 n (word_reverse w :: 'b word) :: 'a word) m \ bit (word_reverse (slice1 k w :: 'a word)) m\ - apply (simp add: bit_slice1_iff bit_word_reverse_iff) - using * ** - apply (cases \n \ LENGTH('a)\; cases \k \ LENGTH('a)\) - apply auto - done -qed - -lemma rev_slice: - "n + k + LENGTH('a::len) = LENGTH('b::len) \ - slice n (word_reverse (w::'b word)) = word_reverse (slice k w :: 'a word)" - apply (unfold slice_def word_size) - apply (rule rev_slice1) - apply arith - done - - -subsubsection \Revcast\ - -definition revcast :: \'a::len word \ 'b::len word\ - where \revcast = slice1 LENGTH('b)\ - -lemma bit_revcast_iff: - \bit (revcast w :: 'b::len word) n \ LENGTH('b) - LENGTH('a) \ n \ n < LENGTH('b) - \ bit w (n + (LENGTH('a) - LENGTH('b)) - (LENGTH('b) - LENGTH('a)))\ - for w :: \'a::len word\ - by (simp add: revcast_def bit_slice1_iff) - -lemma revcast_slice1 [OF refl]: "rc = revcast w \ slice1 (size rc) w = rc" - by (simp add: revcast_def word_size) - -lemma revcast_rev_ucast [OF refl refl refl]: - "cs = [rc, uc] \ rc = revcast (word_reverse w) \ uc = ucast w \ - rc = word_reverse uc" - apply auto - apply (rule bit_word_eqI) - apply (cases \LENGTH('a) \ LENGTH('b)\) - apply (simp_all add: bit_revcast_iff bit_word_reverse_iff bit_ucast_iff not_le - bit_imp_le_length) - using bit_imp_le_length apply fastforce - using bit_imp_le_length apply fastforce - done - -lemma revcast_ucast: "revcast w = word_reverse (ucast (word_reverse w))" - using revcast_rev_ucast [of "word_reverse w"] by simp - -lemma ucast_revcast: "ucast w = word_reverse (revcast (word_reverse w))" - by (fact revcast_rev_ucast [THEN word_rev_gal']) - -lemma ucast_rev_revcast: "ucast (word_reverse w) = word_reverse (revcast w)" - by (fact revcast_ucast [THEN word_rev_gal']) - - -text "linking revcast and cast via shift" - -lemmas wsst_TYs = source_size target_size word_size - -lemmas sym_notr = - not_iff [THEN iffD2, THEN not_sym, THEN not_iff [THEN iffD1]] - - -subsection \Split and cat\ - -lemmas word_split_bin' = word_split_def -lemmas word_cat_bin' = word_cat_eq - -\ \this odd result is analogous to \ucast_id\, - result to the length given by the result type\ - -lemma word_cat_id: "word_cat a b = b" - by transfer (simp add: take_bit_concat_bit_eq) - -lemma word_cat_split_alt: "size w \ size u + size v \ word_split w = (u, v) \ word_cat u v = w" - apply (rule bit_word_eqI) - apply (auto simp add: bit_word_cat_iff not_less word_size word_split_def bit_ucast_iff bit_drop_bit_eq) - done - -lemmas word_cat_split_size = sym [THEN [2] word_cat_split_alt [symmetric]] - - -subsubsection \Split and slice\ - -lemma split_slices: "word_split w = (u, v) \ u = slice (size v) w \ v = slice 0 w" - apply (auto simp add: word_split_def word_size) - apply (rule bit_word_eqI) - apply (simp add: bit_slice_iff bit_ucast_iff bit_drop_bit_eq) - apply (cases \LENGTH('c) \ LENGTH('b)\) - apply (auto simp add: ac_simps dest: bit_imp_le_length) - apply (rule bit_word_eqI) - apply (auto simp add: bit_slice_iff bit_ucast_iff dest: bit_imp_le_length) - done - -lemma slice_cat1 [OF refl]: - "wc = word_cat a b \ size wc >= size a + size b \ slice (size b) wc = a" - apply (rule bit_word_eqI) - apply (auto simp add: bit_slice_iff bit_word_cat_iff word_size) - done - -lemmas slice_cat2 = trans [OF slice_id word_cat_id] - -lemma cat_slices: - "a = slice n c \ b = slice 0 c \ n = size b \ - size a + size b >= size c \ word_cat a b = c" - apply (rule bit_word_eqI) - apply (auto simp add: bit_slice_iff bit_word_cat_iff word_size) - done - -lemma word_split_cat_alt: - "w = word_cat u v \ size u + size v \ size w \ word_split w = (u, v)" - apply (auto simp add: word_split_def word_size) - apply (rule bit_eqI) - apply (auto simp add: bit_ucast_iff bit_drop_bit_eq bit_word_cat_iff dest: bit_imp_le_length) - apply (rule bit_eqI) - apply (auto simp add: bit_ucast_iff bit_drop_bit_eq bit_word_cat_iff dest: bit_imp_le_length) - done - -lemma horner_sum_uint_exp_Cons_eq: - \horner_sum uint (2 ^ LENGTH('a)) (w # ws) = - concat_bit LENGTH('a) (uint w) (horner_sum uint (2 ^ LENGTH('a)) ws)\ - for ws :: \'a::len word list\ - apply (simp add: concat_bit_eq push_bit_eq_mult) - apply transfer - apply simp - done - -lemma bit_horner_sum_uint_exp_iff: - \bit (horner_sum uint (2 ^ LENGTH('a)) ws) n \ - n div LENGTH('a) < length ws \ bit (ws ! (n div LENGTH('a))) (n mod LENGTH('a))\ - for ws :: \'a::len word list\ -proof (induction ws arbitrary: n) - case Nil - then show ?case - by simp -next - case (Cons w ws) - then show ?case - by (cases \n \ LENGTH('a)\) - (simp_all only: horner_sum_uint_exp_Cons_eq, simp_all add: bit_concat_bit_iff le_div_geq le_mod_geq bit_uint_iff Cons) -qed - - -subsection \Rotation\ - -lemma word_rotr_word_rotr_eq: - \word_rotr m (word_rotr n w) = word_rotr (m + n) w\ - by (rule bit_word_eqI) (simp add: bit_word_rotr_iff ac_simps mod_add_right_eq) - -lemma word_rot_rl [simp]: - \word_rotl k (word_rotr k v) = v\ - apply (rule bit_word_eqI) - apply (simp add: word_rotl_eq_word_rotr word_rotr_word_rotr_eq bit_word_rotr_iff algebra_simps) - apply (auto dest: bit_imp_le_length) - apply (metis (no_types, lifting) add.right_neutral add_diff_cancel_right' div_mult_mod_eq mod_add_right_eq mod_add_self2 mod_if mod_mult_self2_is_0) - apply (metis (no_types, lifting) add.right_neutral add_diff_cancel_right' div_mult_mod_eq mod_add_right_eq mod_add_self2 mod_less mod_mult_self2_is_0) - done - -lemma word_rot_lr [simp]: - \word_rotr k (word_rotl k v) = v\ - apply (rule bit_word_eqI) - apply (simp add: word_rotl_eq_word_rotr word_rotr_word_rotr_eq bit_word_rotr_iff algebra_simps) - apply (auto dest: bit_imp_le_length) - apply (metis (no_types, lifting) add.right_neutral add_diff_cancel_right' div_mult_mod_eq mod_add_right_eq mod_add_self2 mod_if mod_mult_self2_is_0) - apply (metis (no_types, lifting) add.right_neutral add_diff_cancel_right' div_mult_mod_eq mod_add_right_eq mod_add_self2 mod_less mod_mult_self2_is_0) - done - -lemma word_rot_gal: - \word_rotr n v = w \ word_rotl n w = v\ - by auto - -lemma word_rot_gal': - \w = word_rotr n v \ v = word_rotl n w\ - by auto - -lemma word_rotr_rev: - \word_rotr n w = word_reverse (word_rotl n (word_reverse w))\ -proof (rule bit_word_eqI) - fix m - assume \m < LENGTH('a)\ - moreover have \1 + - ((int m + int n mod int LENGTH('a)) mod int LENGTH('a) + - ((int LENGTH('a) * 2) mod int LENGTH('a) - (1 + (int m + int n mod int LENGTH('a)))) mod int LENGTH('a)) = - int LENGTH('a)\ - apply (cases \(1 + (int m + int n mod int LENGTH('a))) mod - int LENGTH('a) = 0\) - using zmod_zminus1_eq_if [of \1 + (int m + int n mod int LENGTH('a))\ \int LENGTH('a)\] - apply simp_all - apply (auto simp add: algebra_simps) - apply (simp add: minus_equation_iff [of \int m\]) - apply (drule sym [of _ \int m\]) - apply simp - apply (metis add.commute add_minus_cancel diff_minus_eq_add len_gt_0 less_imp_of_nat_less less_nat_zero_code mod_mult_self3 of_nat_gt_0 zmod_minus1) - apply (metis (no_types, hide_lams) Abs_fnat_hom_add less_not_refl mod_Suc of_nat_Suc of_nat_gt_0 of_nat_mod) - done - then have \int ((m + n) mod LENGTH('a)) = - int (LENGTH('a) - Suc ((LENGTH('a) - Suc m + LENGTH('a) - n mod LENGTH('a)) mod LENGTH('a)))\ - using \m < LENGTH('a)\ - by (simp only: of_nat_mod mod_simps) - (simp add: of_nat_diff of_nat_mod Suc_le_eq add_less_mono algebra_simps mod_simps) - then have \(m + n) mod LENGTH('a) = - LENGTH('a) - Suc ((LENGTH('a) - Suc m + LENGTH('a) - n mod LENGTH('a)) mod LENGTH('a))\ - by simp - ultimately show \bit (word_rotr n w) m \ bit (word_reverse (word_rotl n (word_reverse w))) m\ - by (simp add: word_rotl_eq_word_rotr bit_word_rotr_iff bit_word_reverse_iff) -qed - -lemma word_roti_0 [simp]: "word_roti 0 w = w" - by transfer simp - -lemma word_roti_add: "word_roti (m + n) w = word_roti m (word_roti n w)" - by (rule bit_word_eqI) - (simp add: bit_word_roti_iff nat_less_iff mod_simps ac_simps) - -lemma word_roti_conv_mod': - "word_roti n w = word_roti (n mod int (size w)) w" - by transfer simp - -lemmas word_roti_conv_mod = word_roti_conv_mod' [unfolded word_size] - - -subsubsection \"Word rotation commutes with bit-wise operations\ - -\ \using locale to not pollute lemma namespace\ -locale word_rotate -begin - -lemma word_rot_logs: - "word_rotl n (NOT v) = NOT (word_rotl n v)" - "word_rotr n (NOT v) = NOT (word_rotr n v)" - "word_rotl n (x AND y) = word_rotl n x AND word_rotl n y" - "word_rotr n (x AND y) = word_rotr n x AND word_rotr n y" - "word_rotl n (x OR y) = word_rotl n x OR word_rotl n y" - "word_rotr n (x OR y) = word_rotr n x OR word_rotr n y" - "word_rotl n (x XOR y) = word_rotl n x XOR word_rotl n y" - "word_rotr n (x XOR y) = word_rotr n x XOR word_rotr n y" - apply (rule bit_word_eqI) - apply (auto simp add: bit_word_rotl_iff bit_not_iff algebra_simps exp_eq_zero_iff not_le) - apply (rule bit_word_eqI) - apply (auto simp add: bit_word_rotr_iff bit_not_iff algebra_simps exp_eq_zero_iff not_le) - apply (rule bit_word_eqI) - apply (auto simp add: bit_word_rotl_iff bit_and_iff algebra_simps exp_eq_zero_iff not_le) - apply (rule bit_word_eqI) - apply (auto simp add: bit_word_rotr_iff bit_and_iff algebra_simps exp_eq_zero_iff not_le) - apply (rule bit_word_eqI) - apply (auto simp add: bit_word_rotl_iff bit_or_iff algebra_simps exp_eq_zero_iff not_le) - apply (rule bit_word_eqI) - apply (auto simp add: bit_word_rotr_iff bit_or_iff algebra_simps exp_eq_zero_iff not_le) - apply (rule bit_word_eqI) - apply (auto simp add: bit_word_rotl_iff bit_xor_iff algebra_simps exp_eq_zero_iff not_le) - apply (rule bit_word_eqI) - apply (auto simp add: bit_word_rotr_iff bit_xor_iff algebra_simps exp_eq_zero_iff not_le) - done - -end - -lemmas word_rot_logs = word_rotate.word_rot_logs - -lemma word_rotx_0 [simp] : "word_rotr i 0 = 0 \ word_rotl i 0 = 0" - by transfer simp_all - -lemma word_roti_0' [simp] : "word_roti n 0 = 0" - by transfer simp - -declare word_roti_eq_word_rotr_word_rotl [simp] - - -subsection \Maximum machine word\ - -lemma word_int_cases: - fixes x :: "'a::len word" - obtains n where "x = word_of_int n" and "0 \ n" and "n < 2^LENGTH('a)" - by (rule that [of \uint x\]) simp_all - -lemma word_nat_cases [cases type: word]: - fixes x :: "'a::len word" - obtains n where "x = of_nat n" and "n < 2^LENGTH('a)" - by (rule that [of \unat x\]) simp_all - -lemma max_word_max [intro!]: "n \ max_word" - by (fact word_order.extremum) - -lemma word_of_int_2p_len: "word_of_int (2 ^ LENGTH('a)) = (0::'a::len word)" - by simp - -lemma word_pow_0: "(2::'a::len word) ^ LENGTH('a) = 0" - by (fact word_exp_length_eq_0) - -lemma max_word_wrap: "x + 1 = 0 \ x = max_word" - by (simp add: eq_neg_iff_add_eq_0) - -lemma word_and_max: "x AND max_word = x" - by (fact word_log_esimps) - -lemma word_or_max: "x OR max_word = max_word" - by (fact word_log_esimps) - -lemma word_ao_dist2: "x AND (y OR z) = x AND y OR x AND z" - for x y z :: "'a::len word" - by (fact bit.conj_disj_distrib) - -lemma word_oa_dist2: "x OR y AND z = (x OR y) AND (x OR z)" - for x y z :: "'a::len word" - by (fact bit.disj_conj_distrib) - -lemma word_and_not [simp]: "x AND NOT x = 0" - for x :: "'a::len word" - by (fact bit.conj_cancel_right) - -lemma word_or_not [simp]: "x OR NOT x = max_word" - by (fact bit.disj_cancel_right) - -lemma word_xor_and_or: "x XOR y = x AND NOT y OR NOT x AND y" - for x y :: "'a::len word" - by (fact bit.xor_def) - -lemma uint_lt_0 [simp]: "uint x < 0 = False" - by (simp add: linorder_not_less) - -lemma shiftr1_1 [simp]: "shiftr1 (1::'a::len word) = 0" - by transfer simp - -lemma word_less_1 [simp]: "x < 1 \ x = 0" - for x :: "'a::len word" - by (simp add: word_less_nat_alt unat_0_iff) - -lemma uint_plus_if_size: - "uint (x + y) = - (if uint x + uint y < 2^size x - then uint x + uint y - else uint x + uint y - 2^size x)" - apply (simp only: word_arith_wis word_size uint_word_of_int_eq) - apply (auto simp add: not_less take_bit_int_eq_self_iff) - apply (metis not_less take_bit_eq_mod uint_plus_if' uint_word_ariths(1)) - done - -lemma unat_plus_if_size: - "unat (x + y) = - (if unat x + unat y < 2^size x - then unat x + unat y - else unat x + unat y - 2^size x)" - for x y :: "'a::len word" - apply (subst word_arith_nat_defs) - apply (subst unat_of_nat) - apply (auto simp add: not_less word_size) - apply (metis not_le unat_plus_if' unat_word_ariths(1)) - done - -lemma word_neq_0_conv: "w \ 0 \ 0 < w" - for w :: "'a::len word" - by (fact word_coorder.not_eq_extremum) - -lemma max_lt: "unat (max a b div c) = unat (max a b) div unat c" - for c :: "'a::len word" - by (fact unat_div) - -lemma uint_sub_if_size: - "uint (x - y) = - (if uint y \ uint x - then uint x - uint y - else uint x - uint y + 2^size x)" - apply (simp only: word_arith_wis word_size uint_word_of_int_eq) - apply (auto simp add: take_bit_int_eq_self_iff not_le) - apply (metis not_less uint_sub_if' uint_word_arith_bintrs(2)) - done - -lemma unat_sub: - \unat (a - b) = unat a - unat b\ - if \b \ a\ -proof - - from that have \unat b \ unat a\ - by transfer simp - with that show ?thesis - apply transfer - apply simp - apply (subst take_bit_diff [symmetric]) - apply (subst nat_take_bit_eq) - apply (simp add: nat_le_eq_zle) - apply (simp add: nat_diff_distrib take_bit_nat_eq_self_iff less_imp_diff_less) - done -qed - -lemmas word_less_sub1_numberof [simp] = word_less_sub1 [of "numeral w"] for w -lemmas word_le_sub1_numberof [simp] = word_le_sub1 [of "numeral w"] for w - -lemma word_of_int_minus: "word_of_int (2^LENGTH('a) - i) = (word_of_int (-i)::'a::len word)" - apply transfer - apply (subst take_bit_diff [symmetric]) - apply (simp add: take_bit_minus) - done - -lemma word_of_int_inj: - \(word_of_int x :: 'a::len word) = word_of_int y \ x = y\ - if \0 \ x \ x < 2 ^ LENGTH('a)\ \0 \ y \ y < 2 ^ LENGTH('a)\ - using that by (transfer fixing: x y) (simp add: take_bit_int_eq_self) - -lemma word_le_less_eq: "x \ y \ x = y \ x < y" - for x y :: "'z::len word" - by (auto simp add: order_class.le_less) - -lemma mod_plus_cong: - fixes b b' :: int - assumes 1: "b = b'" - and 2: "x mod b' = x' mod b'" - and 3: "y mod b' = y' mod b'" - and 4: "x' + y' = z'" - shows "(x + y) mod b = z' mod b'" -proof - - from 1 2[symmetric] 3[symmetric] have "(x + y) mod b = (x' mod b' + y' mod b') mod b'" - by (simp add: mod_add_eq) - also have "\ = (x' + y') mod b'" - by (simp add: mod_add_eq) - finally show ?thesis - by (simp add: 4) -qed - -lemma mod_minus_cong: - fixes b b' :: int - assumes "b = b'" - and "x mod b' = x' mod b'" - and "y mod b' = y' mod b'" - and "x' - y' = z'" - shows "(x - y) mod b = z' mod b'" - using assms [symmetric] by (auto intro: mod_diff_cong) - -lemma word_induct_less: - \P m\ if zero: \P 0\ and less: \\n. n < m \ P n \ P (1 + n)\ - for m :: \'a::len word\ -proof - - define q where \q = unat m\ - with less have \\n. n < word_of_nat q \ P n \ P (1 + n)\ - by simp - then have \P (word_of_nat q :: 'a word)\ - proof (induction q) - case 0 - show ?case - by (simp add: zero) - next - case (Suc q) - show ?case - proof (cases \1 + word_of_nat q = (0 :: 'a word)\) - case True - then show ?thesis - by (simp add: zero) - next - case False - then have *: \word_of_nat q < (word_of_nat (Suc q) :: 'a word)\ - by (simp add: unatSuc word_less_nat_alt) - then have **: \n < (1 + word_of_nat q :: 'a word) \ n \ (word_of_nat q :: 'a word)\ for n - by (metis (no_types, lifting) add.commute inc_le le_less_trans not_less of_nat_Suc) - have \P (word_of_nat q)\ - apply (rule Suc.IH) - apply (rule Suc.prems) - apply (erule less_trans) - apply (rule *) - apply assumption - done - with * have \P (1 + word_of_nat q)\ - by (rule Suc.prems) - then show ?thesis - by simp - qed - qed - with \q = unat m\ show ?thesis - by simp -qed - -lemma word_induct: "P 0 \ (\n. P n \ P (1 + n)) \ P m" - for P :: "'a::len word \ bool" - by (rule word_induct_less) - -lemma word_induct2 [induct type]: "P 0 \ (\n. 1 + n \ 0 \ P n \ P (1 + n)) \ P n" - for P :: "'b::len word \ bool" - apply (rule word_induct_less) - apply simp_all - apply (case_tac "1 + na = 0") - apply auto - done - - -subsection \Recursion combinator for words\ - -definition word_rec :: "'a \ ('b::len word \ 'a \ 'a) \ 'b word \ 'a" - where "word_rec forZero forSuc n = rec_nat forZero (forSuc \ of_nat) (unat n)" - -lemma word_rec_0: "word_rec z s 0 = z" - by (simp add: word_rec_def) - -lemma word_rec_Suc: "1 + n \ 0 \ word_rec z s (1 + n) = s n (word_rec z s n)" - for n :: "'a::len word" - apply (auto simp add: word_rec_def unat_word_ariths) - apply (metis (mono_tags, lifting) Abs_fnat_hom_add add_diff_cancel_left' o_def of_nat_1 old.nat.simps(7) plus_1_eq_Suc unatSuc unat_word_ariths(1) unsigned_1 word_arith_nat_add) - done - -lemma word_rec_Pred: "n \ 0 \ word_rec z s n = s (n - 1) (word_rec z s (n - 1))" - apply (rule subst[where t="n" and s="1 + (n - 1)"]) - apply simp - apply (subst word_rec_Suc) - apply simp - apply simp - done - -lemma word_rec_in: "f (word_rec z (\_. f) n) = word_rec (f z) (\_. f) n" - by (induct n) (simp_all add: word_rec_0 word_rec_Suc) - -lemma word_rec_in2: "f n (word_rec z f n) = word_rec (f 0 z) (f \ (+) 1) n" - by (induct n) (simp_all add: word_rec_0 word_rec_Suc) - -lemma word_rec_twice: - "m \ n \ word_rec z f n = word_rec (word_rec z f (n - m)) (f \ (+) (n - m)) m" - apply (erule rev_mp) - apply (rule_tac x=z in spec) - apply (rule_tac x=f in spec) - apply (induct n) - apply (simp add: word_rec_0) - apply clarsimp - apply (rule_tac t="1 + n - m" and s="1 + (n - m)" in subst) - apply simp - apply (case_tac "1 + (n - m) = 0") - apply (simp add: word_rec_0) - apply (rule_tac f = "word_rec a b" for a b in arg_cong) - apply (rule_tac t="m" and s="m + (1 + (n - m))" in subst) - apply simp - apply (simp (no_asm_use)) - apply (simp add: word_rec_Suc word_rec_in2) - apply (erule impE) - apply uint_arith - apply (drule_tac x="x \ (+) 1" in spec) - apply (drule_tac x="x 0 xa" in spec) - apply simp - apply (rule_tac t="\a. x (1 + (n - m + a))" and s="\a. x (1 + (n - m) + a)" in subst) - apply (clarsimp simp add: fun_eq_iff) - apply (rule_tac t="(1 + (n - m + xb))" and s="1 + (n - m) + xb" in subst) - apply simp - apply (rule refl) - apply (rule refl) - done - -lemma word_rec_id: "word_rec z (\_. id) n = z" - by (induct n) (auto simp add: word_rec_0 word_rec_Suc) - -lemma word_rec_id_eq: "\m < n. f m = id \ word_rec z f n = z" - apply (erule rev_mp) - apply (induct n) - apply (auto simp add: word_rec_0 word_rec_Suc) - apply (drule spec, erule mp) - apply uint_arith - apply (drule_tac x=n in spec, erule impE) - apply uint_arith - apply simp - done - -lemma word_rec_max: - "\m\n. m \ - 1 \ f m = id \ word_rec z f (- 1) = word_rec z f n" - apply (subst word_rec_twice[where n="-1" and m="-1 - n"]) - apply simp - apply simp - apply (rule word_rec_id_eq) - apply clarsimp - apply (drule spec, rule mp, erule mp) - apply (rule word_plus_mono_right2[OF _ order_less_imp_le]) - prefer 2 - apply assumption - apply simp - apply (erule contrapos_pn) - apply simp - apply (drule arg_cong[where f="\x. x - n"]) - apply simp - done - - -subsection \More\ - -lemma mask_1: "mask 1 = 1" - by simp - -lemma mask_Suc_0: "mask (Suc 0) = 1" - by simp - -lemma bin_last_bintrunc: "odd (take_bit l n) \ l > 0 \ odd n" - by simp - - -lemma push_bit_word_beyond [simp]: - \push_bit n w = 0\ if \LENGTH('a) \ n\ for w :: \'a::len word\ - using that by (transfer fixing: n) (simp add: take_bit_push_bit) - -lemma drop_bit_word_beyond [simp]: - \drop_bit n w = 0\ if \LENGTH('a) \ n\ for w :: \'a::len word\ - using that by (transfer fixing: n) (simp add: drop_bit_take_bit) - -lemma signed_drop_bit_beyond: - \signed_drop_bit n w = (if bit w (LENGTH('a) - Suc 0) then - 1 else 0)\ - if \LENGTH('a) \ n\ for w :: \'a::len word\ - by (rule bit_word_eqI) (simp add: bit_signed_drop_bit_iff that) - - -subsection \SMT support\ - -ML_file \Tools/smt_word.ML\ - -end diff -r d8661799afb2 -r c7038c397ae3 src/HOL/Word/Word_Examples.thy --- a/src/HOL/Word/Word_Examples.thy Thu Oct 29 09:59:40 2020 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,164 +0,0 @@ -(* Title: HOL/Word/Word_Examples.thy - Authors: Gerwin Klein and Thomas Sewell, NICTA - -Examples demonstrating and testing various word operations. -*) - -section "Examples of word operations" - -theory Word_Examples - imports Word Misc_lsb Misc_set_bit -begin - -type_synonym word32 = "32 word" -type_synonym word8 = "8 word" -type_synonym byte = word8 - -text "modulus" - -lemma "(27 :: 4 word) = -5" by simp - -lemma "(27 :: 4 word) = 11" by simp - -lemma "27 \ (11 :: 6 word)" by simp - -text "signed" - -lemma "(127 :: 6 word) = -1" by simp - -text "number ring simps" - -lemma - "27 + 11 = (38::'a::len word)" - "27 + 11 = (6::5 word)" - "7 * 3 = (21::'a::len word)" - "11 - 27 = (-16::'a::len word)" - "- (- 11) = (11::'a::len word)" - "-40 + 1 = (-39::'a::len word)" - by simp_all - -lemma "word_pred 2 = 1" by simp - -lemma "word_succ (- 3) = -2" by simp - -lemma "23 < (27::8 word)" by simp -lemma "23 \ (27::8 word)" by simp -lemma "\ 23 < (27::2 word)" by simp -lemma "0 < (4::3 word)" by simp -lemma "1 < (4::3 word)" by simp -lemma "0 < (1::3 word)" by simp - -text "ring operations" - -lemma "a + 2 * b + c - b = (b + c) + (a :: 32 word)" by simp - -text "casting" - -lemma "uint (234567 :: 10 word) = 71" by simp -lemma "uint (-234567 :: 10 word) = 953" by simp -lemma "sint (234567 :: 10 word) = 71" by simp -lemma "sint (-234567 :: 10 word) = -71" by simp -lemma "uint (1 :: 10 word) = 1" by simp - -lemma "unat (-234567 :: 10 word) = 953" by simp -lemma "unat (1 :: 10 word) = 1" by simp - -lemma "ucast (0b1010 :: 4 word) = (0b10 :: 2 word)" by simp -lemma "ucast (0b1010 :: 4 word) = (0b1010 :: 10 word)" by simp -lemma "scast (0b1010 :: 4 word) = (0b111010 :: 6 word)" by simp -lemma "ucast (1 :: 4 word) = (1 :: 2 word)" by simp - -text "reducing goals to nat or int and arith:" -lemma "i < x \ i < i + 1" for i x :: "'a::len word" - by unat_arith -lemma "i < x \ i < i + 1" for i x :: "'a::len word" - by unat_arith - -text "bool lists" - -lemma "of_bl [True, False, True, True] = (0b1011::'a::len word)" by simp - -lemma "to_bl (0b110::4 word) = [False, True, True, False]" by simp - -lemma "of_bl (replicate 32 True) = (0xFFFFFFFF::32 word)" - by (simp add: numeral_eq_Suc) - - -text "bit operations" - -lemma "0b110 AND 0b101 = (0b100 :: 32 word)" by simp -lemma "0b110 OR 0b011 = (0b111 :: 8 word)" by simp -lemma "0xF0 XOR 0xFF = (0x0F :: byte)" by simp -lemma "NOT (0xF0 :: 16 word) = 0xFF0F" by simp -lemma "0 AND 5 = (0 :: byte)" by simp -lemma "1 AND 1 = (1 :: byte)" by simp -lemma "1 AND 0 = (0 :: byte)" by simp -lemma "1 AND 5 = (1 :: byte)" by simp -lemma "1 OR 6 = (7 :: byte)" by simp -lemma "1 OR 1 = (1 :: byte)" by simp -lemma "1 XOR 7 = (6 :: byte)" by simp -lemma "1 XOR 1 = (0 :: byte)" by simp -lemma "NOT 1 = (254 :: byte)" by simp -lemma "NOT 0 = (255 :: byte)" by simp - -lemma "(-1 :: 32 word) = 0xFFFFFFFF" by simp - -lemma "(0b0010 :: 4 word) !! 1" by simp -lemma "\ (0b0010 :: 4 word) !! 0" by simp -lemma "\ (0b1000 :: 3 word) !! 4" by simp -lemma "\ (1 :: 3 word) !! 2" by simp - -lemma "(0b11000 :: 10 word) !! n = (n = 4 \ n = 3)" - by (auto simp add: bin_nth_Bit0 bin_nth_Bit1) - -lemma "set_bit 55 7 True = (183::'a::len word)" by simp -lemma "set_bit 0b0010 7 True = (0b10000010::'a::len word)" by simp -lemma "set_bit 0b0010 1 False = (0::'a::len word)" by simp -lemma "set_bit 1 3 True = (0b1001::'a::len word)" by simp -lemma "set_bit 1 0 False = (0::'a::len word)" by simp -lemma "set_bit 0 3 True = (0b1000::'a::len word)" by simp -lemma "set_bit 0 3 False = (0::'a::len word)" by simp - -lemma "lsb (0b0101::'a::len word)" by simp -lemma "\ lsb (0b1000::'a::len word)" by simp -lemma "lsb (1::'a::len word)" by simp -lemma "\ lsb (0::'a::len word)" by simp - -lemma "\ msb (0b0101::4 word)" by simp -lemma "msb (0b1000::4 word)" by simp -lemma "\ msb (1::4 word)" by simp -lemma "\ msb (0::4 word)" by simp - -lemma "word_cat (27::4 word) (27::8 word) = (2843::'a::len word)" by simp -lemma "word_cat (0b0011::4 word) (0b1111::6word) = (0b0011001111 :: 10 word)" - by simp - -lemma "0b1011 << 2 = (0b101100::'a::len word)" by simp -lemma "0b1011 >> 2 = (0b10::8 word)" by simp -lemma "0b1011 >>> 2 = (0b10::8 word)" by simp -lemma "1 << 2 = (0b100::'a::len word)" apply simp? oops - -lemma "slice 3 (0b101111::6 word) = (0b101::3 word)" by simp -lemma "slice 3 (1::6 word) = (0::3 word)" apply simp? oops - -lemma "word_rotr 2 0b0110 = (0b1001::4 word)" by simp -lemma "word_rotl 1 0b1110 = (0b1101::4 word)" by simp -lemma "word_roti 2 0b1110 = (0b1011::4 word)" by simp -lemma "word_roti (- 2) 0b0110 = (0b1001::4 word)" by simp -lemma "word_rotr 2 0 = (0::4 word)" by simp -lemma "word_rotr 2 1 = (0b0100::4 word)" apply simp? oops -lemma "word_rotl 2 1 = (0b0100::4 word)" apply simp? oops -lemma "word_roti (- 2) 1 = (0b0100::4 word)" apply simp? oops - -lemma "(x AND 0xff00) OR (x AND 0x00ff) = (x::16 word)" -proof - - have "(x AND 0xff00) OR (x AND 0x00ff) = x AND (0xff00 OR 0x00ff)" - by (simp only: word_ao_dist2) - also have "0xff00 OR 0x00ff = (-1::16 word)" - by simp - also have "x AND -1 = x" - by simp - finally show ?thesis . -qed - -end diff -r d8661799afb2 -r c7038c397ae3 src/HOL/Word/Word_rsplit.thy --- a/src/HOL/Word/Word_rsplit.thy Thu Oct 29 09:59:40 2020 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,149 +0,0 @@ -(* Title: HOL/Word/Word_rsplit.thy - Author: Jeremy Dawson and Gerwin Klein, NICTA -*) - -theory Word_rsplit - imports Bits_Int Word -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 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 - -end \ No newline at end of file diff -r d8661799afb2 -r c7038c397ae3 src/HOL/Word/document/root.bib --- a/src/HOL/Word/document/root.bib Thu Oct 29 09:59:40 2020 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,14 +0,0 @@ -@InProceedings{dawson-avocs07, - author = {Jeremy Dawson}, - title = {Isabelle Theories for Machine Words}, - booktitle = {Seventh International Workshop on Automated Verification of Critical Systems (AVOCS'07)}, - pages = {15}, - year = {2007}, - editor = {Michael Goldsmith and Bill Roscoe}, - series = {Electronic Notes in Theoretical Computer Science}, - address = {Oxford}, - month = Sep, - publisher = {Elsevier}, - note = {to appear}, -} - diff -r d8661799afb2 -r c7038c397ae3 src/HOL/Word/document/root.tex --- a/src/HOL/Word/document/root.tex Thu Oct 29 09:59:40 2020 +0000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,38 +0,0 @@ -\documentclass[11pt,a4paper]{article} -\usepackage{graphicx,isabelle,isabellesym,amssymb} -\usepackage{pdfsetup} - -\urlstyle{rm} -\isabellestyle{it} -\pagestyle{myheadings} - -\begin{document} - -\title{Machine Words in Isabelle/HOL} - -\author{Jeremy Dawson, Paul Graunke, Brian Huffman, Gerwin Klein, and John Matthews} - -\maketitle - -\begin{abstract} -A formalisation of generic, fixed size machine words in Isabelle/HOL. An earlier version of this -formalisation is described in \cite{dawson-avocs07}. -\end{abstract} - -\tableofcontents - -\begin{center} - \includegraphics[width=\textwidth,height=\textheight,keepaspectratio]{session_graph} -\end{center} - -\newpage - -\renewcommand{\setisabellecontext}[1]{\markright{THEORY~``#1''}} - -\parindent 0pt\parskip 0.5ex -\input{session} - -\bibliographystyle{plain} -\bibliography{root} - -\end{document} diff -r d8661799afb2 -r c7038c397ae3 src/HOL/ex/Bit_Lists.thy --- a/src/HOL/ex/Bit_Lists.thy Thu Oct 29 09:59:40 2020 +0000 +++ b/src/HOL/ex/Bit_Lists.thy Thu Oct 29 10:03:03 2020 +0000 @@ -5,7 +5,7 @@ theory Bit_Lists imports - "HOL-Word.Word" "HOL-Library.More_List" + "HOL-Library.Word" "HOL-Library.More_List" begin subsection \Fragments of algebraic bit representations\