diff -r fb9c119e5b49 -r d804e93ae9ff src/HOL/Bit_Operations.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bit_Operations.thy Mon Aug 02 10:01:06 2021 +0000 @@ -0,0 +1,3574 @@ +(* 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\] that + apply simp + by (metis local.add.left_commute local.add_left_cancel local.div_mult_mod_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 [simp]: + \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) + +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) + +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) + 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) + +lemma bit_eqI: + \a = b\ if \\n. 2 ^ n \ 0 \ 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) + 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 + 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 + 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. 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 \ 2 ^ m \ 0 \ m = n\ + by (auto simp add: bit_iff_odd exp_div_exp_eq) + +lemma bit_1_iff [bit_simps]: + \bit 1 n \ 1 \ 0 \ n = 0\ + using bit_exp_iff [of 0 n] by simp + +lemma bit_2_iff [bit_simps]: + \bit 2 n \ 2 \ 0 \ 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 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_double_iff [bit_simps]: + \bit (2 * a) n \ bit a (n - 1) \ n \ 0 \ 2 ^ n \ 0\ + using even_mult_exp_div_exp_iff [of a 1 n] + by (cases n, auto simp add: bit_iff_odd ac_simps) + +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 + 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_iff: + \bit (2 ^ m - 1) n \ 2 ^ n \ 0 \ n < m\ + by (simp add: bit_iff_odd even_mask_div_iff not_le) + +lemma bit_Numeral1_iff [simp]: + \bit (numeral Num.One) n \ n = 0\ + by (simp add: bit_rec) + +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 + next + case (Suc n) + from Suc.prems(1) [of 0] have even: \even a \ even b\ + by auto + 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) + 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 + +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 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 + +context semiring_bits +begin + +lemma bit_of_bool_iff [bit_simps]: + \bit (of_bool b) n \ b \ n = 0\ + by (simp add: bit_1_iff) + +lemma even_of_nat_iff: + \even (of_nat n) \ even n\ + by (induction n rule: nat_bit_induct) simp_all + +lemma bit_of_nat_iff [bit_simps]: + \bit (of_nat m) n \ (2::'a) ^ n \ 0 \ bit m n\ +proof (cases \(2::'a) ^ n = 0\) + case True + then show ?thesis + by (simp add: exp_eq_0_imp_not_bit) +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 dest: mult_not_zero) + next + case (odd m) + then show ?case + by (cases n) + (auto simp add: bit_double_iff even_bit_succ_iff Bit_Operations.bit_Suc dest: mult_not_zero) + qed + with False show ?thesis + by simp +qed + +end + +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 + +class semiring_bit_shifts = semiring_bits + + fixes push_bit :: \nat \ 'a \ 'a\ + assumes push_bit_eq_mult: \push_bit n a = a * 2 ^ n\ + fixes drop_bit :: \nat \ 'a \ 'a\ + assumes drop_bit_eq_div: \drop_bit n a = a div 2 ^ n\ + fixes take_bit :: \nat \ 'a \ 'a\ + assumes take_bit_eq_mod: \take_bit n a = a mod 2 ^ n\ +begin + +text \ + 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. + Having + them as definitional class operations + takes into account that specific instances of these can be implemented + differently wrt. code generation. +\ + +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: + "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 \ 2 ^ n \ 0 \ bit a (n - m)\ + by (auto simp add: bit_iff_odd push_bit_eq_mult even_mult_exp_div_exp_iff) + +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) + 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_iff 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\ +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) + 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 \ 2 ^ n \ 0) * 2 ^ (n - m)\ + by (rule bit_eqI) (auto simp add: bit_simps) + +end + +instantiation nat :: semiring_bit_shifts +begin + +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\ + +instance + by standard (simp_all add: push_bit_nat_def drop_bit_nat_def take_bit_nat_def) + +end + +context semiring_bit_shifts +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 semiring_bit_shifts_class.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 semiring_bit_shifts_class.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) + +end + +instantiation int :: semiring_bit_shifts +begin + +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\ + +instance + by standard (simp_all add: push_bit_int_def drop_bit_int_def take_bit_int_def) + +end + +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 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_nat_less_exp [simp]: + \take_bit n m < 2 ^ n\ for n m ::nat + by (simp add: take_bit_eq_mod) + +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_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_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 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 + +class unique_euclidean_semiring_with_bit_shifts = + unique_euclidean_semiring_with_nat + semiring_bit_shifts +begin + +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 take_bit_of_mask: + \take_bit m (2 ^ n - 1) = 2 ^ min m n - 1\ + by (simp add: take_bit_eq_mod mask_mod_exp) + +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 take_bit_Suc_1 [simp]: + \take_bit (Suc n) 1 = 1\ + by (simp add: take_bit_Suc) + +lemma take_bit_Suc_bit0 [simp]: + \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 [simp]: + \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 [simp]: + \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 [simp]: + \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 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 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 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 of_nat_drop_bit: + \of_nat (drop_bit m n) = drop_bit m (of_nat n)\ + by (simp add: drop_bit_eq_div semiring_bit_shifts_class.drop_bit_eq_div of_nat_div) + +lemma bit_push_bit_iff_of_nat_iff [bit_simps]: + \bit (push_bit m (of_nat r)) n \ m \ n \ bit (of_nat r) (n - m)\ + by (auto simp add: bit_push_bit_iff) + +lemma take_bit_sum: + "take_bit n a = (\k = 0..k = 0..k = Suc 0..k = Suc 0..k = 0..bit (numeral m :: int) n \ bit (numeral m :: nat) n\ + using bit_of_nat_iff_bit [of \numeral m\ n] by simp + +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 +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_minus_int_iff [bit_simps]: + \bit (- k) n \ \ bit (k - 1) n\ + for k :: int + using bit_not_int_iff' [of \k - 1\] by simp + +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 + +lemma bit_numeral_int_simps [simp]: + \bit (1 :: int) (numeral n) \ bit (0 :: int) (pred_numeral n)\ + \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)\ + \bit (numeral (Num.BitM w) :: int) (numeral n) \ \ bit (- numeral w :: int) (pred_numeral n)\ + \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)\ + \bit (- numeral (Num.BitM w) :: int) (numeral n) \ bit (- (numeral w) :: int) (pred_numeral n)\ + by (simp_all add: bit_1_iff numeral_eq_Suc bit_Suc add_One sub_inc_One_eq bit_minus_int_iff) + +lemma bit_numeral_Bit0_Suc_iff [simp]: + \bit (numeral (Num.Bit0 m) :: int) (Suc n) \ bit (numeral m :: int) n\ + by (simp add: bit_Suc) + +lemma bit_numeral_Bit1_Suc_iff [simp]: + \bit (numeral (Num.Bit1 m) :: int) (Suc n) \ bit (numeral m :: int) n\ + by (simp add: bit_Suc) + +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) + +lemma not_exp_less_eq_0_int [simp]: + \\ 2 ^ n \ (0::int)\ + by (simp add: power_le_zero_eq) + +lemma half_nonnegative_int_iff [simp]: + \k div 2 \ 0 \ k \ 0\ for k :: int +proof (cases \k \ 0\) + case True + then show ?thesis + by (auto simp add: divide_int_def sgn_1_pos) +next + case False + then show ?thesis + by (auto simp add: divide_int_def not_le elim!: evenE) +qed + +lemma half_negative_int_iff [simp]: + \k div 2 < 0 \ k < 0\ for k :: int + by (subst Not_eq_iff [symmetric]) (simp add: not_less) + +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 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) + +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) (simp_all 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) + + +subsection \Bit operations\ + +class semiring_bit_operations = semiring_bit_shifts + + fixes "and" :: \'a \ 'a \ 'a\ (infixr \AND\ 64) + and or :: \'a \ 'a \ 'a\ (infixr \OR\ 59) + and xor :: \'a \ 'a \ 'a\ (infixr \XOR\ 59) + and mask :: \nat \ 'a\ + and set_bit :: \nat \ 'a \ 'a\ + and unset_bit :: \nat \ 'a \ 'a\ + and flip_bit :: \nat \ 'a \ 'a\ + assumes bit_and_iff [bit_simps]: \bit (a AND b) n \ bit a n \ bit b n\ + and bit_or_iff [bit_simps]: \bit (a OR b) n \ bit a n \ bit b n\ + and bit_xor_iff [bit_simps]: \bit (a XOR b) n \ bit a n \ bit b n\ + 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 bit_unset_bit_iff [bit_simps]: \bit (unset_bit m a) n \ bit a n \ m \ n\ + and flip_bit_eq_xor: \flip_bit n a = a XOR push_bit n 1\ +begin + +text \ + We want the bitwise operations to bind slightly weaker + than \+\ and \-\. + For the sake of code generation + the operations \<^const>\and\, \<^const>\or\ and \<^const>\xor\ + are specified as definitional class operations. +\ + +sublocale "and": semilattice \(AND)\ + by standard (auto simp add: bit_eq_iff bit_and_iff) + +sublocale or: semilattice_neutr \(OR)\ 0 + by standard (auto simp add: bit_eq_iff bit_or_iff) + +sublocale xor: comm_monoid \(XOR)\ 0 + by standard (auto simp add: bit_eq_iff bit_xor_iff) + +lemma even_and_iff: + \even (a AND b) \ even a \ even b\ + using bit_and_iff [of a b 0] by auto + +lemma even_or_iff: + \even (a OR b) \ even a \ even b\ + using bit_or_iff [of a b 0] by auto + +lemma even_xor_iff: + \even (a XOR b) \ (even a \ even b)\ + using bit_xor_iff [of a b 0] by auto + +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) + +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) + +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 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 take_bit_and [simp]: + \take_bit n (a AND b) = take_bit n a AND take_bit n b\ + by (auto simp add: bit_eq_iff bit_take_bit_iff bit_and_iff) + +lemma take_bit_or [simp]: + \take_bit n (a OR b) = take_bit n a OR take_bit n b\ + by (auto simp add: bit_eq_iff bit_take_bit_iff bit_or_iff) + +lemma take_bit_xor [simp]: + \take_bit n (a XOR b) = take_bit n a XOR take_bit n b\ + by (auto simp add: bit_eq_iff bit_take_bit_iff bit_xor_iff) + +lemma push_bit_and [simp]: + \push_bit n (a AND b) = push_bit n a AND push_bit n b\ + by (rule bit_eqI) (auto simp add: bit_push_bit_iff bit_and_iff) + +lemma push_bit_or [simp]: + \push_bit n (a OR b) = push_bit n a OR push_bit n b\ + by (rule bit_eqI) (auto simp add: bit_push_bit_iff bit_or_iff) + +lemma push_bit_xor [simp]: + \push_bit n (a XOR b) = push_bit n a XOR push_bit n b\ + by (rule bit_eqI) (auto simp add: bit_push_bit_iff bit_xor_iff) + +lemma drop_bit_and [simp]: + \drop_bit n (a AND b) = drop_bit n a AND drop_bit n b\ + by (rule bit_eqI) (auto simp add: bit_drop_bit_eq bit_and_iff) + +lemma drop_bit_or [simp]: + \drop_bit n (a OR b) = drop_bit n a OR drop_bit n b\ + by (rule bit_eqI) (auto simp add: bit_drop_bit_eq bit_or_iff) + +lemma drop_bit_xor [simp]: + \drop_bit n (a XOR b) = drop_bit n a XOR drop_bit n b\ + by (rule bit_eqI) (auto simp add: bit_drop_bit_eq bit_xor_iff) + +lemma bit_mask_iff [bit_simps]: + \bit (mask m) n \ 2 ^ n \ 0 \ n < m\ + by (simp add: mask_eq_exp_minus_1 bit_mask_iff) + +lemma even_mask_iff: + \even (mask n) \ n = 0\ + using bit_mask_iff [of n 0] by auto + +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 (rule bit_eqI) + (auto simp add: bit_or_iff bit_mask_iff bit_exp_iff not_less le_less_Suc_eq) + +lemma mask_Suc_double: + \mask (Suc n) = 1 OR 2 * mask n\ +proof (rule bit_eqI) + fix q + assume \2 ^ q \ 0\ + show \bit (mask (Suc n)) q \ bit (1 OR 2 * mask n) q\ + by (cases q) + (simp_all add: even_mask_iff even_or_iff bit_or_iff bit_mask_iff bit_exp_iff bit_double_iff not_less le_less_Suc_eq bit_1_iff, auto simp add: mult_2) +qed + +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_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 (rule bit_eqI) + (auto simp add: bit_take_bit_iff bit_and_iff bit_mask_iff) + +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: push_bit_of_1 bit_eq_iff bit_and_iff bit_push_bit_iff exp_eq_0_imp_not_bit) + apply (simp_all add: bit_exp_iff) + done + +lemmas set_bit_def = set_bit_eq_or + +lemma bit_set_bit_iff [bit_simps]: + \bit (set_bit m a) n \ bit a n \ (m = n \ 2 ^ n \ 0)\ + by (auto simp add: set_bit_def push_bit_of_1 bit_or_iff bit_exp_iff) + +lemma even_set_bit_iff: + \even (set_bit m a) \ even a \ m \ 0\ + using bit_set_bit_iff [of m a 0] by auto + +lemma even_unset_bit_iff: + \even (unset_bit m a) \ even a \ m = 0\ + using bit_unset_bit_iff [of m a 0] by auto + +lemma and_exp_eq_0_iff_not_bit: + \a AND 2 ^ n = 0 \ \ bit a n\ (is \?P \ ?Q\) +proof + assume ?Q + then show ?P + by (auto intro: bit_eqI simp add: bit_simps) +next + assume ?P + show ?Q + proof (rule notI) + assume \bit a n\ + then have \a AND 2 ^ n = 2 ^ n\ + by (auto intro: bit_eqI simp add: bit_simps) + with \?P\ show False + using \bit a n\ exp_eq_0_imp_not_bit by auto + qed +qed + +lemmas flip_bit_def = flip_bit_eq_xor + +lemma bit_flip_bit_iff [bit_simps]: + \bit (flip_bit m a) n \ (m = n \ \ bit a n) \ 2 ^ n \ 0\ + by (auto simp add: flip_bit_def push_bit_of_1 bit_xor_iff bit_exp_iff exp_eq_0_imp_not_bit) + +lemma even_flip_bit_iff: + \even (flip_bit m a) \ \ (even a \ m = 0)\ + using bit_flip_bit_iff [of m a 0] by auto + +lemma set_bit_0 [simp]: + \set_bit 0 a = 1 + 2 * (a div 2)\ +proof (rule bit_eqI) + fix m + assume *: \2 ^ m \ 0\ + then show \bit (set_bit 0 a) m = bit (1 + 2 * (a div 2)) m\ + by (simp add: bit_set_bit_iff bit_double_iff even_bit_succ_iff) + (cases m, simp_all add: bit_Suc) +qed + +lemma set_bit_Suc: + \set_bit (Suc n) a = a mod 2 + 2 * set_bit n (a div 2)\ +proof (rule bit_eqI) + fix m + assume *: \2 ^ m \ 0\ + show \bit (set_bit (Suc n) a) m \ bit (a mod 2 + 2 * set_bit n (a div 2)) m\ + proof (cases m) + case 0 + then show ?thesis + by (simp add: even_set_bit_iff) + next + case (Suc m) + with * have \2 ^ m \ 0\ + using mult_2 by auto + show ?thesis + by (cases a rule: parity_cases) + (simp_all add: bit_set_bit_iff bit_double_iff even_bit_succ_iff *, + simp_all add: Suc \2 ^ m \ 0\ bit_Suc) + qed +qed + +lemma unset_bit_0 [simp]: + \unset_bit 0 a = 2 * (a div 2)\ +proof (rule bit_eqI) + fix m + assume *: \2 ^ m \ 0\ + then show \bit (unset_bit 0 a) m = bit (2 * (a div 2)) m\ + by (simp add: bit_unset_bit_iff bit_double_iff) + (cases m, simp_all add: bit_Suc) +qed + +lemma unset_bit_Suc: + \unset_bit (Suc n) a = a mod 2 + 2 * unset_bit n (a div 2)\ +proof (rule bit_eqI) + fix m + assume *: \2 ^ m \ 0\ + then show \bit (unset_bit (Suc n) a) m \ bit (a mod 2 + 2 * unset_bit n (a div 2)) m\ + proof (cases m) + case 0 + then show ?thesis + by (simp add: even_unset_bit_iff) + next + case (Suc m) + show ?thesis + by (cases a rule: parity_cases) + (simp_all add: bit_unset_bit_iff bit_double_iff even_bit_succ_iff *, + simp_all add: Suc bit_Suc) + qed +qed + +lemma flip_bit_0 [simp]: + \flip_bit 0 a = of_bool (even a) + 2 * (a div 2)\ +proof (rule bit_eqI) + fix m + assume *: \2 ^ m \ 0\ + then show \bit (flip_bit 0 a) m = bit (of_bool (even a) + 2 * (a div 2)) m\ + by (simp add: bit_flip_bit_iff bit_double_iff even_bit_succ_iff) + (cases m, simp_all add: bit_Suc) +qed + +lemma flip_bit_Suc: + \flip_bit (Suc n) a = a mod 2 + 2 * flip_bit n (a div 2)\ +proof (rule bit_eqI) + fix m + assume *: \2 ^ m \ 0\ + show \bit (flip_bit (Suc n) a) m \ bit (a mod 2 + 2 * flip_bit n (a div 2)) m\ + proof (cases m) + case 0 + then show ?thesis + by (simp add: even_flip_bit_iff) + next + case (Suc m) + with * have \2 ^ m \ 0\ + using mult_2 by auto + show ?thesis + by (cases a rule: parity_cases) + (simp_all add: bit_flip_bit_iff bit_double_iff even_bit_succ_iff, + simp_all add: Suc \2 ^ m \ 0\ bit_Suc) + qed +qed + +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) + + +end + +class ring_bit_operations = semiring_bit_operations + ring_parity + + fixes not :: \'a \ 'a\ (\NOT\) + assumes bit_not_iff [bit_simps]: \\n. bit (NOT a) n \ 2 ^ n \ 0 \ \ bit a n\ + assumes minus_eq_not_minus_1: \- a = NOT (a - 1)\ +begin + +text \ + For the sake of code generation \<^const>\not\ is specified as + definitional class operation. Note that \<^const>\not\ has no + sensible definition for unlimited but only positive bit strings + (type \<^typ>\nat\). +\ + +lemma bits_minus_1_mod_2_eq [simp]: + \(- 1) mod 2 = 1\ + by (simp add: mod_2_eq_odd) + +lemma not_eq_complement: + \NOT a = - a - 1\ + using minus_eq_not_minus_1 [of \a + 1\] by simp + +lemma minus_eq_not_plus_1: + \- a = NOT a + 1\ + using not_eq_complement [of a] by simp + +lemma bit_minus_iff [bit_simps]: + \bit (- a) n \ 2 ^ n \ 0 \ \ bit (a - 1) n\ + by (simp add: minus_eq_not_minus_1 bit_not_iff) + +lemma even_not_iff [simp]: + \even (NOT a) \ odd a\ + using bit_not_iff [of a 0] by auto + +lemma bit_not_exp_iff [bit_simps]: + \bit (NOT (2 ^ m)) n \ 2 ^ n \ 0 \ n \ m\ + by (auto simp add: bit_not_iff bit_exp_iff) + +lemma bit_minus_1_iff [simp]: + \bit (- 1) n \ 2 ^ n \ 0\ + by (simp add: bit_minus_iff) + +lemma bit_minus_exp_iff [bit_simps]: + \bit (- (2 ^ m)) n \ 2 ^ n \ 0 \ n \ m\ + by (auto simp add: bit_simps simp flip: mask_eq_exp_minus_1) + +lemma bit_minus_2_iff [simp]: + \bit (- 2) n \ 2 ^ n \ 0 \ n > 0\ + by (simp add: bit_minus_iff bit_1_iff) + +lemma not_one [simp]: + \NOT 1 = - 2\ + by (simp add: bit_eq_iff bit_not_iff) (simp add: bit_1_iff) + +sublocale "and": semilattice_neutr \(AND)\ \- 1\ + by standard (rule bit_eqI, simp add: bit_and_iff) + +sublocale bit: boolean_algebra \(AND)\ \(OR)\ NOT 0 \- 1\ + rewrites \bit.xor = (XOR)\ +proof - + interpret bit: boolean_algebra \(AND)\ \(OR)\ NOT 0 \- 1\ + by standard (auto simp add: bit_and_iff bit_or_iff bit_not_iff intro: bit_eqI) + show \boolean_algebra (AND) (OR) NOT 0 (- 1)\ + by standard + show \boolean_algebra.xor (AND) (OR) NOT = (XOR)\ + by (rule ext, rule ext, rule bit_eqI) + (auto simp add: bit.xor_def bit_and_iff bit_or_iff bit_xor_iff bit_not_iff) +qed + +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 (in ring_bit_operations) and_eq_minus_1_iff: + \a AND b = - 1 \ a = - 1 \ b = - 1\ +proof + assume \a = - 1 \ b = - 1\ + then show \a AND b = - 1\ + by simp +next + assume \a AND b = - 1\ + have *: \bit a n\ \bit b n\ if \2 ^ n \ 0\ for n + proof - + from \a AND b = - 1\ + have \bit (a AND b) n = bit (- 1) n\ + by (simp add: bit_eq_iff) + then show \bit a n\ \bit b n\ + using that by (simp_all add: bit_and_iff) + qed + have \a = - 1\ + by (rule bit_eqI) (simp add: *) + moreover have \b = - 1\ + by (rule bit_eqI) (simp add: *) + ultimately show \a = - 1 \ b = - 1\ + by simp +qed + +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 exp_eq_0_imp_not_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: + \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: + \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) + +lemmas unset_bit_def = unset_bit_eq_and_not + +end + + +subsection \Instance \<^typ>\int\\ + +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 + +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 (cases k) (simp_all add: not_int_def divide_int_def nat_add_distrib) + +lemma bit_not_int_iff [bit_simps]: + \bit (NOT k) n \ \ bit k n\ + for k :: int + by (simp add: bit_not_int_iff' not_int_def) + +function and_int :: \int \ int \ int\ + where \(k::int) AND l = (if k \ {0, - 1} \ l \ {0, - 1} + then - of_bool (odd k \ odd l) + else of_bool (odd k \ odd l) + 2 * ((k div 2) AND (l div 2)))\ + by auto + +termination proof (relation \measure (\(k, l). nat (\k\ + \l\))\) + 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 - + have less_eq: \\k div 2\ \ \k\\ for k :: int + by (cases k) (simp_all add: divide_int_def nat_add_distrib) + have less: \\k div 2\ < \k\\ if \k \ {0, - 1}\ for k :: int + proof (cases k) + case (nonneg n) + with that show ?thesis + by (simp add: int_div_less_self) + next + case (neg n) + with that have \n \ 0\ + by simp + then have \n div 2 < n\ + by (simp add: div_less_iff_less_mult) + with neg that show ?thesis + by (simp add: divide_int_def nat_add_distrib) + qed + 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 + qed +qed + +declare and_int.simps [simp del] + +lemma and_int_rec: + \k AND l = of_bool (odd k \ odd l) + 2 * ((k div 2) AND (l div 2))\ + for k l :: int +proof (cases \k \ {0, - 1} \ l \ {0, - 1}\) + case True + then show ?thesis + by auto (simp_all add: and_int.simps) +next + case False + then show ?thesis + by (auto simp add: ac_simps and_int.simps [of k l]) +qed + +lemma bit_and_int_iff: + \bit (k AND l) n \ bit k n \ bit l n\ for k l :: int +proof (induction n arbitrary: k l) + case 0 + then show ?case + by (simp add: and_int_rec [of k l]) +next + case (Suc n) + then show ?case + by (simp add: and_int_rec [of k l] bit_Suc) +qed + +lemma even_and_iff_int: + \even (k AND l) \ even k \ even l\ for k l :: int + using bit_and_int_iff [of k l 0] by auto + +definition or_int :: \int \ int \ int\ + where \k OR l = NOT (NOT k AND NOT l)\ for k l :: int + +lemma or_int_rec: + \k OR l = of_bool (odd k \ odd l) + 2 * ((k div 2) OR (l div 2))\ + for k l :: int + using and_int_rec [of \NOT k\ \NOT l\] + by (simp add: or_int_def even_not_iff_int not_int_div_2) + (simp_all add: not_int_def) + +lemma bit_or_int_iff: + \bit (k OR l) n \ bit k n \ bit l n\ for k l :: int + by (simp add: or_int_def bit_not_int_iff bit_and_int_iff) + +definition xor_int :: \int \ int \ int\ + where \k XOR l = k AND NOT l OR NOT k AND l\ for k l :: int + +lemma xor_int_rec: + \k XOR l = of_bool (odd k \ odd l) + 2 * ((k div 2) XOR (l div 2))\ + for k l :: int + by (simp add: xor_int_def or_int_rec [of \k AND NOT l\ \NOT k AND l\] even_and_iff_int even_not_iff_int) + (simp add: and_int_rec [of \NOT k\ \l\] and_int_rec [of \k\ \NOT l\] not_int_div_2) + +lemma bit_xor_int_iff: + \bit (k XOR l) n \ bit k n \ bit l n\ for k l :: int + by (auto simp add: xor_int_def bit_or_int_iff bit_and_int_iff bit_not_int_iff) + +definition mask_int :: \nat \ int\ + where \mask n = (2 :: int) ^ n - 1\ + +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 + +instance proof + fix k l :: int and m n :: nat + show \- k = NOT (k - 1)\ + by (simp add: not_int_def) + show \bit (k AND l) n \ bit k n \ bit l n\ + by (fact bit_and_int_iff) + show \bit (k OR l) n \ bit k n \ bit l n\ + by (fact bit_or_int_iff) + show \bit (k XOR l) n \ bit k n \ bit l n\ + by (fact bit_xor_int_iff) + show \bit (unset_bit m k) n \ bit k n \ m \ n\ + proof - + have \unset_bit m k = k AND NOT (push_bit m 1)\ + by (simp add: unset_bit_int_def) + also have \NOT (push_bit m 1 :: int) = - (push_bit m 1 + 1)\ + by (simp add: not_int_def) + finally show ?thesis by (simp only: bit_simps bit_and_int_iff) (auto simp add: bit_simps) + qed +qed (simp_all add: bit_not_int_iff mask_int_def set_bit_int_def flip_bit_int_def) + +end + +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] + 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] + 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]) +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]) +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]) +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]) +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\\ + fixes x y :: int + assumes \0 \ x\ \x < 2 ^ n\ \y < 2 ^ n\ + shows \x OR y < 2 ^ n\ +using assms 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) +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) +qed + +lemma XOR_upper: \<^marker>\contributor \Stefan Berghofer\\ + fixes x y :: int + assumes \0 \ x\ \x < 2 ^ n\ \y < 2 ^ n\ + shows \x XOR y < 2 ^ n\ +using assms 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) +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\]) +qed + +lemma AND_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ + fixes x y :: int + assumes \0 \ x\ + shows \0 \ x AND y\ + using assms by simp + +lemma OR_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ + fixes x y :: int + assumes \0 \ x\ \0 \ y\ + shows \0 \ x OR y\ + using assms by simp + +lemma XOR_lower [simp]: \<^marker>\contributor \Stefan Berghofer\\ + fixes x y :: int + assumes \0 \ x\ \0 \ y\ + shows \0 \ x XOR y\ + using assms by simp + +lemma AND_upper1 [simp]: \<^marker>\contributor \Stefan Berghofer\\ + fixes x y :: int + assumes \0 \ x\ + shows \x AND y \ x\ +using assms 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\]) + +lemmas AND_upper1' [simp] = order_trans [OF AND_upper1] \<^marker>\contributor \Stefan Berghofer\\ +lemmas AND_upper1'' [simp] = order_le_less_trans [OF AND_upper1] \<^marker>\contributor \Stefan Berghofer\\ + +lemma AND_upper2 [simp]: \<^marker>\contributor \Stefan Berghofer\\ + fixes x y :: int + assumes \0 \ y\ + shows \x AND y \ y\ + using assms AND_upper1 [of y x] by (simp add: ac_simps) + +lemmas AND_upper2' [simp] = order_trans [OF AND_upper2] \<^marker>\contributor \Stefan Berghofer\\ +lemmas AND_upper2'' [simp] = order_le_less_trans [OF AND_upper2] \<^marker>\contributor \Stefan Berghofer\\ + +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) +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) +qed + +lemma set_bit_nonnegative_int_iff [simp]: + \set_bit n k \ 0 \ k \ 0\ for k :: int + by (simp add: set_bit_def) + +lemma set_bit_negative_int_iff [simp]: + \set_bit n k < 0 \ k < 0\ for k :: int + by (simp add: set_bit_def) + +lemma unset_bit_nonnegative_int_iff [simp]: + \unset_bit n k \ 0 \ k \ 0\ for k :: int + by (simp add: unset_bit_def) + +lemma unset_bit_negative_int_iff [simp]: + \unset_bit n k < 0 \ k < 0\ for k :: int + by (simp add: unset_bit_def) + +lemma flip_bit_nonnegative_int_iff [simp]: + \flip_bit n k \ 0 \ k \ 0\ for k :: int + by (simp add: flip_bit_def) + +lemma flip_bit_negative_int_iff [simp]: + \flip_bit n k < 0 \ k < 0\ for k :: int + by (simp add: flip_bit_def) + +lemma set_bit_greater_eq: + \set_bit n k \ k\ for k :: int + by (simp add: set_bit_def 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) + +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 +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 +qed + +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) + 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 add: take_bit_minus_one_eq_mask) +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) + +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 \ (2::'a) ^ n \ 0 \ bit k n\ +proof (cases \(2::'a) ^ n = 0\) + case True + then show ?thesis + by (simp add: exp_eq_0_imp_not_bit) +next + case False + 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 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_Suc dest: mult_not_zero) + qed + with False 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 semiring_bit_shifts_class.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 semiring_bit_shifts_class.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_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 + +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 minus_numeral_inc_eq: + \- numeral (Num.inc n) = NOT (numeral n :: int)\ + by (simp add: not_int_def sub_inc_One_eq add_One) + +lemma sub_one_eq_not_neg: + \Num.sub n num.One = NOT (- numeral n :: int)\ + by (simp add: not_int_def) + +lemma int_not_numerals [simp]: + \NOT (numeral (Num.Bit0 n) :: int) = - numeral (Num.Bit1 n)\ + \NOT (numeral (Num.Bit1 n) :: int) = - numeral (Num.inc (num.Bit1 n))\ + \NOT (numeral (Num.BitM n) :: int) = - numeral (num.Bit0 n)\ + \NOT (- numeral (Num.Bit0 n) :: int) = numeral (Num.BitM n)\ + \NOT (- numeral (Num.Bit1 n) :: int) = numeral (Num.Bit0 n)\ + by (simp_all add: not_int_def add_One inc_BitM_eq) + +text \FIXME: The rule sets below are very large (24 rules for each + operator). Is there a simpler way to do this?\ + +context +begin + +private lemma eqI: + \k = l\ + if num: \\n. bit k (numeral n) \ bit l (numeral n)\ + and even: \even k \ even l\ + for k l :: int +proof (rule bit_eqI) + fix n + show \bit k n \ bit l n\ + proof (cases n) + case 0 + with even show ?thesis + by simp + next + case (Suc n) + with num [of \num_of_nat (Suc n)\] show ?thesis + by (simp only: numeral_num_of_nat) + qed +qed + +lemma int_and_numerals [simp]: + \numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND numeral y)\ + \numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (2 :: int) * (numeral x AND numeral y)\ + \numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND numeral y)\ + \numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x AND numeral y)\ + \numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND - numeral y)\ + \numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (2 :: int) * (numeral x AND - numeral (y + Num.One))\ + \numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND - numeral y)\ + \numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x AND - numeral (y + Num.One))\ + \- numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (2 :: int) * (- numeral x AND numeral y)\ + \- numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (2 :: int) * (- numeral x AND numeral y)\ + \- numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (2 :: int) * (- numeral (x + Num.One) AND numeral y)\ + \- numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) AND numeral y)\ + \- numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x AND - numeral y)\ + \- numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (2 :: int) * (- numeral x AND - numeral (y + Num.One))\ + \- numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (- numeral (x + Num.One) AND - numeral y)\ + \- numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) AND - numeral (y + Num.One))\ + \(1::int) AND numeral (Num.Bit0 y) = 0\ + \(1::int) AND numeral (Num.Bit1 y) = 1\ + \(1::int) AND - numeral (Num.Bit0 y) = 0\ + \(1::int) AND - numeral (Num.Bit1 y) = 1\ + \numeral (Num.Bit0 x) AND (1::int) = 0\ + \numeral (Num.Bit1 x) AND (1::int) = 1\ + \- numeral (Num.Bit0 x) AND (1::int) = 0\ + \- numeral (Num.Bit1 x) AND (1::int) = 1\ + by (auto simp add: bit_and_iff bit_minus_iff even_and_iff bit_double_iff even_bit_succ_iff add_One sub_inc_One_eq intro: eqI) + +lemma int_or_numerals [simp]: + \numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (2 :: int) * (numeral x OR numeral y)\ + \numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR numeral y)\ + \numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x OR numeral y)\ + \numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR numeral y)\ + \numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (2 :: int) * (numeral x OR - numeral y)\ + \numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR - numeral (y + Num.One))\ + \numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x OR - numeral y)\ + \numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR - numeral (y + Num.One))\ + \- numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (2 :: int) * (- numeral x OR numeral y)\ + \- numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x OR numeral y)\ + \- numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR numeral y)\ + \- numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR numeral y)\ + \- numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x OR - numeral y)\ + \- numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x OR - numeral (y + Num.One))\ + \- numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR - numeral y)\ + \- numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR - numeral (y + Num.One))\ + \(1::int) OR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)\ + \(1::int) OR numeral (Num.Bit1 y) = numeral (Num.Bit1 y)\ + \(1::int) OR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)\ + \(1::int) OR - numeral (Num.Bit1 y) = - numeral (Num.Bit1 y)\ + \numeral (Num.Bit0 x) OR (1::int) = numeral (Num.Bit1 x)\ + \numeral (Num.Bit1 x) OR (1::int) = numeral (Num.Bit1 x)\ + \- numeral (Num.Bit0 x) OR (1::int) = - numeral (Num.BitM x)\ + \- numeral (Num.Bit1 x) OR (1::int) = - numeral (Num.Bit1 x)\ + by (auto simp add: bit_or_iff bit_minus_iff even_or_iff bit_double_iff even_bit_succ_iff add_One sub_inc_One_eq sub_BitM_One_eq intro: eqI) + +lemma int_xor_numerals [simp]: + \numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (2 :: int) * (numeral x XOR numeral y)\ + \numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x XOR numeral y)\ + \numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x XOR numeral y)\ + \numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (2 :: int) * (numeral x XOR numeral y)\ + \numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (2 :: int) * (numeral x XOR - numeral y)\ + \numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x XOR - numeral (y + Num.One))\ + \numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x XOR - numeral y)\ + \numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (2 :: int) * (numeral x XOR - numeral (y + Num.One))\ + \- numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (2 :: int) * (- numeral x XOR numeral y)\ + \- numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x XOR numeral y)\ + \- numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) XOR numeral y)\ + \- numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (2 :: int) * (- numeral (x + Num.One) XOR numeral y)\ + \- numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x XOR - numeral y)\ + \- numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x XOR - numeral (y + Num.One))\ + \- numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) XOR - numeral y)\ + \- numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (2 :: int) * (- numeral (x + Num.One) XOR - numeral (y + Num.One))\ + \(1::int) XOR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)\ + \(1::int) XOR numeral (Num.Bit1 y) = numeral (Num.Bit0 y)\ + \(1::int) XOR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)\ + \(1::int) XOR - numeral (Num.Bit1 y) = - numeral (Num.Bit0 (y + Num.One))\ + \numeral (Num.Bit0 x) XOR (1::int) = numeral (Num.Bit1 x)\ + \numeral (Num.Bit1 x) XOR (1::int) = numeral (Num.Bit0 x)\ + \- numeral (Num.Bit0 x) XOR (1::int) = - numeral (Num.BitM x)\ + \- numeral (Num.Bit1 x) XOR (1::int) = - numeral (Num.Bit0 (x + Num.One))\ + by (auto simp add: bit_xor_iff bit_minus_iff even_xor_iff bit_double_iff even_bit_succ_iff add_One sub_inc_One_eq sub_BitM_One_eq intro: eqI) + +end + + +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) + +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: 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 \ 2 ^ n \ 0 \ 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) + (use exp_eq_0_imp_not_bit in blast) + +lemma signed_take_bit_0 [simp]: + \signed_take_bit 0 a = - (a mod 2)\ + by (simp add: 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)\ +proof (rule bit_eqI) + fix m + assume *: \2 ^ m \ 0\ + show \bit (signed_take_bit (Suc n) a) m \ + bit (a mod 2 + 2 * signed_take_bit n (a div 2)) m\ + proof (cases m) + case 0 + then show ?thesis + by (simp add: even_signed_take_bit_iff) + next + case (Suc m) + with * have \2 ^ m \ 0\ + by (metis mult_not_zero power_Suc) + with Suc show ?thesis + by (simp add: bit_signed_take_bit_iff mod2_eq_if bit_double_iff even_bit_succ_iff + ac_simps flip: bit_Suc) + qed +qed + +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 take_bit_minus_one_eq_mask mask_eq_exp_minus_1) + +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_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 exp_eq_0_imp_not_bit in fastforce) + then show ?thesis + by (simp add: bit_eq_iff fun_eq_iff) +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\ +proof (rule bit_eqI) + fix q + show \bit (signed_take_bit m (signed_take_bit n a)) q \ + bit (signed_take_bit (min m n) a) q\ + by (simp add: bit_signed_take_bit_iff min_def bit_or_iff bit_not_iff bit_mask_iff bit_take_bit_iff) + (use le_Suc_ex exp_add_not_zero_imp in blast) +qed + +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 push_bit_minus_one_eq_not_mask) + +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) +qed + + +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 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 = (if bit n m then n - push_bit m 1 else 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 q :: nat + show \bit (m AND n) q \ bit m q \ bit n q\ + by (simp add: and_nat_def bit_simps) + show \bit (m OR n) q \ bit m q \ bit n q\ + by (simp add: or_nat_def bit_simps) + show \bit (m XOR n) q \ bit m q \ bit n q\ + by (simp add: xor_nat_def bit_simps) + show \bit (unset_bit m n) q \ bit n q \ m \ q\ + proof (cases \bit n m\) + case False + then show ?thesis by (auto simp add: unset_bit_nat_def) + next + case True + have \push_bit m (drop_bit m n) + take_bit m n = n\ + by (fact bits_ident) + also from \bit n m\ have \drop_bit m n = 2 * drop_bit (Suc m) n + 1\ + by (simp add: drop_bit_Suc drop_bit_half even_drop_bit_iff_not_bit ac_simps) + finally have \push_bit m (2 * drop_bit (Suc m) n) + take_bit m n + push_bit m 1 = n\ + by (simp only: push_bit_add ac_simps) + then have \n - push_bit m 1 = push_bit m (2 * drop_bit (Suc m) n) + take_bit m n\ + by simp + then have \n - push_bit m 1 = push_bit m (2 * drop_bit (Suc m) n) OR take_bit m n\ + by (simp add: or_nat_def bit_simps flip: disjunctive_add) + with \bit n m\ show ?thesis + by (auto simp add: unset_bit_nat_def or_nat_def bit_simps) + qed +qed (simp_all add: mask_nat_def set_bit_nat_def flip_bit_nat_def) + +end + +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) + +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) + +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) + +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 + +context semiring_bit_operations +begin + +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) + +end + +context ring_bit_operations +begin + +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 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 + + +subsection \Horner sums\ + +context semiring_bit_shifts +begin + +lemma horner_sum_bit_eq_take_bit: + \horner_sum of_bool 2 (map (bit a) [0.. +proof (induction a arbitrary: n rule: bits_induct) + case (stable a) + moreover have \bit a = (\_. odd a)\ + using stable by (simp add: stable_imp_bit_iff_odd fun_eq_iff) + moreover have \{q. q < n} = {0.. + by auto + ultimately show ?case + by (simp add: stable_imp_take_bit_eq horner_sum_eq_sum mask_eq_sum_exp) +next + case (rec a b) + show ?case + proof (cases n) + case 0 + then show ?thesis + by simp + next + case (Suc m) + have \map (bit (of_bool b + 2 * a)) [0.. + by (simp only: upt_conv_Cons) simp + also have \\ = b # map (bit a) [0.. + by (simp only: flip: map_Suc_upt) (simp add: bit_Suc rec.hyps) + finally show ?thesis + using Suc rec.IH [of m] by (simp add: take_bit_Suc rec.hyps) + (simp_all add: ac_simps mod_2_eq_odd) + qed +qed + +end + +context unique_euclidean_semiring_with_bit_shifts +begin + +lemma bit_horner_sum_bit_iff [bit_simps]: + \bit (horner_sum of_bool 2 bs) 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 + next + case (Suc m) + with bit_rec [of _ n] Cons.prems Cons.IH [of m] + show ?thesis by simp + qed +qed + +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) + +end + +lemma horner_sum_of_bool_2_less: + \(horner_sum of_bool 2 bs :: int) < 2 ^ length bs\ +proof - + have \(\n = 0.. (\n = 0.. + by (rule sum_mono) simp + also have \\ = 2 ^ length bs - 1\ + by (induction bs) simp_all + finally show ?thesis + by (simp add: horner_sum_eq_sum) +qed + + +subsection \Symbolic computations on numeral expressions\ \<^marker>\contributor \Andreas Lochbihler\\ + +fun and_num :: \num \ num \ num option\ +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'))\ + +fun and_not_num :: \num \ num \ num option\ +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)\ + +fun or_num :: \num \ num \ num\ +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)\ + +fun or_not_num_neg :: \num \ num \ num\ +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)\ + +fun xor_num :: \num \ num \ num option\ +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 int_numeral_and_num: + \numeral m AND numeral n = (case and_num m n of None \ 0 :: int | Some n' \ numeral n')\ + by (induction m n rule: and_num.induct) (simp_all split: option.split) + +lemma and_num_eq_None_iff: + \and_num m n = None \ numeral m AND numeral n = (0::int)\ + by (simp add: int_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 :: int)\ + by (simp add: int_numeral_and_num split: option.split) + +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 add: add_One BitM_inc_eq not_int_def split: option.split) + +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 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 add: int_numeral_and_not_num split: option.split) + +lemma int_numeral_or_num: + \numeral m OR numeral n = (numeral (or_num m n) :: int)\ + 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 :: int)\ + by (simp add: int_numeral_or_num) + +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 add: add_One BitM_inc_eq not_int_def) + +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 int_numeral_xor_num: + \numeral m XOR numeral n = (case xor_num m n of None \ 0 :: int | 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::int)\ + by (simp add: int_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 :: int)\ + by (simp add: int_numeral_xor_num split: option.split) + + +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> 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]} +\ + +no_notation + "and" (infixr \AND\ 64) + and or (infixr \OR\ 59) + and xor (infixr \XOR\ 59) + +bundle bit_operations_syntax +begin + +notation + "and" (infixr \AND\ 64) + and or (infixr \OR\ 59) + and xor (infixr \XOR\ 59) + +end + +end