diff --git a/src/HOL/ex/Bit_Lists.thy b/src/HOL/ex/Bit_Lists.thy --- a/src/HOL/ex/Bit_Lists.thy +++ b/src/HOL/ex/Bit_Lists.thy @@ -1,375 +1,380 @@ (* Author: Florian Haftmann, TUM *) section \Proof(s) of concept for algebraically founded lists of bits\ theory Bit_Lists imports Word "HOL-Library.More_List" begin subsection \Fragments of algebraic bit representations\ context comm_semiring_1 begin primrec radix_value :: "('b \ 'a) \ 'a \ 'b list \ 'a" where "radix_value f b [] = 0" | "radix_value f b (a # as) = f a + radix_value f b as * b" abbreviation (input) unsigned_of_bits :: "bool list \ 'a" where "unsigned_of_bits \ radix_value of_bool 2" lemma unsigned_of_bits_replicate_False [simp]: "unsigned_of_bits (replicate n False) = 0" by (induction n) simp_all end context unique_euclidean_semiring_with_bit_shifts begin lemma unsigned_of_bits_append [simp]: "unsigned_of_bits (bs @ cs) = unsigned_of_bits bs + push_bit (length bs) (unsigned_of_bits cs)" by (induction bs) (simp_all add: push_bit_double, simp_all add: algebra_simps) lemma unsigned_of_bits_take [simp]: "unsigned_of_bits (take n bs) = take_bit n (unsigned_of_bits bs)" proof (induction bs arbitrary: n) case Nil then show ?case by simp next case (Cons b bs) then show ?case by (cases n) (simp_all add: ac_simps) qed lemma unsigned_of_bits_drop [simp]: "unsigned_of_bits (drop n bs) = Parity.drop_bit n (unsigned_of_bits bs)" proof (induction bs arbitrary: n) case Nil then show ?case by simp next case (Cons b bs) then show ?case by (cases n) simp_all qed lemma bit_unsigned_of_bits_iff: \bit (unsigned_of_bits bs) n \ nth_default False bs n\ proof (induction bs arbitrary: n) case Nil then show ?case by simp next case (Cons b bs) then show ?case by (cases n) simp_all qed primrec n_bits_of :: "nat \ 'a \ bool list" where "n_bits_of 0 a = []" | "n_bits_of (Suc n) a = odd a # n_bits_of n (a div 2)" lemma n_bits_of_eq_iff: "n_bits_of n a = n_bits_of n b \ Parity.take_bit n a = Parity.take_bit n b" apply (induction n arbitrary: a b) apply (auto elim!: evenE oddE) apply (metis dvd_triv_right even_plus_one_iff) apply (metis dvd_triv_right even_plus_one_iff) done lemma take_n_bits_of [simp]: "take m (n_bits_of n a) = n_bits_of (min m n) a" proof - define q and v and w where "q = min m n" and "v = m - q" and "w = n - q" then have "v = 0 \ w = 0" by auto then have "take (q + v) (n_bits_of (q + w) a) = n_bits_of q a" by (induction q arbitrary: a) auto with q_def v_def w_def show ?thesis by simp qed lemma unsigned_of_bits_n_bits_of [simp]: "unsigned_of_bits (n_bits_of n a) = Parity.take_bit n a" by (induction n arbitrary: a) (simp_all add: ac_simps) end subsection \Syntactic bit representation\ class bit_representation = fixes bits_of :: "'a \ bool list" and of_bits :: "bool list \ 'a" assumes of_bits_of [simp]: "of_bits (bits_of a) = a" text \Unclear whether a \<^typ>\bool\ instantiation is needed or not\ instantiation nat :: bit_representation begin fun bits_of_nat :: "nat \ bool list" where "bits_of (n::nat) = (if n = 0 then [] else odd n # bits_of (n div 2))" lemma bits_of_nat_simps [simp]: "bits_of (0::nat) = []" "n > 0 \ bits_of n = odd n # bits_of (n div 2)" for n :: nat by simp_all declare bits_of_nat.simps [simp del] definition of_bits_nat :: "bool list \ nat" where [simp]: "of_bits_nat = unsigned_of_bits" \ \remove simp\ instance proof show "of_bits (bits_of n) = n" for n :: nat by (induction n rule: nat_bit_induct) simp_all qed end lemma bit_of_bits_nat_iff: \bit (of_bits bs :: nat) n \ nth_default False bs n\ by (simp add: bit_unsigned_of_bits_iff) lemma bits_of_Suc_0 [simp]: "bits_of (Suc 0) = [True]" by simp lemma bits_of_1_nat [simp]: "bits_of (1 :: nat) = [True]" by simp lemma bits_of_nat_numeral_simps [simp]: "bits_of (numeral Num.One :: nat) = [True]" (is ?One) "bits_of (numeral (Num.Bit0 n) :: nat) = False # bits_of (numeral n :: nat)" (is ?Bit0) "bits_of (numeral (Num.Bit1 n) :: nat) = True # bits_of (numeral n :: nat)" (is ?Bit1) proof - show ?One by simp define m :: nat where "m = numeral n" then have "m > 0" and *: "numeral n = m" "numeral (Num.Bit0 n) = 2 * m" "numeral (Num.Bit1 n) = Suc (2 * m)" by simp_all from \m > 0\ show ?Bit0 ?Bit1 by (simp_all add: *) qed lemma unsigned_of_bits_of_nat [simp]: "unsigned_of_bits (bits_of n) = n" for n :: nat using of_bits_of [of n] by simp instantiation int :: bit_representation begin fun bits_of_int :: "int \ bool list" where "bits_of_int k = odd k # (if k = 0 \ k = - 1 then [] else bits_of_int (k div 2))" lemma bits_of_int_simps [simp]: "bits_of (0 :: int) = [False]" "bits_of (- 1 :: int) = [True]" "k \ 0 \ k \ - 1 \ bits_of k = odd k # bits_of (k div 2)" for k :: int by simp_all lemma bits_of_not_Nil [simp]: "bits_of k \ []" for k :: int by simp declare bits_of_int.simps [simp del] definition of_bits_int :: "bool list \ int" where "of_bits_int bs = (if bs = [] \ \ last bs then unsigned_of_bits bs else unsigned_of_bits bs - 2 ^ length bs)" lemma of_bits_int_simps [simp]: "of_bits [] = (0 :: int)" "of_bits [False] = (0 :: int)" "of_bits [True] = (- 1 :: int)" "of_bits (bs @ [b]) = (unsigned_of_bits bs :: int) - (2 ^ length bs) * of_bool b" "of_bits (False # bs) = 2 * (of_bits bs :: int)" "bs \ [] \ of_bits (True # bs) = 1 + 2 * (of_bits bs :: int)" by (simp_all add: of_bits_int_def push_bit_of_1) instance proof show "of_bits (bits_of k) = k" for k :: int by (induction k rule: int_bit_induct) simp_all qed lemma bits_of_1_int [simp]: "bits_of (1 :: int) = [True, False]" by simp lemma bits_of_int_numeral_simps [simp]: "bits_of (numeral Num.One :: int) = [True, False]" (is ?One) "bits_of (numeral (Num.Bit0 n) :: int) = False # bits_of (numeral n :: int)" (is ?Bit0) "bits_of (numeral (Num.Bit1 n) :: int) = True # bits_of (numeral n :: int)" (is ?Bit1) "bits_of (- numeral (Num.Bit0 n) :: int) = False # bits_of (- numeral n :: int)" (is ?nBit0) "bits_of (- numeral (Num.Bit1 n) :: int) = True # bits_of (- numeral (Num.inc n) :: int)" (is ?nBit1) proof - show ?One by simp define k :: int where "k = numeral n" then have "k > 0" and *: "numeral n = k" "numeral (Num.Bit0 n) = 2 * k" "numeral (Num.Bit1 n) = 2 * k + 1" "numeral (Num.inc n) = k + 1" by (simp_all add: add_One) have "- (2 * k) div 2 = - k" "(- (2 * k) - 1) div 2 = - k - 1" by simp_all with \k > 0\ show ?Bit0 ?Bit1 ?nBit0 ?nBit1 by (simp_all add: *) qed lemma bit_of_bits_int_iff: \bit (of_bits bs :: int) n \ nth_default (bs \ [] \ last bs) bs n\ proof (induction bs arbitrary: n) case Nil then show ?case by simp next case (Cons b bs) then show ?case by (cases n; cases b; cases bs) simp_all qed lemma of_bits_append [simp]: "of_bits (bs @ cs) = of_bits bs + push_bit (length bs) (of_bits cs :: int)" if "bs \ []" "\ last bs" using that proof (induction bs rule: list_nonempty_induct) case (single b) then show ?case by simp next case (cons b bs) then show ?case by (cases b) (simp_all add: push_bit_double) qed lemma of_bits_replicate_False [simp]: "of_bits (replicate n False) = (0 :: int)" by (auto simp add: of_bits_int_def) lemma of_bits_drop [simp]: "of_bits (drop n bs) = Parity.drop_bit n (of_bits bs :: int)" if "n < length bs" using that proof (induction bs arbitrary: n) case Nil then show ?case by simp next case (Cons b bs) show ?case proof (cases n) case 0 then show ?thesis by simp next case (Suc n) with Cons.prems have "bs \ []" by auto with Suc Cons.IH [of n] Cons.prems show ?thesis by (cases b) simp_all qed qed end lemma unsigned_of_bits_eq_of_bits: "unsigned_of_bits bs = (of_bits (bs @ [False]) :: int)" by (simp add: of_bits_int_def) +unbundle word.lifting + instantiation word :: (len) bit_representation begin lift_definition bits_of_word :: "'a word \ bool list" is "n_bits_of LENGTH('a)" by (simp add: n_bits_of_eq_iff) lift_definition of_bits_word :: "bool list \ 'a word" is unsigned_of_bits . instance proof fix a :: "'a word" show "of_bits (bits_of a) = a" by transfer simp qed end +lifting_update word.lifting +lifting_forget word.lifting + subsection \Bit representations with bit operations\ class semiring_bit_representation = semiring_bit_operations + bit_representation + assumes and_eq: "length bs = length cs \ of_bits bs AND of_bits cs = of_bits (map2 (\) bs cs)" and or_eq: "length bs = length cs \ of_bits bs OR of_bits cs = of_bits (map2 (\) bs cs)" and xor_eq: "length bs = length cs \ of_bits bs XOR of_bits cs = of_bits (map2 (\) bs cs)" and push_bit_eq: "push_bit n a = of_bits (replicate n False @ bits_of a)" and drop_bit_eq: "n < length (bits_of a) \ drop_bit n a = of_bits (drop n (bits_of a))" class ring_bit_representation = ring_bit_operations + semiring_bit_representation + assumes not_eq: "not = of_bits \ map Not \ bits_of" context zip_nat begin lemma of_bits: "of_bits bs \<^bold>\ of_bits cs = (of_bits (map2 (\<^bold>*) bs cs) :: nat)" if "length bs = length cs" for bs cs by (simp add: Parity.bit_eq_iff bit_unsigned_of_bits_iff bit_eq_iff) (simp add: that end_of_bits nth_default_map2 [of _ _ _ False False]) end instance nat :: semiring_bit_representation apply standard apply (simp_all only: and_nat.of_bits or_nat.of_bits xor_nat.of_bits) apply simp_all done context zip_int begin lemma of_bits: \of_bits bs \<^bold>\ of_bits cs = (of_bits (map2 (\<^bold>*) bs cs) :: int)\ if \length bs = length cs\ and \\ False \<^bold>* False\ for bs cs proof (cases \bs = []\) case True moreover have \cs = []\ using True that by simp ultimately show ?thesis by (simp add: Parity.bit_eq_iff bit_eq_iff that) next case False moreover have \cs \ []\ using False that by auto ultimately show ?thesis apply (simp add: Parity.bit_eq_iff bit_of_bits_int_iff bit_eq_iff zip_eq_Nil_iff last_map last_zip that) apply (simp add: that nth_default_map2 [of _ _ _ \last bs\ \last cs\]) done qed end instance int :: ring_bit_representation proof show "(not :: int \ _) = of_bits \ map Not \ bits_of" proof (rule sym, rule ext) fix k :: int show "(of_bits \ map Not \ bits_of) k = NOT k" by (induction k rule: int_bit_induct) (simp_all add: not_int_def) qed show "push_bit n k = of_bits (replicate n False @ bits_of k)" for k :: int and n :: nat by (cases "n = 0") simp_all qed (simp_all add: and_int.of_bits or_int.of_bits xor_int.of_bits) end diff --git a/src/HOL/ex/Word.thy b/src/HOL/ex/Word.thy --- a/src/HOL/ex/Word.thy +++ b/src/HOL/ex/Word.thy @@ -1,718 +1,721 @@ (* Author: Florian Haftmann, TUM *) section \Proof of concept for algebraically founded bit word types\ theory Word imports Main "HOL-Library.Type_Length" "HOL-ex.Bit_Operations" begin subsection \Preliminaries\ definition signed_take_bit :: "nat \ int \ int" where signed_take_bit_eq_take_bit: "signed_take_bit n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n" lemma signed_take_bit_eq_take_bit': "signed_take_bit (n - Suc 0) k = take_bit n (k + 2 ^ (n - 1)) - 2 ^ (n - 1)" if "n > 0" using that by (simp add: signed_take_bit_eq_take_bit) lemma signed_take_bit_0 [simp]: "signed_take_bit 0 k = - (k mod 2)" proof (cases "even k") case True then have "odd (k + 1)" by simp then have "(k + 1) mod 2 = 1" by (simp add: even_iff_mod_2_eq_zero) with True show ?thesis by (simp add: signed_take_bit_eq_take_bit) next case False then show ?thesis by (simp add: signed_take_bit_eq_take_bit odd_iff_mod_2_eq_one) qed lemma signed_take_bit_Suc [simp]: "signed_take_bit (Suc n) k = signed_take_bit n (k div 2) * 2 + k mod 2" by (simp add: odd_iff_mod_2_eq_one signed_take_bit_eq_take_bit algebra_simps) lemma signed_take_bit_of_0 [simp]: "signed_take_bit n 0 = 0" by (simp add: signed_take_bit_eq_take_bit take_bit_eq_mod) lemma signed_take_bit_of_minus_1 [simp]: "signed_take_bit n (- 1) = - 1" by (induct n) simp_all lemma signed_take_bit_eq_iff_take_bit_eq: "signed_take_bit (n - Suc 0) k = signed_take_bit (n - Suc 0) l \ take_bit n k = take_bit n l" (is "?P \ ?Q") if "n > 0" proof - from that obtain m where m: "n = Suc m" by (cases n) auto show ?thesis proof assume ?Q have "take_bit (Suc m) (k + 2 ^ m) = take_bit (Suc m) (take_bit (Suc m) k + take_bit (Suc m) (2 ^ m))" by (simp only: take_bit_add) also have "\ = take_bit (Suc m) (take_bit (Suc m) l + take_bit (Suc m) (2 ^ m))" by (simp only: \?Q\ m [symmetric]) also have "\ = take_bit (Suc m) (l + 2 ^ m)" by (simp only: take_bit_add) finally show ?P by (simp only: signed_take_bit_eq_take_bit m) simp next assume ?P with that have "(k + 2 ^ (n - Suc 0)) mod 2 ^ n = (l + 2 ^ (n - Suc 0)) mod 2 ^ n" by (simp add: signed_take_bit_eq_take_bit' take_bit_eq_mod) then have "(i + (k + 2 ^ (n - Suc 0))) mod 2 ^ n = (i + (l + 2 ^ (n - Suc 0))) mod 2 ^ n" for i by (metis mod_add_eq) then have "k mod 2 ^ n = l mod 2 ^ n" by (metis add_diff_cancel_right' uminus_add_conv_diff) then show ?Q by (simp add: take_bit_eq_mod) qed qed subsection \Bit strings as quotient type\ subsubsection \Basic properties\ quotient_type (overloaded) 'a word = int / "\k l. take_bit LENGTH('a) k = take_bit LENGTH('a::len0) l" by (auto intro!: equivpI reflpI sympI transpI) instantiation word :: (len0) "{semiring_numeral, comm_semiring_0, comm_ring}" 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 plus by (subst take_bit_add [symmetric]) (simp add: take_bit_add) lift_definition uminus_word :: "'a word \ 'a word" is uminus by (subst take_bit_minus [symmetric]) (simp add: take_bit_minus) lift_definition minus_word :: "'a word \ 'a word \ 'a word" is minus by (subst take_bit_diff [symmetric]) (simp add: take_bit_diff) lift_definition times_word :: "'a word \ 'a word \ 'a word" is times by (auto simp add: take_bit_eq_mod intro: mod_mult_cong) instance by standard (transfer; simp add: algebra_simps)+ end instance word :: (len) comm_ring_1 by standard (transfer; simp)+ quickcheck_generator word constructors: "zero_class.zero :: ('a::len0) word", "numeral :: num \ ('a::len0) word", "uminus :: ('a::len0) word \ ('a::len0) word" context includes lifting_syntax notes power_transfer [transfer_rule] begin lemma power_transfer_word [transfer_rule]: \(pcr_word ===> (=) ===> pcr_word) (^) (^)\ by transfer_prover end subsubsection \Conversions\ context includes lifting_syntax notes 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 [transfer_rule]: "((=) ===> (pcr_word :: int \ 'a::len word \ bool)) of_bool of_bool" by transfer_prover lemma [transfer_rule]: "((=) ===> (pcr_word :: int \ 'a::len word \ bool)) 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 end lemma abs_word_eq: "abs_word = of_int" by (rule ext) (transfer, rule) context semiring_1 begin lift_definition unsigned :: "'b::len0 word \ 'a" is "of_nat \ nat \ take_bit LENGTH('b)" by simp lemma unsigned_0 [simp]: "unsigned 0 = 0" by transfer simp end context semiring_char_0 begin lemma word_eq_iff_unsigned: "a = b \ unsigned a = unsigned b" by safe (transfer; simp add: eq_nat_nat_iff) end instantiation word :: (len0) equal begin definition equal_word :: "'a word \ 'a word \ bool" where "equal_word a b \ (unsigned a :: int) = unsigned b" instance proof fix a b :: "'a word" show "HOL.equal a b \ a = b" using word_eq_iff_unsigned [of a b] by (auto simp add: equal_word_def) qed end context ring_1 begin lift_definition signed :: "'b::len word \ 'a" is "of_int \ signed_take_bit (LENGTH('b) - 1)" by (simp add: signed_take_bit_eq_iff_take_bit_eq [symmetric]) lemma signed_0 [simp]: "signed 0 = 0" by transfer simp end lemma unsigned_of_nat [simp]: "unsigned (of_nat n :: 'a word) = take_bit LENGTH('a::len) n" by transfer (simp add: nat_eq_iff take_bit_eq_mod zmod_int) lemma of_nat_unsigned [simp]: "of_nat (unsigned a) = a" by transfer simp lemma of_int_unsigned [simp]: "of_int (unsigned a) = a" by transfer simp lemma unsigned_nat_less: \unsigned a < (2 ^ LENGTH('a) :: nat)\ for a :: \'a::len0 word\ by transfer (simp add: take_bit_eq_mod) lemma unsigned_int_less: \unsigned a < (2 ^ LENGTH('a) :: int)\ for a :: \'a::len0 word\ by transfer (simp add: take_bit_eq_mod) context ring_char_0 begin lemma word_eq_iff_signed: "a = b \ signed a = signed b" by safe (transfer; auto simp add: signed_take_bit_eq_iff_take_bit_eq) end lemma signed_of_int [simp]: "signed (of_int k :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) k" by transfer simp lemma of_int_signed [simp]: "of_int (signed a) = a" by transfer (simp add: signed_take_bit_eq_take_bit take_bit_eq_mod mod_simps) subsubsection \Properties\ lemma exp_eq_zero_iff: \(2 :: 'a::len word) ^ n = 0 \ LENGTH('a) \ n\ by transfer simp subsubsection \Division\ instantiation word :: (len0) 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 .. end lemma zero_word_div_eq [simp]: \0 div a = 0\ for a :: \'a::len0 word\ by transfer simp lemma div_zero_word_eq [simp]: \a div 0 = 0\ for a :: \'a::len0 word\ by transfer simp context includes lifting_syntax begin 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 instance word :: (len) semiring_modulo 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 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) show "\ 2 dvd a \ a mod 2 = 1" for a :: "'a word" by transfer (simp_all add: mod_2_eq_odd) qed subsubsection \Orderings\ instantiation word :: (len0) 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 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 lemma word_greater_zero_iff: \a > 0 \ a \ 0\ for a :: \'a::len0 word\ by transfer (simp add: less_le) lemma of_nat_word_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_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_nat_word_eq_0_iff: \of_nat n = (0 :: 'a::len word) \ 2 ^ LENGTH('a) dvd n\ using of_nat_word_eq_iff [where ?'a = 'a, of n 0] by (simp add: take_bit_eq_0_iff) lemma of_int_word_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_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 lemma of_int_word_eq_0_iff: \of_int k = (0 :: 'a::len word) \ 2 ^ LENGTH('a) dvd k\ using of_int_word_eq_iff [where ?'a = 'a, of k 0] by (simp add: take_bit_eq_0_iff) subsection \Bit structure on \<^typ>\'a word\\ 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) - 1) \ P (2 * a)\ and word_odd: \\a. P a \ a < 2 ^ (LENGTH('a) - 1) \ P (1 + 2 * a)\ for P and a :: \'a::len word\ proof - define m :: nat where \m = LENGTH('a) - 1\ then have l: \LENGTH('a) = Suc m\ by simp define n :: nat where \n = unsigned a\ then have \n < 2 ^ LENGTH('a)\ by (simp add: unsigned_nat_less) 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 of_nat_word_eq_0_iff l) moreover from \n < 2 ^ m\ have \(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - 1)\ using of_nat_word_less_iff [where ?'a = 'a, of n \2 ^ m\] by (cases \m = 0\) (simp_all add: not_less take_bit_eq_self ac_simps l) 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) - 1)\ using of_nat_word_less_iff [where ?'a = 'a, of n \2 ^ m\] by (cases \m = 0\) (simp_all add: not_less take_bit_eq_self ac_simps l) ultimately have \P (1 + 2 * of_nat n)\ by (rule word_odd) then show ?case by simp qed then 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 (auto; transfer) simp_all 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 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) instance word :: (len) semiring_bits 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 from 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 \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) 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 context includes lifting_syntax begin lemma transfer_rule_bit_word [transfer_rule]: \((pcr_word :: int \ 'a::len word \ bool) ===> (=)) (\k n. n < LENGTH('a) \ bit k n) bit\ proof - let ?t = \\a n. odd (take_bit LENGTH('a) a div take_bit LENGTH('a) ((2::int) ^ n))\ have \((pcr_word :: int \ 'a word \ bool) ===> (=)) ?t bit\ by (unfold bit_def) transfer_prover also have \?t = (\k n. n < LENGTH('a) \ bit k n)\ by (simp add: fun_eq_iff bit_take_bit_iff flip: bit_def) finally show ?thesis . qed end 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) 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) qed end 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 instance proof fix a b :: \'a word\ and n :: nat show \- a = NOT (a - 1)\ by transfer (simp add: minus_eq_not_minus_1) show \bit (NOT a) n \ (2 :: 'a word) ^ n \ 0 \ \ bit a n\ by transfer (simp add: bit_not_iff) show \bit (a AND b) n \ bit a n \ bit b n\ by transfer (auto simp add: bit_and_iff) show \bit (a OR b) n \ bit a n \ bit b n\ by transfer (auto simp add: bit_or_iff) show \bit (a XOR b) n \ bit a n \ bit b n\ by transfer (auto simp add: bit_xor_iff) qed end +lifting_update word.lifting +lifting_forget word.lifting + end