diff --git a/src/HOL/ex/Bit_Operations.thy b/src/HOL/ex/Bit_Operations.thy --- a/src/HOL/ex/Bit_Operations.thy +++ b/src/HOL/ex/Bit_Operations.thy @@ -1,778 +1,801 @@ (* Author: Florian Haftmann, TUM *) section \Proof of concept for purely algebraically founded lists of bits\ theory Bit_Operations imports "HOL-Library.Boolean_Algebra" Main begin +lemma minus_1_div_exp_eq_int [simp]: + \- 1 div (2 :: int) ^ n = - 1\ + for n :: nat + by (induction n) (use div_exp_eq [symmetric, of \- 1 :: int\ 1] in \simp_all add: ac_simps\) + +context semiring_bits +begin + +lemma bits_div_by_0 [simp]: + \a div 0 = 0\ + by (metis local.add_cancel_right_right local.bit_mod_div_trivial local.mod_mult_div_eq local.mult_not_zero) + +lemma bit_0_eq [simp]: + \bit 0 = bot\ + by (simp add: fun_eq_iff bit_def) + +end + +context semiring_bit_shifts +begin + +end + + subsection \Bit operations in suitable algebraic structures\ class semiring_bit_operations = semiring_bit_shifts + - fixes "and" :: "'a \ 'a \ 'a" (infixr "AND" 64) - and or :: "'a \ 'a \ 'a" (infixr "OR" 59) - and xor :: "'a \ 'a \ 'a" (infixr "XOR" 59) + fixes "and" :: \'a \ 'a \ 'a\ (infixr "AND" 64) + and or :: \'a \ 'a \ 'a\ (infixr "OR" 59) + and xor :: \'a \ 'a \ 'a\ (infixr "XOR" 59) + assumes bit_and_iff: \\n. bit (a AND b) n \ bit a n \ bit b n\ + and bit_or_iff: \\n. bit (a OR b) n \ bit a n \ bit b n\ + and bit_xor_iff: \\n. bit (a XOR b) n \ bit a n \ bit b n\ begin text \ We want the bitwise operations to bind slightly weaker than \+\ and \-\. For the sake of code generation the operations \<^const>\and\, \<^const>\or\ and \<^const>\xor\ are specified as definitional class operations. \ definition map_bit :: \nat \ (bool \ bool) \ 'a \ 'a\ where \map_bit n f a = take_bit n a + push_bit n (of_bool (f (bit a n)) + 2 * drop_bit (Suc n) a)\ definition set_bit :: \nat \ 'a \ 'a\ where \set_bit n = map_bit n top\ definition unset_bit :: \nat \ 'a \ 'a\ where \unset_bit n = map_bit n bot\ definition flip_bit :: \nat \ 'a \ 'a\ where \flip_bit n = map_bit n Not\ text \ Having - <^const>\set_bit\, \<^const>\unset_bit\ and \<^const>\flip_bit\ as separate + \<^const>\set_bit\, \<^const>\unset_bit\ and \<^const>\flip_bit\ as separate operations allows to implement them using bit masks later. \ lemma stable_imp_drop_eq: \drop_bit n a = a\ if \a div 2 = a\ by (induction n) (simp_all add: that) lemma map_bit_0 [simp]: \map_bit 0 f a = of_bool (f (odd a)) + 2 * (a div 2)\ by (simp add: map_bit_def) lemma map_bit_Suc [simp]: \map_bit (Suc n) f a = a mod 2 + 2 * map_bit n f (a div 2)\ by (auto simp add: map_bit_def algebra_simps mod2_eq_if push_bit_add mult_2 elim: evenE oddE) lemma set_bit_0 [simp]: \set_bit 0 a = 1 + 2 * (a div 2)\ by (simp add: set_bit_def) lemma set_bit_Suc [simp]: \set_bit (Suc n) a = a mod 2 + 2 * set_bit n (a div 2)\ by (simp add: set_bit_def) lemma unset_bit_0 [simp]: \unset_bit 0 a = 2 * (a div 2)\ by (simp add: unset_bit_def) lemma unset_bit_Suc [simp]: \unset_bit (Suc n) a = a mod 2 + 2 * unset_bit n (a div 2)\ by (simp add: unset_bit_def) lemma flip_bit_0 [simp]: \flip_bit 0 a = of_bool (even a) + 2 * (a div 2)\ by (simp add: flip_bit_def) lemma flip_bit_Suc [simp]: \flip_bit (Suc n) a = a mod 2 + 2 * flip_bit n (a div 2)\ by (simp add: flip_bit_def) end class ring_bit_operations = semiring_bit_operations + ring_parity + fixes not :: \'a \ 'a\ (\NOT\) - assumes boolean_algebra: \boolean_algebra (AND) (OR) NOT 0 (- 1)\ - and boolean_algebra_xor_eq: \boolean_algebra.xor (AND) (OR) NOT = (XOR)\ + assumes bits_even_minus_1_div_exp_iff [simp]: \even (- 1 div 2 ^ n) \ 2 ^ n = 0\ + assumes bit_not_iff: \\n. bit (NOT a) n \ 2 ^ n \ 0 \ \ bit a n\ begin +lemma bits_minus_1_mod_2_eq [simp]: + \(- 1) mod 2 = 1\ + by (simp add: mod_2_eq_odd) + +lemma bit_minus_1_iff [simp]: + \bit (- 1) n \ 2 ^ n \ 0\ + by (simp add: bit_def) + sublocale bit: boolean_algebra \(AND)\ \(OR)\ NOT 0 \- 1\ rewrites \bit.xor = (XOR)\ proof - interpret bit: boolean_algebra \(AND)\ \(OR)\ NOT 0 \- 1\ - by (fact boolean_algebra) + apply standard + apply (auto simp add: bit_eq_iff bit_and_iff bit_or_iff bit_not_iff) + apply (metis local.bit_def local.bit_exp_iff local.bits_div_by_0) + apply (metis local.bit_def local.bit_exp_iff local.bits_div_by_0) + done show \boolean_algebra (AND) (OR) NOT 0 (- 1)\ by standard - show \boolean_algebra.xor (AND) (OR) NOT = (XOR)\ - by (fact boolean_algebra_xor_eq) + show \boolean_algebra.xor (AND) (OR) NOT = (XOR)\ + apply (auto simp add: fun_eq_iff bit.xor_def bit_eq_iff bit_and_iff bit_or_iff bit_not_iff bit_xor_iff) + apply (metis local.bit_def local.bit_exp_iff local.bits_div_by_0) + apply (metis local.bit_def local.bit_exp_iff local.bits_div_by_0) + apply (metis local.bit_def local.bit_exp_iff local.bits_div_by_0) + apply (metis local.bit_def local.bit_exp_iff local.bits_div_by_0) + apply (metis local.bit_def local.bit_exp_iff local.bits_div_by_0) + apply (metis local.bit_def local.bit_exp_iff local.bits_div_by_0) + done qed text \ For the sake of code generation \<^const>\not\ is specified as definitional class operation. Note that \<^const>\not\ has no sensible definition for unlimited but only positive bit strings (type \<^typ>\nat\). \ end subsubsection \Instance \<^typ>\nat\\ locale zip_nat = single: abel_semigroup f for f :: "bool \ bool \ bool" (infixl "\<^bold>*" 70) + assumes end_of_bits: "\ False \<^bold>* False" begin lemma False_P_imp: "False \<^bold>* True \ P" if "False \<^bold>* P" using that end_of_bits by (cases P) simp_all function F :: "nat \ nat \ nat" (infixl "\<^bold>\" 70) where "m \<^bold>\ n = (if m = 0 \ n = 0 then 0 else of_bool (odd m \<^bold>* odd n) + (m div 2) \<^bold>\ (n div 2) * 2)" by auto termination by (relation "measure (case_prod (+))") auto lemma zero_left_eq: "0 \<^bold>\ n = of_bool (False \<^bold>* True) * n" by (induction n rule: nat_bit_induct) (simp_all add: end_of_bits) lemma zero_right_eq: "m \<^bold>\ 0 = of_bool (True \<^bold>* False) * m" by (induction m rule: nat_bit_induct) (simp_all add: end_of_bits) lemma simps [simp]: "0 \<^bold>\ 0 = 0" "0 \<^bold>\ n = of_bool (False \<^bold>* True) * n" "m \<^bold>\ 0 = of_bool (True \<^bold>* False) * m" "m > 0 \ n > 0 \ m \<^bold>\ n = of_bool (odd m \<^bold>* odd n) + (m div 2) \<^bold>\ (n div 2) * 2" by (simp_all only: zero_left_eq zero_right_eq) simp lemma rec: "m \<^bold>\ n = of_bool (odd m \<^bold>* odd n) + (m div 2) \<^bold>\ (n div 2) * 2" by (cases "m = 0 \ n = 0") (auto simp add: end_of_bits) declare F.simps [simp del] sublocale abel_semigroup F proof show "m \<^bold>\ n \<^bold>\ q = m \<^bold>\ (n \<^bold>\ q)" for m n q :: nat proof (induction m arbitrary: n q rule: nat_bit_induct) case zero show ?case by simp next case (even m) with rec [of "2 * m"] rec [of _ q] show ?case by (cases "even n") (auto simp add: ac_simps dest: False_P_imp) next case (odd m) with rec [of "Suc (2 * m)"] rec [of _ q] show ?case by (cases "even n"; cases "even q") (auto dest: False_P_imp simp add: ac_simps) qed show "m \<^bold>\ n = n \<^bold>\ m" for m n :: nat proof (induction m arbitrary: n rule: nat_bit_induct) case zero show ?case by (simp add: ac_simps) next case (even m) with rec [of "2 * m" n] rec [of n "2 * m"] show ?case by (simp add: ac_simps) next case (odd m) with rec [of "Suc (2 * m)" n] rec [of n "Suc (2 * m)"] show ?case by (simp add: ac_simps) qed qed lemma self [simp]: "n \<^bold>\ n = of_bool (True \<^bold>* True) * n" by (induction n rule: nat_bit_induct) (simp_all add: end_of_bits) lemma even_iff [simp]: "even (m \<^bold>\ n) \ \ (odd m \<^bold>* odd n)" proof (induction m arbitrary: n rule: nat_bit_induct) case zero show ?case by (cases "even n") (simp_all add: end_of_bits) next case (even m) then show ?case by (simp add: rec [of "2 * m"]) next case (odd m) then show ?case by (simp add: rec [of "Suc (2 * m)"]) qed end instantiation nat :: semiring_bit_operations begin global_interpretation and_nat: zip_nat "(\)" defines and_nat = and_nat.F by standard auto global_interpretation and_nat: semilattice "(AND) :: nat \ nat \ nat" proof (rule semilattice.intro, fact and_nat.abel_semigroup_axioms, standard) show "n AND n = n" for n :: nat by (simp add: and_nat.self) qed declare and_nat.simps [simp] \ \inconsistent declaration handling by \global_interpretation\ in \instantiation\\ lemma zero_nat_and_eq [simp]: "0 AND n = 0" for n :: nat by simp lemma and_zero_nat_eq [simp]: "n AND 0 = 0" for n :: nat by simp global_interpretation or_nat: zip_nat "(\)" defines or_nat = or_nat.F by standard auto global_interpretation or_nat: semilattice "(OR) :: nat \ nat \ nat" proof (rule semilattice.intro, fact or_nat.abel_semigroup_axioms, standard) show "n OR n = n" for n :: nat by (simp add: or_nat.self) qed declare or_nat.simps [simp] \ \inconsistent declaration handling by \global_interpretation\ in \instantiation\\ lemma zero_nat_or_eq [simp]: "0 OR n = n" for n :: nat by simp lemma or_zero_nat_eq [simp]: "n OR 0 = n" for n :: nat by simp global_interpretation xor_nat: zip_nat "(\)" defines xor_nat = xor_nat.F by standard auto declare xor_nat.simps [simp] \ \inconsistent declaration handling by \global_interpretation\ in \instantiation\\ lemma zero_nat_xor_eq [simp]: "0 XOR n = n" for n :: nat by simp lemma xor_zero_nat_eq [simp]: "n XOR 0 = n" for n :: nat by simp -instance .. +instance proof + fix m n q :: nat + show \bit (m AND n) q \ bit m q \ bit n q\ + proof (rule sym, induction q arbitrary: m n) + case 0 + then show ?case + by (simp add: and_nat.even_iff) + next + case (Suc q) + with and_nat.rec [of m n] show ?case + by simp + qed + show \bit (m OR n) q \ bit m q \ bit n q\ + proof (rule sym, induction q arbitrary: m n) + case 0 + then show ?case + by (simp add: or_nat.even_iff) + next + case (Suc q) + with or_nat.rec [of m n] show ?case + by simp + qed + show \bit (m XOR n) q \ bit m q \ bit n q\ + proof (rule sym, induction q arbitrary: m n) + case 0 + then show ?case + by (simp add: xor_nat.even_iff) + next + case (Suc q) + with xor_nat.rec [of m n] show ?case + by simp + qed +qed end global_interpretation or_nat: semilattice_neutr "(OR)" "0 :: nat" by standard simp global_interpretation xor_nat: comm_monoid "(XOR)" "0 :: nat" by standard simp lemma Suc_0_and_eq [simp]: "Suc 0 AND n = n mod 2" by (cases n) auto lemma and_Suc_0_eq [simp]: "n AND Suc 0 = n mod 2" using Suc_0_and_eq [of n] by (simp add: ac_simps) lemma Suc_0_or_eq [simp]: "Suc 0 OR n = n + of_bool (even n)" by (cases n) (simp_all add: ac_simps) lemma or_Suc_0_eq [simp]: "n OR Suc 0 = n + of_bool (even n)" using Suc_0_or_eq [of n] by (simp add: ac_simps) lemma Suc_0_xor_eq [simp]: "Suc 0 XOR n = n + of_bool (even n) - of_bool (odd n)" by (cases n) (simp_all add: ac_simps) lemma xor_Suc_0_eq [simp]: "n XOR Suc 0 = n + of_bool (even n) - of_bool (odd n)" using Suc_0_xor_eq [of n] by (simp add: ac_simps) subsubsection \Instance \<^typ>\int\\ abbreviation (input) complement :: "int \ int" where "complement k \ - k - 1" lemma complement_half: "complement (k * 2) div 2 = complement k" by simp lemma complement_div_2: "complement (k div 2) = complement k div 2" by linarith locale zip_int = single: abel_semigroup f for f :: "bool \ bool \ bool" (infixl "\<^bold>*" 70) begin lemma False_False_imp_True_True: "True \<^bold>* True" if "False \<^bold>* False" proof (rule ccontr) assume "\ True \<^bold>* True" with that show False using single.assoc [of False True True] by (cases "False \<^bold>* True") simp_all qed function F :: "int \ int \ int" (infixl "\<^bold>\" 70) where "k \<^bold>\ l = (if k \ {0, - 1} \ l \ {0, - 1} then - of_bool (odd k \<^bold>* odd l) else of_bool (odd k \<^bold>* odd l) + (k div 2) \<^bold>\ (l div 2) * 2)" by auto termination by (relation "measure (\(k, l). nat (\k\ + \l\))") auto lemma zero_left_eq: "0 \<^bold>\ l = (case (False \<^bold>* False, False \<^bold>* True) of (False, False) \ 0 | (False, True) \ l | (True, False) \ complement l | (True, True) \ - 1)" by (induction l rule: int_bit_induct) (simp_all split: bool.split) lemma minus_left_eq: "- 1 \<^bold>\ l = (case (True \<^bold>* False, True \<^bold>* True) of (False, False) \ 0 | (False, True) \ l | (True, False) \ complement l | (True, True) \ - 1)" by (induction l rule: int_bit_induct) (simp_all split: bool.split) lemma zero_right_eq: "k \<^bold>\ 0 = (case (False \<^bold>* False, False \<^bold>* True) of (False, False) \ 0 | (False, True) \ k | (True, False) \ complement k | (True, True) \ - 1)" by (induction k rule: int_bit_induct) (simp_all add: ac_simps split: bool.split) lemma minus_right_eq: "k \<^bold>\ - 1 = (case (True \<^bold>* False, True \<^bold>* True) of (False, False) \ 0 | (False, True) \ k | (True, False) \ complement k | (True, True) \ - 1)" by (induction k rule: int_bit_induct) (simp_all add: ac_simps split: bool.split) lemma simps [simp]: "0 \<^bold>\ 0 = - of_bool (False \<^bold>* False)" "- 1 \<^bold>\ 0 = - of_bool (True \<^bold>* False)" "0 \<^bold>\ - 1 = - of_bool (False \<^bold>* True)" "- 1 \<^bold>\ - 1 = - of_bool (True \<^bold>* True)" "0 \<^bold>\ l = (case (False \<^bold>* False, False \<^bold>* True) of (False, False) \ 0 | (False, True) \ l | (True, False) \ complement l | (True, True) \ - 1)" "- 1 \<^bold>\ l = (case (True \<^bold>* False, True \<^bold>* True) of (False, False) \ 0 | (False, True) \ l | (True, False) \ complement l | (True, True) \ - 1)" "k \<^bold>\ 0 = (case (False \<^bold>* False, False \<^bold>* True) of (False, False) \ 0 | (False, True) \ k | (True, False) \ complement k | (True, True) \ - 1)" "k \<^bold>\ - 1 = (case (True \<^bold>* False, True \<^bold>* True) of (False, False) \ 0 | (False, True) \ k | (True, False) \ complement k | (True, True) \ - 1)" "k \ 0 \ k \ - 1 \ l \ 0 \ l \ - 1 \ k \<^bold>\ l = of_bool (odd k \<^bold>* odd l) + (k div 2) \<^bold>\ (l div 2) * 2" by simp_all[4] (simp_all only: zero_left_eq minus_left_eq zero_right_eq minus_right_eq, simp) declare F.simps [simp del] lemma rec: "k \<^bold>\ l = of_bool (odd k \<^bold>* odd l) + (k div 2) \<^bold>\ (l div 2) * 2" by (cases "k \ {0, - 1} \ l \ {0, - 1}") (auto simp add: ac_simps F.simps [of k l] split: bool.split) sublocale abel_semigroup F proof show "k \<^bold>\ l \<^bold>\ r = k \<^bold>\ (l \<^bold>\ r)" for k l r :: int proof (induction k arbitrary: l r rule: int_bit_induct) case zero have "complement l \<^bold>\ r = complement (l \<^bold>\ r)" if "False \<^bold>* False" "\ False \<^bold>* True" proof (induction l arbitrary: r rule: int_bit_induct) case zero from that show ?case by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits) next case minus from that show ?case by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits) next case (even l) with that rec [of _ r] show ?case by (cases "even r") (auto simp add: complement_half ac_simps False_False_imp_True_True split: bool.splits) next case (odd l) moreover have "- l - 1 = - 1 - l" by simp ultimately show ?case using that rec [of _ r] by (cases "even r") (auto simp add: ac_simps False_False_imp_True_True split: bool.splits) qed then show ?case by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits) next case minus have "complement l \<^bold>\ r = complement (l \<^bold>\ r)" if "\ True \<^bold>* True" "False \<^bold>* True" proof (induction l arbitrary: r rule: int_bit_induct) case zero from that show ?case by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits) next case minus from that show ?case by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits) next case (even l) with that rec [of _ r] show ?case by (cases "even r") (auto simp add: complement_half ac_simps False_False_imp_True_True split: bool.splits) next case (odd l) moreover have "- l - 1 = - 1 - l" by simp ultimately show ?case using that rec [of _ r] by (cases "even r") (auto simp add: ac_simps False_False_imp_True_True split: bool.splits) qed then show ?case by (auto simp add: ac_simps False_False_imp_True_True split: bool.splits) next case (even k) with rec [of "k * 2"] rec [of _ r] show ?case by (cases "even r"; cases "even l") (auto simp add: ac_simps False_False_imp_True_True) next case (odd k) with rec [of "1 + k * 2"] rec [of _ r] show ?case by (cases "even r"; cases "even l") (auto simp add: ac_simps False_False_imp_True_True) qed show "k \<^bold>\ l = l \<^bold>\ k" for k l :: int proof (induction k arbitrary: l rule: int_bit_induct) case zero show ?case by simp next case minus show ?case by simp next case (even k) with rec [of "k * 2" l] rec [of l "k * 2"] show ?case by (simp add: ac_simps) next case (odd k) with rec [of "k * 2 + 1" l] rec [of l "k * 2 + 1"] show ?case by (simp add: ac_simps) qed qed lemma self [simp]: "k \<^bold>\ k = (case (False \<^bold>* False, True \<^bold>* True) of (False, False) \ 0 | (False, True) \ k | (True, True) \ - 1)" by (induction k rule: int_bit_induct) (auto simp add: False_False_imp_True_True split: bool.split) lemma even_iff [simp]: "even (k \<^bold>\ l) \ \ (odd k \<^bold>* odd l)" proof (induction k arbitrary: l rule: int_bit_induct) case zero show ?case by (cases "even l") (simp_all split: bool.splits) next case minus show ?case by (cases "even l") (simp_all split: bool.splits) next case (even k) then show ?case by (simp add: rec [of "k * 2"]) next case (odd k) then show ?case by (simp add: rec [of "1 + k * 2"]) qed end instantiation int :: ring_bit_operations begin definition not_int :: "int \ int" where "not_int = complement" global_interpretation and_int: zip_int "(\)" defines and_int = and_int.F by standard declare and_int.simps [simp] \ \inconsistent declaration handling by \global_interpretation\ in \instantiation\\ global_interpretation and_int: semilattice "(AND) :: int \ int \ int" proof (rule semilattice.intro, fact and_int.abel_semigroup_axioms, standard) show "k AND k = k" for k :: int by (simp add: and_int.self) qed lemma zero_int_and_eq [simp]: "0 AND k = 0" for k :: int by simp lemma and_zero_int_eq [simp]: "k AND 0 = 0" for k :: int by simp lemma minus_int_and_eq [simp]: "- 1 AND k = k" for k :: int by simp lemma and_minus_int_eq [simp]: "k AND - 1 = k" for k :: int by simp global_interpretation or_int: zip_int "(\)" defines or_int = or_int.F by standard declare or_int.simps [simp] \ \inconsistent declaration handling by \global_interpretation\ in \instantiation\\ global_interpretation or_int: semilattice "(OR) :: int \ int \ int" proof (rule semilattice.intro, fact or_int.abel_semigroup_axioms, standard) show "k OR k = k" for k :: int by (simp add: or_int.self) qed lemma zero_int_or_eq [simp]: "0 OR k = k" for k :: int by simp lemma and_zero_or_eq [simp]: "k OR 0 = k" for k :: int by simp lemma minus_int_or_eq [simp]: "- 1 OR k = - 1" for k :: int by simp lemma or_minus_int_eq [simp]: "k OR - 1 = - 1" for k :: int by simp global_interpretation xor_int: zip_int "(\)" defines xor_int = xor_int.F by standard declare xor_int.simps [simp] \ \inconsistent declaration handling by \global_interpretation\ in \instantiation\\ lemma zero_int_xor_eq [simp]: "0 XOR k = k" for k :: int by simp lemma and_zero_xor_eq [simp]: "k XOR 0 = k" for k :: int by simp lemma minus_int_xor_eq [simp]: "- 1 XOR k = complement k" for k :: int by simp lemma xor_minus_int_eq [simp]: "k XOR - 1 = complement k" for k :: int by simp lemma not_div_2: "NOT k div 2 = NOT (k div 2)" for k :: int by (simp add: complement_div_2 not_int_def) lemma not_int_simps [simp]: "NOT 0 = (- 1 :: int)" "NOT (- 1) = (0 :: int)" "k \ 0 \ k \ - 1 \ NOT k = of_bool (even k) + 2 * NOT (k div 2)" for k :: int by (auto simp add: not_int_def elim: oddE) lemma not_one_int [simp]: "NOT 1 = (- 2 :: int)" by simp lemma even_not_iff [simp]: "even (NOT k) \ odd k" - for k :: int + for k :: int by (simp add: not_int_def) +lemma bit_not_iff_int: + \bit (NOT k) n \ \ bit k n\ + for k :: int + by (induction n arbitrary: k) + (simp_all add: not_int_def flip: complement_div_2) + + instance proof - interpret bit_int: boolean_algebra "(AND)" "(OR)" NOT 0 "- 1 :: int" - proof - show "k AND (l OR r) = k AND l OR k AND r" - for k l r :: int - proof (induction k arbitrary: l r rule: int_bit_induct) - case zero - show ?case - by simp - next - case minus - show ?case - by simp - next - case (even k) - then show ?case by (simp add: and_int.rec [of "k * 2"] - or_int.rec [of "(k AND l div 2) * 2"] or_int.rec [of l]) - next - case (odd k) - then show ?case by (simp add: and_int.rec [of "1 + k * 2"] - or_int.rec [of "(k AND l div 2) * 2"] or_int.rec [of "1 + (k AND l div 2) * 2"] or_int.rec [of l]) - qed - show "k OR l AND r = (k OR l) AND (k OR r)" - for k l r :: int - proof (induction k arbitrary: l r rule: int_bit_induct) - case zero - then show ?case - by simp - next - case minus - then show ?case - by simp - next - case (even k) - then show ?case by (simp add: or_int.rec [of "k * 2"] - and_int.rec [of "(k OR l div 2) * 2"] and_int.rec [of "1 + (k OR l div 2) * 2"] and_int.rec [of l]) - next - case (odd k) - then show ?case by (simp add: or_int.rec [of "1 + k * 2"] - and_int.rec [of "1 + (k OR l div 2) * 2"] and_int.rec [of l]) - qed - show "k AND NOT k = 0" for k :: int - by (induction k rule: int_bit_induct) - (simp_all add: not_int_def complement_half minus_diff_commute [of 1]) - show "k OR NOT k = - 1" for k :: int - by (induction k rule: int_bit_induct) - (simp_all add: not_int_def complement_half minus_diff_commute [of 1]) - qed (simp_all add: and_int.assoc or_int.assoc, - simp_all add: and_int.commute or_int.commute) - show "boolean_algebra (AND) (OR) NOT 0 (- 1 :: int)" - by (fact bit_int.boolean_algebra_axioms) - show "bit_int.xor = ((XOR) :: int \ _)" - proof (rule ext)+ - fix k l :: int - have "k XOR l = k AND NOT l OR NOT k AND l" - proof (induction k arbitrary: l rule: int_bit_induct) - case zero - show ?case - by simp - next - case minus - show ?case - by (simp add: not_int_def) - next - case (even k) - then show ?case - by (simp add: xor_int.rec [of "k * 2"] and_int.rec [of "k * 2"] - or_int.rec [of _ "1 + 2 * NOT k AND l"] not_div_2) - (simp add: and_int.rec [of _ l]) - next - case (odd k) - then show ?case - by (simp add: xor_int.rec [of "1 + k * 2"] and_int.rec [of "1 + k * 2"] - or_int.rec [of _ "2 * NOT k AND l"] not_div_2) - (simp add: and_int.rec [of _ l]) - qed - then show "bit_int.xor k l = k XOR l" - by (simp add: bit_int.xor_def) + fix k l :: int and n :: nat + show \bit (k AND l) n \ bit k n \ bit l n\ + proof (rule sym, induction n arbitrary: k l) + case 0 + then show ?case + by (simp add: and_int.even_iff) + next + case (Suc n) + with and_int.rec [of k l] show ?case + by simp qed -qed + show \bit (k OR l) n \ bit k n \ bit l n\ + proof (rule sym, induction n arbitrary: k l) + case 0 + then show ?case + by (simp add: or_int.even_iff) + next + case (Suc n) + with or_int.rec [of k l] show ?case + by simp + qed + show \bit (k XOR l) n \ bit k n \ bit l n\ + proof (rule sym, induction n arbitrary: k l) + case 0 + then show ?case + by (simp add: xor_int.even_iff) + next + case (Suc n) + with xor_int.rec [of k l] show ?case + by simp + qed +qed (simp_all add: bit_not_iff_int) end lemma one_and_int_eq [simp]: "1 AND k = k mod 2" for k :: int using and_int.rec [of 1] by (simp add: mod2_eq_if) lemma and_one_int_eq [simp]: "k AND 1 = k mod 2" for k :: int using one_and_int_eq [of 1] by (simp add: ac_simps) lemma one_or_int_eq [simp]: "1 OR k = k + of_bool (even k)" for k :: int using or_int.rec [of 1] by (auto elim: oddE) lemma or_one_int_eq [simp]: "k OR 1 = k + of_bool (even k)" for k :: int using one_or_int_eq [of k] by (simp add: ac_simps) lemma one_xor_int_eq [simp]: "1 XOR k = k + of_bool (even k) - of_bool (odd k)" for k :: int using xor_int.rec [of 1] by (auto elim: oddE) lemma xor_one_int_eq [simp]: "k XOR 1 = k + of_bool (even k) - of_bool (odd k)" for k :: int using one_xor_int_eq [of k] by (simp add: ac_simps) lemma take_bit_complement_iff: "Parity.take_bit n (complement k) = Parity.take_bit n (complement l) \ Parity.take_bit n k = Parity.take_bit n l" for k l :: int by (simp add: Parity.take_bit_eq_mod mod_eq_dvd_iff dvd_diff_commute) lemma take_bit_not_iff: "Parity.take_bit n (NOT k) = Parity.take_bit n (NOT l) \ Parity.take_bit n k = Parity.take_bit n l" for k l :: int - by (simp add: not_int_def take_bit_complement_iff) + by (auto simp add: bit_eq_iff bit_take_bit_iff bit_not_iff_int) lemma take_bit_and [simp]: "Parity.take_bit n (k AND l) = Parity.take_bit n k AND Parity.take_bit n l" for k l :: int - apply (induction n arbitrary: k l) - apply simp - apply (subst and_int.rec) - apply (subst (2) and_int.rec) - apply simp - done + by (auto simp add: bit_eq_iff bit_take_bit_iff bit_and_iff) lemma take_bit_or [simp]: "Parity.take_bit n (k OR l) = Parity.take_bit n k OR Parity.take_bit n l" for k l :: int - apply (induction n arbitrary: k l) - apply simp - apply (subst or_int.rec) - apply (subst (2) or_int.rec) - apply simp - done + by (auto simp add: bit_eq_iff bit_take_bit_iff bit_or_iff) lemma take_bit_xor [simp]: "Parity.take_bit n (k XOR l) = Parity.take_bit n k XOR Parity.take_bit n l" for k l :: int - apply (induction n arbitrary: k l) - apply simp - apply (subst xor_int.rec) - apply (subst (2) xor_int.rec) - apply simp - done + by (auto simp add: bit_eq_iff bit_take_bit_iff bit_xor_iff) 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,754 +1,769 @@ (* 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 +context + includes lifting_syntax +begin + +lemma transfer_rule_of_bool: + \((\) ===> (\)) of_bool of_bool\ + if [transfer_rule]: \0 \ 0\ \1 \ 1\ + for R :: \'a::zero_neq_one \ 'b::zero_neq_one \ bool\ (infix \\\ 50) + by (unfold of_bool_def [abs_def]) transfer_prover + +end + + subsection \Preliminaries\ lemma length_not_greater_eq_2_iff [simp]: \\ 2 \ LENGTH('a::len) \ LENGTH('a) = 1\ by (auto simp add: not_le dest: less_2_cases) lemma take_bit_uminus: "take_bit n (- (take_bit n k)) = take_bit n (- k)" for k :: int by (simp add: take_bit_eq_mod mod_minus_eq) lemma take_bit_minus: "take_bit n (take_bit n k - take_bit n l) = take_bit n (k - l)" for k l :: int by (simp add: take_bit_eq_mod mod_diff_eq) lemma take_bit_nonnegative [simp]: "take_bit n k \ 0" for k :: int by (simp add: take_bit_eq_mod) 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_uminus [symmetric]) (simp add: take_bit_uminus) lift_definition minus_word :: "'a word \ 'a word \ 'a word" is minus by (subst take_bit_minus [symmetric]) (simp add: take_bit_minus) 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 length_cases: \ \TODO get rid of\ obtains (triv) "LENGTH('a::len) = 1" "take_bit LENGTH('a) 2 = (0 :: int)" | (take_bit_2) "take_bit LENGTH('a) 2 = (2 :: int)" proof (cases "LENGTH('a) \ 2") case False then have "LENGTH('a) = 1" by (auto simp add: not_le dest: less_2_cases) then have "take_bit LENGTH('a) 2 = (0 :: int)" by simp with \LENGTH('a) = 1\ triv show ?thesis by simp next case True then obtain n where "LENGTH('a) = Suc (Suc n)" by (auto dest: le_Suc_ex) then have "take_bit LENGTH('a) 2 = (2 :: int)" by simp with take_bit_2 show ?thesis by simp qed 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 +(*lemma + \a div a = of_bool (a \ 0)\ for a :: \'a::len 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; cases rule: length_cases [where ?'a = 'a]) (simp_all add: mod_2_eq_odd) show "\ 2 dvd a \ a mod 2 = 1" for a :: "'a word" by (transfer; cases rule: length_cases [where ?'a = 'a]) (simp_all add: mod_2_eq_odd) qed +(*lemma + \2 ^ n = (0 :: 'a word) \ LENGTH('a::len) \ n\ + apply transfer*) + + 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 rule: length_cases [where ?'a = 'a]) case triv have \of_bool (odd k) < (1 :: int) \ even k\ for k :: int by auto with triv that show ?thesis by (auto; transfer) simp_all next case take_bit_2 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 take_bit_2 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 take_bit_2 show \take_bit LENGTH('a) (take_bit LENGTH('a) (1 + k * 2) div take_bit LENGTH('a) 2) = take_bit LENGTH('a) k\ by simp qed ultimately show ?thesis by simp qed qed 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 apply transfer apply (auto simp flip: take_bit_eq_mod) apply (simp add: ac_simps) done 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 apply transfer apply (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) done qed context includes lifting_syntax begin -lemma transfer_rule_bit_word: +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 \Parity.take_bit LENGTH('a) (push_bit n k) = Parity.take_bit LENGTH('a) (push_bit n l)\ if \Parity.take_bit LENGTH('a) k = Parity.take_bit LENGTH('a) l\ for k l :: int and n :: nat proof - from that have \Parity.take_bit (LENGTH('a) - n) (Parity.take_bit LENGTH('a) k) = Parity.take_bit (LENGTH('a) - n) (Parity.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" proof (cases \n < LENGTH('a)\) case True then show ?thesis by transfer (simp add: take_bit_eq_mod drop_bit_eq_div) next case False then obtain m where n: \n = LENGTH('a) + m\ by (auto simp add: not_less dest: le_Suc_ex) then show ?thesis by transfer (simp add: take_bit_eq_mod drop_bit_eq_div power_add zdiv_zmult2_eq) qed 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 - interpret bit_word: boolean_algebra "(AND)" "(OR)" NOT 0 "- 1 :: 'a word" - proof - show "a AND (b OR c) = a AND b OR a AND c" - for a b c :: "'a word" - by transfer (simp add: bit.conj_disj_distrib) - show "a OR b AND c = (a OR b) AND (a OR c)" - for a b c :: "'a word" - by transfer (simp add: bit.disj_conj_distrib) - qed (transfer; simp add: ac_simps)+ - show "boolean_algebra (AND) (OR) NOT 0 (- 1 :: 'a word)" - by (fact bit_word.boolean_algebra_axioms) - show "bit_word.xor = ((XOR) :: 'a word \ _)" - proof (rule ext)+ - fix a b :: "'a word" - have "a XOR b = a AND NOT b OR NOT a AND b" - by transfer (simp add: bit.xor_def) - then show "bit_word.xor a b = a XOR b" - by (simp add: bit_word.xor_def) - qed + fix a b :: \'a word\ and n :: nat + show \even (- 1 div (2 :: 'a word) ^ n) \ (2 :: 'a word) ^ n = 0\ + by transfer + (simp flip: drop_bit_eq_div add: drop_bit_take_bit, simp add: drop_bit_eq_div) + 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 end