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,609 +1,677 @@ (* 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 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) 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. \ sublocale "and": semilattice \(AND)\ by standard (auto simp add: bit_eq_iff bit_and_iff) sublocale or: semilattice_neutr \(OR)\ 0 by standard (auto simp add: bit_eq_iff bit_or_iff) sublocale xor: comm_monoid \(XOR)\ 0 by standard (auto simp add: bit_eq_iff bit_xor_iff) lemma zero_and_eq [simp]: "0 AND a = 0" by (simp add: bit_eq_iff bit_and_iff) lemma and_zero_eq [simp]: "a AND 0 = 0" by (simp add: bit_eq_iff bit_and_iff) lemma one_and_eq [simp]: "1 AND a = of_bool (odd a)" by (simp add: bit_eq_iff bit_and_iff) (auto simp add: bit_1_iff) lemma and_one_eq [simp]: "a AND 1 = of_bool (odd a)" using one_and_eq [of a] by (simp add: ac_simps) lemma one_or_eq [simp]: "1 OR a = a + of_bool (even a)" by (simp add: bit_eq_iff bit_or_iff add.commute [of _ 1] even_bit_succ_iff) (auto simp add: bit_1_iff) lemma or_one_eq [simp]: "a OR 1 = a + of_bool (even a)" using one_or_eq [of a] by (simp add: ac_simps) lemma one_xor_eq [simp]: "1 XOR a = a + of_bool (even a) - of_bool (odd a)" by (simp add: bit_eq_iff bit_xor_iff add.commute [of _ 1] even_bit_succ_iff) (auto simp add: bit_1_iff odd_bit_iff_bit_pred elim: oddE) lemma xor_one_eq [simp]: "a XOR 1 = a + of_bool (even a) - of_bool (odd a)" using one_xor_eq [of a] by (simp add: ac_simps) lemma take_bit_and [simp]: \take_bit n (a AND b) = take_bit n a AND take_bit n b\ by (auto simp add: bit_eq_iff bit_take_bit_iff bit_and_iff) lemma take_bit_or [simp]: \take_bit n (a OR b) = take_bit n a OR take_bit n b\ by (auto simp add: bit_eq_iff bit_take_bit_iff bit_or_iff) lemma take_bit_xor [simp]: \take_bit n (a XOR b) = take_bit n a XOR take_bit n b\ by (auto simp add: bit_eq_iff bit_take_bit_iff bit_xor_iff) end class ring_bit_operations = semiring_bit_operations + ring_parity + fixes not :: \'a \ 'a\ (\NOT\) assumes bit_not_iff: \\n. bit (NOT a) n \ 2 ^ n \ 0 \ \ bit a n\ assumes minus_eq_not_minus_1: \- a = NOT (a - 1)\ begin 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\). \ lemma bits_minus_1_mod_2_eq [simp]: \(- 1) mod 2 = 1\ by (simp add: mod_2_eq_odd) lemma not_eq_complement: \NOT a = - a - 1\ using minus_eq_not_minus_1 [of \a + 1\] by simp lemma minus_eq_not_plus_1: \- a = NOT a + 1\ using not_eq_complement [of a] by simp lemma bit_minus_iff: \bit (- a) n \ 2 ^ n \ 0 \ \ bit (a - 1) n\ by (simp add: minus_eq_not_minus_1 bit_not_iff) lemma even_not_iff [simp]: "even (NOT a) \ odd a" using bit_not_iff [of a 0] by auto lemma bit_not_exp_iff: \bit (NOT (2 ^ m)) n \ 2 ^ n \ 0 \ n \ m\ by (auto simp add: bit_not_iff bit_exp_iff) lemma bit_minus_1_iff [simp]: \bit (- 1) n \ 2 ^ n \ 0\ by (simp add: bit_minus_iff) lemma bit_minus_exp_iff: \bit (- (2 ^ m)) n \ 2 ^ n \ 0 \ n \ m\ oops lemma bit_minus_2_iff [simp]: \bit (- 2) n \ 2 ^ n \ 0 \ n > 0\ by (simp add: bit_minus_iff bit_1_iff) lemma not_one [simp]: "NOT 1 = - 2" by (simp add: bit_eq_iff bit_not_iff) (simp add: bit_1_iff) sublocale "and": semilattice_neutr \(AND)\ \- 1\ apply standard apply (simp add: bit_eq_iff bit_and_iff) apply (auto simp add: exp_eq_0_imp_not_bit bit_exp_iff) done sublocale bit: boolean_algebra \(AND)\ \(OR)\ NOT 0 \- 1\ rewrites \bit.xor = (XOR)\ proof - interpret bit: boolean_algebra \(AND)\ \(OR)\ NOT 0 \- 1\ apply standard apply (simp_all add: bit_eq_iff bit_and_iff bit_or_iff bit_not_iff) apply (auto simp add: exp_eq_0_imp_not_bit bit_exp_iff) done show \boolean_algebra (AND) (OR) NOT 0 (- 1)\ by standard 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 (simp_all add: bit_exp_iff, simp_all add: bit_def) apply (metis local.bit_exp_iff local.bits_div_by_0) apply (metis local.bit_exp_iff local.bits_div_by_0) done qed lemma push_bit_minus: \push_bit n (- a) = - push_bit n a\ by (simp add: push_bit_eq_mult) lemma take_bit_not_take_bit: \take_bit n (NOT (take_bit n a)) = take_bit n (NOT a)\ by (auto simp add: bit_eq_iff bit_take_bit_iff bit_not_iff) lemma take_bit_not_iff: "take_bit n (NOT a) = take_bit n (NOT b) \ take_bit n a = take_bit n b" apply (simp add: bit_eq_iff bit_not_iff bit_take_bit_iff) apply (simp add: bit_exp_iff) apply (use local.exp_eq_0_imp_not_bit in blast) done definition set_bit :: \nat \ 'a \ 'a\ where \set_bit n a = a OR 2 ^ n\ definition unset_bit :: \nat \ 'a \ 'a\ where \unset_bit n a = a AND NOT (2 ^ n)\ definition flip_bit :: \nat \ 'a \ 'a\ where \flip_bit n a = a XOR 2 ^ n\ lemma bit_set_bit_iff: \bit (set_bit m a) n \ bit a n \ (m = n \ 2 ^ n \ 0)\ by (auto simp add: set_bit_def bit_or_iff bit_exp_iff) lemma even_set_bit_iff: \even (set_bit m a) \ even a \ m \ 0\ using bit_set_bit_iff [of m a 0] by auto lemma bit_unset_bit_iff: \bit (unset_bit m a) n \ bit a n \ m \ n\ by (auto simp add: unset_bit_def bit_and_iff bit_not_iff bit_exp_iff exp_eq_0_imp_not_bit) lemma even_unset_bit_iff: \even (unset_bit m a) \ even a \ m = 0\ using bit_unset_bit_iff [of m a 0] by auto lemma bit_flip_bit_iff: \bit (flip_bit m a) n \ (m = n \ \ bit a n) \ 2 ^ n \ 0\ by (auto simp add: flip_bit_def bit_xor_iff bit_exp_iff exp_eq_0_imp_not_bit) lemma even_flip_bit_iff: \even (flip_bit m a) \ \ (even a \ m = 0)\ using bit_flip_bit_iff [of m a 0] by auto lemma set_bit_0 [simp]: \set_bit 0 a = 1 + 2 * (a div 2)\ proof (rule bit_eqI) fix m assume *: \2 ^ m \ 0\ then show \bit (set_bit 0 a) m = bit (1 + 2 * (a div 2)) m\ by (simp add: bit_set_bit_iff bit_double_iff even_bit_succ_iff) (cases m, simp_all add: bit_Suc) qed lemma set_bit_Suc [simp]: \set_bit (Suc n) a = a mod 2 + 2 * set_bit n (a div 2)\ proof (rule bit_eqI) fix m assume *: \2 ^ m \ 0\ show \bit (set_bit (Suc n) a) m \ bit (a mod 2 + 2 * set_bit n (a div 2)) m\ proof (cases m) case 0 then show ?thesis by (simp add: even_set_bit_iff) next case (Suc m) with * have \2 ^ m \ 0\ using mult_2 by auto show ?thesis by (cases a rule: parity_cases) (simp_all add: bit_set_bit_iff bit_double_iff even_bit_succ_iff *, simp_all add: Suc \2 ^ m \ 0\ bit_Suc) qed qed lemma unset_bit_0 [simp]: \unset_bit 0 a = 2 * (a div 2)\ proof (rule bit_eqI) fix m assume *: \2 ^ m \ 0\ then show \bit (unset_bit 0 a) m = bit (2 * (a div 2)) m\ by (simp add: bit_unset_bit_iff bit_double_iff) (cases m, simp_all add: bit_Suc) qed lemma unset_bit_Suc [simp]: \unset_bit (Suc n) a = a mod 2 + 2 * unset_bit n (a div 2)\ proof (rule bit_eqI) fix m assume *: \2 ^ m \ 0\ then show \bit (unset_bit (Suc n) a) m \ bit (a mod 2 + 2 * unset_bit n (a div 2)) m\ proof (cases m) case 0 then show ?thesis by (simp add: even_unset_bit_iff) next case (Suc m) show ?thesis by (cases a rule: parity_cases) (simp_all add: bit_unset_bit_iff bit_double_iff even_bit_succ_iff *, simp_all add: Suc bit_Suc) qed qed lemma flip_bit_0 [simp]: \flip_bit 0 a = of_bool (even a) + 2 * (a div 2)\ proof (rule bit_eqI) fix m assume *: \2 ^ m \ 0\ then show \bit (flip_bit 0 a) m = bit (of_bool (even a) + 2 * (a div 2)) m\ by (simp add: bit_flip_bit_iff bit_double_iff even_bit_succ_iff) (cases m, simp_all add: bit_Suc) qed lemma flip_bit_Suc [simp]: \flip_bit (Suc n) a = a mod 2 + 2 * flip_bit n (a div 2)\ proof (rule bit_eqI) fix m assume *: \2 ^ m \ 0\ show \bit (flip_bit (Suc n) a) m \ bit (a mod 2 + 2 * flip_bit n (a div 2)) m\ proof (cases m) case 0 then show ?thesis by (simp add: even_flip_bit_iff) next case (Suc m) with * have \2 ^ m \ 0\ using mult_2 by auto show ?thesis by (cases a rule: parity_cases) (simp_all add: bit_flip_bit_iff bit_double_iff even_bit_succ_iff, simp_all add: Suc \2 ^ m \ 0\ bit_Suc) qed qed 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 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) + 2 * ((m div 2) \<^bold>\ (n div 2)))\ by auto termination by (relation "measure (case_prod (+))") auto declare F.simps [simp del] lemma rec: "m \<^bold>\ n = of_bool (odd m \<^bold>* odd n) + (m div 2) \<^bold>\ (n div 2) * 2" proof (cases \m = 0 \ n = 0\) case True then have \m \<^bold>\ n = 0\ using True by (simp add: F.simps [of 0 0]) moreover have \(m div 2) \<^bold>\ (n div 2) = m \<^bold>\ n\ using True by simp ultimately show ?thesis using True by (simp add: end_of_bits) next case False then show ?thesis by (auto simp add: ac_simps F.simps [of m n]) qed lemma bit_eq_iff: \bit (m \<^bold>\ n) q \ bit m q \<^bold>* bit n q\ proof (induction q arbitrary: m n) case 0 then show ?case by (simp add: rec [of m n]) next case (Suc n) then show ?case by (simp add: rec [of m n] bit_Suc) qed sublocale abel_semigroup F by standard (simp_all add: Parity.bit_eq_iff bit_eq_iff ac_simps) 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: bit_eq_iff and_nat.bit_eq_iff) qed 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: bit_eq_iff or_nat.bit_eq_iff) qed global_interpretation xor_nat: zip_nat \(\)\ defines xor_nat = xor_nat.F by standard auto instance proof fix m n q :: nat show \bit (m AND n) q \ bit m q \ bit n q\ by (fact and_nat.bit_eq_iff) show \bit (m OR n) q \ bit m q \ bit n q\ by (fact or_nat.bit_eq_iff) show \bit (m XOR n) q \ bit m q \ bit n q\ by (fact xor_nat.bit_eq_iff) qed end lemma Suc_0_and_eq [simp]: \Suc 0 AND n = of_bool (odd n)\ using one_and_eq [of n] by simp lemma and_Suc_0_eq [simp]: \n AND Suc 0 = of_bool (odd n)\ using and_one_eq [of n] by simp lemma Suc_0_or_eq [simp]: \Suc 0 OR n = n + of_bool (even n)\ using one_or_eq [of n] by simp lemma or_Suc_0_eq [simp]: \n OR Suc 0 = n + of_bool (even n)\ using or_one_eq [of n] by simp lemma Suc_0_xor_eq [simp]: \Suc 0 XOR n = n + of_bool (even n) - of_bool (odd n)\ using one_xor_eq [of n] by simp lemma xor_Suc_0_eq [simp]: \n XOR Suc 0 = n + of_bool (even n) - of_bool (odd n)\ using xor_one_eq [of n] by simp subsubsection \Instance \<^typ>\int\\ locale zip_int = single: abel_semigroup f for f :: \bool \ bool \ bool\ (infixl \\<^bold>*\ 70) begin 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) + 2 * ((k div 2) \<^bold>\ (l div 2)))\ by auto termination by (relation "measure (\(k, l). nat (\k\ + \l\))") auto declare F.simps [simp del] lemma rec: \k \<^bold>\ l = of_bool (odd k \<^bold>* odd l) + 2 * ((k div 2) \<^bold>\ (l div 2))\ proof (cases \k \ {0, - 1} \ l \ {0, - 1}\) case True then have \(k div 2) \<^bold>\ (l div 2) = k \<^bold>\ l\ by auto moreover have \of_bool (odd k \<^bold>* odd l) = - (k \<^bold>\ l)\ using True by (simp add: F.simps [of k l]) ultimately show ?thesis by simp next case False then show ?thesis by (auto simp add: ac_simps F.simps [of k l]) qed lemma bit_eq_iff: \bit (k \<^bold>\ l) n \ bit k n \<^bold>* bit l n\ proof (induction n arbitrary: k l) case 0 then show ?case by (simp add: rec [of k l]) next case (Suc n) then show ?case by (simp add: rec [of k l] bit_Suc) qed sublocale abel_semigroup F by standard (simp_all add: Parity.bit_eq_iff bit_eq_iff ac_simps) end instantiation int :: ring_bit_operations begin global_interpretation and_int: zip_int "(\)" defines and_int = and_int.F by standard 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: bit_eq_iff and_int.bit_eq_iff) qed global_interpretation or_int: zip_int "(\)" defines or_int = or_int.F by standard 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: bit_eq_iff or_int.bit_eq_iff) qed global_interpretation xor_int: zip_int "(\)" defines xor_int = xor_int.F by standard definition not_int :: \int \ int\ where \not_int k = - k - 1\ lemma not_int_rec: "NOT k = of_bool (even k) + 2 * NOT (k div 2)" for k :: int by (auto simp add: not_int_def elim: oddE) lemma even_not_iff_int: \even (NOT k) \ odd k\ for k :: int by (simp add: not_int_def) lemma not_int_div_2: \NOT k div 2 = NOT (k div 2)\ 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_div_2 even_not_iff_int bit_Suc) instance proof fix k l :: int and n :: nat show \- k = NOT (k - 1)\ by (simp add: not_int_def) show \bit (k AND l) n \ bit k n \ bit l n\ by (fact and_int.bit_eq_iff) show \bit (k OR l) n \ bit k n \ bit l n\ by (fact or_int.bit_eq_iff) show \bit (k XOR l) n \ bit k n \ bit l n\ by (fact xor_int.bit_eq_iff) qed (simp_all add: bit_not_iff_int) end subsubsection \Instances for \<^typ>\integer\ and \<^typ>\natural\\ unbundle integer.lifting natural.lifting context includes lifting_syntax begin lemma transfer_rule_bit_integer [transfer_rule]: \((pcr_integer :: int \ integer \ bool) ===> (=)) bit bit\ by (unfold bit_def) transfer_prover lemma transfer_rule_bit_natural [transfer_rule]: \((pcr_natural :: nat \ natural \ bool) ===> (=)) bit bit\ by (unfold bit_def) transfer_prover end instantiation integer :: ring_bit_operations begin lift_definition not_integer :: \integer \ integer\ is not . lift_definition and_integer :: \integer \ integer \ integer\ is \and\ . lift_definition or_integer :: \integer \ integer \ integer\ is or . lift_definition xor_integer :: \integer \ integer \ integer\ is xor . instance proof fix k l :: \integer\ and n :: nat show \- k = NOT (k - 1)\ by transfer (simp add: minus_eq_not_minus_1) show \bit (NOT k) n \ (2 :: integer) ^ n \ 0 \ \ bit k n\ by transfer (fact bit_not_iff) show \bit (k AND l) n \ bit k n \ bit l n\ by transfer (fact and_int.bit_eq_iff) show \bit (k OR l) n \ bit k n \ bit l n\ by transfer (fact or_int.bit_eq_iff) show \bit (k XOR l) n \ bit k n \ bit l n\ by transfer (fact xor_int.bit_eq_iff) qed end instantiation natural :: semiring_bit_operations begin lift_definition and_natural :: \natural \ natural \ natural\ is \and\ . lift_definition or_natural :: \natural \ natural \ natural\ is or . lift_definition xor_natural :: \natural \ natural \ natural\ is xor . instance proof fix m n :: \natural\ and q :: nat show \bit (m AND n) q \ bit m q \ bit n q\ by transfer (fact and_nat.bit_eq_iff) show \bit (m OR n) q \ bit m q \ bit n q\ by transfer (fact or_nat.bit_eq_iff) show \bit (m XOR n) q \ bit m q \ bit n q\ by transfer (fact xor_nat.bit_eq_iff) qed end lifting_update integer.lifting lifting_forget integer.lifting lifting_update natural.lifting lifting_forget natural.lifting + +subsection \Key ideas of bit operations\ + +subsection \Key ideas of bit operations\ + +text \ + When formalizing bit operations, it is tempting to represent + bit values as explicit lists over a binary type. This however + is a bad idea, mainly due to the inherent ambiguities in + representation concerning repeating leading bits. + + Hence this approach avoids such explicit lists altogether + following an algebraic path: + + \<^item> Bit values are represented by numeric types: idealized + unbounded bit values can be represented by type \<^typ>\int\, + bounded bit values by quotient types over \<^typ>\int\. + + \<^item> (A special case are idealized unbounded bit values ending + in @{term [source] 0} which can be represented by type \<^typ>\nat\ but + only support a restricted set of operations). + + \<^item> From this idea follows that + + \<^item> multiplication by \<^term>\2 :: int\ is a bit shift to the left and + + \<^item> division by \<^term>\2 :: int\ is a bit shift to the right. + + \<^item> Concerning bounded bit values, iterated shifts to the left + may result in eliminating all bits by shifting them all + beyond the boundary. The property \<^prop>\(2 :: int) ^ n \ 0\ + represents that \<^term>\n\ is \<^emph>\not\ beyond that boundary. + + \<^item> The projection on a single bit is then @{thm bit_def [where ?'a = int, no_vars]}. + + \<^item> This leads to the most fundamental properties of bit values: + + \<^item> Equality rule: @{thm bit_eqI [where ?'a = int, no_vars]} + + \<^item> Induction rule: @{thm bits_induct [where ?'a = int, no_vars]} + + \<^item> Typical operations are characterized as follows: + + \<^item> Singleton \<^term>\n\th bit: \<^term>\(2 :: int) ^ n\ + + \<^item> Bit mask upto bit \<^term>\n\: \<^term>\(2 :: int) ^ n - 1\ + + \<^item> Left shift: @{thm push_bit_eq_mult [where ?'a = int, no_vars]} + + \<^item> Right shift: @{thm drop_bit_eq_div [where ?'a = int, no_vars]} + + \<^item> Truncation: @{thm take_bit_eq_mod [where ?'a = int, no_vars]} + + \<^item> Negation: @{thm bit_not_iff [where ?'a = int, no_vars]} + + \<^item> And: @{thm bit_and_iff [where ?'a = int, no_vars]} + + \<^item> Or: @{thm bit_or_iff [where ?'a = int, no_vars]} + + \<^item> Xor: @{thm bit_xor_iff [where ?'a = int, no_vars]} + + \<^item> Set a single bit: @{thm set_bit_def [where ?'a = int, no_vars]} + + \<^item> Unset a single bit: @{thm unset_bit_def [where ?'a = int, no_vars]} + + \<^item> Flip a single bit: @{thm flip_bit_def [where ?'a = int, no_vars]} +\ + end