diff --git a/src/HOL/Bit_Operations.thy b/src/HOL/Bit_Operations.thy --- a/src/HOL/Bit_Operations.thy +++ b/src/HOL/Bit_Operations.thy @@ -1,3861 +1,3836 @@ (* Author: Florian Haftmann, TUM *) section \Bit operations in suitable algebraic structures\ theory Bit_Operations imports Presburger Groups_List begin subsection \Abstract bit structures\ class semiring_bits = semiring_parity + assumes bits_induct [case_names stable rec]: \(\a. a div 2 = a \ P a) \ (\a b. P a \ (of_bool b + 2 * a) div 2 = a \ P (of_bool b + 2 * a)) \ P a\ assumes bits_div_0 [simp]: \0 div a = 0\ and bits_div_by_1 [simp]: \a div 1 = a\ and bits_mod_div_trivial [simp]: \a mod b div b = 0\ and even_succ_div_2 [simp]: \even a \ (1 + a) div 2 = a div 2\ and even_mask_div_iff: \even ((2 ^ m - 1) div 2 ^ n) \ 2 ^ n = 0 \ m \ n\ and exp_div_exp_eq: \2 ^ m div 2 ^ n = of_bool (2 ^ m \ 0 \ m \ n) * 2 ^ (m - n)\ and div_exp_eq: \a div 2 ^ m div 2 ^ n = a div 2 ^ (m + n)\ and mod_exp_eq: \a mod 2 ^ m mod 2 ^ n = a mod 2 ^ min m n\ and mult_exp_mod_exp_eq: \m \ n \ (a * 2 ^ m) mod (2 ^ n) = (a mod 2 ^ (n - m)) * 2 ^ m\ and div_exp_mod_exp_eq: \a div 2 ^ n mod 2 ^ m = a mod (2 ^ (n + m)) div 2 ^ n\ and even_mult_exp_div_exp_iff: \even (a * 2 ^ m div 2 ^ n) \ m > n \ 2 ^ n = 0 \ (m \ n \ even (a div 2 ^ (n - m)))\ fixes bit :: \'a \ nat \ bool\ assumes bit_iff_odd: \bit a n \ odd (a div 2 ^ n)\ begin text \ Having \<^const>\bit\ as definitional class operation takes into account that specific instances can be implemented differently wrt. code generation. \ lemma bits_div_by_0 [simp]: \a div 0 = 0\ by (metis add_cancel_right_right bits_mod_div_trivial mod_mult_div_eq mult_not_zero) lemma bits_1_div_2 [simp]: \1 div 2 = 0\ using even_succ_div_2 [of 0] by simp lemma bits_1_div_exp [simp]: \1 div 2 ^ n = of_bool (n = 0)\ using div_exp_eq [of 1 1] by (cases n) simp_all lemma even_succ_div_exp [simp]: \(1 + a) div 2 ^ n = a div 2 ^ n\ if \even a\ and \n > 0\ proof (cases n) case 0 with that show ?thesis by simp next case (Suc n) with \even a\ have \(1 + a) div 2 ^ Suc n = a div 2 ^ Suc n\ proof (induction n) case 0 then show ?case by simp next case (Suc n) then show ?case using div_exp_eq [of _ 1 \Suc n\, symmetric] by simp qed with Suc show ?thesis by simp qed lemma even_succ_mod_exp [simp]: \(1 + a) mod 2 ^ n = 1 + (a mod 2 ^ n)\ if \even a\ and \n > 0\ using div_mult_mod_eq [of \1 + a\ \2 ^ n\] div_mult_mod_eq [of a \2 ^ n\] that by simp (metis (full_types) add.left_commute add_left_imp_eq) lemma bits_mod_by_1 [simp]: \a mod 1 = 0\ using div_mult_mod_eq [of a 1] by simp lemma bits_mod_0 [simp]: \0 mod a = 0\ using div_mult_mod_eq [of 0 a] by simp lemma bits_one_mod_two_eq_one [simp]: \1 mod 2 = 1\ by (simp add: mod2_eq_if) lemma bit_0: \bit a 0 \ odd a\ by (simp add: bit_iff_odd) lemma bit_Suc: \bit a (Suc n) \ bit (a div 2) n\ using div_exp_eq [of a 1 n] by (simp add: bit_iff_odd) lemma bit_rec: \bit a n \ (if n = 0 then odd a else bit (a div 2) (n - 1))\ by (cases n) (simp_all add: bit_Suc bit_0) lemma bit_0_eq [simp]: \bit 0 = bot\ by (simp add: fun_eq_iff bit_iff_odd) context fixes a assumes stable: \a div 2 = a\ begin lemma bits_stable_imp_add_self: \a + a mod 2 = 0\ proof - have \a div 2 * 2 + a mod 2 = a\ by (fact div_mult_mod_eq) then have \a * 2 + a mod 2 = a\ by (simp add: stable) then show ?thesis by (simp add: mult_2_right ac_simps) qed lemma stable_imp_bit_iff_odd: \bit a n \ odd a\ by (induction n) (simp_all add: stable bit_Suc bit_0) end lemma bit_iff_idd_imp_stable: \a div 2 = a\ if \\n. bit a n \ odd a\ using that proof (induction a rule: bits_induct) case (stable a) then show ?case by simp next case (rec a b) from rec.prems [of 1] have [simp]: \b = odd a\ by (simp add: rec.hyps bit_Suc bit_0) from rec.hyps have hyp: \(of_bool (odd a) + 2 * a) div 2 = a\ by simp have \bit a n \ odd a\ for n using rec.prems [of \Suc n\] by (simp add: hyp bit_Suc) then have \a div 2 = a\ by (rule rec.IH) then have \of_bool (odd a) + 2 * a = 2 * (a div 2) + of_bool (odd a)\ by (simp add: ac_simps) also have \\ = a\ using mult_div_mod_eq [of 2 a] by (simp add: of_bool_odd_eq_mod_2) finally show ?case using \a div 2 = a\ by (simp add: hyp) qed lemma exp_eq_0_imp_not_bit: \\ bit a n\ if \2 ^ n = 0\ using that by (simp add: bit_iff_odd) definition possible_bit :: \'a itself \ nat \ bool\ where \possible_bit TYPE('a) n \ 2 ^ n \ 0\ \ \This auxiliary avoids non-termination with extensionality.\ lemma possible_bit_0 [simp]: \possible_bit TYPE('a) 0\ by (simp add: possible_bit_def) lemma fold_possible_bit: \2 ^ n = 0 \ \ possible_bit TYPE('a) n\ by (simp add: possible_bit_def) lemma bit_imp_possible_bit: \possible_bit TYPE('a) n\ if \bit a n\ using that by (auto simp add: possible_bit_def exp_eq_0_imp_not_bit) lemma impossible_bit: \\ bit a n\ if \\ possible_bit TYPE('a) n\ using that by (blast dest: bit_imp_possible_bit) lemma possible_bit_less_imp: \possible_bit TYPE('a) j\ if \possible_bit TYPE('a) i\ \j \ i\ using power_add [of 2 j \i - j\] that mult_not_zero [of \2 ^ j\ \2 ^ (i - j)\] by (simp add: possible_bit_def) lemma possible_bit_min [simp]: \possible_bit TYPE('a) (min i j) \ possible_bit TYPE('a) i \ possible_bit TYPE('a) j\ by (auto simp add: min_def elim: possible_bit_less_imp) lemma bit_eqI: \a = b\ if \\n. possible_bit TYPE('a) n \ bit a n \ bit b n\ proof - have \bit a n \ bit b n\ for n proof (cases \2 ^ n = 0\) case True then show ?thesis by (simp add: exp_eq_0_imp_not_bit) next case False then show ?thesis by (rule that[unfolded possible_bit_def]) qed then show ?thesis proof (induction a arbitrary: b rule: bits_induct) case (stable a) from stable(2) [of 0] have **: \even b \ even a\ by (simp add: bit_0) have \b div 2 = b\ proof (rule bit_iff_idd_imp_stable) fix n from stable have *: \bit b n \ bit a n\ by simp also have \bit a n \ odd a\ using stable by (simp add: stable_imp_bit_iff_odd) finally show \bit b n \ odd b\ by (simp add: **) qed from ** have \a mod 2 = b mod 2\ by (simp add: mod2_eq_if) then have \a mod 2 + (a + b) = b mod 2 + (a + b)\ by simp then have \a + a mod 2 + b = b + b mod 2 + a\ by (simp add: ac_simps) with \a div 2 = a\ \b div 2 = b\ show ?case by (simp add: bits_stable_imp_add_self) next case (rec a p) from rec.prems [of 0] have [simp]: \p = odd b\ by (simp add: bit_0) from rec.hyps have \bit a n \ bit (b div 2) n\ for n using rec.prems [of \Suc n\] by (simp add: bit_Suc) then have \a = b div 2\ by (rule rec.IH) then have \2 * a = 2 * (b div 2)\ by simp then have \b mod 2 + 2 * a = b mod 2 + 2 * (b div 2)\ by simp also have \\ = b\ by (fact mod_mult_div_eq) finally show ?case by (auto simp add: mod2_eq_if) qed qed lemma bit_eq_iff: \a = b \ (\n. possible_bit TYPE('a) n \ bit a n \ bit b n)\ by (auto intro: bit_eqI) named_theorems bit_simps \Simplification rules for \<^const>\bit\\ lemma bit_exp_iff [bit_simps]: \bit (2 ^ m) n \ possible_bit TYPE('a) n \ m = n\ by (auto simp add: bit_iff_odd exp_div_exp_eq possible_bit_def) lemma bit_1_iff [bit_simps]: \bit 1 n \ n = 0\ using bit_exp_iff [of 0 n] by auto lemma bit_2_iff [bit_simps]: \bit 2 n \ possible_bit TYPE('a) 1 \ n = 1\ using bit_exp_iff [of 1 n] by auto lemma even_bit_succ_iff: \bit (1 + a) n \ bit a n \ n = 0\ if \even a\ using that by (cases \n = 0\) (simp_all add: bit_iff_odd) lemma bit_double_iff [bit_simps]: \bit (2 * a) n \ bit a (n - 1) \ n \ 0 \ possible_bit TYPE('a) n\ using even_mult_exp_div_exp_iff [of a 1 n] by (cases n, auto simp add: bit_iff_odd ac_simps possible_bit_def) lemma odd_bit_iff_bit_pred: \bit a n \ bit (a - 1) n \ n = 0\ if \odd a\ proof - from \odd a\ obtain b where \a = 2 * b + 1\ .. moreover have \bit (2 * b) n \ n = 0 \ bit (1 + 2 * b) n\ using even_bit_succ_iff by simp ultimately show ?thesis by (simp add: ac_simps) qed lemma bit_eq_rec: \a = b \ (even a \ even b) \ a div 2 = b div 2\ (is \?P = ?Q\) proof assume ?P then show ?Q by simp next assume ?Q then have \even a \ even b\ and \a div 2 = b div 2\ by simp_all show ?P proof (rule bit_eqI) fix n show \bit a n \ bit b n\ proof (cases n) case 0 with \even a \ even b\ show ?thesis by (simp add: bit_0) next case (Suc n) moreover from \a div 2 = b div 2\ have \bit (a div 2) n = bit (b div 2) n\ by simp ultimately show ?thesis by (simp add: bit_Suc) qed qed qed lemma bit_mod_2_iff [simp]: \bit (a mod 2) n \ n = 0 \ odd a\ by (cases a rule: parity_cases) (simp_all add: bit_iff_odd) lemma bit_mask_sub_iff: \bit (2 ^ m - 1) n \ possible_bit TYPE('a) n \ n < m\ by (simp add: bit_iff_odd even_mask_div_iff not_le possible_bit_def) lemma exp_add_not_zero_imp: \2 ^ m \ 0\ and \2 ^ n \ 0\ if \2 ^ (m + n) \ 0\ proof - have \\ (2 ^ m = 0 \ 2 ^ n = 0)\ proof (rule notI) assume \2 ^ m = 0 \ 2 ^ n = 0\ then have \2 ^ (m + n) = 0\ by (rule disjE) (simp_all add: power_add) with that show False .. qed then show \2 ^ m \ 0\ and \2 ^ n \ 0\ by simp_all qed lemma bit_disjunctive_add_iff: \bit (a + b) n \ bit a n \ bit b n\ if \\n. \ bit a n \ \ bit b n\ proof (cases \2 ^ n = 0\) case True then show ?thesis by (simp add: exp_eq_0_imp_not_bit) next case False with that show ?thesis proof (induction n arbitrary: a b) case 0 from "0.prems"(1) [of 0] show ?case by (auto simp add: bit_0) next case (Suc n) from Suc.prems(1) [of 0] have even: \even a \ even b\ by (auto simp add: bit_0) have bit: \\ bit (a div 2) n \ \ bit (b div 2) n\ for n using Suc.prems(1) [of \Suc n\] by (simp add: bit_Suc) from Suc.prems(2) have \2 * 2 ^ n \ 0\ \2 ^ n \ 0\ by (auto simp add: mult_2) have \a + b = (a div 2 * 2 + a mod 2) + (b div 2 * 2 + b mod 2)\ using div_mult_mod_eq [of a 2] div_mult_mod_eq [of b 2] by simp also have \\ = of_bool (odd a \ odd b) + 2 * (a div 2 + b div 2)\ using even by (auto simp add: algebra_simps mod2_eq_if) finally have \bit ((a + b) div 2) n \ bit (a div 2 + b div 2) n\ using \2 * 2 ^ n \ 0\ by simp (simp_all flip: bit_Suc add: bit_double_iff possible_bit_def) also have \\ \ bit (a div 2) n \ bit (b div 2) n\ using bit \2 ^ n \ 0\ by (rule Suc.IH) finally show ?case by (simp add: bit_Suc) qed qed lemma exp_add_not_zero_imp_left: \2 ^ m \ 0\ and exp_add_not_zero_imp_right: \2 ^ n \ 0\ if \2 ^ (m + n) \ 0\ proof - have \\ (2 ^ m = 0 \ 2 ^ n = 0)\ proof (rule notI) assume \2 ^ m = 0 \ 2 ^ n = 0\ then have \2 ^ (m + n) = 0\ by (rule disjE) (simp_all add: power_add) with that show False .. qed then show \2 ^ m \ 0\ and \2 ^ n \ 0\ by simp_all qed lemma exp_not_zero_imp_exp_diff_not_zero: \2 ^ (n - m) \ 0\ if \2 ^ n \ 0\ proof (cases \m \ n\) case True moreover define q where \q = n - m\ ultimately have \n = m + q\ by simp with that show ?thesis by (simp add: exp_add_not_zero_imp_right) next case False with that show ?thesis by simp qed lemma bit_of_bool_iff [bit_simps]: \bit (of_bool b) n \ b \ n = 0\ by (simp add: bit_1_iff) end lemma nat_bit_induct [case_names zero even odd]: \P n\ if zero: \P 0\ and even: \\n. P n \ n > 0 \ P (2 * n)\ and odd: \\n. P n \ P (Suc (2 * n))\ proof (induction n rule: less_induct) case (less n) show \P n\ proof (cases \n = 0\) case True with zero show ?thesis by simp next case False with less have hyp: \P (n div 2)\ by simp show ?thesis proof (cases \even n\) case True then have \n \ 1\ by auto with \n \ 0\ have \n div 2 > 0\ by simp with \even n\ hyp even [of \n div 2\] show ?thesis by simp next case False with hyp odd [of \n div 2\] show ?thesis by simp qed qed qed instantiation nat :: semiring_bits begin definition bit_nat :: \nat \ nat \ bool\ where \bit_nat m n \ odd (m div 2 ^ n)\ instance proof show \P n\ if stable: \\n. n div 2 = n \ P n\ and rec: \\n b. P n \ (of_bool b + 2 * n) div 2 = n \ P (of_bool b + 2 * n)\ for P and n :: nat proof (induction n rule: nat_bit_induct) case zero from stable [of 0] show ?case by simp next case (even n) with rec [of n False] show ?case by simp next case (odd n) with rec [of n True] show ?case by simp qed show \q mod 2 ^ m mod 2 ^ n = q mod 2 ^ min m n\ for q m n :: nat apply (auto simp add: less_iff_Suc_add power_add mod_mod_cancel split: split_min_lin) apply (metis div_mult2_eq mod_div_trivial mod_eq_self_iff_div_eq_0 mod_mult_self2_is_0 power_commutes) done show \(q * 2 ^ m) mod (2 ^ n) = (q mod 2 ^ (n - m)) * 2 ^ m\ if \m \ n\ for q m n :: nat using that apply (auto simp add: mod_mod_cancel div_mult2_eq power_add mod_mult2_eq le_iff_add split: split_min_lin) done show \even ((2 ^ m - (1::nat)) div 2 ^ n) \ 2 ^ n = (0::nat) \ m \ n\ for m n :: nat using even_mask_div_iff' [where ?'a = nat, of m n] by simp show \even (q * 2 ^ m div 2 ^ n) \ n < m \ (2::nat) ^ n = 0 \ m \ n \ even (q div 2 ^ (n - m))\ for m n q r :: nat apply (auto simp add: not_less power_add ac_simps dest!: le_Suc_ex) apply (metis (full_types) dvd_mult dvd_mult_imp_div dvd_power_iff_le not_less not_less_eq order_refl power_Suc) done qed (auto simp add: div_mult2_eq mod_mult2_eq power_add power_diff bit_nat_def) end lemma possible_bit_nat [simp]: \possible_bit TYPE(nat) n\ by (simp add: possible_bit_def) lemma not_bit_Suc_0_Suc [simp]: \\ bit (Suc 0) (Suc n)\ by (simp add: bit_Suc) lemma not_bit_Suc_0_numeral [simp]: \\ bit (Suc 0) (numeral n)\ by (simp add: numeral_eq_Suc) context semiring_bits begin lemma bit_of_nat_iff [bit_simps]: \bit (of_nat m) n \ possible_bit TYPE('a) n \ bit m n\ proof (cases \(2::'a) ^ n = 0\) case True then show ?thesis by (simp add: exp_eq_0_imp_not_bit possible_bit_def) next case False then have \bit (of_nat m) n \ bit m n\ proof (induction m arbitrary: n rule: nat_bit_induct) case zero then show ?case by simp next case (even m) then show ?case by (cases n) (auto simp add: bit_double_iff Bit_Operations.bit_double_iff possible_bit_def bit_0 dest: mult_not_zero) next case (odd m) then show ?case by (cases n) (auto simp add: bit_double_iff even_bit_succ_iff possible_bit_def Bit_Operations.bit_Suc Bit_Operations.bit_0 dest: mult_not_zero) qed with False show ?thesis by (simp add: possible_bit_def) qed end lemma int_bit_induct [case_names zero minus even odd]: \P k\ if zero_int: \P 0\ and minus_int: \P (- 1)\ and even_int: \\k. P k \ k \ 0 \ P (k * 2)\ and odd_int: \\k. P k \ k \ - 1 \ P (1 + (k * 2))\ for k :: int proof (cases \k \ 0\) case True define n where \n = nat k\ with True have \k = int n\ by simp then show \P k\ proof (induction n arbitrary: k rule: nat_bit_induct) case zero then show ?case by (simp add: zero_int) next case (even n) have \P (int n * 2)\ by (rule even_int) (use even in simp_all) with even show ?case by (simp add: ac_simps) next case (odd n) have \P (1 + (int n * 2))\ by (rule odd_int) (use odd in simp_all) with odd show ?case by (simp add: ac_simps) qed next case False define n where \n = nat (- k - 1)\ with False have \k = - int n - 1\ by simp then show \P k\ proof (induction n arbitrary: k rule: nat_bit_induct) case zero then show ?case by (simp add: minus_int) next case (even n) have \P (1 + (- int (Suc n) * 2))\ by (rule odd_int) (use even in \simp_all add: algebra_simps\) also have \\ = - int (2 * n) - 1\ by (simp add: algebra_simps) finally show ?case using even.prems by simp next case (odd n) have \P (- int (Suc n) * 2)\ by (rule even_int) (use odd in \simp_all add: algebra_simps\) also have \\ = - int (Suc (2 * n)) - 1\ by (simp add: algebra_simps) finally show ?case using odd.prems by simp qed qed instantiation int :: semiring_bits begin definition bit_int :: \int \ nat \ bool\ where \bit_int k n \ odd (k div 2 ^ n)\ instance proof show \P k\ if stable: \\k. k div 2 = k \ P k\ and rec: \\k b. P k \ (of_bool b + 2 * k) div 2 = k \ P (of_bool b + 2 * k)\ for P and k :: int proof (induction k rule: int_bit_induct) case zero from stable [of 0] show ?case by simp next case minus from stable [of \- 1\] show ?case by simp next case (even k) with rec [of k False] show ?case by (simp add: ac_simps) next case (odd k) with rec [of k True] show ?case by (simp add: ac_simps) qed show \(2::int) ^ m div 2 ^ n = of_bool ((2::int) ^ m \ 0 \ n \ m) * 2 ^ (m - n)\ for m n :: nat proof (cases \m < n\) case True then have \n = m + (n - m)\ by simp then have \(2::int) ^ m div 2 ^ n = (2::int) ^ m div 2 ^ (m + (n - m))\ by simp also have \\ = (2::int) ^ m div (2 ^ m * 2 ^ (n - m))\ by (simp add: power_add) also have \\ = (2::int) ^ m div 2 ^ m div 2 ^ (n - m)\ by (simp add: zdiv_zmult2_eq) finally show ?thesis using \m < n\ by simp next case False then show ?thesis by (simp add: power_diff) qed show \k mod 2 ^ m mod 2 ^ n = k mod 2 ^ min m n\ for m n :: nat and k :: int using mod_exp_eq [of \nat k\ m n] apply (auto simp add: mod_mod_cancel zdiv_zmult2_eq power_add zmod_zmult2_eq le_iff_add split: split_min_lin) apply (auto simp add: less_iff_Suc_add mod_mod_cancel power_add) apply (simp only: flip: mult.left_commute [of \2 ^ m\]) apply (subst zmod_zmult2_eq) apply simp_all done show \(k * 2 ^ m) mod (2 ^ n) = (k mod 2 ^ (n - m)) * 2 ^ m\ if \m \ n\ for m n :: nat and k :: int using that apply (auto simp add: power_add zmod_zmult2_eq le_iff_add split: split_min_lin) done show \even ((2 ^ m - (1::int)) div 2 ^ n) \ 2 ^ n = (0::int) \ m \ n\ for m n :: nat using even_mask_div_iff' [where ?'a = int, of m n] by simp show \even (k * 2 ^ m div 2 ^ n) \ n < m \ (2::int) ^ n = 0 \ m \ n \ even (k div 2 ^ (n - m))\ for m n :: nat and k l :: int apply (auto simp add: not_less power_add ac_simps dest!: le_Suc_ex) apply (metis Suc_leI dvd_mult dvd_mult_imp_div dvd_power_le dvd_refl power.simps(2)) done qed (auto simp add: zdiv_zmult2_eq zmod_zmult2_eq power_add power_diff not_le bit_int_def) end lemma possible_bit_int [simp]: \possible_bit TYPE(int) n\ by (simp add: possible_bit_def) lemma bit_not_int_iff': \bit (- k - 1) n \ \ bit k n\ for k :: int proof (induction n arbitrary: k) case 0 show ?case by (simp add: bit_0) next case (Suc n) have \- k - 1 = - (k + 2) + 1\ by simp also have \(- (k + 2) + 1) div 2 = - (k div 2) - 1\ proof (cases \even k\) case True then have \- k div 2 = - (k div 2)\ by rule (simp flip: mult_minus_right) with True show ?thesis by simp next case False have \4 = 2 * (2::int)\ by simp also have \2 * 2 div 2 = (2::int)\ by (simp only: nonzero_mult_div_cancel_left) finally have *: \4 div 2 = (2::int)\ . from False obtain l where k: \k = 2 * l + 1\ .. then have \- k - 2 = 2 * - (l + 2) + 1\ by simp then have \(- k - 2) div 2 + 1 = - (k div 2) - 1\ by (simp flip: mult_minus_right add: *) (simp add: k) with False show ?thesis by simp qed finally have \(- k - 1) div 2 = - (k div 2) - 1\ . with Suc show ?case by (simp add: bit_Suc) qed lemma bit_nat_iff [bit_simps]: \bit (nat k) n \ k \ 0 \ bit k n\ proof (cases \k \ 0\) case True moreover define m where \m = nat k\ ultimately have \k = int m\ by simp then show ?thesis by (simp add: bit_simps) next case False then show ?thesis by simp qed subsection \Bit operations\ class semiring_bit_operations = semiring_bits + fixes "and" :: \'a \ 'a \ 'a\ (infixr \AND\ 64) and or :: \'a \ 'a \ 'a\ (infixr \OR\ 59) and xor :: \'a \ 'a \ 'a\ (infixr \XOR\ 59) and mask :: \nat \ 'a\ and set_bit :: \nat \ 'a \ 'a\ and unset_bit :: \nat \ 'a \ 'a\ and flip_bit :: \nat \ 'a \ 'a\ and push_bit :: \nat \ 'a \ 'a\ and drop_bit :: \nat \ 'a \ 'a\ and take_bit :: \nat \ 'a \ 'a\ assumes and_rec: \a AND b = of_bool (odd a \ odd b) + 2 * ((a div 2) AND (b div 2))\ and or_rec: \a OR b = of_bool (odd a \ odd b) + 2 * ((a div 2) OR (b div 2))\ and xor_rec: \a XOR b = of_bool (odd a \ odd b) + 2 * ((a div 2) XOR (b div 2))\ and mask_eq_exp_minus_1: \mask n = 2 ^ n - 1\ and set_bit_eq_or: \set_bit n a = a OR push_bit n 1\ and unset_bit_0 [simp]: \unset_bit 0 a = 2 * (a div 2)\ and unset_bit_Suc: \unset_bit (Suc n) a = a mod 2 + 2 * unset_bit n (a div 2)\ and flip_bit_eq_xor: \flip_bit n a = a XOR push_bit n 1\ and push_bit_eq_mult: \push_bit n a = a * 2 ^ n\ and drop_bit_eq_div: \drop_bit n a = a div 2 ^ n\ and take_bit_eq_mod: \take_bit n a = a mod 2 ^ n\ begin text \ We want the bitwise operations to bind slightly weaker than \+\ and \-\. Logically, \<^const>\push_bit\, \<^const>\drop_bit\ and \<^const>\take_bit\ are just aliases; having them as separate operations makes proofs easier, otherwise proof automation would fiddle with concrete expressions \<^term>\2 ^ n\ in a way obfuscating the basic algebraic relationships between those operations. - For the sake of code generation operations + For the sake of code generation operations are specified as definitional class operations, taking into account that specific instances of these can be implemented differently wrt. code generation. \ lemma bit_and_iff [bit_simps]: \bit (a AND b) n \ bit a n \ bit b n\ proof (induction n arbitrary: a b) case 0 show ?case by (simp add: bit_0 and_rec [of a b] even_bit_succ_iff) next case (Suc n) from Suc [of \a div 2\ \b div 2\] show ?case by (simp add: and_rec [of a b] bit_Suc) (auto simp flip: bit_Suc simp add: bit_double_iff dest: bit_imp_possible_bit) qed lemma bit_or_iff [bit_simps]: \bit (a OR b) n \ bit a n \ bit b n\ proof (induction n arbitrary: a b) case 0 show ?case by (simp add: bit_0 or_rec [of a b] even_bit_succ_iff) next case (Suc n) from Suc [of \a div 2\ \b div 2\] show ?case by (simp add: or_rec [of a b] bit_Suc) (auto simp flip: bit_Suc simp add: bit_double_iff dest: bit_imp_possible_bit) qed lemma bit_xor_iff [bit_simps]: \bit (a XOR b) n \ bit a n \ bit b n\ proof (induction n arbitrary: a b) case 0 show ?case by (simp add: bit_0 xor_rec [of a b] even_bit_succ_iff) next case (Suc n) from Suc [of \a div 2\ \b div 2\] show ?case by (simp add: xor_rec [of a b] bit_Suc) (auto simp flip: bit_Suc simp add: bit_double_iff dest: bit_imp_possible_bit) qed 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 even_and_iff: \even (a AND b) \ even a \ even b\ using bit_and_iff [of a b 0] by (auto simp add: bit_0) lemma even_or_iff: \even (a OR b) \ even a \ even b\ using bit_or_iff [of a b 0] by (auto simp add: bit_0) lemma even_xor_iff: \even (a XOR b) \ (even a \ even b)\ using bit_xor_iff [of a b 0] by (auto simp add: bit_0) 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: \1 AND a = a mod 2\ by (simp add: bit_eq_iff bit_and_iff) (auto simp add: bit_1_iff bit_0) lemma and_one_eq: \a AND 1 = a mod 2\ using one_and_eq [of a] by (simp add: ac_simps) lemma one_or_eq: \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 bit_0) lemma or_one_eq: \a OR 1 = a + of_bool (even a)\ using one_or_eq [of a] by (simp add: ac_simps) lemma one_xor_eq: \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 bit_0 elim: oddE) lemma xor_one_eq: \a XOR 1 = a + of_bool (even a) - of_bool (odd a)\ using one_xor_eq [of a] by (simp add: ac_simps) lemma xor_self_eq [simp]: \a XOR a = 0\ by (rule bit_eqI) (simp add: bit_simps) lemma bit_iff_odd_drop_bit: \bit a n \ odd (drop_bit n a)\ by (simp add: bit_iff_odd drop_bit_eq_div) lemma even_drop_bit_iff_not_bit: \even (drop_bit n a) \ \ bit a n\ by (simp add: bit_iff_odd_drop_bit) lemma div_push_bit_of_1_eq_drop_bit: \a div push_bit n 1 = drop_bit n a\ by (simp add: push_bit_eq_mult drop_bit_eq_div) lemma bits_ident: \push_bit n (drop_bit n a) + take_bit n a = a\ using div_mult_mod_eq by (simp add: push_bit_eq_mult take_bit_eq_mod drop_bit_eq_div) lemma push_bit_push_bit [simp]: \push_bit m (push_bit n a) = push_bit (m + n) a\ by (simp add: push_bit_eq_mult power_add ac_simps) lemma push_bit_0_id [simp]: \push_bit 0 = id\ by (simp add: fun_eq_iff push_bit_eq_mult) lemma push_bit_of_0 [simp]: \push_bit n 0 = 0\ by (simp add: push_bit_eq_mult) lemma push_bit_of_1 [simp]: \push_bit n 1 = 2 ^ n\ by (simp add: push_bit_eq_mult) lemma push_bit_Suc [simp]: \push_bit (Suc n) a = push_bit n (a * 2)\ by (simp add: push_bit_eq_mult ac_simps) lemma push_bit_double: \push_bit n (a * 2) = push_bit n a * 2\ by (simp add: push_bit_eq_mult ac_simps) lemma push_bit_add: \push_bit n (a + b) = push_bit n a + push_bit n b\ by (simp add: push_bit_eq_mult algebra_simps) lemma push_bit_numeral [simp]: \push_bit (numeral l) (numeral k) = push_bit (pred_numeral l) (numeral (Num.Bit0 k))\ by (simp add: numeral_eq_Suc mult_2_right) (simp add: numeral_Bit0) lemma take_bit_0 [simp]: "take_bit 0 a = 0" by (simp add: take_bit_eq_mod) lemma take_bit_Suc: \take_bit (Suc n) a = take_bit n (a div 2) * 2 + a mod 2\ proof - have \take_bit (Suc n) (a div 2 * 2 + of_bool (odd a)) = take_bit n (a div 2) * 2 + of_bool (odd a)\ using even_succ_mod_exp [of \2 * (a div 2)\ \Suc n\] mult_exp_mod_exp_eq [of 1 \Suc n\ \a div 2\] by (auto simp add: take_bit_eq_mod ac_simps) then show ?thesis using div_mult_mod_eq [of a 2] by (simp add: mod_2_eq_odd) qed lemma take_bit_rec: \take_bit n a = (if n = 0 then 0 else take_bit (n - 1) (a div 2) * 2 + a mod 2)\ by (cases n) (simp_all add: take_bit_Suc) lemma take_bit_Suc_0 [simp]: \take_bit (Suc 0) a = a mod 2\ by (simp add: take_bit_eq_mod) lemma take_bit_of_0 [simp]: \take_bit n 0 = 0\ by (simp add: take_bit_eq_mod) lemma take_bit_of_1 [simp]: \take_bit n 1 = of_bool (n > 0)\ by (cases n) (simp_all add: take_bit_Suc) lemma drop_bit_of_0 [simp]: \drop_bit n 0 = 0\ by (simp add: drop_bit_eq_div) lemma drop_bit_of_1 [simp]: \drop_bit n 1 = of_bool (n = 0)\ by (simp add: drop_bit_eq_div) lemma drop_bit_0 [simp]: \drop_bit 0 = id\ by (simp add: fun_eq_iff drop_bit_eq_div) lemma drop_bit_Suc: \drop_bit (Suc n) a = drop_bit n (a div 2)\ using div_exp_eq [of a 1] by (simp add: drop_bit_eq_div) lemma drop_bit_rec: \drop_bit n a = (if n = 0 then a else drop_bit (n - 1) (a div 2))\ by (cases n) (simp_all add: drop_bit_Suc) lemma drop_bit_half: \drop_bit n (a div 2) = drop_bit n a div 2\ by (induction n arbitrary: a) (simp_all add: drop_bit_Suc) lemma drop_bit_of_bool [simp]: \drop_bit n (of_bool b) = of_bool (n = 0 \ b)\ by (cases n) simp_all lemma even_take_bit_eq [simp]: \even (take_bit n a) \ n = 0 \ even a\ by (simp add: take_bit_rec [of n a]) lemma take_bit_take_bit [simp]: \take_bit m (take_bit n a) = take_bit (min m n) a\ by (simp add: take_bit_eq_mod mod_exp_eq ac_simps) lemma drop_bit_drop_bit [simp]: \drop_bit m (drop_bit n a) = drop_bit (m + n) a\ by (simp add: drop_bit_eq_div power_add div_exp_eq ac_simps) lemma push_bit_take_bit: \push_bit m (take_bit n a) = take_bit (m + n) (push_bit m a)\ apply (simp add: push_bit_eq_mult take_bit_eq_mod power_add ac_simps) using mult_exp_mod_exp_eq [of m \m + n\ a] apply (simp add: ac_simps power_add) done lemma take_bit_push_bit: \take_bit m (push_bit n a) = push_bit n (take_bit (m - n) a)\ proof (cases \m \ n\) case True then show ?thesis apply (simp add:) apply (simp_all add: push_bit_eq_mult take_bit_eq_mod) apply (auto dest!: le_Suc_ex simp add: power_add ac_simps) using mult_exp_mod_exp_eq [of m m \a * 2 ^ n\ for n] apply (simp add: ac_simps) done next case False then show ?thesis using push_bit_take_bit [of n \m - n\ a] by simp qed lemma take_bit_drop_bit: \take_bit m (drop_bit n a) = drop_bit n (take_bit (m + n) a)\ by (simp add: drop_bit_eq_div take_bit_eq_mod ac_simps div_exp_mod_exp_eq) lemma drop_bit_take_bit: \drop_bit m (take_bit n a) = take_bit (n - m) (drop_bit m a)\ proof (cases "m \ n") case True then show ?thesis using take_bit_drop_bit [of "n - m" m a] by simp next case False then obtain q where \m = n + q\ by (auto simp add: not_le dest: less_imp_Suc_add) then have \drop_bit m (take_bit n a) = 0\ using div_exp_eq [of \a mod 2 ^ n\ n q] by (simp add: take_bit_eq_mod drop_bit_eq_div) with False show ?thesis by simp qed lemma even_push_bit_iff [simp]: \even (push_bit n a) \ n \ 0 \ even a\ by (simp add: push_bit_eq_mult) auto lemma bit_push_bit_iff [bit_simps]: \bit (push_bit m a) n \ m \ n \ possible_bit TYPE('a) n \ bit a (n - m)\ by (auto simp add: bit_iff_odd push_bit_eq_mult even_mult_exp_div_exp_iff possible_bit_def) lemma bit_drop_bit_eq [bit_simps]: \bit (drop_bit n a) = bit a \ (+) n\ by (simp add: bit_iff_odd fun_eq_iff ac_simps flip: drop_bit_eq_div) lemma bit_take_bit_iff [bit_simps]: \bit (take_bit m a) n \ n < m \ bit a n\ by (simp add: bit_iff_odd drop_bit_take_bit not_le flip: drop_bit_eq_div) lemma stable_imp_drop_bit_eq: \drop_bit n a = a\ if \a div 2 = a\ by (induction n) (simp_all add: that drop_bit_Suc) lemma stable_imp_take_bit_eq: \take_bit n a = (if even a then 0 else 2 ^ n - 1)\ if \a div 2 = a\ proof (rule bit_eqI[unfolded possible_bit_def]) fix m assume \2 ^ m \ 0\ with that show \bit (take_bit n a) m \ bit (if even a then 0 else 2 ^ n - 1) m\ by (simp add: bit_take_bit_iff bit_mask_sub_iff possible_bit_def stable_imp_bit_iff_odd) qed lemma exp_dvdE: assumes \2 ^ n dvd a\ obtains b where \a = push_bit n b\ proof - from assms obtain b where \a = 2 ^ n * b\ .. then have \a = push_bit n b\ by (simp add: push_bit_eq_mult ac_simps) with that show thesis . qed lemma take_bit_eq_0_iff: \take_bit n a = 0 \ 2 ^ n dvd a\ (is \?P \ ?Q\) proof assume ?P then show ?Q by (simp add: take_bit_eq_mod mod_0_imp_dvd) next assume ?Q then obtain b where \a = push_bit n b\ by (rule exp_dvdE) then show ?P by (simp add: take_bit_push_bit) qed lemma take_bit_tightened: - \take_bit m a = take_bit m b\ if \take_bit n a = take_bit n b\ and \m \ n\ + \take_bit m a = take_bit m b\ if \take_bit n a = take_bit n b\ and \m \ n\ proof - from that have \take_bit m (take_bit n a) = take_bit m (take_bit n b)\ by simp then have \take_bit (min m n) a = take_bit (min m n) b\ by simp with that show ?thesis by (simp add: min_def) qed lemma take_bit_eq_self_iff_drop_bit_eq_0: \take_bit n a = a \ drop_bit n a = 0\ (is \?P \ ?Q\) proof assume ?P show ?Q proof (rule bit_eqI) fix m from \?P\ have \a = take_bit n a\ .. also have \\ bit (take_bit n a) (n + m)\ unfolding bit_simps - by (simp add: bit_simps) + by (simp add: bit_simps) finally show \bit (drop_bit n a) m \ bit 0 m\ by (simp add: bit_simps) qed next assume ?Q show ?P proof (rule bit_eqI) fix m from \?Q\ have \\ bit (drop_bit n a) (m - n)\ by simp then have \ \ bit a (n + (m - n))\ by (simp add: bit_simps) then show \bit (take_bit n a) m \ bit a m\ by (cases \m < n\) (auto simp add: bit_simps) qed qed lemma drop_bit_exp_eq: \drop_bit m (2 ^ n) = of_bool (m \ n \ possible_bit TYPE('a) n) * 2 ^ (n - m)\ by (auto simp add: bit_eq_iff bit_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_simps) 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_simps) 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_simps) lemma push_bit_and [simp]: \push_bit n (a AND b) = push_bit n a AND push_bit n b\ by (auto simp add: bit_eq_iff bit_simps) lemma push_bit_or [simp]: \push_bit n (a OR b) = push_bit n a OR push_bit n b\ by (auto simp add: bit_eq_iff bit_simps) lemma push_bit_xor [simp]: \push_bit n (a XOR b) = push_bit n a XOR push_bit n b\ by (auto simp add: bit_eq_iff bit_simps) lemma drop_bit_and [simp]: \drop_bit n (a AND b) = drop_bit n a AND drop_bit n b\ by (auto simp add: bit_eq_iff bit_simps) lemma drop_bit_or [simp]: \drop_bit n (a OR b) = drop_bit n a OR drop_bit n b\ by (auto simp add: bit_eq_iff bit_simps) lemma drop_bit_xor [simp]: \drop_bit n (a XOR b) = drop_bit n a XOR drop_bit n b\ by (auto simp add: bit_eq_iff bit_simps) lemma bit_mask_iff [bit_simps]: \bit (mask m) n \ possible_bit TYPE('a) n \ n < m\ by (simp add: mask_eq_exp_minus_1 bit_mask_sub_iff) lemma even_mask_iff: \even (mask n) \ n = 0\ using bit_mask_iff [of n 0] by (auto simp add: bit_0) lemma mask_0 [simp]: \mask 0 = 0\ by (simp add: mask_eq_exp_minus_1) lemma mask_Suc_0 [simp]: \mask (Suc 0) = 1\ by (simp add: mask_eq_exp_minus_1 add_implies_diff sym) lemma mask_Suc_exp: \mask (Suc n) = 2 ^ n OR mask n\ by (auto simp add: bit_eq_iff bit_simps) lemma mask_Suc_double: \mask (Suc n) = 1 OR 2 * mask n\ by (auto simp add: bit_eq_iff bit_simps elim: possible_bit_less_imp) lemma mask_numeral: \mask (numeral n) = 1 + 2 * mask (pred_numeral n)\ by (simp add: numeral_eq_Suc mask_Suc_double one_or_eq ac_simps) lemma take_bit_of_mask [simp]: \take_bit m (mask n) = mask (min m n)\ by (rule bit_eqI) (simp add: bit_simps) lemma take_bit_eq_mask: \take_bit n a = a AND mask n\ by (auto simp add: bit_eq_iff bit_simps) lemma or_eq_0_iff: \a OR b = 0 \ a = 0 \ b = 0\ by (auto simp add: bit_eq_iff bit_or_iff) lemma disjunctive_add: \a + b = a OR b\ if \\n. \ bit a n \ \ bit b n\ by (rule bit_eqI) (use that in \simp add: bit_disjunctive_add_iff bit_or_iff\) lemma bit_iff_and_drop_bit_eq_1: \bit a n \ drop_bit n a AND 1 = 1\ by (simp add: bit_iff_odd_drop_bit and_one_eq odd_iff_mod_2_eq_one) lemma bit_iff_and_push_bit_not_eq_0: \bit a n \ a AND push_bit n 1 \ 0\ apply (cases \2 ^ n = 0\) apply (simp_all add: bit_eq_iff bit_and_iff bit_push_bit_iff exp_eq_0_imp_not_bit) apply (simp_all add: bit_exp_iff) done lemma bit_set_bit_iff [bit_simps]: \bit (set_bit m a) n \ bit a n \ (m = n \ possible_bit TYPE('a) n)\ by (auto simp add: set_bit_eq_or 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 simp add: bit_0) lemma bit_unset_bit_iff [bit_simps]: \bit (unset_bit m a) n \ bit a n \ m \ n\ proof (induction m arbitrary: a n) case 0 then show ?case by (auto simp add: bit_simps simp flip: bit_Suc dest: bit_imp_possible_bit) next case (Suc m) show ?case proof (cases n) case 0 then show ?thesis by (cases m) (simp_all add: bit_0 unset_bit_Suc) next case (Suc n) with Suc.IH [of \a div 2\ n] show ?thesis by (auto simp add: unset_bit_Suc mod_2_eq_odd bit_simps even_bit_succ_iff simp flip: bit_Suc dest: bit_imp_possible_bit) qed qed 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 simp add: bit_0) lemma and_exp_eq_0_iff_not_bit: \a AND 2 ^ n = 0 \ \ bit a n\ (is \?P \ ?Q\) using bit_imp_possible_bit[of a n] by (auto simp add: bit_eq_iff bit_simps) lemma bit_flip_bit_iff [bit_simps]: \bit (flip_bit m a) n \ (m = n \ \ bit a n) \ possible_bit TYPE('a) n\ by (auto simp add: bit_eq_iff bit_simps flip_bit_eq_xor bit_imp_possible_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 simp: possible_bit_def bit_0) lemma set_bit_0 [simp]: \set_bit 0 a = 1 + 2 * (a div 2)\ by (auto simp add: bit_eq_iff bit_simps even_bit_succ_iff simp flip: bit_Suc) lemma bit_sum_mult_2_cases: assumes a: "\j. \ bit a (Suc j)" shows "bit (a + 2 * b) n = (if n = 0 then odd a else bit (2 * b) n)" proof - have a_eq: "bit a i \ i = 0 \ odd a" for i by (cases i) (simp_all add: a bit_0) show ?thesis by (simp add: disjunctive_add[simplified disj_imp] a_eq bit_simps) qed lemma set_bit_Suc: \set_bit (Suc n) a = a mod 2 + 2 * set_bit n (a div 2)\ by (auto simp add: bit_eq_iff bit_sum_mult_2_cases bit_simps bit_0 simp flip: bit_Suc elim: possible_bit_less_imp) lemma flip_bit_0 [simp]: \flip_bit 0 a = of_bool (even a) + 2 * (a div 2)\ by (auto simp add: bit_eq_iff bit_simps even_bit_succ_iff bit_0 simp flip: bit_Suc) lemma flip_bit_Suc: \flip_bit (Suc n) a = a mod 2 + 2 * flip_bit n (a div 2)\ by (auto simp add: bit_eq_iff bit_sum_mult_2_cases bit_simps bit_0 simp flip: bit_Suc elim: possible_bit_less_imp) lemma flip_bit_eq_if: \flip_bit n a = (if bit a n then unset_bit else set_bit) n a\ by (rule bit_eqI) (auto simp add: bit_set_bit_iff bit_unset_bit_iff bit_flip_bit_iff) lemma take_bit_set_bit_eq: \take_bit n (set_bit m a) = (if n \ m then take_bit n a else set_bit m (take_bit n a))\ by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_set_bit_iff) lemma take_bit_unset_bit_eq: \take_bit n (unset_bit m a) = (if n \ m then take_bit n a else unset_bit m (take_bit n a))\ by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_unset_bit_iff) lemma take_bit_flip_bit_eq: \take_bit n (flip_bit m a) = (if n \ m then take_bit n a else flip_bit m (take_bit n a))\ by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_flip_bit_iff) lemma bit_1_0 [simp]: \bit 1 0\ by (simp add: bit_0) lemma not_bit_1_Suc [simp]: \\ bit 1 (Suc n)\ by (simp add: bit_Suc) lemma push_bit_Suc_numeral [simp]: \push_bit (Suc n) (numeral k) = push_bit n (numeral (Num.Bit0 k))\ by (simp add: numeral_eq_Suc mult_2_right) (simp add: numeral_Bit0) lemma mask_eq_0_iff [simp]: \mask n = 0 \ n = 0\ by (cases n) (simp_all add: mask_Suc_double or_eq_0_iff) lemma bit_horner_sum_bit_iff [bit_simps]: \bit (horner_sum of_bool 2 bs) n \ possible_bit TYPE('a) n \ n < length bs \ bs ! n\ 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 add: bit_0) next case (Suc m) with bit_rec [of _ n] Cons.prems Cons.IH [of m] show ?thesis by (simp add: bit_simps) (auto simp add: possible_bit_less_imp bit_simps simp flip: bit_Suc) qed qed lemma horner_sum_bit_eq_take_bit: \horner_sum of_bool 2 (map (bit a) [0.. by (rule bit_eqI) (auto simp add: bit_simps) lemma take_bit_horner_sum_bit_eq: \take_bit n (horner_sum of_bool 2 bs) = horner_sum of_bool 2 (take n bs)\ by (auto simp add: bit_eq_iff bit_take_bit_iff bit_horner_sum_bit_iff) lemma take_bit_sum: \take_bit n a = (\k = 0.. by (simp flip: horner_sum_bit_eq_take_bit add: horner_sum_eq_sum push_bit_eq_mult) -lemmas set_bit_def = set_bit_eq_or - -lemmas flip_bit_def = flip_bit_eq_xor - end class ring_bit_operations = semiring_bit_operations + ring_parity + fixes not :: \'a \ 'a\ (\NOT\) assumes not_rec: \NOT a = of_bool (even a) + 2 * NOT (a div 2)\ 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 even_not_iff [simp]: \even (NOT a) \ odd a\ by (simp add: not_rec [of a]) lemma bit_not_iff [bit_simps]: \bit (NOT a) n \ possible_bit TYPE('a) n \ \ bit a n\ proof (cases \possible_bit TYPE('a) n\) case False then show ?thesis - by (auto dest: bit_imp_possible_bit) + by (auto dest: bit_imp_possible_bit) next case True moreover have \bit (NOT a) n \ \ bit a n\ using \possible_bit TYPE('a) n\ proof (induction n arbitrary: a) case 0 then show ?case by (simp add: bit_0) next case (Suc n) from Suc.prems Suc.IH [of \a div 2\] show ?case by (simp add: impossible_bit possible_bit_less_imp not_rec [of a] even_bit_succ_iff bit_double_iff flip: bit_Suc) qed ultimately show ?thesis by simp qed lemma bit_not_exp_iff [bit_simps]: \bit (NOT (2 ^ m)) n \ possible_bit TYPE('a) n \ n \ m\ by (auto simp add: bit_not_iff bit_exp_iff) lemma bit_minus_iff [bit_simps]: \bit (- a) n \ possible_bit TYPE('a) n \ \ bit (a - 1) n\ by (simp add: minus_eq_not_minus_1 bit_not_iff) lemma bit_minus_1_iff [simp]: \bit (- 1) n \ possible_bit TYPE('a) n\ by (simp add: bit_minus_iff) lemma bit_minus_exp_iff [bit_simps]: \bit (- (2 ^ m)) n \ possible_bit TYPE('a) n \ n \ m\ by (auto simp add: bit_simps simp flip: mask_eq_exp_minus_1) lemma bit_minus_2_iff [simp]: \bit (- 2) n \ possible_bit TYPE('a) n \ n > 0\ by (simp add: bit_minus_iff bit_1_iff) lemma bit_not_iff_eq: \bit (NOT a) n \ 2 ^ n \ 0 \ \ bit a n\ by (simp add: bit_simps possible_bit_def) lemma not_one_eq [simp]: \NOT 1 = - 2\ by (simp add: bit_eq_iff bit_not_iff) (simp add: bit_1_iff) sublocale "and": semilattice_neutr \(AND)\ \- 1\ by standard (rule bit_eqI, simp add: bit_and_iff) sublocale bit: abstract_boolean_algebra \(AND)\ \(OR)\ NOT 0 \- 1\ by standard (auto simp add: bit_and_iff bit_or_iff bit_not_iff intro: bit_eqI) sublocale bit: abstract_boolean_algebra_sym_diff \(AND)\ \(OR)\ NOT 0 \- 1\ \(XOR)\ apply standard apply (rule bit_eqI) apply (auto simp add: bit_simps) done lemma and_eq_not_not_or: \a AND b = NOT (NOT a OR NOT b)\ by simp lemma or_eq_not_not_and: \a OR b = NOT (NOT a AND NOT b)\ by simp lemma not_add_distrib: \NOT (a + b) = NOT a - b\ by (simp add: not_eq_complement algebra_simps) lemma not_diff_distrib: \NOT (a - b) = NOT a + b\ using not_add_distrib [of a \- b\] by simp lemma and_eq_minus_1_iff: \a AND b = - 1 \ a = - 1 \ b = - 1\ by (auto simp: bit_eq_iff bit_simps) lemma disjunctive_diff: \a - b = a AND NOT b\ if \\n. bit b n \ bit a n\ proof - have \NOT a + b = NOT a OR b\ by (rule disjunctive_add) (auto simp add: bit_not_iff dest: that) then have \NOT (NOT a + b) = NOT (NOT a OR b)\ by simp then show ?thesis by (simp add: not_add_distrib) 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) apply (simp add: bit_not_iff bit_take_bit_iff bit_exp_iff) apply (use exp_eq_0_imp_not_bit in blast) done lemma take_bit_not_eq_mask_diff: \take_bit n (NOT a) = mask n - take_bit n a\ proof - have \take_bit n (NOT a) = take_bit n (NOT (take_bit n a))\ by (simp add: take_bit_not_take_bit) also have \\ = mask n AND NOT (take_bit n a)\ by (simp add: take_bit_eq_mask ac_simps) also have \\ = mask n - take_bit n a\ by (subst disjunctive_diff) (auto simp add: bit_take_bit_iff bit_mask_iff bit_imp_possible_bit) finally show ?thesis by simp qed lemma mask_eq_take_bit_minus_one: \mask n = take_bit n (- 1)\ by (simp add: bit_eq_iff bit_mask_iff bit_take_bit_iff conj_commute) lemma take_bit_minus_one_eq_mask [simp]: \take_bit n (- 1) = mask n\ by (simp add: mask_eq_take_bit_minus_one) lemma minus_exp_eq_not_mask: \- (2 ^ n) = NOT (mask n)\ by (rule bit_eqI) (simp add: bit_minus_iff bit_not_iff flip: mask_eq_exp_minus_1) lemma push_bit_minus_one_eq_not_mask [simp]: \push_bit n (- 1) = NOT (mask n)\ by (simp add: push_bit_eq_mult minus_exp_eq_not_mask) lemma take_bit_not_mask_eq_0: \take_bit m (NOT (mask n)) = 0\ if \n \ m\ by (rule bit_eqI) (use that in \simp add: bit_take_bit_iff bit_not_iff bit_mask_iff\) lemma unset_bit_eq_and_not: \unset_bit n a = a AND NOT (push_bit n 1)\ by (rule bit_eqI) (auto simp add: bit_simps) lemma push_bit_Suc_minus_numeral [simp]: \push_bit (Suc n) (- numeral k) = push_bit n (- numeral (Num.Bit0 k))\ apply (simp only: numeral_Bit0) apply simp apply (simp only: numeral_mult mult_2_right numeral_add) done lemma push_bit_minus_numeral [simp]: \push_bit (numeral l) (- numeral k) = push_bit (pred_numeral l) (- numeral (Num.Bit0 k))\ by (simp only: numeral_eq_Suc push_bit_Suc_minus_numeral) lemma take_bit_Suc_minus_1_eq: \take_bit (Suc n) (- 1) = 2 ^ Suc n - 1\ by (simp add: mask_eq_exp_minus_1) lemma take_bit_numeral_minus_1_eq: \take_bit (numeral k) (- 1) = 2 ^ numeral k - 1\ by (simp add: mask_eq_exp_minus_1) lemma push_bit_mask_eq: \push_bit m (mask n) = mask (n + m) AND NOT (mask m)\ apply (rule bit_eqI) apply (auto simp add: bit_simps not_less possible_bit_def) apply (drule sym [of 0]) apply (simp only:) using exp_not_zero_imp_exp_diff_not_zero apply (blast dest: exp_not_zero_imp_exp_diff_not_zero) done lemma slice_eq_mask: \push_bit n (take_bit m (drop_bit n a)) = a AND mask (m + n) AND NOT (mask n)\ by (rule bit_eqI) (auto simp add: bit_simps) lemma push_bit_numeral_minus_1 [simp]: \push_bit (numeral n) (- 1) = - (2 ^ numeral n)\ by (simp add: push_bit_eq_mult) -lemmas unset_bit_def = unset_bit_eq_and_not - end subsection \Instance \<^typ>\int\\ locale fold2_bit_int = fixes f :: \bool \ bool \ bool\ begin context begin function F :: \int \ int \ int\ where \F k l = (if k \ {0, - 1} \ l \ {0, - 1} then - of_bool (f (odd k) (odd l)) else of_bool (f (odd k) (odd l)) + 2 * (F (k div 2) (l div 2)))\ by auto private termination proof (relation \measure (\(k, l). nat (\k\ + \l\))\) have less_eq: \\k div 2\ \ \k\\ for k :: int by (cases k) (simp_all add: divide_int_def nat_add_distrib) then have less: \\k div 2\ < \k\\ if \k \ {0, - 1}\ for k :: int using that by (auto simp add: less_le [of k]) show \wf (measure (\(k, l). nat (\k\ + \l\)))\ by simp show \((k div 2, l div 2), k, l) \ measure (\(k, l). nat (\k\ + \l\))\ if \\ (k \ {0, - 1} \ l \ {0, - 1})\ for k l proof - from that have *: \k \ {0, - 1} \ l \ {0, - 1}\ by simp then have \0 < \k\ + \l\\ by auto moreover from * have \\k div 2\ + \l div 2\ < \k\ + \l\\ proof assume \k \ {0, - 1}\ then have \\k div 2\ < \k\\ by (rule less) with less_eq [of l] show ?thesis by auto next assume \l \ {0, - 1}\ then have \\l div 2\ < \l\\ by (rule less) with less_eq [of k] show ?thesis by auto qed ultimately show ?thesis by (simp only: in_measure split_def fst_conv snd_conv nat_mono_iff) qed qed declare F.simps [simp del] lemma rec: \F k l = of_bool (f (odd k) (odd l)) + 2 * (F (k div 2) (l div 2))\ for k l :: int proof (cases \k \ {0, - 1} \ l \ {0, - 1}\) case True then show ?thesis by (auto simp add: F.simps [of 0] F.simps [of \- 1\]) next case False then show ?thesis by (auto simp add: ac_simps F.simps [of k l]) qed lemma bit_iff: \bit (F k l) n \ f (bit k n) (bit l n)\ for k l :: int proof (induction n arbitrary: k l) case 0 then show ?case by (simp add: rec [of k l] bit_0) next case (Suc n) then show ?case by (simp add: rec [of k l] bit_Suc) qed end end instantiation int :: ring_bit_operations begin 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_int_iff: - \bit (NOT k) n \ \ bit k n\ - for k :: int - by (simp add: bit_not_int_iff' not_int_def) - global_interpretation and_int: fold2_bit_int \(\)\ defines and_int = and_int.F . global_interpretation or_int: fold2_bit_int \(\)\ defines or_int = or_int.F . global_interpretation xor_int: fold2_bit_int \(\)\ defines xor_int = xor_int.F . definition mask_int :: \nat \ int\ where \mask n = (2 :: int) ^ n - 1\ definition push_bit_int :: \nat \ int \ int\ where \push_bit_int n k = k * 2 ^ n\ definition drop_bit_int :: \nat \ int \ int\ where \drop_bit_int n k = k div 2 ^ n\ definition take_bit_int :: \nat \ int \ int\ where \take_bit_int n k = k mod 2 ^ n\ definition set_bit_int :: \nat \ int \ int\ where \set_bit n k = k OR push_bit n 1\ for k :: int definition unset_bit_int :: \nat \ int \ int\ where \unset_bit n k = k AND NOT (push_bit n 1)\ for k :: int definition flip_bit_int :: \nat \ int \ int\ where \flip_bit n k = k XOR push_bit n 1\ for k :: int +lemma bit_not_int_iff: + \bit (NOT k) n \ \ bit k n\ + for k :: int + by (simp add: bit_not_int_iff' not_int_def) + instance proof fix k l :: int and m n :: nat show \- k = NOT (k - 1)\ by (simp add: not_int_def) show \NOT k = of_bool (even k) + 2 * NOT (k div 2)\ by (auto simp add: not_int_def elim: oddE) show \unset_bit 0 k = 2 * (k div 2)\ by (rule bit_eqI) (auto simp add: unset_bit_int_def push_bit_int_def and_int.bit_iff bit_not_int_iff bit_simps simp flip: bit_Suc) show \unset_bit (Suc n) k = k mod 2 + 2 * unset_bit n (k div 2)\ by (rule bit_eqI) (auto simp add: unset_bit_int_def push_bit_int_def and_int.bit_iff bit_not_int_iff bit_simps mod_2_eq_odd even_bit_succ_iff bit_0 simp flip: bit_Suc) qed (fact and_int.rec or_int.rec xor_int.rec mask_int_def set_bit_int_def flip_bit_int_def push_bit_int_def drop_bit_int_def take_bit_int_def)+ end -lemmas and_int_rec = and_int.rec - -lemmas bit_and_int_iff = and_int.bit_iff - -lemmas or_int_rec = or_int.rec - -lemmas bit_or_int_iff = or_int.bit_iff - -lemmas xor_int_rec = xor_int.rec - -lemmas bit_xor_int_iff = xor_int.bit_iff - -lemma even_and_iff_int: - \even (k AND l) \ even k \ even l\ for k l :: int - by (fact even_and_iff) +lemma not_int_div_2: + \NOT k div 2 = NOT (k div 2)\ for k :: int + by (simp add: not_int_def) lemma bit_push_bit_iff_int: \bit (push_bit m k) n \ m \ n \ bit k (n - m)\ for k :: int by (auto simp add: bit_push_bit_iff) lemma take_bit_nonnegative [simp]: \take_bit n k \ 0\ for k :: int by (simp add: take_bit_eq_mod) lemma not_take_bit_negative [simp]: \\ take_bit n k < 0\ for k :: int by (simp add: not_less) lemma take_bit_int_less_exp [simp]: \take_bit n k < 2 ^ n\ for k :: int by (simp add: take_bit_eq_mod) lemma take_bit_int_eq_self_iff: \take_bit n k = k \ 0 \ k \ k < 2 ^ n\ (is \?P \ ?Q\) for k :: int proof assume ?P moreover note take_bit_int_less_exp [of n k] take_bit_nonnegative [of n k] ultimately show ?Q by simp next assume ?Q then show ?P by (simp add: take_bit_eq_mod) qed lemma take_bit_int_eq_self: \take_bit n k = k\ if \0 \ k\ \k < 2 ^ n\ for k :: int using that by (simp add: take_bit_int_eq_self_iff) lemma mask_half_int: \mask n div 2 = (mask (n - 1) :: int)\ by (cases n) (simp_all add: mask_eq_exp_minus_1 algebra_simps) lemma mask_nonnegative_int [simp]: \mask n \ (0::int)\ by (simp add: mask_eq_exp_minus_1) lemma not_mask_negative_int [simp]: \\ mask n < (0::int)\ by (simp add: not_less) lemma not_nonnegative_int_iff [simp]: \NOT k \ 0 \ k < 0\ for k :: int by (simp add: not_int_def) lemma not_negative_int_iff [simp]: \NOT k < 0 \ k \ 0\ for k :: int by (subst Not_eq_iff [symmetric]) (simp add: not_less not_le) lemma and_nonnegative_int_iff [simp]: \k AND l \ 0 \ k \ 0 \ l \ 0\ for k l :: int proof (induction k arbitrary: l 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 - using and_int_rec [of \k * 2\ l] + using and_int.rec [of \k * 2\ l] by (simp add: pos_imp_zdiv_nonneg_iff zero_le_mult_iff) next case (odd k) from odd have \0 \ k AND l div 2 \ 0 \ k \ 0 \ l div 2\ by simp then have \0 \ (1 + k * 2) div 2 AND l div 2 \ 0 \ (1 + k * 2) div 2 \ 0 \ l div 2\ by simp - with and_int_rec [of \1 + k * 2\ l] + with and_int.rec [of \1 + k * 2\ l] show ?case by (auto simp add: zero_le_mult_iff not_le) qed lemma and_negative_int_iff [simp]: \k AND l < 0 \ k < 0 \ l < 0\ for k l :: int by (subst Not_eq_iff [symmetric]) (simp add: not_less) lemma and_less_eq: \k AND l \ k\ if \l < 0\ for k l :: int using that proof (induction k arbitrary: l rule: int_bit_induct) case zero then show ?case by simp next case minus then show ?case by simp next case (even k) from even.IH [of \l div 2\] even.hyps even.prems show ?case - by (simp add: and_int_rec [of _ l]) + by (simp add: and_int.rec [of _ l]) next case (odd k) from odd.IH [of \l div 2\] odd.hyps odd.prems show ?case - by (simp add: and_int_rec [of _ l]) + by (simp add: and_int.rec [of _ l]) qed lemma or_nonnegative_int_iff [simp]: \k OR l \ 0 \ k \ 0 \ l \ 0\ for k l :: int by (simp only: or_eq_not_not_and not_nonnegative_int_iff) simp lemma or_negative_int_iff [simp]: \k OR l < 0 \ k < 0 \ l < 0\ for k l :: int by (subst Not_eq_iff [symmetric]) (simp add: not_less) lemma or_greater_eq: \k OR l \ k\ if \l \ 0\ for k l :: int using that proof (induction k arbitrary: l rule: int_bit_induct) case zero then show ?case by simp next case minus then show ?case by simp next case (even k) from even.IH [of \l div 2\] even.hyps even.prems show ?case - by (simp add: or_int_rec [of _ l]) + by (simp add: or_int.rec [of _ l]) next case (odd k) from odd.IH [of \l div 2\] odd.hyps odd.prems show ?case - by (simp add: or_int_rec [of _ l]) + by (simp add: or_int.rec [of _ l]) qed lemma xor_nonnegative_int_iff [simp]: \k XOR l \ 0 \ (k \ 0 \ l \ 0)\ for k l :: int by (simp only: bit.xor_def or_nonnegative_int_iff) auto lemma xor_negative_int_iff [simp]: \k XOR l < 0 \ (k < 0) \ (l < 0)\ for k l :: int by (subst Not_eq_iff [symmetric]) (auto simp add: not_less) lemma OR_upper: \<^marker>\contributor \Stefan Berghofer\\ \x OR y < 2 ^ n\ if \0 \ x\ \x < 2 ^ n\ \y < 2 ^ n\ for x y :: int using that proof (induction x arbitrary: y n rule: int_bit_induct) case zero then show ?case by simp next case minus then show ?case by simp next case (even x) from even.IH [of \n - 1\ \y div 2\] even.prems even.hyps - show ?case - by (cases n) (auto simp add: or_int_rec [of \_ * 2\] elim: oddE) + show ?case + by (cases n) (auto simp add: or_int.rec [of \_ * 2\] elim: oddE) next case (odd x) from odd.IH [of \n - 1\ \y div 2\] odd.prems odd.hyps show ?case - by (cases n) (auto simp add: or_int_rec [of \1 + _ * 2\], linarith) + by (cases n) (auto simp add: or_int.rec [of \1 + _ * 2\], linarith) qed lemma XOR_upper: \<^marker>\contributor \Stefan Berghofer\\ \x XOR y < 2 ^ n\ if \0 \ x\ \x < 2 ^ n\ \y < 2 ^ n\ for x y :: int using that proof (induction x arbitrary: y n rule: int_bit_induct) case zero then show ?case by simp next case minus then show ?case by simp next case (even x) from even.IH [of \n - 1\ \y div 2\] even.prems even.hyps - show ?case - by (cases n) (auto simp add: xor_int_rec [of \_ * 2\] elim: oddE) + show ?case + by (cases n) (auto simp add: xor_int.rec [of \_ * 2\] elim: oddE) next case (odd x) from odd.IH [of \n - 1\ \y div 2\] odd.prems odd.hyps show ?case - by (cases n) (auto simp add: xor_int_rec [of \1 + _ * 2\]) + by (cases n) (auto simp add: xor_int.rec [of \1 + _ * 2\]) qed lemma AND_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ \0 \ x AND y\ if \0 \ x\ for x y :: int using that by simp lemma OR_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ \0 \ x OR y\ if \0 \ x\ \0 \ y\ for x y :: int using that by simp lemma XOR_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ \0 \ x XOR y\ if \0 \ x\ \0 \ y\ for x y :: int using that by simp lemma AND_upper1 [simp]: \<^marker>\contributor \Stefan Berghofer\\ \x AND y \ x\ if \0 \ x\ for x y :: int using that proof (induction x arbitrary: y rule: int_bit_induct) case (odd k) then have \k AND y div 2 \ k\ by simp - then show ?case - by (simp add: and_int_rec [of \1 + _ * 2\]) -qed (simp_all add: and_int_rec [of \_ * 2\]) + then show ?case + by (simp add: and_int.rec [of \1 + _ * 2\]) +qed (simp_all add: and_int.rec [of \_ * 2\]) lemma AND_upper1' [simp]: \<^marker>\contributor \Stefan Berghofer\\ \y AND x \ z\ if \0 \ y\ \y \ z\ for x y z :: int using _ \y \ z\ by (rule order_trans) (use \0 \ y\ in simp) lemma AND_upper1'' [simp]: \<^marker>\contributor \Stefan Berghofer\\ \y AND x < z\ if \0 \ y\ \y < z\ for x y z :: int using _ \y < z\ by (rule order_le_less_trans) (use \0 \ y\ in simp) lemma AND_upper2 [simp]: \<^marker>\contributor \Stefan Berghofer\\ \x AND y \ y\ if \0 \ y\ for x y :: int using that AND_upper1 [of y x] by (simp add: ac_simps) lemma AND_upper2' [simp]: \<^marker>\contributor \Stefan Berghofer\\ \x AND y \ z\ if \0 \ y\ \y \ z\ for x y :: int using that AND_upper1' [of y z x] by (simp add: ac_simps) lemma AND_upper2'' [simp]: \<^marker>\contributor \Stefan Berghofer\\ \x AND y < z\ if \0 \ y\ \y < z\ for x y :: int using that AND_upper1'' [of y z x] by (simp add: ac_simps) lemma plus_and_or: \(x AND y) + (x OR y) = x + y\ for x y :: int proof (induction x arbitrary: y rule: int_bit_induct) case zero then show ?case by simp next case minus then show ?case by simp next case (even x) from even.IH [of \y div 2\] show ?case - by (auto simp add: and_int_rec [of _ y] or_int_rec [of _ y] elim: oddE) + by (auto simp add: and_int.rec [of _ y] or_int.rec [of _ y] elim: oddE) next case (odd x) from odd.IH [of \y div 2\] show ?case - by (auto simp add: and_int_rec [of _ y] or_int_rec [of _ y] elim: oddE) + by (auto simp add: and_int.rec [of _ y] or_int.rec [of _ y] elim: oddE) qed lemma push_bit_minus_one: \push_bit n (- 1 :: int) = - (2 ^ n)\ by (simp add: push_bit_eq_mult) lemma minus_1_div_exp_eq_int: \- 1 div (2 :: int) ^ n = - 1\ by (induction n) (use div_exp_eq [symmetric, of \- 1 :: int\ 1] in \simp_all add: ac_simps\) lemma drop_bit_minus_one [simp]: \drop_bit n (- 1 :: int) = - 1\ by (simp add: drop_bit_eq_div minus_1_div_exp_eq_int) lemma take_bit_Suc_from_most: \take_bit (Suc n) k = 2 ^ n * of_bool (bit k n) + take_bit n k\ for k :: int by (simp only: take_bit_eq_mod power_Suc2) (simp_all add: bit_iff_odd odd_iff_mod_2_eq_one zmod_zmult2_eq) lemma take_bit_minus: \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_diff: \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 bit_imp_take_bit_positive: \0 < take_bit m k\ if \n < m\ and \bit k n\ for k :: int proof (rule ccontr) assume \\ 0 < take_bit m k\ then have \take_bit m k = 0\ by (auto simp add: not_less intro: order_antisym) then have \bit (take_bit m k) n = bit 0 n\ by simp with that show False by (simp add: bit_take_bit_iff) qed lemma take_bit_mult: \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_mult_eq) lemma (in ring_1) of_nat_nat_take_bit_eq [simp]: \of_nat (nat (take_bit n k)) = of_int (take_bit n k)\ by simp lemma take_bit_minus_small_eq: \take_bit n (- k) = 2 ^ n - k\ if \0 < k\ \k \ 2 ^ n\ for k :: int proof - define m where \m = nat k\ with that have \k = int m\ and \0 < m\ and \m \ 2 ^ n\ by simp_all have \(2 ^ n - m) mod 2 ^ n = 2 ^ n - m\ using \0 < m\ by simp then have \int ((2 ^ n - m) mod 2 ^ n) = int (2 ^ n - m)\ by simp then have \(2 ^ n - int m) mod 2 ^ n = 2 ^ n - int m\ using \m \ 2 ^ n\ by (simp only: of_nat_mod of_nat_diff) simp with \k = int m\ have \(2 ^ n - k) mod 2 ^ n = 2 ^ n - k\ by simp then show ?thesis by (simp add: take_bit_eq_mod) qed lemma drop_bit_push_bit_int: \drop_bit m (push_bit n k) = drop_bit (m - n) (push_bit (n - m) k)\ for k :: int by (cases \m \ n\) (auto simp add: mult.left_commute [of _ \2 ^ n\] mult.commute [of _ \2 ^ n\] mult.assoc mult.commute [of k] drop_bit_eq_div push_bit_eq_mult not_le power_add dest!: le_Suc_ex less_imp_Suc_add) lemma push_bit_nonnegative_int_iff [simp]: \push_bit n k \ 0 \ k \ 0\ for k :: int by (simp add: push_bit_eq_mult zero_le_mult_iff power_le_zero_eq) lemma push_bit_negative_int_iff [simp]: \push_bit n k < 0 \ k < 0\ for k :: int by (subst Not_eq_iff [symmetric]) (simp add: not_less) lemma drop_bit_nonnegative_int_iff [simp]: \drop_bit n k \ 0 \ k \ 0\ for k :: int by (induction n) (auto simp add: drop_bit_Suc drop_bit_half) lemma drop_bit_negative_int_iff [simp]: \drop_bit n k < 0 \ k < 0\ for k :: int by (subst Not_eq_iff [symmetric]) (simp add: not_less) lemma set_bit_nonnegative_int_iff [simp]: \set_bit n k \ 0 \ k \ 0\ for k :: int - by (simp add: set_bit_def) + by (simp add: set_bit_eq_or) lemma set_bit_negative_int_iff [simp]: \set_bit n k < 0 \ k < 0\ for k :: int - by (simp add: set_bit_def) + by (simp add: set_bit_eq_or) lemma unset_bit_nonnegative_int_iff [simp]: \unset_bit n k \ 0 \ k \ 0\ for k :: int - by (simp add: unset_bit_def) + by (simp add: unset_bit_eq_and_not) lemma unset_bit_negative_int_iff [simp]: \unset_bit n k < 0 \ k < 0\ for k :: int - by (simp add: unset_bit_def) + by (simp add: unset_bit_eq_and_not) lemma flip_bit_nonnegative_int_iff [simp]: \flip_bit n k \ 0 \ k \ 0\ for k :: int - by (simp add: flip_bit_def) + by (simp add: flip_bit_eq_xor) lemma flip_bit_negative_int_iff [simp]: \flip_bit n k < 0 \ k < 0\ for k :: int - by (simp add: flip_bit_def) + by (simp add: flip_bit_eq_xor) lemma set_bit_greater_eq: \set_bit n k \ k\ for k :: int - by (simp add: set_bit_def or_greater_eq) + by (simp add: set_bit_eq_or or_greater_eq) lemma unset_bit_less_eq: \unset_bit n k \ k\ for k :: int - by (simp add: unset_bit_def and_less_eq) + by (simp add: unset_bit_eq_and_not and_less_eq) lemma set_bit_eq: \set_bit n k = k + of_bool (\ bit k n) * 2 ^ n\ for k :: int -proof (rule bit_eqI) - fix m - show \bit (set_bit n k) m \ bit (k + of_bool (\ bit k n) * 2 ^ n) m\ - proof (cases \m = n\) - case True - then show ?thesis - apply (simp add: bit_set_bit_iff) - apply (simp add: bit_iff_odd div_plus_div_distrib_dvd_right) - done - next - case False - then show ?thesis - apply (clarsimp simp add: bit_set_bit_iff) - apply (subst disjunctive_add) - apply (clarsimp simp add: bit_exp_iff) - apply (clarsimp simp add: bit_or_iff bit_exp_iff) - done - qed +proof - + have \set_bit n k = k OR of_bool (\ bit k n) * 2 ^ n\ + by (rule bit_eqI) (auto simp add: bit_simps) + then show ?thesis + by (subst disjunctive_add) (auto simp add: bit_simps) qed lemma unset_bit_eq: \unset_bit n k = k - of_bool (bit k n) * 2 ^ n\ for k :: int -proof (rule bit_eqI) - fix m - show \bit (unset_bit n k) m \ bit (k - of_bool (bit k n) * 2 ^ n) m\ - proof (cases \m = n\) - case True - then show ?thesis - apply (simp add: bit_unset_bit_iff) - apply (simp add: bit_iff_odd) - using div_plus_div_distrib_dvd_right [of \2 ^ n\ \- (2 ^ n)\ k] - apply (simp add: dvd_neg_div) - done - next - case False - then show ?thesis - apply (clarsimp simp add: bit_unset_bit_iff) - apply (subst disjunctive_diff) - apply (clarsimp simp add: bit_exp_iff) - apply (clarsimp simp add: bit_and_iff bit_not_iff bit_exp_iff) - done - qed +proof - + have \unset_bit n k = k AND NOT (of_bool (bit k n) * 2 ^ n)\ + by (rule bit_eqI) (auto simp add: bit_simps) + then show ?thesis + by (subst disjunctive_diff) (auto simp add: bit_simps simp flip: push_bit_eq_mult) qed lemma and_int_unfold: \k AND l = (if k = 0 \ l = 0 then 0 else if k = - 1 then l else if l = - 1 then k else (k mod 2) * (l mod 2) + 2 * ((k div 2) AND (l div 2)))\ for k l :: int - by (auto simp add: and_int_rec [of k l] zmult_eq_1_iff elim: oddE) + by (auto simp add: and_int.rec [of k l] zmult_eq_1_iff elim: oddE) lemma or_int_unfold: \k OR l = (if k = - 1 \ l = - 1 then - 1 else if k = 0 then l else if l = 0 then k else max (k mod 2) (l mod 2) + 2 * ((k div 2) OR (l div 2)))\ for k l :: int - by (auto simp add: or_int_rec [of k l] elim: oddE) + by (auto simp add: or_int.rec [of k l] elim: oddE) lemma xor_int_unfold: \k XOR l = (if k = - 1 then NOT l else if l = - 1 then NOT k else if k = 0 then l else if l = 0 then k else \k mod 2 - l mod 2\ + 2 * ((k div 2) XOR (l div 2)))\ for k l :: int - by (auto simp add: xor_int_rec [of k l] not_int_def elim!: oddE) + by (auto simp add: xor_int.rec [of k l] not_int_def elim!: oddE) lemma bit_minus_int_iff: \bit (- k) n \ bit (NOT (k - 1)) n\ for k :: int by (simp add: bit_simps) lemma take_bit_incr_eq: \take_bit n (k + 1) = 1 + take_bit n k\ if \take_bit n k \ 2 ^ n - 1\ for k :: int proof - from that have \2 ^ n \ k mod 2 ^ n + 1\ by (simp add: take_bit_eq_mod) moreover have \k mod 2 ^ n < 2 ^ n\ by simp ultimately have *: \k mod 2 ^ n + 1 < 2 ^ n\ by linarith have \(k + 1) mod 2 ^ n = (k mod 2 ^ n + 1) mod 2 ^ n\ by (simp add: mod_simps) also have \\ = k mod 2 ^ n + 1\ using * by (simp add: zmod_trivial_iff) finally have \(k + 1) mod 2 ^ n = k mod 2 ^ n + 1\ . then show ?thesis by (simp add: take_bit_eq_mod) qed lemma take_bit_decr_eq: \take_bit n (k - 1) = take_bit n k - 1\ if \take_bit n k \ 0\ for k :: int proof - from that have \k mod 2 ^ n \ 0\ by (simp add: take_bit_eq_mod) moreover have \k mod 2 ^ n \ 0\ \k mod 2 ^ n < 2 ^ n\ by simp_all ultimately have *: \k mod 2 ^ n > 0\ by linarith have \(k - 1) mod 2 ^ n = (k mod 2 ^ n - 1) mod 2 ^ n\ by (simp add: mod_simps) also have \\ = k mod 2 ^ n - 1\ by (simp add: zmod_trivial_iff) (use \k mod 2 ^ n < 2 ^ n\ * in linarith) finally have \(k - 1) mod 2 ^ n = k mod 2 ^ n - 1\ . then show ?thesis by (simp add: take_bit_eq_mod) qed lemma take_bit_int_greater_eq: \k + 2 ^ n \ take_bit n k\ if \k < 0\ for k :: int proof - have \k + 2 ^ n \ take_bit n (k + 2 ^ n)\ proof (cases \k > - (2 ^ n)\) case False then have \k + 2 ^ n \ 0\ by simp also note take_bit_nonnegative finally show ?thesis . next case True with that have \0 \ k + 2 ^ n\ and \k + 2 ^ n < 2 ^ n\ by simp_all then show ?thesis by (simp only: take_bit_eq_mod mod_pos_pos_trivial) qed then show ?thesis by (simp add: take_bit_eq_mod) qed lemma take_bit_int_less_eq: \take_bit n k \ k - 2 ^ n\ if \2 ^ n \ k\ and \n > 0\ for k :: int using that zmod_le_nonneg_dividend [of \k - 2 ^ n\ \2 ^ n\] by (simp add: take_bit_eq_mod) lemma take_bit_int_less_eq_self_iff: \take_bit n k \ k \ 0 \ k\ (is \?P \ ?Q\) for k :: int proof assume ?P show ?Q proof (rule ccontr) assume \\ 0 \ k\ then have \k < 0\ by simp with \?P\ have \take_bit n k < 0\ by (rule le_less_trans) then show False by simp qed next assume ?Q then show ?P by (simp add: take_bit_eq_mod zmod_le_nonneg_dividend) qed lemma take_bit_int_less_self_iff: \take_bit n k < k \ 2 ^ n \ k\ for k :: int by (auto simp add: less_le take_bit_int_less_eq_self_iff take_bit_int_eq_self_iff intro: order_trans [of 0 \2 ^ n\ k]) lemma take_bit_int_greater_self_iff: \k < take_bit n k \ k < 0\ for k :: int using take_bit_int_less_eq_self_iff [of n k] by auto lemma take_bit_int_greater_eq_self_iff: \k \ take_bit n k \ k < 2 ^ n\ for k :: int by (auto simp add: le_less take_bit_int_greater_self_iff take_bit_int_eq_self_iff dest: sym not_sym intro: less_trans [of k 0 \2 ^ n\]) lemma not_exp_less_eq_0_int [simp]: \\ 2 ^ n \ (0::int)\ by (simp add: power_le_zero_eq) lemma int_bit_bound: fixes k :: int obtains n where \\m. n \ m \ bit k m \ bit k n\ and \n > 0 \ bit k (n - 1) \ bit k n\ proof - obtain q where *: \\m. q \ m \ bit k m \ bit k q\ proof (cases \k \ 0\) case True moreover from power_gt_expt [of 2 \nat k\] have \nat k < 2 ^ nat k\ by simp then have \int (nat k) < int (2 ^ nat k)\ by (simp only: of_nat_less_iff) ultimately have *: \k div 2 ^ nat k = 0\ by simp show thesis proof (rule that [of \nat k\]) fix m assume \nat k \ m\ then show \bit k m \ bit k (nat k)\ by (auto simp add: * bit_iff_odd power_add zdiv_zmult2_eq dest!: le_Suc_ex) qed next case False moreover from power_gt_expt [of 2 \nat (- k)\] have \nat (- k) < 2 ^ nat (- k)\ by simp then have \int (nat (- k)) < int (2 ^ nat (- k))\ by (simp only: of_nat_less_iff) ultimately have \- k div - (2 ^ nat (- k)) = - 1\ by (subst div_pos_neg_trivial) simp_all then have *: \k div 2 ^ nat (- k) = - 1\ by simp show thesis proof (rule that [of \nat (- k)\]) fix m assume \nat (- k) \ m\ then show \bit k m \ bit k (nat (- k))\ by (auto simp add: * bit_iff_odd power_add zdiv_zmult2_eq minus_1_div_exp_eq_int dest!: le_Suc_ex) qed qed show thesis proof (cases \\m. bit k m \ bit k q\) case True then have \bit k 0 \ bit k q\ by blast with True that [of 0] show thesis by simp next case False then obtain r where **: \bit k r \ bit k q\ by blast have \r < q\ by (rule ccontr) (use * [of r] ** in simp) define N where \N = {n. n < q \ bit k n \ bit k q}\ moreover have \finite N\ \r \ N\ using ** N_def \r < q\ by auto moreover define n where \n = Suc (Max N)\ ultimately have \\m. n \ m \ bit k m \ bit k n\ apply auto apply (metis (full_types, lifting) "*" Max_ge_iff Suc_n_not_le_n \finite N\ all_not_in_conv mem_Collect_eq not_le) apply (metis "*" Max_ge Suc_n_not_le_n \finite N\ linorder_not_less mem_Collect_eq) apply (metis "*" Max_ge Suc_n_not_le_n \finite N\ linorder_not_less mem_Collect_eq) apply (metis (full_types, lifting) "*" Max_ge_iff Suc_n_not_le_n \finite N\ all_not_in_conv mem_Collect_eq not_le) done have \bit k (Max N) \ bit k n\ by (metis (mono_tags, lifting) "*" Max_in N_def \\m. n \ m \ bit k m = bit k n\ \finite N\ \r \ N\ empty_iff le_cases mem_Collect_eq) show thesis apply (rule that [of n]) using \\m. n \ m \ bit k m = bit k n\ apply blast using \bit k (Max N) \ bit k n\ n_def by auto qed qed lemma take_bit_tightened_less_eq_int: \take_bit m k \ take_bit n k\ if \m \ n\ for k :: int proof - have \take_bit m (take_bit n k) \ take_bit n k\ by (simp only: take_bit_int_less_eq_self_iff take_bit_nonnegative) with that show ?thesis by simp qed context ring_bit_operations begin lemma even_of_int_iff: \even (of_int k) \ even k\ by (induction k rule: int_bit_induct) simp_all lemma bit_of_int_iff [bit_simps]: \bit (of_int k) n \ possible_bit TYPE('a) n \ bit k n\ proof (cases \possible_bit TYPE('a) n\) case False then show ?thesis by (simp add: impossible_bit) next case True then have \bit (of_int k) n \ bit k n\ proof (induction k arbitrary: n 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 using bit_double_iff [of \of_int k\ n] Bit_Operations.bit_double_iff [of k n] by (cases n) (auto simp add: ac_simps possible_bit_def dest: mult_not_zero) next case (odd k) then show ?case using bit_double_iff [of \of_int k\ n] by (cases n) (auto simp add: ac_simps bit_double_iff even_bit_succ_iff Bit_Operations.bit_0 Bit_Operations.bit_Suc possible_bit_def dest: mult_not_zero) qed with True show ?thesis by simp qed lemma push_bit_of_int: \push_bit n (of_int k) = of_int (push_bit n k)\ by (simp add: push_bit_eq_mult Bit_Operations.push_bit_eq_mult) lemma of_int_push_bit: \of_int (push_bit n k) = push_bit n (of_int k)\ by (simp add: push_bit_eq_mult Bit_Operations.push_bit_eq_mult) lemma take_bit_of_int: \take_bit n (of_int k) = of_int (take_bit n k)\ by (rule bit_eqI) (simp add: bit_take_bit_iff Bit_Operations.bit_take_bit_iff bit_of_int_iff) lemma of_int_take_bit: \of_int (take_bit n k) = take_bit n (of_int k)\ by (rule bit_eqI) (simp add: bit_take_bit_iff Bit_Operations.bit_take_bit_iff bit_of_int_iff) lemma of_int_not_eq: \of_int (NOT k) = NOT (of_int k)\ by (rule bit_eqI) (simp add: bit_not_iff Bit_Operations.bit_not_iff bit_of_int_iff) lemma of_int_not_numeral: \of_int (NOT (numeral k)) = NOT (numeral k)\ by (simp add: local.of_int_not_eq) lemma of_int_and_eq: \of_int (k AND l) = of_int k AND of_int l\ by (rule bit_eqI) (simp add: bit_of_int_iff bit_and_iff Bit_Operations.bit_and_iff) lemma of_int_or_eq: \of_int (k OR l) = of_int k OR of_int l\ by (rule bit_eqI) (simp add: bit_of_int_iff bit_or_iff Bit_Operations.bit_or_iff) lemma of_int_xor_eq: \of_int (k XOR l) = of_int k XOR of_int l\ by (rule bit_eqI) (simp add: bit_of_int_iff bit_xor_iff Bit_Operations.bit_xor_iff) lemma of_int_mask_eq: \of_int (mask n) = mask n\ by (induction n) (simp_all add: mask_Suc_double Bit_Operations.mask_Suc_double of_int_or_eq) end subsection \Instance \<^typ>\nat\\ instantiation nat :: semiring_bit_operations begin definition and_nat :: \nat \ nat \ nat\ where \m AND n = nat (int m AND int n)\ for m n :: nat definition or_nat :: \nat \ nat \ nat\ where \m OR n = nat (int m OR int n)\ for m n :: nat definition xor_nat :: \nat \ nat \ nat\ where \m XOR n = nat (int m XOR int n)\ for m n :: nat definition mask_nat :: \nat \ nat\ where \mask n = (2 :: nat) ^ n - 1\ definition push_bit_nat :: \nat \ nat \ nat\ where \push_bit_nat n m = m * 2 ^ n\ definition drop_bit_nat :: \nat \ nat \ nat\ where \drop_bit_nat n m = m div 2 ^ n\ definition take_bit_nat :: \nat \ nat \ nat\ where \take_bit_nat n m = m mod 2 ^ n\ definition set_bit_nat :: \nat \ nat \ nat\ where \set_bit m n = n OR push_bit m 1\ for m n :: nat definition unset_bit_nat :: \nat \ nat \ nat\ where \unset_bit m n = nat (unset_bit m (int n))\ for m n :: nat definition flip_bit_nat :: \nat \ nat \ nat\ where \flip_bit m n = n XOR push_bit m 1\ for m n :: nat instance proof fix m n :: nat show \m AND n = of_bool (odd m \ odd n) + 2 * (m div 2 AND n div 2)\ by (simp add: and_nat_def and_rec [of \int m\ \int n\] nat_add_distrib of_nat_div) show \m OR n = of_bool (odd m \ odd n) + 2 * (m div 2 OR n div 2)\ by (simp add: or_nat_def or_rec [of \int m\ \int n\] nat_add_distrib of_nat_div) show \m XOR n = of_bool (odd m \ odd n) + 2 * (m div 2 XOR n div 2)\ by (simp add: xor_nat_def xor_rec [of \int m\ \int n\] nat_add_distrib of_nat_div) show \unset_bit 0 n = 2 * (n div 2)\ by (simp add: unset_bit_nat_def nat_mult_distrib) show \unset_bit (Suc m) n = n mod 2 + 2 * unset_bit m (n div 2)\ by (simp add: unset_bit_nat_def unset_bit_Suc nat_add_distrib nat_mult_distrib nat_mod_distrib of_nat_div) qed (simp_all add: mask_nat_def set_bit_nat_def flip_bit_nat_def push_bit_nat_def drop_bit_nat_def take_bit_nat_def) end lemma take_bit_nat_less_exp [simp]: - \take_bit n m < 2 ^ n\ for n m :: nat + \take_bit n m < 2 ^ n\ for n m :: nat by (simp add: take_bit_eq_mod) lemma take_bit_nat_eq_self_iff: \take_bit n m = m \ m < 2 ^ n\ (is \?P \ ?Q\) for n m :: nat proof assume ?P moreover note take_bit_nat_less_exp [of n m] ultimately show ?Q by simp next assume ?Q then show ?P by (simp add: take_bit_eq_mod) qed lemma take_bit_nat_eq_self: \take_bit n m = m\ if \m < 2 ^ n\ for m n :: nat using that by (simp add: take_bit_nat_eq_self_iff) lemma take_bit_nat_less_eq_self [simp]: \take_bit n m \ m\ for n m :: nat by (simp add: take_bit_eq_mod) lemma take_bit_nat_less_self_iff: \take_bit n m < m \ 2 ^ n \ m\ (is \?P \ ?Q\) for m n :: nat proof assume ?P then have \take_bit n m \ m\ by simp then show \?Q\ by (simp add: take_bit_nat_eq_self_iff) next have \take_bit n m < 2 ^ n\ by (fact take_bit_nat_less_exp) also assume ?Q finally show ?P . qed lemma bit_push_bit_iff_nat: \bit (push_bit m q) n \ m \ n \ bit q (n - m)\ for q :: nat by (auto simp add: bit_push_bit_iff) lemma and_nat_rec: \m AND n = of_bool (odd m \ odd n) + 2 * ((m div 2) AND (n div 2))\ for m n :: nat - by (simp add: and_nat_def and_int_rec [of \int m\ \int n\] zdiv_int nat_add_distrib nat_mult_distrib) + by (simp add: and_nat_def and_int.rec [of \int m\ \int n\] zdiv_int nat_add_distrib nat_mult_distrib) lemma or_nat_rec: \m OR n = of_bool (odd m \ odd n) + 2 * ((m div 2) OR (n div 2))\ for m n :: nat - by (simp add: or_nat_def or_int_rec [of \int m\ \int n\] zdiv_int nat_add_distrib nat_mult_distrib) + by (simp add: or_nat_def or_int.rec [of \int m\ \int n\] zdiv_int nat_add_distrib nat_mult_distrib) lemma xor_nat_rec: \m XOR n = of_bool (odd m \ odd n) + 2 * ((m div 2) XOR (n div 2))\ for m n :: nat - by (simp add: xor_nat_def xor_int_rec [of \int m\ \int n\] zdiv_int nat_add_distrib nat_mult_distrib) + by (simp add: xor_nat_def xor_int.rec [of \int m\ \int n\] zdiv_int nat_add_distrib nat_mult_distrib) lemma Suc_0_and_eq [simp]: \Suc 0 AND n = n mod 2\ using one_and_eq [of n] by simp lemma and_Suc_0_eq [simp]: \n AND Suc 0 = n mod 2\ using and_one_eq [of n] by simp lemma Suc_0_or_eq: \Suc 0 OR n = n + of_bool (even n)\ using one_or_eq [of n] by simp lemma or_Suc_0_eq: \n OR Suc 0 = n + of_bool (even n)\ using or_one_eq [of n] by simp lemma Suc_0_xor_eq: \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: \n XOR Suc 0 = n + of_bool (even n) - of_bool (odd n)\ using xor_one_eq [of n] by simp lemma and_nat_unfold [code]: \m AND n = (if m = 0 \ n = 0 then 0 else (m mod 2) * (n mod 2) + 2 * ((m div 2) AND (n div 2)))\ for m n :: nat by (auto simp add: and_nat_rec [of m n] elim: oddE) lemma or_nat_unfold [code]: \m OR n = (if m = 0 then n else if n = 0 then m else max (m mod 2) (n mod 2) + 2 * ((m div 2) OR (n div 2)))\ for m n :: nat by (auto simp add: or_nat_rec [of m n] elim: oddE) lemma xor_nat_unfold [code]: \m XOR n = (if m = 0 then n else if n = 0 then m else (m mod 2 + n mod 2) mod 2 + 2 * ((m div 2) XOR (n div 2)))\ for m n :: nat by (auto simp add: xor_nat_rec [of m n] elim!: oddE) lemma [code]: \unset_bit 0 m = 2 * (m div 2)\ \unset_bit (Suc n) m = m mod 2 + 2 * unset_bit n (m div 2)\ for m n :: nat by (simp_all add: unset_bit_Suc) lemma push_bit_of_Suc_0 [simp]: \push_bit n (Suc 0) = 2 ^ n\ using push_bit_of_1 [where ?'a = nat] by simp lemma take_bit_of_Suc_0 [simp]: \take_bit n (Suc 0) = of_bool (0 < n)\ using take_bit_of_1 [where ?'a = nat] by simp lemma drop_bit_of_Suc_0 [simp]: \drop_bit n (Suc 0) = of_bool (n = 0)\ using drop_bit_of_1 [where ?'a = nat] by simp lemma Suc_mask_eq_exp: \Suc (mask n) = 2 ^ n\ by (simp add: mask_eq_exp_minus_1) lemma less_eq_mask: \n \ mask n\ by (simp add: mask_eq_exp_minus_1 le_diff_conv2) (metis Suc_mask_eq_exp diff_Suc_1 diff_le_diff_pow diff_zero le_refl not_less_eq_eq power_0) lemma less_mask: \n < mask n\ if \Suc 0 < n\ proof - define m where \m = n - 2\ with that have *: \n = m + 2\ by simp have \Suc (Suc (Suc m)) < 4 * 2 ^ m\ by (induction m) simp_all then have \Suc (m + 2) < Suc (mask (m + 2))\ by (simp add: Suc_mask_eq_exp) then have \m + 2 < mask (m + 2)\ by (simp add: less_le) with * show ?thesis by simp qed lemma mask_nat_less_exp [simp]: \(mask n :: nat) < 2 ^ n\ by (simp add: mask_eq_exp_minus_1) lemma mask_nat_positive_iff [simp]: \(0::nat) < mask n \ 0 < n\ proof (cases \n = 0\) case True then show ?thesis by simp next case False then have \0 < n\ by simp then have \(0::nat) < mask n\ using less_eq_mask [of n] by (rule order_less_le_trans) with \0 < n\ show ?thesis by simp qed lemma take_bit_tightened_less_eq_nat: \take_bit m q \ take_bit n q\ if \m \ n\ for q :: nat proof - have \take_bit m (take_bit n q) \ take_bit n q\ by (rule take_bit_nat_less_eq_self) with that show ?thesis by simp qed lemma push_bit_nat_eq: \push_bit n (nat k) = nat (push_bit n k)\ by (cases \k \ 0\) (simp_all add: push_bit_eq_mult nat_mult_distrib not_le mult_nonneg_nonpos2) lemma drop_bit_nat_eq: \drop_bit n (nat k) = nat (drop_bit n k)\ apply (cases \k \ 0\) apply (simp_all add: drop_bit_eq_div nat_div_distrib nat_power_eq not_le) apply (simp add: divide_int_def) done lemma take_bit_nat_eq: \take_bit n (nat k) = nat (take_bit n k)\ if \k \ 0\ using that by (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq) lemma nat_take_bit_eq: \nat (take_bit n k) = take_bit n (nat k)\ if \k \ 0\ using that by (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq) context semiring_bit_operations begin lemma push_bit_of_nat: \push_bit n (of_nat m) = of_nat (push_bit n m)\ by (simp add: push_bit_eq_mult Bit_Operations.push_bit_eq_mult) lemma of_nat_push_bit: \of_nat (push_bit m n) = push_bit m (of_nat n)\ by (simp add: push_bit_eq_mult Bit_Operations.push_bit_eq_mult) lemma take_bit_of_nat: \take_bit n (of_nat m) = of_nat (take_bit n m)\ by (rule bit_eqI) (simp add: bit_take_bit_iff Bit_Operations.bit_take_bit_iff bit_of_nat_iff) lemma of_nat_take_bit: \of_nat (take_bit n m) = take_bit n (of_nat m)\ by (rule bit_eqI) (simp add: bit_take_bit_iff Bit_Operations.bit_take_bit_iff bit_of_nat_iff) lemma of_nat_and_eq: \of_nat (m AND n) = of_nat m AND of_nat n\ by (rule bit_eqI) (simp add: bit_of_nat_iff bit_and_iff Bit_Operations.bit_and_iff) lemma of_nat_or_eq: \of_nat (m OR n) = of_nat m OR of_nat n\ by (rule bit_eqI) (simp add: bit_of_nat_iff bit_or_iff Bit_Operations.bit_or_iff) lemma of_nat_xor_eq: \of_nat (m XOR n) = of_nat m XOR of_nat n\ by (rule bit_eqI) (simp add: bit_of_nat_iff bit_xor_iff Bit_Operations.bit_xor_iff) lemma of_nat_mask_eq: \of_nat (mask n) = mask n\ by (induction n) (simp_all add: mask_Suc_double Bit_Operations.mask_Suc_double of_nat_or_eq) end lemma nat_mask_eq: \nat (mask n) = mask n\ by (simp add: nat_eq_iff of_nat_mask_eq) subsection \Common algebraic structure\ class linordered_euclidean_semiring_bit_operations = linordered_euclidean_semiring + semiring_bit_operations begin lemma possible_bit [simp]: \possible_bit TYPE('a) n\ by (simp add: possible_bit_def) lemma take_bit_of_exp [simp]: \take_bit m (2 ^ n) = of_bool (n < m) * 2 ^ n\ by (simp add: take_bit_eq_mod exp_mod_exp) lemma take_bit_of_2 [simp]: \take_bit n 2 = of_bool (2 \ n) * 2\ using take_bit_of_exp [of n 1] by simp lemma push_bit_eq_0_iff [simp]: \push_bit n a = 0 \ a = 0\ by (simp add: push_bit_eq_mult) lemma take_bit_add: \take_bit n (take_bit n a + take_bit n b) = take_bit n (a + b)\ by (simp add: take_bit_eq_mod mod_simps) lemma take_bit_of_1_eq_0_iff [simp]: \take_bit n 1 = 0 \ n = 0\ by (simp add: take_bit_eq_mod) lemma drop_bit_Suc_bit0 [simp]: \drop_bit (Suc n) (numeral (Num.Bit0 k)) = drop_bit n (numeral k)\ by (simp add: drop_bit_Suc numeral_Bit0_div_2) lemma drop_bit_Suc_bit1 [simp]: \drop_bit (Suc n) (numeral (Num.Bit1 k)) = drop_bit n (numeral k)\ by (simp add: drop_bit_Suc numeral_Bit1_div_2) lemma drop_bit_numeral_bit0 [simp]: \drop_bit (numeral l) (numeral (Num.Bit0 k)) = drop_bit (pred_numeral l) (numeral k)\ by (simp add: drop_bit_rec numeral_Bit0_div_2) lemma drop_bit_numeral_bit1 [simp]: \drop_bit (numeral l) (numeral (Num.Bit1 k)) = drop_bit (pred_numeral l) (numeral k)\ by (simp add: drop_bit_rec numeral_Bit1_div_2) lemma take_bit_Suc_1 [simp]: \take_bit (Suc n) 1 = 1\ by (simp add: take_bit_Suc) lemma take_bit_Suc_bit0: \take_bit (Suc n) (numeral (Num.Bit0 k)) = take_bit n (numeral k) * 2\ by (simp add: take_bit_Suc numeral_Bit0_div_2) lemma take_bit_Suc_bit1: \take_bit (Suc n) (numeral (Num.Bit1 k)) = take_bit n (numeral k) * 2 + 1\ by (simp add: take_bit_Suc numeral_Bit1_div_2 mod_2_eq_odd) lemma take_bit_numeral_1 [simp]: \take_bit (numeral l) 1 = 1\ by (simp add: take_bit_rec [of \numeral l\ 1]) lemma take_bit_numeral_bit0: \take_bit (numeral l) (numeral (Num.Bit0 k)) = take_bit (pred_numeral l) (numeral k) * 2\ by (simp add: take_bit_rec numeral_Bit0_div_2) lemma take_bit_numeral_bit1: \take_bit (numeral l) (numeral (Num.Bit1 k)) = take_bit (pred_numeral l) (numeral k) * 2 + 1\ by (simp add: take_bit_rec numeral_Bit1_div_2 mod_2_eq_odd) lemma bit_of_nat_iff_bit [bit_simps]: \bit (of_nat m) n \ bit m n\ proof - have \even (m div 2 ^ n) \ even (of_nat (m div 2 ^ n))\ by simp also have \of_nat (m div 2 ^ n) = of_nat m div of_nat (2 ^ n)\ by (simp add: of_nat_div) finally show ?thesis by (simp add: bit_iff_odd semiring_bits_class.bit_iff_odd) qed lemma drop_bit_mask_eq: \drop_bit m (mask n) = mask (n - m)\ by (rule bit_eqI) (auto simp add: bit_simps possible_bit_def) lemma drop_bit_of_nat: "drop_bit n (of_nat m) = of_nat (drop_bit n m)" by (simp add: drop_bit_eq_div Bit_Operations.drop_bit_eq_div of_nat_div [of m "2 ^ n"]) lemma of_nat_drop_bit: \of_nat (drop_bit m n) = drop_bit m (of_nat n)\ by (simp add: drop_bit_eq_div Bit_Operations.drop_bit_eq_div of_nat_div) end instance nat :: linordered_euclidean_semiring_bit_operations .. instance int :: linordered_euclidean_semiring_bit_operations .. subsection \Symbolic computations on numeral expressions\ context semiring_bits begin lemma not_bit_numeral_Bit0_0 [simp]: \\ bit (numeral (Num.Bit0 m)) 0\ by (simp add: bit_0) lemma bit_numeral_Bit1_0 [simp]: \bit (numeral (Num.Bit1 m)) 0\ by (simp add: bit_0) end context ring_bit_operations begin lemma not_bit_minus_numeral_Bit0_0 [simp]: \\ bit (- numeral (Num.Bit0 m)) 0\ by (simp add: bit_0) lemma bit_minus_numeral_Bit1_0 [simp]: \bit (- numeral (Num.Bit1 m)) 0\ by (simp add: bit_0) end context linordered_euclidean_semiring_bit_operations begin lemma bit_numeral_iff: \bit (numeral m) n \ bit (numeral m :: nat) n\ using bit_of_nat_iff_bit [of \numeral m\ n] by simp lemma bit_numeral_Bit0_Suc_iff [simp]: \bit (numeral (Num.Bit0 m)) (Suc n) \ bit (numeral m) n\ by (simp add: bit_Suc numeral_Bit0_div_2) lemma bit_numeral_Bit1_Suc_iff [simp]: \bit (numeral (Num.Bit1 m)) (Suc n) \ bit (numeral m) n\ by (simp add: bit_Suc numeral_Bit1_div_2) lemma bit_numeral_rec: \bit (numeral (Num.Bit0 w)) n \ (case n of 0 \ False | Suc m \ bit (numeral w) m)\ \bit (numeral (Num.Bit1 w)) n \ (case n of 0 \ True | Suc m \ bit (numeral w) m)\ by (cases n; simp add: bit_0)+ lemma bit_numeral_simps [simp]: \\ bit 1 (numeral n)\ \bit (numeral (Num.Bit0 w)) (numeral n) \ bit (numeral w) (pred_numeral n)\ \bit (numeral (Num.Bit1 w)) (numeral n) \ bit (numeral w) (pred_numeral n)\ by (simp_all add: bit_1_iff numeral_eq_Suc) lemma and_numerals [simp]: \1 AND numeral (Num.Bit0 y) = 0\ \1 AND numeral (Num.Bit1 y) = 1\ \numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = 2 * (numeral x AND numeral y)\ \numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = 2 * (numeral x AND numeral y)\ \numeral (Num.Bit0 x) AND 1 = 0\ \numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = 2 * (numeral x AND numeral y)\ \numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + 2 * (numeral x AND numeral y)\ \numeral (Num.Bit1 x) AND 1 = 1\ by (simp_all add: bit_eq_iff) (simp_all add: bit_0 bit_simps bit_Suc bit_numeral_rec split: nat.splits) fun and_num :: \num \ num \ num option\ \<^marker>\contributor \Andreas Lochbihler\\ where \and_num num.One num.One = Some num.One\ | \and_num num.One (num.Bit0 n) = None\ | \and_num num.One (num.Bit1 n) = Some num.One\ | \and_num (num.Bit0 m) num.One = None\ | \and_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (and_num m n)\ | \and_num (num.Bit0 m) (num.Bit1 n) = map_option num.Bit0 (and_num m n)\ | \and_num (num.Bit1 m) num.One = Some num.One\ | \and_num (num.Bit1 m) (num.Bit0 n) = map_option num.Bit0 (and_num m n)\ | \and_num (num.Bit1 m) (num.Bit1 n) = (case and_num m n of None \ Some num.One | Some n' \ Some (num.Bit1 n'))\ lemma numeral_and_num: \numeral m AND numeral n = (case and_num m n of None \ 0 | Some n' \ numeral n')\ by (induction m n rule: and_num.induct) (simp_all add: split: option.split) lemma and_num_eq_None_iff: \and_num m n = None \ numeral m AND numeral n = 0\ by (simp add: numeral_and_num split: option.split) lemma and_num_eq_Some_iff: \and_num m n = Some q \ numeral m AND numeral n = numeral q\ by (simp add: numeral_and_num split: option.split) lemma or_numerals [simp]: \1 OR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)\ \1 OR numeral (Num.Bit1 y) = numeral (Num.Bit1 y)\ \numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = 2 * (numeral x OR numeral y)\ \numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + 2 * (numeral x OR numeral y)\ \numeral (Num.Bit0 x) OR 1 = numeral (Num.Bit1 x)\ \numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + 2 * (numeral x OR numeral y)\ \numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + 2 * (numeral x OR numeral y)\ \numeral (Num.Bit1 x) OR 1 = numeral (Num.Bit1 x)\ by (simp_all add: bit_eq_iff) (simp_all add: bit_0 bit_simps bit_Suc bit_numeral_rec split: nat.splits) fun or_num :: \num \ num \ num\ \<^marker>\contributor \Andreas Lochbihler\\ where \or_num num.One num.One = num.One\ | \or_num num.One (num.Bit0 n) = num.Bit1 n\ | \or_num num.One (num.Bit1 n) = num.Bit1 n\ | \or_num (num.Bit0 m) num.One = num.Bit1 m\ | \or_num (num.Bit0 m) (num.Bit0 n) = num.Bit0 (or_num m n)\ | \or_num (num.Bit0 m) (num.Bit1 n) = num.Bit1 (or_num m n)\ | \or_num (num.Bit1 m) num.One = num.Bit1 m\ | \or_num (num.Bit1 m) (num.Bit0 n) = num.Bit1 (or_num m n)\ | \or_num (num.Bit1 m) (num.Bit1 n) = num.Bit1 (or_num m n)\ lemma numeral_or_num: \numeral m OR numeral n = numeral (or_num m n)\ by (induction m n rule: or_num.induct) simp_all lemma numeral_or_num_eq: \numeral (or_num m n) = numeral m OR numeral n\ by (simp add: numeral_or_num) lemma xor_numerals [simp]: \1 XOR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)\ \1 XOR numeral (Num.Bit1 y) = numeral (Num.Bit0 y)\ \numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = 2 * (numeral x XOR numeral y)\ \numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + 2 * (numeral x XOR numeral y)\ \numeral (Num.Bit0 x) XOR 1 = numeral (Num.Bit1 x)\ \numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + 2 * (numeral x XOR numeral y)\ \numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = 2 * (numeral x XOR numeral y)\ \numeral (Num.Bit1 x) XOR 1 = numeral (Num.Bit0 x)\ by (simp_all add: bit_eq_iff) (simp_all add: bit_0 bit_simps bit_Suc bit_numeral_rec split: nat.splits) fun xor_num :: \num \ num \ num option\ \<^marker>\contributor \Andreas Lochbihler\\ where \xor_num num.One num.One = None\ | \xor_num num.One (num.Bit0 n) = Some (num.Bit1 n)\ | \xor_num num.One (num.Bit1 n) = Some (num.Bit0 n)\ | \xor_num (num.Bit0 m) num.One = Some (num.Bit1 m)\ | \xor_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (xor_num m n)\ | \xor_num (num.Bit0 m) (num.Bit1 n) = Some (case xor_num m n of None \ num.One | Some n' \ num.Bit1 n')\ | \xor_num (num.Bit1 m) num.One = Some (num.Bit0 m)\ | \xor_num (num.Bit1 m) (num.Bit0 n) = Some (case xor_num m n of None \ num.One | Some n' \ num.Bit1 n')\ | \xor_num (num.Bit1 m) (num.Bit1 n) = map_option num.Bit0 (xor_num m n)\ lemma numeral_xor_num: \numeral m XOR numeral n = (case xor_num m n of None \ 0 | Some n' \ numeral n')\ by (induction m n rule: xor_num.induct) (simp_all split: option.split) lemma xor_num_eq_None_iff: \xor_num m n = None \ numeral m XOR numeral n = 0\ by (simp add: numeral_xor_num split: option.split) lemma xor_num_eq_Some_iff: \xor_num m n = Some q \ numeral m XOR numeral n = numeral q\ by (simp add: numeral_xor_num split: option.split) end lemma drop_bit_Suc_minus_bit0 [simp]: \drop_bit (Suc n) (- numeral (Num.Bit0 k)) = drop_bit n (- numeral k :: int)\ by (simp add: drop_bit_Suc numeral_Bit0_div_2) lemma drop_bit_Suc_minus_bit1 [simp]: \drop_bit (Suc n) (- numeral (Num.Bit1 k)) = drop_bit n (- numeral (Num.inc k) :: int)\ by (simp add: drop_bit_Suc numeral_Bit1_div_2 add_One) lemma drop_bit_numeral_minus_bit0 [simp]: \drop_bit (numeral l) (- numeral (Num.Bit0 k)) = drop_bit (pred_numeral l) (- numeral k :: int)\ by (simp add: numeral_eq_Suc numeral_Bit0_div_2) lemma drop_bit_numeral_minus_bit1 [simp]: \drop_bit (numeral l) (- numeral (Num.Bit1 k)) = drop_bit (pred_numeral l) (- numeral (Num.inc k) :: int)\ by (simp add: numeral_eq_Suc numeral_Bit1_div_2) lemma take_bit_Suc_minus_bit0: \take_bit (Suc n) (- numeral (Num.Bit0 k)) = take_bit n (- numeral k) * (2 :: int)\ by (simp add: take_bit_Suc numeral_Bit0_div_2) lemma take_bit_Suc_minus_bit1: \take_bit (Suc n) (- numeral (Num.Bit1 k)) = take_bit n (- numeral (Num.inc k)) * 2 + (1 :: int)\ by (simp add: take_bit_Suc numeral_Bit1_div_2 add_One) lemma take_bit_numeral_minus_bit0: \take_bit (numeral l) (- numeral (Num.Bit0 k)) = take_bit (pred_numeral l) (- numeral k) * (2 :: int)\ by (simp add: numeral_eq_Suc numeral_Bit0_div_2 take_bit_Suc_minus_bit0) lemma take_bit_numeral_minus_bit1: \take_bit (numeral l) (- numeral (Num.Bit1 k)) = take_bit (pred_numeral l) (- numeral (Num.inc k)) * 2 + (1 :: int)\ by (simp add: numeral_eq_Suc numeral_Bit1_div_2 take_bit_Suc_minus_bit1) lemma bit_Suc_0_iff [bit_simps]: \bit (Suc 0) n \ n = 0\ using bit_1_iff [of n, where ?'a = nat] by simp lemma and_nat_numerals [simp]: \Suc 0 AND numeral (Num.Bit0 y) = 0\ \Suc 0 AND numeral (Num.Bit1 y) = 1\ \numeral (Num.Bit0 x) AND Suc 0 = 0\ \numeral (Num.Bit1 x) AND Suc 0 = 1\ by (simp_all only: and_numerals flip: One_nat_def) lemma or_nat_numerals [simp]: \Suc 0 OR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)\ \Suc 0 OR numeral (Num.Bit1 y) = numeral (Num.Bit1 y)\ \numeral (Num.Bit0 x) OR Suc 0 = numeral (Num.Bit1 x)\ \numeral (Num.Bit1 x) OR Suc 0 = numeral (Num.Bit1 x)\ by (simp_all only: or_numerals flip: One_nat_def) lemma xor_nat_numerals [simp]: \Suc 0 XOR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)\ \Suc 0 XOR numeral (Num.Bit1 y) = numeral (Num.Bit0 y)\ \numeral (Num.Bit0 x) XOR Suc 0 = numeral (Num.Bit1 x)\ \numeral (Num.Bit1 x) XOR Suc 0 = numeral (Num.Bit0 x)\ by (simp_all only: xor_numerals flip: One_nat_def) context ring_bit_operations begin lemma minus_numeral_inc_eq: \- numeral (Num.inc n) = NOT (numeral n)\ by (simp add: not_eq_complement sub_inc_One_eq add_One) lemma sub_one_eq_not_neg: \Num.sub n num.One = NOT (- numeral n)\ by (simp add: not_eq_complement) lemma minus_numeral_eq_not_sub_one: \- numeral n = NOT (Num.sub n num.One)\ by (simp add: not_eq_complement) lemma not_numeral_eq [simp]: \NOT (numeral n) = - numeral (Num.inc n)\ by (simp add: minus_numeral_inc_eq) lemma not_minus_numeral_eq [simp]: \NOT (- numeral n) = Num.sub n num.One\ by (simp add: sub_one_eq_not_neg) lemma minus_not_numeral_eq [simp]: \- (NOT (numeral n)) = numeral (Num.inc n)\ by simp lemma not_numeral_BitM_eq: \NOT (numeral (Num.BitM n)) = - numeral (num.Bit0 n)\ - by (simp add: inc_BitM_eq) + by (simp add: inc_BitM_eq) lemma not_numeral_Bit0_eq: \NOT (numeral (Num.Bit0 n)) = - numeral (num.Bit1 n)\ by simp end lemma bit_minus_numeral_int [simp]: \bit (- numeral (num.Bit0 w) :: int) (numeral n) \ bit (- numeral w :: int) (pred_numeral n)\ \bit (- numeral (num.Bit1 w) :: int) (numeral n) \ \ bit (numeral w :: int) (pred_numeral n)\ by (simp_all add: bit_minus_iff bit_not_iff numeral_eq_Suc bit_Suc add_One sub_inc_One_eq) lemma bit_minus_numeral_Bit0_Suc_iff [simp]: \bit (- numeral (num.Bit0 w) :: int) (Suc n) \ bit (- numeral w :: int) n\ by (simp add: bit_Suc) lemma bit_minus_numeral_Bit1_Suc_iff [simp]: \bit (- numeral (num.Bit1 w) :: int) (Suc n) \ \ bit (numeral w :: int) n\ by (simp add: bit_Suc add_One flip: bit_not_int_iff) lemma and_not_numerals: \1 AND NOT 1 = (0 :: int)\ \1 AND NOT (numeral (Num.Bit0 n)) = (1 :: int)\ \1 AND NOT (numeral (Num.Bit1 n)) = (0 :: int)\ \numeral (Num.Bit0 m) AND NOT (1 :: int) = numeral (Num.Bit0 m)\ \numeral (Num.Bit0 m) AND NOT (numeral (Num.Bit0 n)) = (2 :: int) * (numeral m AND NOT (numeral n))\ \numeral (Num.Bit0 m) AND NOT (numeral (Num.Bit1 n)) = (2 :: int) * (numeral m AND NOT (numeral n))\ \numeral (Num.Bit1 m) AND NOT (1 :: int) = numeral (Num.Bit0 m)\ \numeral (Num.Bit1 m) AND NOT (numeral (Num.Bit0 n)) = 1 + (2 :: int) * (numeral m AND NOT (numeral n))\ \numeral (Num.Bit1 m) AND NOT (numeral (Num.Bit1 n)) = (2 :: int) * (numeral m AND NOT (numeral n))\ by (simp_all add: bit_eq_iff) (auto simp add: bit_0 bit_simps bit_Suc bit_numeral_rec BitM_inc_eq sub_inc_One_eq split: nat.split) fun and_not_num :: \num \ num \ num option\ \<^marker>\contributor \Andreas Lochbihler\\ where \and_not_num num.One num.One = None\ | \and_not_num num.One (num.Bit0 n) = Some num.One\ | \and_not_num num.One (num.Bit1 n) = None\ | \and_not_num (num.Bit0 m) num.One = Some (num.Bit0 m)\ | \and_not_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (and_not_num m n)\ | \and_not_num (num.Bit0 m) (num.Bit1 n) = map_option num.Bit0 (and_not_num m n)\ | \and_not_num (num.Bit1 m) num.One = Some (num.Bit0 m)\ | \and_not_num (num.Bit1 m) (num.Bit0 n) = (case and_not_num m n of None \ Some num.One | Some n' \ Some (num.Bit1 n'))\ | \and_not_num (num.Bit1 m) (num.Bit1 n) = map_option num.Bit0 (and_not_num m n)\ lemma int_numeral_and_not_num: \numeral m AND NOT (numeral n) = (case and_not_num m n of None \ 0 :: int | Some n' \ numeral n')\ by (induction m n rule: and_not_num.induct) (simp_all del: not_numeral_eq not_one_eq add: and_not_numerals split: option.splits) lemma int_numeral_not_and_num: \NOT (numeral m) AND numeral n = (case and_not_num n m of None \ 0 :: int | Some n' \ numeral n')\ using int_numeral_and_not_num [of n m] by (simp add: ac_simps) lemma and_not_num_eq_None_iff: \and_not_num m n = None \ numeral m AND NOT (numeral n) = (0 :: int)\ by (simp del: not_numeral_eq add: int_numeral_and_not_num split: option.split) lemma and_not_num_eq_Some_iff: \and_not_num m n = Some q \ numeral m AND NOT (numeral n) = (numeral q :: int)\ by (simp del: not_numeral_eq add: int_numeral_and_not_num split: option.split) lemma and_minus_numerals [simp]: \1 AND - (numeral (num.Bit0 n)) = (0::int)\ \1 AND - (numeral (num.Bit1 n)) = (1::int)\ \numeral m AND - (numeral (num.Bit0 n)) = (case and_not_num m (Num.BitM n) of None \ 0 :: int | Some n' \ numeral n')\ \numeral m AND - (numeral (num.Bit1 n)) = (case and_not_num m (Num.Bit0 n) of None \ 0 :: int | Some n' \ numeral n')\ \- (numeral (num.Bit0 n)) AND 1 = (0::int)\ \- (numeral (num.Bit1 n)) AND 1 = (1::int)\ \- (numeral (num.Bit0 n)) AND numeral m = (case and_not_num m (Num.BitM n) of None \ 0 :: int | Some n' \ numeral n')\ \- (numeral (num.Bit1 n)) AND numeral m = (case and_not_num m (Num.Bit0 n) of None \ 0 :: int | Some n' \ numeral n')\ by (simp_all del: not_numeral_eq add: ac_simps and_not_numerals one_and_eq not_numeral_BitM_eq not_numeral_Bit0_eq and_not_num_eq_None_iff and_not_num_eq_Some_iff split: option.split) lemma and_minus_minus_numerals [simp]: \- (numeral m :: int) AND - (numeral n :: int) = NOT ((numeral m - 1) OR (numeral n - 1))\ by (simp add: minus_numeral_eq_not_sub_one) lemma or_not_numerals: \1 OR NOT 1 = NOT (0 :: int)\ \1 OR NOT (numeral (Num.Bit0 n)) = NOT (numeral (Num.Bit0 n) :: int)\ \1 OR NOT (numeral (Num.Bit1 n)) = NOT (numeral (Num.Bit0 n) :: int)\ \numeral (Num.Bit0 m) OR NOT (1 :: int) = NOT (1 :: int)\ \numeral (Num.Bit0 m) OR NOT (numeral (Num.Bit0 n)) = 1 + (2 :: int) * (numeral m OR NOT (numeral n))\ \numeral (Num.Bit0 m) OR NOT (numeral (Num.Bit1 n)) = (2 :: int) * (numeral m OR NOT (numeral n))\ \numeral (Num.Bit1 m) OR NOT (1 :: int) = NOT (0 :: int)\ \numeral (Num.Bit1 m) OR NOT (numeral (Num.Bit0 n)) = 1 + (2 :: int) * (numeral m OR NOT (numeral n))\ \numeral (Num.Bit1 m) OR NOT (numeral (Num.Bit1 n)) = 1 + (2 :: int) * (numeral m OR NOT (numeral n))\ by (simp_all add: bit_eq_iff) (auto simp add: bit_0 bit_simps bit_Suc bit_numeral_rec sub_inc_One_eq split: nat.split) fun or_not_num_neg :: \num \ num \ num\ \<^marker>\contributor \Andreas Lochbihler\\ where \or_not_num_neg num.One num.One = num.One\ | \or_not_num_neg num.One (num.Bit0 m) = num.Bit1 m\ | \or_not_num_neg num.One (num.Bit1 m) = num.Bit1 m\ | \or_not_num_neg (num.Bit0 n) num.One = num.Bit0 num.One\ | \or_not_num_neg (num.Bit0 n) (num.Bit0 m) = Num.BitM (or_not_num_neg n m)\ | \or_not_num_neg (num.Bit0 n) (num.Bit1 m) = num.Bit0 (or_not_num_neg n m)\ | \or_not_num_neg (num.Bit1 n) num.One = num.One\ | \or_not_num_neg (num.Bit1 n) (num.Bit0 m) = Num.BitM (or_not_num_neg n m)\ | \or_not_num_neg (num.Bit1 n) (num.Bit1 m) = Num.BitM (or_not_num_neg n m)\ lemma int_numeral_or_not_num_neg: \numeral m OR NOT (numeral n :: int) = - numeral (or_not_num_neg m n)\ by (induction m n rule: or_not_num_neg.induct) (simp_all del: not_numeral_eq not_one_eq add: or_not_numerals, simp_all) lemma int_numeral_not_or_num_neg: \NOT (numeral m) OR (numeral n :: int) = - numeral (or_not_num_neg n m)\ using int_numeral_or_not_num_neg [of n m] by (simp add: ac_simps) lemma numeral_or_not_num_eq: \numeral (or_not_num_neg m n) = - (numeral m OR NOT (numeral n :: int))\ using int_numeral_or_not_num_neg [of m n] by simp lemma or_minus_numerals [simp]: \1 OR - (numeral (num.Bit0 n)) = - (numeral (or_not_num_neg num.One (Num.BitM n)) :: int)\ \1 OR - (numeral (num.Bit1 n)) = - (numeral (num.Bit1 n) :: int)\ \numeral m OR - (numeral (num.Bit0 n)) = - (numeral (or_not_num_neg m (Num.BitM n)) :: int)\ \numeral m OR - (numeral (num.Bit1 n)) = - (numeral (or_not_num_neg m (Num.Bit0 n)) :: int)\ \- (numeral (num.Bit0 n)) OR 1 = - (numeral (or_not_num_neg num.One (Num.BitM n)) :: int)\ \- (numeral (num.Bit1 n)) OR 1 = - (numeral (num.Bit1 n) :: int)\ \- (numeral (num.Bit0 n)) OR numeral m = - (numeral (or_not_num_neg m (Num.BitM n)) :: int)\ \- (numeral (num.Bit1 n)) OR numeral m = - (numeral (or_not_num_neg m (Num.Bit0 n)) :: int)\ by (simp_all only: or.commute [of _ 1] or.commute [of _ \numeral m\] minus_numeral_eq_not_sub_one or_not_numerals numeral_or_not_num_eq arith_simps minus_minus numeral_One) lemma or_minus_minus_numerals [simp]: \- (numeral m :: int) OR - (numeral n :: int) = NOT ((numeral m - 1) AND (numeral n - 1))\ by (simp add: minus_numeral_eq_not_sub_one) lemma xor_minus_numerals [simp]: \- numeral n XOR k = NOT (neg_numeral_class.sub n num.One XOR k)\ \k XOR - numeral n = NOT (k XOR (neg_numeral_class.sub n num.One))\ for k :: int by (simp_all add: minus_numeral_eq_not_sub_one) definition take_bit_num :: \nat \ num \ num option\ where \take_bit_num n m = (if take_bit n (numeral m :: nat) = 0 then None else Some (num_of_nat (take_bit n (numeral m :: nat))))\ lemma take_bit_num_simps: \take_bit_num 0 m = None\ \take_bit_num (Suc n) Num.One = Some Num.One\ \take_bit_num (Suc n) (Num.Bit0 m) = (case take_bit_num n m of None \ None | Some q \ Some (Num.Bit0 q))\ \take_bit_num (Suc n) (Num.Bit1 m) = Some (case take_bit_num n m of None \ Num.One | Some q \ Num.Bit1 q)\ \take_bit_num (numeral r) Num.One = Some Num.One\ \take_bit_num (numeral r) (Num.Bit0 m) = (case take_bit_num (pred_numeral r) m of None \ None | Some q \ Some (Num.Bit0 q))\ \take_bit_num (numeral r) (Num.Bit1 m) = Some (case take_bit_num (pred_numeral r) m of None \ Num.One | Some q \ Num.Bit1 q)\ by (auto simp add: take_bit_num_def ac_simps mult_2 num_of_nat_double take_bit_Suc_bit0 take_bit_Suc_bit1 take_bit_numeral_bit0 take_bit_numeral_bit1) lemma take_bit_num_code [code]: \ \Ocaml-style pattern matching is more robust wrt. different representations of \<^typ>\nat\\ \take_bit_num n m = (case (n, m) of (0, _) \ None | (Suc n, Num.One) \ Some Num.One | (Suc n, Num.Bit0 m) \ (case take_bit_num n m of None \ None | Some q \ Some (Num.Bit0 q)) | (Suc n, Num.Bit1 m) \ Some (case take_bit_num n m of None \ Num.One | Some q \ Num.Bit1 q))\ by (cases n; cases m) (simp_all add: take_bit_num_simps) context semiring_bit_operations begin lemma take_bit_num_eq_None_imp: \take_bit m (numeral n) = 0\ if \take_bit_num m n = None\ proof - from that have \take_bit m (numeral n :: nat) = 0\ by (simp add: take_bit_num_def split: if_splits) then have \of_nat (take_bit m (numeral n)) = of_nat 0\ by simp then show ?thesis by (simp add: of_nat_take_bit) qed - + lemma take_bit_num_eq_Some_imp: \take_bit m (numeral n) = numeral q\ if \take_bit_num m n = Some q\ proof - from that have \take_bit m (numeral n :: nat) = numeral q\ by (auto simp add: take_bit_num_def Num.numeral_num_of_nat_unfold split: if_splits) then have \of_nat (take_bit m (numeral n)) = of_nat (numeral q)\ by simp then show ?thesis by (simp add: of_nat_take_bit) qed lemma take_bit_numeral_numeral: \take_bit (numeral m) (numeral n) = (case take_bit_num (numeral m) n of None \ 0 | Some q \ numeral q)\ by (auto split: option.split dest: take_bit_num_eq_None_imp take_bit_num_eq_Some_imp) end lemma take_bit_numeral_minus_numeral_int: \take_bit (numeral m) (- numeral n :: int) = (case take_bit_num (numeral m) n of None \ 0 | Some q \ take_bit (numeral m) (2 ^ numeral m - numeral q))\ (is \?lhs = ?rhs\) proof (cases \take_bit_num (numeral m) n\) case None then show ?thesis by (auto dest: take_bit_num_eq_None_imp [where ?'a = int] simp add: take_bit_eq_0_iff) next case (Some q) then have q: \take_bit (numeral m) (numeral n :: int) = numeral q\ by (auto dest: take_bit_num_eq_Some_imp) let ?T = \take_bit (numeral m) :: int \ int\ have *: \?T (2 ^ numeral m) = ?T (?T 0)\ by (simp add: take_bit_eq_0_iff) have \?lhs = ?T (0 - numeral n)\ by simp also have \\ = ?T (?T (?T 0) - ?T (?T (numeral n)))\ by (simp only: take_bit_diff) also have \\ = ?T (2 ^ numeral m - ?T (numeral n))\ by (simp only: take_bit_diff flip: *) also have \\ = ?rhs\ by (simp add: q Some) finally show ?thesis . qed declare take_bit_num_simps [simp] take_bit_numeral_numeral [simp] take_bit_numeral_minus_numeral_int [simp] subsection \More properties\ lemma take_bit_eq_mask_iff: \take_bit n k = mask n \ take_bit n (k + 1) = 0\ (is \?P \ ?Q\) for k :: int proof assume ?P then have \take_bit n (take_bit n k + take_bit n 1) = 0\ by (simp add: mask_eq_exp_minus_1 take_bit_eq_0_iff) then show ?Q by (simp only: take_bit_add) next assume ?Q then have \take_bit n (k + 1) - 1 = - 1\ by simp then have \take_bit n (take_bit n (k + 1) - 1) = take_bit n (- 1)\ by simp moreover have \take_bit n (take_bit n (k + 1) - 1) = take_bit n k\ by (simp add: take_bit_eq_mod mod_simps) ultimately show ?P by simp qed lemma take_bit_eq_mask_iff_exp_dvd: \take_bit n k = mask n \ 2 ^ n dvd k + 1\ for k :: int by (simp add: take_bit_eq_mask_iff flip: take_bit_eq_0_iff) subsection \Bit concatenation\ definition concat_bit :: \nat \ int \ int \ int\ where \concat_bit n k l = take_bit n k OR push_bit n l\ lemma bit_concat_bit_iff [bit_simps]: \bit (concat_bit m k l) n \ n < m \ bit k n \ m \ n \ bit l (n - m)\ by (simp add: concat_bit_def bit_or_iff bit_and_iff bit_take_bit_iff bit_push_bit_iff ac_simps) lemma concat_bit_eq: \concat_bit n k l = take_bit n k + push_bit n l\ by (simp add: concat_bit_def take_bit_eq_mask bit_and_iff bit_mask_iff bit_push_bit_iff disjunctive_add) lemma concat_bit_0 [simp]: \concat_bit 0 k l = l\ by (simp add: concat_bit_def) lemma concat_bit_Suc: \concat_bit (Suc n) k l = k mod 2 + 2 * concat_bit n (k div 2) l\ by (simp add: concat_bit_eq take_bit_Suc push_bit_double) lemma concat_bit_of_zero_1 [simp]: \concat_bit n 0 l = push_bit n l\ by (simp add: concat_bit_def) lemma concat_bit_of_zero_2 [simp]: \concat_bit n k 0 = take_bit n k\ by (simp add: concat_bit_def take_bit_eq_mask) lemma concat_bit_nonnegative_iff [simp]: \concat_bit n k l \ 0 \ l \ 0\ by (simp add: concat_bit_def) lemma concat_bit_negative_iff [simp]: \concat_bit n k l < 0 \ l < 0\ by (simp add: concat_bit_def) lemma concat_bit_assoc: \concat_bit n k (concat_bit m l r) = concat_bit (m + n) (concat_bit n k l) r\ by (rule bit_eqI) (auto simp add: bit_concat_bit_iff ac_simps) lemma concat_bit_assoc_sym: \concat_bit m (concat_bit n k l) r = concat_bit (min m n) k (concat_bit (m - n) l r)\ by (rule bit_eqI) (auto simp add: bit_concat_bit_iff ac_simps min_def) lemma concat_bit_eq_iff: \concat_bit n k l = concat_bit n r s \ take_bit n k = take_bit n r \ l = s\ (is \?P \ ?Q\) proof assume ?Q then show ?P by (simp add: concat_bit_def) next assume ?P then have *: \bit (concat_bit n k l) m = bit (concat_bit n r s) m\ for m by (simp add: bit_eq_iff) have \take_bit n k = take_bit n r\ proof (rule bit_eqI) fix m from * [of m] show \bit (take_bit n k) m \ bit (take_bit n r) m\ by (auto simp add: bit_take_bit_iff bit_concat_bit_iff) qed moreover have \push_bit n l = push_bit n s\ proof (rule bit_eqI) fix m from * [of m] show \bit (push_bit n l) m \ bit (push_bit n s) m\ by (auto simp add: bit_push_bit_iff bit_concat_bit_iff) qed then have \l = s\ by (simp add: push_bit_eq_mult) ultimately show ?Q by (simp add: concat_bit_def) qed lemma take_bit_concat_bit_eq: \take_bit m (concat_bit n k l) = concat_bit (min m n) k (take_bit (m - n) l)\ by (rule bit_eqI) - (auto simp add: bit_take_bit_iff bit_concat_bit_iff min_def) + (auto simp add: bit_take_bit_iff bit_concat_bit_iff min_def) lemma concat_bit_take_bit_eq: \concat_bit n (take_bit n b) = concat_bit n b\ by (simp add: concat_bit_def [abs_def]) subsection \Taking bits with sign propagation\ context ring_bit_operations begin definition signed_take_bit :: \nat \ 'a \ 'a\ where \signed_take_bit n a = take_bit n a OR (of_bool (bit a n) * NOT (mask n))\ lemma signed_take_bit_eq_if_positive: \signed_take_bit n a = take_bit n a\ if \\ bit a n\ using that by (simp add: signed_take_bit_def) lemma signed_take_bit_eq_if_negative: \signed_take_bit n a = take_bit n a OR NOT (mask n)\ if \bit a n\ using that by (simp add: signed_take_bit_def) lemma even_signed_take_bit_iff: \even (signed_take_bit m a) \ even a\ by (auto simp add: bit_0 signed_take_bit_def even_or_iff even_mask_iff bit_double_iff) lemma bit_signed_take_bit_iff [bit_simps]: \bit (signed_take_bit m a) n \ possible_bit TYPE('a) n \ bit a (min m n)\ by (simp add: signed_take_bit_def bit_take_bit_iff bit_or_iff bit_not_iff bit_mask_iff min_def not_le) (blast dest: bit_imp_possible_bit) lemma signed_take_bit_0 [simp]: \signed_take_bit 0 a = - (a mod 2)\ by (simp add: bit_0 signed_take_bit_def odd_iff_mod_2_eq_one) lemma signed_take_bit_Suc: \signed_take_bit (Suc n) a = a mod 2 + 2 * signed_take_bit n (a div 2)\ by (simp add: bit_eq_iff bit_sum_mult_2_cases bit_simps bit_0 possible_bit_less_imp flip: bit_Suc min_Suc_Suc) lemma signed_take_bit_of_0 [simp]: \signed_take_bit n 0 = 0\ by (simp add: signed_take_bit_def) lemma signed_take_bit_of_minus_1 [simp]: \signed_take_bit n (- 1) = - 1\ by (simp add: signed_take_bit_def mask_eq_exp_minus_1 possible_bit_def) lemma signed_take_bit_Suc_1 [simp]: \signed_take_bit (Suc n) 1 = 1\ by (simp add: signed_take_bit_Suc) lemma signed_take_bit_numeral_of_1 [simp]: \signed_take_bit (numeral k) 1 = 1\ by (simp add: bit_1_iff signed_take_bit_eq_if_positive) lemma signed_take_bit_rec: \signed_take_bit n a = (if n = 0 then - (a mod 2) else a mod 2 + 2 * signed_take_bit (n - 1) (a div 2))\ by (cases n) (simp_all add: signed_take_bit_Suc) lemma signed_take_bit_eq_iff_take_bit_eq: \signed_take_bit n a = signed_take_bit n b \ take_bit (Suc n) a = take_bit (Suc n) b\ proof - have \bit (signed_take_bit n a) = bit (signed_take_bit n b) \ bit (take_bit (Suc n) a) = bit (take_bit (Suc n) b)\ by (simp add: fun_eq_iff bit_signed_take_bit_iff bit_take_bit_iff not_le less_Suc_eq_le min_def) (use bit_imp_possible_bit in fastforce) then show ?thesis by (auto simp add: fun_eq_iff intro: bit_eqI) qed lemma signed_take_bit_signed_take_bit [simp]: \signed_take_bit m (signed_take_bit n a) = signed_take_bit (min m n) a\ by (auto simp add: bit_eq_iff bit_simps ac_simps) lemma signed_take_bit_take_bit: \signed_take_bit m (take_bit n a) = (if n \ m then take_bit n else signed_take_bit m) a\ by (rule bit_eqI) (auto simp add: bit_signed_take_bit_iff min_def bit_take_bit_iff) lemma take_bit_signed_take_bit: \take_bit m (signed_take_bit n a) = take_bit m a\ if \m \ Suc n\ using that by (rule le_SucE; intro bit_eqI) (auto simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def less_Suc_eq) end text \Modulus centered around 0\ lemma signed_take_bit_eq_concat_bit: \signed_take_bit n k = concat_bit n k (- of_bool (bit k n))\ by (simp add: concat_bit_def signed_take_bit_def) lemma signed_take_bit_add: \signed_take_bit n (signed_take_bit n k + signed_take_bit n l) = signed_take_bit n (k + l)\ for k l :: int proof - have \take_bit (Suc n) (take_bit (Suc n) (signed_take_bit n k) + take_bit (Suc n) (signed_take_bit n l)) = take_bit (Suc n) (k + l)\ by (simp add: take_bit_signed_take_bit take_bit_add) then show ?thesis by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_add) qed lemma signed_take_bit_diff: \signed_take_bit n (signed_take_bit n k - signed_take_bit n l) = signed_take_bit n (k - l)\ for k l :: int proof - have \take_bit (Suc n) (take_bit (Suc n) (signed_take_bit n k) - take_bit (Suc n) (signed_take_bit n l)) = take_bit (Suc n) (k - l)\ by (simp add: take_bit_signed_take_bit take_bit_diff) then show ?thesis by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_diff) qed lemma signed_take_bit_minus: \signed_take_bit n (- signed_take_bit n k) = signed_take_bit n (- k)\ for k :: int proof - have \take_bit (Suc n) (- take_bit (Suc n) (signed_take_bit n k)) = take_bit (Suc n) (- k)\ by (simp add: take_bit_signed_take_bit take_bit_minus) then show ?thesis by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_minus) qed lemma signed_take_bit_mult: \signed_take_bit n (signed_take_bit n k * signed_take_bit n l) = signed_take_bit n (k * l)\ for k l :: int proof - have \take_bit (Suc n) (take_bit (Suc n) (signed_take_bit n k) * take_bit (Suc n) (signed_take_bit n l)) = take_bit (Suc n) (k * l)\ by (simp add: take_bit_signed_take_bit take_bit_mult) then show ?thesis by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_mult) qed lemma signed_take_bit_eq_take_bit_minus: \signed_take_bit n k = take_bit (Suc n) k - 2 ^ Suc n * of_bool (bit k n)\ for k :: int proof (cases \bit k n\) case True have \signed_take_bit n k = take_bit (Suc n) k OR NOT (mask (Suc n))\ by (rule bit_eqI) (auto simp add: bit_signed_take_bit_iff min_def bit_take_bit_iff bit_or_iff bit_not_iff bit_mask_iff less_Suc_eq True) then have \signed_take_bit n k = take_bit (Suc n) k + NOT (mask (Suc n))\ by (simp add: disjunctive_add bit_take_bit_iff bit_not_iff bit_mask_iff) with True show ?thesis by (simp flip: minus_exp_eq_not_mask) next case False show ?thesis by (rule bit_eqI) (simp add: False bit_signed_take_bit_iff bit_take_bit_iff min_def less_Suc_eq) qed lemma signed_take_bit_eq_take_bit_shift: \signed_take_bit n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n\ for k :: int proof - have *: \take_bit n k OR 2 ^ n = take_bit n k + 2 ^ n\ by (simp add: disjunctive_add bit_exp_iff bit_take_bit_iff) have \take_bit n k - 2 ^ n = take_bit n k + NOT (mask n)\ by (simp add: minus_exp_eq_not_mask) also have \\ = take_bit n k OR NOT (mask n)\ by (rule disjunctive_add) (simp add: bit_exp_iff bit_take_bit_iff bit_not_iff bit_mask_iff) finally have **: \take_bit n k - 2 ^ n = take_bit n k OR NOT (mask n)\ . have \take_bit (Suc n) (k + 2 ^ n) = take_bit (Suc n) (take_bit (Suc n) k + take_bit (Suc n) (2 ^ n))\ by (simp only: take_bit_add) also have \take_bit (Suc n) k = 2 ^ n * of_bool (bit k n) + take_bit n k\ by (simp add: take_bit_Suc_from_most) finally have \take_bit (Suc n) (k + 2 ^ n) = take_bit (Suc n) (2 ^ (n + of_bool (bit k n)) + take_bit n k)\ by (simp add: ac_simps) also have \2 ^ (n + of_bool (bit k n)) + take_bit n k = 2 ^ (n + of_bool (bit k n)) OR take_bit n k\ by (rule disjunctive_add) (auto simp add: disjunctive_add bit_take_bit_iff bit_double_iff bit_exp_iff) finally show ?thesis using * ** by (simp add: signed_take_bit_def concat_bit_Suc min_def ac_simps) qed lemma signed_take_bit_nonnegative_iff [simp]: \0 \ signed_take_bit n k \ \ bit k n\ for k :: int by (simp add: signed_take_bit_def not_less concat_bit_def) lemma signed_take_bit_negative_iff [simp]: \signed_take_bit n k < 0 \ bit k n\ for k :: int by (simp add: signed_take_bit_def not_less concat_bit_def) lemma signed_take_bit_int_greater_eq_minus_exp [simp]: \- (2 ^ n) \ signed_take_bit n k\ for k :: int by (simp add: signed_take_bit_eq_take_bit_shift) lemma signed_take_bit_int_less_exp [simp]: \signed_take_bit n k < 2 ^ n\ for k :: int using take_bit_int_less_exp [of \Suc n\] by (simp add: signed_take_bit_eq_take_bit_shift) lemma signed_take_bit_int_eq_self_iff: \signed_take_bit n k = k \ - (2 ^ n) \ k \ k < 2 ^ n\ for k :: int by (auto simp add: signed_take_bit_eq_take_bit_shift take_bit_int_eq_self_iff algebra_simps) lemma signed_take_bit_int_eq_self: \signed_take_bit n k = k\ if \- (2 ^ n) \ k\ \k < 2 ^ n\ for k :: int using that by (simp add: signed_take_bit_int_eq_self_iff) lemma signed_take_bit_int_less_eq_self_iff: \signed_take_bit n k \ k \ - (2 ^ n) \ k\ for k :: int by (simp add: signed_take_bit_eq_take_bit_shift take_bit_int_less_eq_self_iff algebra_simps) linarith lemma signed_take_bit_int_less_self_iff: \signed_take_bit n k < k \ 2 ^ n \ k\ for k :: int by (simp add: signed_take_bit_eq_take_bit_shift take_bit_int_less_self_iff algebra_simps) lemma signed_take_bit_int_greater_self_iff: \k < signed_take_bit n k \ k < - (2 ^ n)\ for k :: int by (simp add: signed_take_bit_eq_take_bit_shift take_bit_int_greater_self_iff algebra_simps) linarith lemma signed_take_bit_int_greater_eq_self_iff: \k \ signed_take_bit n k \ k < 2 ^ n\ for k :: int by (simp add: signed_take_bit_eq_take_bit_shift take_bit_int_greater_eq_self_iff algebra_simps) lemma signed_take_bit_int_greater_eq: \k + 2 ^ Suc n \ signed_take_bit n k\ if \k < - (2 ^ n)\ for k :: int using that take_bit_int_greater_eq [of \k + 2 ^ n\ \Suc n\] by (simp add: signed_take_bit_eq_take_bit_shift) lemma signed_take_bit_int_less_eq: \signed_take_bit n k \ k - 2 ^ Suc n\ if \k \ 2 ^ n\ for k :: int using that take_bit_int_less_eq [of \Suc n\ \k + 2 ^ n\] by (simp add: signed_take_bit_eq_take_bit_shift) lemma signed_take_bit_Suc_bit0 [simp]: \signed_take_bit (Suc n) (numeral (Num.Bit0 k)) = signed_take_bit n (numeral k) * (2 :: int)\ by (simp add: signed_take_bit_Suc) lemma signed_take_bit_Suc_bit1 [simp]: \signed_take_bit (Suc n) (numeral (Num.Bit1 k)) = signed_take_bit n (numeral k) * 2 + (1 :: int)\ by (simp add: signed_take_bit_Suc) lemma signed_take_bit_Suc_minus_bit0 [simp]: \signed_take_bit (Suc n) (- numeral (Num.Bit0 k)) = signed_take_bit n (- numeral k) * (2 :: int)\ by (simp add: signed_take_bit_Suc) lemma signed_take_bit_Suc_minus_bit1 [simp]: \signed_take_bit (Suc n) (- numeral (Num.Bit1 k)) = signed_take_bit n (- numeral k - 1) * 2 + (1 :: int)\ by (simp add: signed_take_bit_Suc) lemma signed_take_bit_numeral_bit0 [simp]: \signed_take_bit (numeral l) (numeral (Num.Bit0 k)) = signed_take_bit (pred_numeral l) (numeral k) * (2 :: int)\ by (simp add: signed_take_bit_rec) lemma signed_take_bit_numeral_bit1 [simp]: \signed_take_bit (numeral l) (numeral (Num.Bit1 k)) = signed_take_bit (pred_numeral l) (numeral k) * 2 + (1 :: int)\ by (simp add: signed_take_bit_rec) lemma signed_take_bit_numeral_minus_bit0 [simp]: \signed_take_bit (numeral l) (- numeral (Num.Bit0 k)) = signed_take_bit (pred_numeral l) (- numeral k) * (2 :: int)\ by (simp add: signed_take_bit_rec) lemma signed_take_bit_numeral_minus_bit1 [simp]: \signed_take_bit (numeral l) (- numeral (Num.Bit1 k)) = signed_take_bit (pred_numeral l) (- numeral k - 1) * 2 + (1 :: int)\ by (simp add: signed_take_bit_rec) lemma signed_take_bit_code [code]: \signed_take_bit n a = (let l = take_bit (Suc n) a in if bit l n then l + push_bit (Suc n) (- 1) else l)\ proof - have *: \take_bit (Suc n) a + push_bit n (- 2) = take_bit (Suc n) a OR NOT (mask (Suc n))\ by (auto simp add: bit_take_bit_iff bit_push_bit_iff bit_not_iff bit_mask_iff disjunctive_add simp flip: push_bit_minus_one_eq_not_mask) show ?thesis by (rule bit_eqI) (auto simp add: Let_def * bit_signed_take_bit_iff bit_take_bit_iff min_def less_Suc_eq bit_not_iff bit_mask_iff bit_or_iff simp del: push_bit_minus_one_eq_not_mask) qed subsection \Symbolic computations for code generation\ lemma bit_int_code [code]: \bit (0::int) n \ False\ \bit (Int.Neg num.One) n \ True\ \bit (Int.Pos num.One) 0 \ True\ \bit (Int.Pos (num.Bit0 m)) 0 \ False\ \bit (Int.Pos (num.Bit1 m)) 0 \ True\ \bit (Int.Neg (num.Bit0 m)) 0 \ False\ \bit (Int.Neg (num.Bit1 m)) 0 \ True\ \bit (Int.Pos num.One) (Suc n) \ False\ \bit (Int.Pos (num.Bit0 m)) (Suc n) \ bit (Int.Pos m) n\ \bit (Int.Pos (num.Bit1 m)) (Suc n) \ bit (Int.Pos m) n\ \bit (Int.Neg (num.Bit0 m)) (Suc n) \ bit (Int.Neg m) n\ \bit (Int.Neg (num.Bit1 m)) (Suc n) \ bit (Int.Neg (Num.inc m)) n\ by (simp_all add: Num.add_One bit_0 bit_Suc) lemma not_int_code [code]: \NOT (0 :: int) = - 1\ \NOT (Int.Pos n) = Int.Neg (Num.inc n)\ \NOT (Int.Neg n) = Num.sub n num.One\ by (simp_all add: Num.add_One not_int_def) lemma and_int_code [code]: fixes i j :: int shows \0 AND j = 0\ \i AND 0 = 0\ \Int.Pos n AND Int.Pos m = (case and_num n m of None \ 0 | Some n' \ Int.Pos n')\ \Int.Neg n AND Int.Neg m = NOT (Num.sub n num.One OR Num.sub m num.One)\ \Int.Pos n AND Int.Neg num.One = Int.Pos n\ \Int.Pos n AND Int.Neg (num.Bit0 m) = Num.sub (or_not_num_neg (Num.BitM m) n) num.One\ \Int.Pos n AND Int.Neg (num.Bit1 m) = Num.sub (or_not_num_neg (num.Bit0 m) n) num.One\ \Int.Neg num.One AND Int.Pos m = Int.Pos m\ \Int.Neg (num.Bit0 n) AND Int.Pos m = Num.sub (or_not_num_neg (Num.BitM n) m) num.One\ \Int.Neg (num.Bit1 n) AND Int.Pos m = Num.sub (or_not_num_neg (num.Bit0 n) m) num.One\ apply (auto simp add: and_num_eq_None_iff [where ?'a = int] and_num_eq_Some_iff [where ?'a = int] split: option.split) apply (simp_all only: sub_one_eq_not_neg numeral_or_not_num_eq minus_minus and_not_numerals bit.de_Morgan_disj bit.double_compl and_not_num_eq_None_iff and_not_num_eq_Some_iff ac_simps) done lemma or_int_code [code]: fixes i j :: int shows \0 OR j = j\ \i OR 0 = i\ \Int.Pos n OR Int.Pos m = Int.Pos (or_num n m)\ \Int.Neg n OR Int.Neg m = NOT (Num.sub n num.One AND Num.sub m num.One)\ \Int.Pos n OR Int.Neg num.One = Int.Neg num.One\ \Int.Pos n OR Int.Neg (num.Bit0 m) = (case and_not_num (Num.BitM m) n of None \ -1 | Some n' \ Int.Neg (Num.inc n'))\ \Int.Pos n OR Int.Neg (num.Bit1 m) = (case and_not_num (num.Bit0 m) n of None \ -1 | Some n' \ Int.Neg (Num.inc n'))\ \Int.Neg num.One OR Int.Pos m = Int.Neg num.One\ \Int.Neg (num.Bit0 n) OR Int.Pos m = (case and_not_num (Num.BitM n) m of None \ -1 | Some n' \ Int.Neg (Num.inc n'))\ \Int.Neg (num.Bit1 n) OR Int.Pos m = (case and_not_num (num.Bit0 n) m of None \ -1 | Some n' \ Int.Neg (Num.inc n'))\ apply (auto simp add: numeral_or_num_eq split: option.splits) apply (simp_all only: and_not_num_eq_None_iff and_not_num_eq_Some_iff and_not_numerals numeral_or_not_num_eq or_eq_not_not_and bit.double_compl ac_simps flip: numeral_eq_iff [where ?'a = int]) apply simp_all done lemma xor_int_code [code]: fixes i j :: int shows \0 XOR j = j\ \i XOR 0 = i\ \Int.Pos n XOR Int.Pos m = (case xor_num n m of None \ 0 | Some n' \ Int.Pos n')\ \Int.Neg n XOR Int.Neg m = Num.sub n num.One XOR Num.sub m num.One\ \Int.Neg n XOR Int.Pos m = NOT (Num.sub n num.One XOR Int.Pos m)\ \Int.Pos n XOR Int.Neg m = NOT (Int.Pos n XOR Num.sub m num.One)\ by (simp_all add: xor_num_eq_None_iff [where ?'a = int] xor_num_eq_Some_iff [where ?'a = int] split: option.split) lemma push_bit_int_code [code]: \push_bit 0 i = i\ \push_bit (Suc n) i = push_bit n (Int.dup i)\ by (simp_all add: ac_simps) lemma drop_bit_int_code [code]: fixes i :: int shows \drop_bit 0 i = i\ \drop_bit (Suc n) 0 = (0 :: int)\ \drop_bit (Suc n) (Int.Pos num.One) = 0\ \drop_bit (Suc n) (Int.Pos (num.Bit0 m)) = drop_bit n (Int.Pos m)\ \drop_bit (Suc n) (Int.Pos (num.Bit1 m)) = drop_bit n (Int.Pos m)\ \drop_bit (Suc n) (Int.Neg num.One) = - 1\ \drop_bit (Suc n) (Int.Neg (num.Bit0 m)) = drop_bit n (Int.Neg m)\ \drop_bit (Suc n) (Int.Neg (num.Bit1 m)) = drop_bit n (Int.Neg (Num.inc m))\ by (simp_all add: drop_bit_Suc add_One) 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_iff_odd [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\: @{thm mask_eq_exp_minus_1 [where ?'a = int, no_vars]} \<^item> Left shift: @{thm push_bit_eq_mult [where ?'a = int, no_vars]} \<^item> Right shift: @{thm drop_bit_eq_div [where ?'a = int, no_vars]} \<^item> Truncation: @{thm take_bit_eq_mod [where ?'a = int, no_vars]} \<^item> 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]} + \<^item> Set a single bit: @{thm set_bit_eq_or [where ?'a = int, no_vars]} + + \<^item> Unset a single bit: @{thm unset_bit_eq_and_not [where ?'a = int, no_vars]} + + \<^item> Flip a single bit: @{thm flip_bit_eq_xor [where ?'a = int, no_vars]} \<^item> Signed truncation, or modulus centered around \<^term>\0::int\: @{thm signed_take_bit_def [no_vars]} \<^item> Bit concatenation: @{thm concat_bit_def [no_vars]} \<^item> (Bounded) conversion from and to a list of bits: @{thm horner_sum_bit_eq_take_bit [where ?'a = int, no_vars]} \ + +subsection \Lemma duplicates and other\ + +lemmas set_bit_def = set_bit_eq_or + +lemmas unset_bit_def = unset_bit_eq_and_not + +lemmas flip_bit_def = flip_bit_eq_xor + +lemmas and_int_rec = and_int.rec + +lemmas bit_and_int_iff = and_int.bit_iff + +lemmas or_int_rec = or_int.rec + +lemmas bit_or_int_iff = or_int.bit_iff + +lemmas xor_int_rec = xor_int.rec + +lemmas bit_xor_int_iff = xor_int.bit_iff + +lemma not_int_rec: + \NOT k = of_bool (even k) + 2 * NOT (k div 2)\ for k :: int + by (fact not_rec) + +lemma even_not_iff_int: + \even (NOT k) \ odd k\ for k :: int + by (fact even_not_iff) + +lemma even_and_iff_int: + \even (k AND l) \ even k \ even l\ for k l :: int + by (fact even_and_iff) + no_notation not (\NOT\) and "and" (infixr \AND\ 64) and or (infixr \OR\ 59) and xor (infixr \XOR\ 59) bundle bit_operations_syntax begin notation not (\NOT\) and "and" (infixr \AND\ 64) and or (infixr \OR\ 59) and xor (infixr \XOR\ 59) end end