--- a/NEWS Sun Aug 01 23:18:13 2021 +0200
+++ b/NEWS Mon Aug 02 10:01:06 2021 +0000
@@ -184,18 +184,20 @@
separate theory "Transposition" in HOL-Combinatorics. INCOMPATIBILITY.
* Infix syntax for bit operations AND, OR, XOR is now organized in
-bundle bit_operations_syntax. INCOMPATIBILITY.
+bundle bit_operations_syntax. INCOMPATIBILITY.
* Bit operations set_bit, unset_bit and flip_bit are now class
operations. INCOMPATIBILITY.
+* Theory Bit_Operations is now part of HOL-Main. Minor INCOMPATIBILITY.
+
* Abbreviation "max_word" has been moved to session Word_Lib in the AFP,
as also have constants "shiftl1", "shiftr1", "sshiftr1", "bshiftr1",
"setBit", "clearBit". See there further the changelog in theory Guide.
INCOMPATIBILITY.
* New simp rules: less_exp, min.absorb1, min.absorb2, min.absorb3,
-min.absorb4, max.absorb1, max.absorb2, max.absorb3, max.absorb4. Minor
+min.absorb4, max.absorb1, max.absorb2, max.absorb3, max.absorb4. Minor
INCOMPATIBILITY.
* Sledgehammer:
@@ -210,6 +212,7 @@
* Metis:
- Renamed option "hide_lams" to "opaque_lifting". Minor INCOMPATIBILITY.
+
*** ML ***
* ML antiquotations \<^try>\<open>expr\<close> and \<^can>\<open>expr\<close> operate directly on
--- /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 \<open>Bit operations in suitable algebraic structures\<close>
+
+theory Bit_Operations
+ imports Presburger Groups_List
+begin
+
+subsection \<open>Abstract bit structures\<close>
+
+class semiring_bits = semiring_parity +
+ assumes bits_induct [case_names stable rec]:
+ \<open>(\<And>a. a div 2 = a \<Longrightarrow> P a)
+ \<Longrightarrow> (\<And>a b. P a \<Longrightarrow> (of_bool b + 2 * a) div 2 = a \<Longrightarrow> P (of_bool b + 2 * a))
+ \<Longrightarrow> P a\<close>
+ assumes bits_div_0 [simp]: \<open>0 div a = 0\<close>
+ and bits_div_by_1 [simp]: \<open>a div 1 = a\<close>
+ and bits_mod_div_trivial [simp]: \<open>a mod b div b = 0\<close>
+ and even_succ_div_2 [simp]: \<open>even a \<Longrightarrow> (1 + a) div 2 = a div 2\<close>
+ and even_mask_div_iff: \<open>even ((2 ^ m - 1) div 2 ^ n) \<longleftrightarrow> 2 ^ n = 0 \<or> m \<le> n\<close>
+ and exp_div_exp_eq: \<open>2 ^ m div 2 ^ n = of_bool (2 ^ m \<noteq> 0 \<and> m \<ge> n) * 2 ^ (m - n)\<close>
+ and div_exp_eq: \<open>a div 2 ^ m div 2 ^ n = a div 2 ^ (m + n)\<close>
+ and mod_exp_eq: \<open>a mod 2 ^ m mod 2 ^ n = a mod 2 ^ min m n\<close>
+ and mult_exp_mod_exp_eq: \<open>m \<le> n \<Longrightarrow> (a * 2 ^ m) mod (2 ^ n) = (a mod 2 ^ (n - m)) * 2 ^ m\<close>
+ and div_exp_mod_exp_eq: \<open>a div 2 ^ n mod 2 ^ m = a mod (2 ^ (n + m)) div 2 ^ n\<close>
+ and even_mult_exp_div_exp_iff: \<open>even (a * 2 ^ m div 2 ^ n) \<longleftrightarrow> m > n \<or> 2 ^ n = 0 \<or> (m \<le> n \<and> even (a div 2 ^ (n - m)))\<close>
+ fixes bit :: \<open>'a \<Rightarrow> nat \<Rightarrow> bool\<close>
+ assumes bit_iff_odd: \<open>bit a n \<longleftrightarrow> odd (a div 2 ^ n)\<close>
+begin
+
+text \<open>
+ Having \<^const>\<open>bit\<close> as definitional class operation
+ takes into account that specific instances can be implemented
+ differently wrt. code generation.
+\<close>
+
+lemma bits_div_by_0 [simp]:
+ \<open>a div 0 = 0\<close>
+ by (metis add_cancel_right_right bits_mod_div_trivial mod_mult_div_eq mult_not_zero)
+
+lemma bits_1_div_2 [simp]:
+ \<open>1 div 2 = 0\<close>
+ using even_succ_div_2 [of 0] by simp
+
+lemma bits_1_div_exp [simp]:
+ \<open>1 div 2 ^ n = of_bool (n = 0)\<close>
+ using div_exp_eq [of 1 1] by (cases n) simp_all
+
+lemma even_succ_div_exp [simp]:
+ \<open>(1 + a) div 2 ^ n = a div 2 ^ n\<close> if \<open>even a\<close> and \<open>n > 0\<close>
+proof (cases n)
+ case 0
+ with that show ?thesis
+ by simp
+next
+ case (Suc n)
+ with \<open>even a\<close> have \<open>(1 + a) div 2 ^ Suc n = a div 2 ^ Suc n\<close>
+ proof (induction n)
+ case 0
+ then show ?case
+ by simp
+ next
+ case (Suc n)
+ then show ?case
+ using div_exp_eq [of _ 1 \<open>Suc n\<close>, symmetric]
+ by simp
+ qed
+ with Suc show ?thesis
+ by simp
+qed
+
+lemma even_succ_mod_exp [simp]:
+ \<open>(1 + a) mod 2 ^ n = 1 + (a mod 2 ^ n)\<close> if \<open>even a\<close> and \<open>n > 0\<close>
+ using div_mult_mod_eq [of \<open>1 + a\<close> \<open>2 ^ n\<close>] that
+ apply simp
+ by (metis local.add.left_commute local.add_left_cancel local.div_mult_mod_eq)
+
+lemma bits_mod_by_1 [simp]:
+ \<open>a mod 1 = 0\<close>
+ using div_mult_mod_eq [of a 1] by simp
+
+lemma bits_mod_0 [simp]:
+ \<open>0 mod a = 0\<close>
+ using div_mult_mod_eq [of 0 a] by simp
+
+lemma bits_one_mod_two_eq_one [simp]:
+ \<open>1 mod 2 = 1\<close>
+ by (simp add: mod2_eq_if)
+
+lemma bit_0 [simp]:
+ \<open>bit a 0 \<longleftrightarrow> odd a\<close>
+ by (simp add: bit_iff_odd)
+
+lemma bit_Suc:
+ \<open>bit a (Suc n) \<longleftrightarrow> bit (a div 2) n\<close>
+ using div_exp_eq [of a 1 n] by (simp add: bit_iff_odd)
+
+lemma bit_rec:
+ \<open>bit a n \<longleftrightarrow> (if n = 0 then odd a else bit (a div 2) (n - 1))\<close>
+ by (cases n) (simp_all add: bit_Suc)
+
+lemma bit_0_eq [simp]:
+ \<open>bit 0 = bot\<close>
+ by (simp add: fun_eq_iff bit_iff_odd)
+
+context
+ fixes a
+ assumes stable: \<open>a div 2 = a\<close>
+begin
+
+lemma bits_stable_imp_add_self:
+ \<open>a + a mod 2 = 0\<close>
+proof -
+ have \<open>a div 2 * 2 + a mod 2 = a\<close>
+ by (fact div_mult_mod_eq)
+ then have \<open>a * 2 + a mod 2 = a\<close>
+ by (simp add: stable)
+ then show ?thesis
+ by (simp add: mult_2_right ac_simps)
+qed
+
+lemma stable_imp_bit_iff_odd:
+ \<open>bit a n \<longleftrightarrow> odd a\<close>
+ by (induction n) (simp_all add: stable bit_Suc)
+
+end
+
+lemma bit_iff_idd_imp_stable:
+ \<open>a div 2 = a\<close> if \<open>\<And>n. bit a n \<longleftrightarrow> odd a\<close>
+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]: \<open>b = odd a\<close>
+ by (simp add: rec.hyps bit_Suc)
+ from rec.hyps have hyp: \<open>(of_bool (odd a) + 2 * a) div 2 = a\<close>
+ by simp
+ have \<open>bit a n \<longleftrightarrow> odd a\<close> for n
+ using rec.prems [of \<open>Suc n\<close>] by (simp add: hyp bit_Suc)
+ then have \<open>a div 2 = a\<close>
+ by (rule rec.IH)
+ then have \<open>of_bool (odd a) + 2 * a = 2 * (a div 2) + of_bool (odd a)\<close>
+ by (simp add: ac_simps)
+ also have \<open>\<dots> = a\<close>
+ using mult_div_mod_eq [of 2 a]
+ by (simp add: of_bool_odd_eq_mod_2)
+ finally show ?case
+ using \<open>a div 2 = a\<close> by (simp add: hyp)
+qed
+
+lemma exp_eq_0_imp_not_bit:
+ \<open>\<not> bit a n\<close> if \<open>2 ^ n = 0\<close>
+ using that by (simp add: bit_iff_odd)
+
+lemma bit_eqI:
+ \<open>a = b\<close> if \<open>\<And>n. 2 ^ n \<noteq> 0 \<Longrightarrow> bit a n \<longleftrightarrow> bit b n\<close>
+proof -
+ have \<open>bit a n \<longleftrightarrow> bit b n\<close> for n
+ proof (cases \<open>2 ^ n = 0\<close>)
+ 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 **: \<open>even b \<longleftrightarrow> even a\<close>
+ by simp
+ have \<open>b div 2 = b\<close>
+ proof (rule bit_iff_idd_imp_stable)
+ fix n
+ from stable have *: \<open>bit b n \<longleftrightarrow> bit a n\<close>
+ by simp
+ also have \<open>bit a n \<longleftrightarrow> odd a\<close>
+ using stable by (simp add: stable_imp_bit_iff_odd)
+ finally show \<open>bit b n \<longleftrightarrow> odd b\<close>
+ by (simp add: **)
+ qed
+ from ** have \<open>a mod 2 = b mod 2\<close>
+ by (simp add: mod2_eq_if)
+ then have \<open>a mod 2 + (a + b) = b mod 2 + (a + b)\<close>
+ by simp
+ then have \<open>a + a mod 2 + b = b + b mod 2 + a\<close>
+ by (simp add: ac_simps)
+ with \<open>a div 2 = a\<close> \<open>b div 2 = b\<close> show ?case
+ by (simp add: bits_stable_imp_add_self)
+ next
+ case (rec a p)
+ from rec.prems [of 0] have [simp]: \<open>p = odd b\<close>
+ by simp
+ from rec.hyps have \<open>bit a n \<longleftrightarrow> bit (b div 2) n\<close> for n
+ using rec.prems [of \<open>Suc n\<close>] by (simp add: bit_Suc)
+ then have \<open>a = b div 2\<close>
+ by (rule rec.IH)
+ then have \<open>2 * a = 2 * (b div 2)\<close>
+ by simp
+ then have \<open>b mod 2 + 2 * a = b mod 2 + 2 * (b div 2)\<close>
+ by simp
+ also have \<open>\<dots> = b\<close>
+ by (fact mod_mult_div_eq)
+ finally show ?case
+ by (auto simp add: mod2_eq_if)
+ qed
+qed
+
+lemma bit_eq_iff:
+ \<open>a = b \<longleftrightarrow> (\<forall>n. bit a n \<longleftrightarrow> bit b n)\<close>
+ by (auto intro: bit_eqI)
+
+named_theorems bit_simps \<open>Simplification rules for \<^const>\<open>bit\<close>\<close>
+
+lemma bit_exp_iff [bit_simps]:
+ \<open>bit (2 ^ m) n \<longleftrightarrow> 2 ^ m \<noteq> 0 \<and> m = n\<close>
+ by (auto simp add: bit_iff_odd exp_div_exp_eq)
+
+lemma bit_1_iff [bit_simps]:
+ \<open>bit 1 n \<longleftrightarrow> 1 \<noteq> 0 \<and> n = 0\<close>
+ using bit_exp_iff [of 0 n] by simp
+
+lemma bit_2_iff [bit_simps]:
+ \<open>bit 2 n \<longleftrightarrow> 2 \<noteq> 0 \<and> n = 1\<close>
+ using bit_exp_iff [of 1 n] by auto
+
+lemma even_bit_succ_iff:
+ \<open>bit (1 + a) n \<longleftrightarrow> bit a n \<or> n = 0\<close> if \<open>even a\<close>
+ using that by (cases \<open>n = 0\<close>) (simp_all add: bit_iff_odd)
+
+lemma odd_bit_iff_bit_pred:
+ \<open>bit a n \<longleftrightarrow> bit (a - 1) n \<or> n = 0\<close> if \<open>odd a\<close>
+proof -
+ from \<open>odd a\<close> obtain b where \<open>a = 2 * b + 1\<close> ..
+ moreover have \<open>bit (2 * b) n \<or> n = 0 \<longleftrightarrow> bit (1 + 2 * b) n\<close>
+ using even_bit_succ_iff by simp
+ ultimately show ?thesis by (simp add: ac_simps)
+qed
+
+lemma bit_double_iff [bit_simps]:
+ \<open>bit (2 * a) n \<longleftrightarrow> bit a (n - 1) \<and> n \<noteq> 0 \<and> 2 ^ n \<noteq> 0\<close>
+ 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:
+ \<open>a = b \<longleftrightarrow> (even a \<longleftrightarrow> even b) \<and> a div 2 = b div 2\<close> (is \<open>?P = ?Q\<close>)
+proof
+ assume ?P
+ then show ?Q
+ by simp
+next
+ assume ?Q
+ then have \<open>even a \<longleftrightarrow> even b\<close> and \<open>a div 2 = b div 2\<close>
+ by simp_all
+ show ?P
+ proof (rule bit_eqI)
+ fix n
+ show \<open>bit a n \<longleftrightarrow> bit b n\<close>
+ proof (cases n)
+ case 0
+ with \<open>even a \<longleftrightarrow> even b\<close> show ?thesis
+ by simp
+ next
+ case (Suc n)
+ moreover from \<open>a div 2 = b div 2\<close> have \<open>bit (a div 2) n = bit (b div 2) n\<close>
+ by simp
+ ultimately show ?thesis
+ by (simp add: bit_Suc)
+ qed
+ qed
+qed
+
+lemma bit_mod_2_iff [simp]:
+ \<open>bit (a mod 2) n \<longleftrightarrow> n = 0 \<and> odd a\<close>
+ by (cases a rule: parity_cases) (simp_all add: bit_iff_odd)
+
+lemma bit_mask_iff:
+ \<open>bit (2 ^ m - 1) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n < m\<close>
+ by (simp add: bit_iff_odd even_mask_div_iff not_le)
+
+lemma bit_Numeral1_iff [simp]:
+ \<open>bit (numeral Num.One) n \<longleftrightarrow> n = 0\<close>
+ by (simp add: bit_rec)
+
+lemma exp_add_not_zero_imp:
+ \<open>2 ^ m \<noteq> 0\<close> and \<open>2 ^ n \<noteq> 0\<close> if \<open>2 ^ (m + n) \<noteq> 0\<close>
+proof -
+ have \<open>\<not> (2 ^ m = 0 \<or> 2 ^ n = 0)\<close>
+ proof (rule notI)
+ assume \<open>2 ^ m = 0 \<or> 2 ^ n = 0\<close>
+ then have \<open>2 ^ (m + n) = 0\<close>
+ by (rule disjE) (simp_all add: power_add)
+ with that show False ..
+ qed
+ then show \<open>2 ^ m \<noteq> 0\<close> and \<open>2 ^ n \<noteq> 0\<close>
+ by simp_all
+qed
+
+lemma bit_disjunctive_add_iff:
+ \<open>bit (a + b) n \<longleftrightarrow> bit a n \<or> bit b n\<close>
+ if \<open>\<And>n. \<not> bit a n \<or> \<not> bit b n\<close>
+proof (cases \<open>2 ^ n = 0\<close>)
+ 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: \<open>even a \<or> even b\<close>
+ by auto
+ have bit: \<open>\<not> bit (a div 2) n \<or> \<not> bit (b div 2) n\<close> for n
+ using Suc.prems(1) [of \<open>Suc n\<close>] by (simp add: bit_Suc)
+ from Suc.prems(2) have \<open>2 * 2 ^ n \<noteq> 0\<close> \<open>2 ^ n \<noteq> 0\<close>
+ by (auto simp add: mult_2)
+ have \<open>a + b = (a div 2 * 2 + a mod 2) + (b div 2 * 2 + b mod 2)\<close>
+ using div_mult_mod_eq [of a 2] div_mult_mod_eq [of b 2] by simp
+ also have \<open>\<dots> = of_bool (odd a \<or> odd b) + 2 * (a div 2 + b div 2)\<close>
+ using even by (auto simp add: algebra_simps mod2_eq_if)
+ finally have \<open>bit ((a + b) div 2) n \<longleftrightarrow> bit (a div 2 + b div 2) n\<close>
+ using \<open>2 * 2 ^ n \<noteq> 0\<close> by simp (simp_all flip: bit_Suc add: bit_double_iff)
+ also have \<open>\<dots> \<longleftrightarrow> bit (a div 2) n \<or> bit (b div 2) n\<close>
+ using bit \<open>2 ^ n \<noteq> 0\<close> by (rule Suc.IH)
+ finally show ?case
+ by (simp add: bit_Suc)
+ qed
+qed
+
+lemma
+ exp_add_not_zero_imp_left: \<open>2 ^ m \<noteq> 0\<close>
+ and exp_add_not_zero_imp_right: \<open>2 ^ n \<noteq> 0\<close>
+ if \<open>2 ^ (m + n) \<noteq> 0\<close>
+proof -
+ have \<open>\<not> (2 ^ m = 0 \<or> 2 ^ n = 0)\<close>
+ proof (rule notI)
+ assume \<open>2 ^ m = 0 \<or> 2 ^ n = 0\<close>
+ then have \<open>2 ^ (m + n) = 0\<close>
+ by (rule disjE) (simp_all add: power_add)
+ with that show False ..
+ qed
+ then show \<open>2 ^ m \<noteq> 0\<close> and \<open>2 ^ n \<noteq> 0\<close>
+ by simp_all
+qed
+
+lemma exp_not_zero_imp_exp_diff_not_zero:
+ \<open>2 ^ (n - m) \<noteq> 0\<close> if \<open>2 ^ n \<noteq> 0\<close>
+proof (cases \<open>m \<le> n\<close>)
+ case True
+ moreover define q where \<open>q = n - m\<close>
+ ultimately have \<open>n = m + q\<close>
+ 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: "\<And>n. P n \<Longrightarrow> n > 0 \<Longrightarrow> P (2 * n)"
+ and odd: "\<And>n. P n \<Longrightarrow> 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 \<noteq> 1"
+ by auto
+ with \<open>n \<noteq> 0\<close> have "n div 2 > 0"
+ by simp
+ with \<open>even n\<close> 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 :: \<open>nat \<Rightarrow> nat \<Rightarrow> bool\<close>
+ where \<open>bit_nat m n \<longleftrightarrow> odd (m div 2 ^ n)\<close>
+
+instance
+proof
+ show \<open>P n\<close> if stable: \<open>\<And>n. n div 2 = n \<Longrightarrow> P n\<close>
+ and rec: \<open>\<And>n b. P n \<Longrightarrow> (of_bool b + 2 * n) div 2 = n \<Longrightarrow> P (of_bool b + 2 * n)\<close>
+ 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 \<open>q mod 2 ^ m mod 2 ^ n = q mod 2 ^ min m n\<close>
+ 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 \<open>(q * 2 ^ m) mod (2 ^ n) = (q mod 2 ^ (n - m)) * 2 ^ m\<close> if \<open>m \<le> n\<close>
+ 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 \<open>even ((2 ^ m - (1::nat)) div 2 ^ n) \<longleftrightarrow> 2 ^ n = (0::nat) \<or> m \<le> n\<close>
+ for m n :: nat
+ using even_mask_div_iff' [where ?'a = nat, of m n] by simp
+ show \<open>even (q * 2 ^ m div 2 ^ n) \<longleftrightarrow> n < m \<or> (2::nat) ^ n = 0 \<or> m \<le> n \<and> even (q div 2 ^ (n - m))\<close>
+ 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: "\<And>k. P k \<Longrightarrow> k \<noteq> 0 \<Longrightarrow> P (k * 2)"
+ and odd_int: "\<And>k. P k \<Longrightarrow> k \<noteq> - 1 \<Longrightarrow> P (1 + (k * 2))" for k :: int
+proof (cases "k \<ge> 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 \<open>simp_all add: algebra_simps\<close>)
+ also have "\<dots> = - 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 \<open>simp_all add: algebra_simps\<close>)
+ also have "\<dots> = - 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]:
+ \<open>bit (of_bool b) n \<longleftrightarrow> b \<and> n = 0\<close>
+ by (simp add: bit_1_iff)
+
+lemma even_of_nat_iff:
+ \<open>even (of_nat n) \<longleftrightarrow> even n\<close>
+ by (induction n rule: nat_bit_induct) simp_all
+
+lemma bit_of_nat_iff [bit_simps]:
+ \<open>bit (of_nat m) n \<longleftrightarrow> (2::'a) ^ n \<noteq> 0 \<and> bit m n\<close>
+proof (cases \<open>(2::'a) ^ n = 0\<close>)
+ case True
+ then show ?thesis
+ by (simp add: exp_eq_0_imp_not_bit)
+next
+ case False
+ then have \<open>bit (of_nat m) n \<longleftrightarrow> bit m n\<close>
+ 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 :: \<open>int \<Rightarrow> nat \<Rightarrow> bool\<close>
+ where \<open>bit_int k n \<longleftrightarrow> odd (k div 2 ^ n)\<close>
+
+instance
+proof
+ show \<open>P k\<close> if stable: \<open>\<And>k. k div 2 = k \<Longrightarrow> P k\<close>
+ and rec: \<open>\<And>k b. P k \<Longrightarrow> (of_bool b + 2 * k) div 2 = k \<Longrightarrow> P (of_bool b + 2 * k)\<close>
+ 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 \<open>- 1\<close>] 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 \<open>(2::int) ^ m div 2 ^ n = of_bool ((2::int) ^ m \<noteq> 0 \<and> n \<le> m) * 2 ^ (m - n)\<close>
+ for m n :: nat
+ proof (cases \<open>m < n\<close>)
+ case True
+ then have \<open>n = m + (n - m)\<close>
+ by simp
+ then have \<open>(2::int) ^ m div 2 ^ n = (2::int) ^ m div 2 ^ (m + (n - m))\<close>
+ by simp
+ also have \<open>\<dots> = (2::int) ^ m div (2 ^ m * 2 ^ (n - m))\<close>
+ by (simp add: power_add)
+ also have \<open>\<dots> = (2::int) ^ m div 2 ^ m div 2 ^ (n - m)\<close>
+ by (simp add: zdiv_zmult2_eq)
+ finally show ?thesis using \<open>m < n\<close> by simp
+ next
+ case False
+ then show ?thesis
+ by (simp add: power_diff)
+ qed
+ show \<open>k mod 2 ^ m mod 2 ^ n = k mod 2 ^ min m n\<close>
+ for m n :: nat and k :: int
+ using mod_exp_eq [of \<open>nat k\<close> 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 \<open>2 ^ m\<close>])
+ apply (subst zmod_zmult2_eq) apply simp_all
+ done
+ show \<open>(k * 2 ^ m) mod (2 ^ n) = (k mod 2 ^ (n - m)) * 2 ^ m\<close>
+ if \<open>m \<le> n\<close> 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 \<open>even ((2 ^ m - (1::int)) div 2 ^ n) \<longleftrightarrow> 2 ^ n = (0::int) \<or> m \<le> n\<close>
+ for m n :: nat
+ using even_mask_div_iff' [where ?'a = int, of m n] by simp
+ show \<open>even (k * 2 ^ m div 2 ^ n) \<longleftrightarrow> n < m \<or> (2::int) ^ n = 0 \<or> m \<le> n \<and> even (k div 2 ^ (n - m))\<close>
+ 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 :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
+ assumes push_bit_eq_mult: \<open>push_bit n a = a * 2 ^ n\<close>
+ fixes drop_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
+ assumes drop_bit_eq_div: \<open>drop_bit n a = a div 2 ^ n\<close>
+ fixes take_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
+ assumes take_bit_eq_mod: \<open>take_bit n a = a mod 2 ^ n\<close>
+begin
+
+text \<open>
+ Logically, \<^const>\<open>push_bit\<close>,
+ \<^const>\<open>drop_bit\<close> and \<^const>\<open>take_bit\<close> are just aliases; having them
+ as separate operations makes proofs easier, otherwise proof automation
+ would fiddle with concrete expressions \<^term>\<open>2 ^ n\<close> 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.
+\<close>
+
+lemma bit_iff_odd_drop_bit:
+ \<open>bit a n \<longleftrightarrow> odd (drop_bit n a)\<close>
+ by (simp add: bit_iff_odd drop_bit_eq_div)
+
+lemma even_drop_bit_iff_not_bit:
+ \<open>even (drop_bit n a) \<longleftrightarrow> \<not> bit a n\<close>
+ by (simp add: bit_iff_odd_drop_bit)
+
+lemma div_push_bit_of_1_eq_drop_bit:
+ \<open>a div push_bit n 1 = drop_bit n a\<close>
+ 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]:
+ \<open>push_bit (numeral l) (numeral k) = push_bit (pred_numeral l) (numeral (Num.Bit0 k))\<close>
+ 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:
+ \<open>take_bit (Suc n) a = take_bit n (a div 2) * 2 + a mod 2\<close>
+proof -
+ have \<open>take_bit (Suc n) (a div 2 * 2 + of_bool (odd a)) = take_bit n (a div 2) * 2 + of_bool (odd a)\<close>
+ using even_succ_mod_exp [of \<open>2 * (a div 2)\<close> \<open>Suc n\<close>]
+ mult_exp_mod_exp_eq [of 1 \<open>Suc n\<close> \<open>a div 2\<close>]
+ 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:
+ \<open>take_bit n a = (if n = 0 then 0 else take_bit (n - 1) (a div 2) * 2 + a mod 2)\<close>
+ by (cases n) (simp_all add: take_bit_Suc)
+
+lemma take_bit_Suc_0 [simp]:
+ \<open>take_bit (Suc 0) a = a mod 2\<close>
+ 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 \<and> b)"
+ by (cases n) simp_all
+
+lemma even_take_bit_eq [simp]:
+ \<open>even (take_bit n a) \<longleftrightarrow> n = 0 \<or> even a\<close>
+ 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 \<open>m + n\<close> 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 \<le> 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 \<open>a * 2 ^ n\<close> 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 \<le> 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 \<open>m = n + q\<close>
+ by (auto simp add: not_le dest: less_imp_Suc_add)
+ then have \<open>drop_bit m (take_bit n a) = 0\<close>
+ using div_exp_eq [of \<open>a mod 2 ^ n\<close> 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]:
+ \<open>even (push_bit n a) \<longleftrightarrow> n \<noteq> 0 \<or> even a\<close>
+ by (simp add: push_bit_eq_mult) auto
+
+lemma bit_push_bit_iff [bit_simps]:
+ \<open>bit (push_bit m a) n \<longleftrightarrow> m \<le> n \<and> 2 ^ n \<noteq> 0 \<and> bit a (n - m)\<close>
+ by (auto simp add: bit_iff_odd push_bit_eq_mult even_mult_exp_div_exp_iff)
+
+lemma bit_drop_bit_eq [bit_simps]:
+ \<open>bit (drop_bit n a) = bit a \<circ> (+) n\<close>
+ by (simp add: bit_iff_odd fun_eq_iff ac_simps flip: drop_bit_eq_div)
+
+lemma bit_take_bit_iff [bit_simps]:
+ \<open>bit (take_bit m a) n \<longleftrightarrow> n < m \<and> bit a n\<close>
+ by (simp add: bit_iff_odd drop_bit_take_bit not_le flip: drop_bit_eq_div)
+
+lemma stable_imp_drop_bit_eq:
+ \<open>drop_bit n a = a\<close>
+ if \<open>a div 2 = a\<close>
+ by (induction n) (simp_all add: that drop_bit_Suc)
+
+lemma stable_imp_take_bit_eq:
+ \<open>take_bit n a = (if even a then 0 else 2 ^ n - 1)\<close>
+ if \<open>a div 2 = a\<close>
+proof (rule bit_eqI)
+ fix m
+ assume \<open>2 ^ m \<noteq> 0\<close>
+ with that show \<open>bit (take_bit n a) m \<longleftrightarrow> bit (if even a then 0 else 2 ^ n - 1) m\<close>
+ by (simp add: bit_take_bit_iff bit_mask_iff stable_imp_bit_iff_odd)
+qed
+
+lemma exp_dvdE:
+ assumes \<open>2 ^ n dvd a\<close>
+ obtains b where \<open>a = push_bit n b\<close>
+proof -
+ from assms obtain b where \<open>a = 2 ^ n * b\<close> ..
+ then have \<open>a = push_bit n b\<close>
+ by (simp add: push_bit_eq_mult ac_simps)
+ with that show thesis .
+qed
+
+lemma take_bit_eq_0_iff:
+ \<open>take_bit n a = 0 \<longleftrightarrow> 2 ^ n dvd a\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
+proof
+ assume ?P
+ then show ?Q
+ by (simp add: take_bit_eq_mod mod_0_imp_dvd)
+next
+ assume ?Q
+ then obtain b where \<open>a = push_bit n b\<close>
+ by (rule exp_dvdE)
+ then show ?P
+ by (simp add: take_bit_push_bit)
+qed
+
+lemma take_bit_tightened:
+ \<open>take_bit m a = take_bit m b\<close> if \<open>take_bit n a = take_bit n b\<close> and \<open>m \<le> n\<close>
+proof -
+ from that have \<open>take_bit m (take_bit n a) = take_bit m (take_bit n b)\<close>
+ by simp
+ then have \<open>take_bit (min m n) a = take_bit (min m n) b\<close>
+ by simp
+ with that show ?thesis
+ by (simp add: min_def)
+qed
+
+lemma take_bit_eq_self_iff_drop_bit_eq_0:
+ \<open>take_bit n a = a \<longleftrightarrow> drop_bit n a = 0\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
+proof
+ assume ?P
+ show ?Q
+ proof (rule bit_eqI)
+ fix m
+ from \<open>?P\<close> have \<open>a = take_bit n a\<close> ..
+ also have \<open>\<not> bit (take_bit n a) (n + m)\<close>
+ unfolding bit_simps
+ by (simp add: bit_simps)
+ finally show \<open>bit (drop_bit n a) m \<longleftrightarrow> bit 0 m\<close>
+ by (simp add: bit_simps)
+ qed
+next
+ assume ?Q
+ show ?P
+ proof (rule bit_eqI)
+ fix m
+ from \<open>?Q\<close> have \<open>\<not> bit (drop_bit n a) (m - n)\<close>
+ by simp
+ then have \<open> \<not> bit a (n + (m - n))\<close>
+ by (simp add: bit_simps)
+ then show \<open>bit (take_bit n a) m \<longleftrightarrow> bit a m\<close>
+ by (cases \<open>m < n\<close>) (auto simp add: bit_simps)
+ qed
+qed
+
+lemma drop_bit_exp_eq:
+ \<open>drop_bit m (2 ^ n) = of_bool (m \<le> n \<and> 2 ^ n \<noteq> 0) * 2 ^ (n - m)\<close>
+ by (rule bit_eqI) (auto simp add: bit_simps)
+
+end
+
+instantiation nat :: semiring_bit_shifts
+begin
+
+definition push_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
+ where \<open>push_bit_nat n m = m * 2 ^ n\<close>
+
+definition drop_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
+ where \<open>drop_bit_nat n m = m div 2 ^ n\<close>
+
+definition take_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
+ where \<open>take_bit_nat n m = m mod 2 ^ n\<close>
+
+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:
+ \<open>push_bit n (of_nat m) = of_nat (push_bit n m)\<close>
+ by (simp add: push_bit_eq_mult semiring_bit_shifts_class.push_bit_eq_mult)
+
+lemma of_nat_push_bit:
+ \<open>of_nat (push_bit m n) = push_bit m (of_nat n)\<close>
+ by (simp add: push_bit_eq_mult semiring_bit_shifts_class.push_bit_eq_mult)
+
+lemma take_bit_of_nat:
+ \<open>take_bit n (of_nat m) = of_nat (take_bit n m)\<close>
+ 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:
+ \<open>of_nat (take_bit n m) = take_bit n (of_nat m)\<close>
+ 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 :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
+ where \<open>push_bit_int n k = k * 2 ^ n\<close>
+
+definition drop_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
+ where \<open>drop_bit_int n k = k div 2 ^ n\<close>
+
+definition take_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
+ where \<open>take_bit_int n k = k mod 2 ^ n\<close>
+
+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:
+ \<open>bit (push_bit m q) n \<longleftrightarrow> m \<le> n \<and> bit q (n - m)\<close> for q :: nat
+ by (auto simp add: bit_push_bit_iff)
+
+lemma bit_push_bit_iff_int:
+ \<open>bit (push_bit m k) n \<longleftrightarrow> m \<le> n \<and> bit k (n - m)\<close> for k :: int
+ by (auto simp add: bit_push_bit_iff)
+
+lemma take_bit_nat_less_exp [simp]:
+ \<open>take_bit n m < 2 ^ n\<close> for n m ::nat
+ by (simp add: take_bit_eq_mod)
+
+lemma take_bit_nonnegative [simp]:
+ \<open>take_bit n k \<ge> 0\<close> for k :: int
+ by (simp add: take_bit_eq_mod)
+
+lemma not_take_bit_negative [simp]:
+ \<open>\<not> take_bit n k < 0\<close> for k :: int
+ by (simp add: not_less)
+
+lemma take_bit_int_less_exp [simp]:
+ \<open>take_bit n k < 2 ^ n\<close> for k :: int
+ by (simp add: take_bit_eq_mod)
+
+lemma take_bit_nat_eq_self_iff:
+ \<open>take_bit n m = m \<longleftrightarrow> m < 2 ^ n\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
+ 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:
+ \<open>take_bit n m = m\<close> if \<open>m < 2 ^ n\<close> for m n :: nat
+ using that by (simp add: take_bit_nat_eq_self_iff)
+
+lemma take_bit_int_eq_self_iff:
+ \<open>take_bit n k = k \<longleftrightarrow> 0 \<le> k \<and> k < 2 ^ n\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
+ 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:
+ \<open>take_bit n k = k\<close> if \<open>0 \<le> k\<close> \<open>k < 2 ^ n\<close> for k :: int
+ using that by (simp add: take_bit_int_eq_self_iff)
+
+lemma take_bit_nat_less_eq_self [simp]:
+ \<open>take_bit n m \<le> m\<close> for n m :: nat
+ by (simp add: take_bit_eq_mod)
+
+lemma take_bit_nat_less_self_iff:
+ \<open>take_bit n m < m \<longleftrightarrow> 2 ^ n \<le> m\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
+ for m n :: nat
+proof
+ assume ?P
+ then have \<open>take_bit n m \<noteq> m\<close>
+ by simp
+ then show \<open>?Q\<close>
+ by (simp add: take_bit_nat_eq_self_iff)
+next
+ have \<open>take_bit n m < 2 ^ n\<close>
+ 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]:
+ \<open>take_bit m (2 ^ n) = of_bool (n < m) * 2 ^ n\<close>
+ by (simp add: take_bit_eq_mod exp_mod_exp)
+
+lemma take_bit_of_2 [simp]:
+ \<open>take_bit n 2 = of_bool (2 \<le> n) * 2\<close>
+ using take_bit_of_exp [of n 1] by simp
+
+lemma take_bit_of_mask:
+ \<open>take_bit m (2 ^ n - 1) = 2 ^ min m n - 1\<close>
+ by (simp add: take_bit_eq_mod mask_mod_exp)
+
+lemma push_bit_eq_0_iff [simp]:
+ "push_bit n a = 0 \<longleftrightarrow> 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 \<longleftrightarrow> n = 0"
+ by (simp add: take_bit_eq_mod)
+
+lemma take_bit_Suc_1 [simp]:
+ \<open>take_bit (Suc n) 1 = 1\<close>
+ by (simp add: take_bit_Suc)
+
+lemma take_bit_Suc_bit0 [simp]:
+ \<open>take_bit (Suc n) (numeral (Num.Bit0 k)) = take_bit n (numeral k) * 2\<close>
+ by (simp add: take_bit_Suc numeral_Bit0_div_2)
+
+lemma take_bit_Suc_bit1 [simp]:
+ \<open>take_bit (Suc n) (numeral (Num.Bit1 k)) = take_bit n (numeral k) * 2 + 1\<close>
+ by (simp add: take_bit_Suc numeral_Bit1_div_2 mod_2_eq_odd)
+
+lemma take_bit_numeral_1 [simp]:
+ \<open>take_bit (numeral l) 1 = 1\<close>
+ by (simp add: take_bit_rec [of \<open>numeral l\<close> 1])
+
+lemma take_bit_numeral_bit0 [simp]:
+ \<open>take_bit (numeral l) (numeral (Num.Bit0 k)) = take_bit (pred_numeral l) (numeral k) * 2\<close>
+ by (simp add: take_bit_rec numeral_Bit0_div_2)
+
+lemma take_bit_numeral_bit1 [simp]:
+ \<open>take_bit (numeral l) (numeral (Num.Bit1 k)) = take_bit (pred_numeral l) (numeral k) * 2 + 1\<close>
+ by (simp add: take_bit_rec numeral_Bit1_div_2 mod_2_eq_odd)
+
+lemma drop_bit_Suc_bit0 [simp]:
+ \<open>drop_bit (Suc n) (numeral (Num.Bit0 k)) = drop_bit n (numeral k)\<close>
+ by (simp add: drop_bit_Suc numeral_Bit0_div_2)
+
+lemma drop_bit_Suc_bit1 [simp]:
+ \<open>drop_bit (Suc n) (numeral (Num.Bit1 k)) = drop_bit n (numeral k)\<close>
+ by (simp add: drop_bit_Suc numeral_Bit1_div_2)
+
+lemma drop_bit_numeral_bit0 [simp]:
+ \<open>drop_bit (numeral l) (numeral (Num.Bit0 k)) = drop_bit (pred_numeral l) (numeral k)\<close>
+ by (simp add: drop_bit_rec numeral_Bit0_div_2)
+
+lemma drop_bit_numeral_bit1 [simp]:
+ \<open>drop_bit (numeral l) (numeral (Num.Bit1 k)) = drop_bit (pred_numeral l) (numeral k)\<close>
+ 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]:
+ \<open>bit (of_nat m) n \<longleftrightarrow> bit m n\<close>
+proof -
+ have \<open>even (m div 2 ^ n) \<longleftrightarrow> even (of_nat (m div 2 ^ n))\<close>
+ by simp
+ also have \<open>of_nat (m div 2 ^ n) = of_nat m div of_nat (2 ^ n)\<close>
+ 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:
+ \<open>of_nat (drop_bit m n) = drop_bit m (of_nat n)\<close>
+ 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]:
+ \<open>bit (push_bit m (of_nat r)) n \<longleftrightarrow> m \<le> n \<and> bit (of_nat r) (n - m)\<close>
+ by (auto simp add: bit_push_bit_iff)
+
+lemma take_bit_sum:
+ "take_bit n a = (\<Sum>k = 0..<n. push_bit k (of_bool (bit a k)))"
+ for n :: nat
+proof (induction n arbitrary: a)
+ case 0
+ then show ?case
+ by simp
+next
+ case (Suc n)
+ have "(\<Sum>k = 0..<Suc n. push_bit k (of_bool (bit a k))) =
+ of_bool (odd a) + (\<Sum>k = Suc 0..<Suc n. push_bit k (of_bool (bit a k)))"
+ by (simp add: sum.atLeast_Suc_lessThan ac_simps)
+ also have "(\<Sum>k = Suc 0..<Suc n. push_bit k (of_bool (bit a k)))
+ = (\<Sum>k = 0..<n. push_bit k (of_bool (bit (a div 2) k))) * 2"
+ by (simp only: sum.atLeast_Suc_lessThan_Suc_shift) (simp add: sum_distrib_right push_bit_double drop_bit_Suc bit_Suc)
+ finally show ?case
+ using Suc [of "a div 2"] by (simp add: ac_simps take_bit_Suc mod_2_eq_odd)
+qed
+
+end
+
+instance nat :: unique_euclidean_semiring_with_bit_shifts ..
+
+instance int :: unique_euclidean_semiring_with_bit_shifts ..
+
+lemma bit_numeral_int_iff [bit_simps]:
+ \<open>bit (numeral m :: int) n \<longleftrightarrow> bit (numeral m :: nat) n\<close>
+ using bit_of_nat_iff_bit [of \<open>numeral m\<close> n] by simp
+
+lemma bit_not_int_iff':
+ \<open>bit (- k - 1) n \<longleftrightarrow> \<not> bit k n\<close>
+ for k :: int
+proof (induction n arbitrary: k)
+ case 0
+ show ?case
+ by simp
+next
+ case (Suc n)
+ have \<open>- k - 1 = - (k + 2) + 1\<close>
+ by simp
+ also have \<open>(- (k + 2) + 1) div 2 = - (k div 2) - 1\<close>
+ proof (cases \<open>even k\<close>)
+ case True
+ then have \<open>- k div 2 = - (k div 2)\<close>
+ by rule (simp flip: mult_minus_right)
+ with True show ?thesis
+ by simp
+ next
+ case False
+ have \<open>4 = 2 * (2::int)\<close>
+ by simp
+ also have \<open>2 * 2 div 2 = (2::int)\<close>
+ by (simp only: nonzero_mult_div_cancel_left)
+ finally have *: \<open>4 div 2 = (2::int)\<close> .
+ from False obtain l where k: \<open>k = 2 * l + 1\<close> ..
+ then have \<open>- k - 2 = 2 * - (l + 2) + 1\<close>
+ by simp
+ then have \<open>(- k - 2) div 2 + 1 = - (k div 2) - 1\<close>
+ by (simp flip: mult_minus_right add: *) (simp add: k)
+ with False show ?thesis
+ by simp
+ qed
+ finally have \<open>(- k - 1) div 2 = - (k div 2) - 1\<close> .
+ with Suc show ?case
+ by (simp add: bit_Suc)
+qed
+
+lemma bit_minus_int_iff [bit_simps]:
+ \<open>bit (- k) n \<longleftrightarrow> \<not> bit (k - 1) n\<close>
+ for k :: int
+ using bit_not_int_iff' [of \<open>k - 1\<close>] by simp
+
+lemma bit_nat_iff [bit_simps]:
+ \<open>bit (nat k) n \<longleftrightarrow> k \<ge> 0 \<and> bit k n\<close>
+proof (cases \<open>k \<ge> 0\<close>)
+ case True
+ moreover define m where \<open>m = nat k\<close>
+ ultimately have \<open>k = int m\<close>
+ 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]:
+ \<open>bit (1 :: int) (numeral n) \<longleftrightarrow> bit (0 :: int) (pred_numeral n)\<close>
+ \<open>bit (numeral (num.Bit0 w) :: int) (numeral n) \<longleftrightarrow> bit (numeral w :: int) (pred_numeral n)\<close>
+ \<open>bit (numeral (num.Bit1 w) :: int) (numeral n) \<longleftrightarrow> bit (numeral w :: int) (pred_numeral n)\<close>
+ \<open>bit (numeral (Num.BitM w) :: int) (numeral n) \<longleftrightarrow> \<not> bit (- numeral w :: int) (pred_numeral n)\<close>
+ \<open>bit (- numeral (num.Bit0 w) :: int) (numeral n) \<longleftrightarrow> bit (- numeral w :: int) (pred_numeral n)\<close>
+ \<open>bit (- numeral (num.Bit1 w) :: int) (numeral n) \<longleftrightarrow> \<not> bit (numeral w :: int) (pred_numeral n)\<close>
+ \<open>bit (- numeral (Num.BitM w) :: int) (numeral n) \<longleftrightarrow> bit (- (numeral w) :: int) (pred_numeral n)\<close>
+ 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]:
+ \<open>bit (numeral (Num.Bit0 m) :: int) (Suc n) \<longleftrightarrow> bit (numeral m :: int) n\<close>
+ by (simp add: bit_Suc)
+
+lemma bit_numeral_Bit1_Suc_iff [simp]:
+ \<open>bit (numeral (Num.Bit1 m) :: int) (Suc n) \<longleftrightarrow> bit (numeral m :: int) n\<close>
+ by (simp add: bit_Suc)
+
+lemma push_bit_nat_eq:
+ \<open>push_bit n (nat k) = nat (push_bit n k)\<close>
+ by (cases \<open>k \<ge> 0\<close>) (simp_all add: push_bit_eq_mult nat_mult_distrib not_le mult_nonneg_nonpos2)
+
+lemma drop_bit_nat_eq:
+ \<open>drop_bit n (nat k) = nat (drop_bit n k)\<close>
+ apply (cases \<open>k \<ge> 0\<close>)
+ 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:
+ \<open>take_bit n (nat k) = nat (take_bit n k)\<close> if \<open>k \<ge> 0\<close>
+ using that by (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq)
+
+lemma nat_take_bit_eq:
+ \<open>nat (take_bit n k) = take_bit n (nat k)\<close>
+ if \<open>k \<ge> 0\<close>
+ using that by (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq)
+
+lemma not_exp_less_eq_0_int [simp]:
+ \<open>\<not> 2 ^ n \<le> (0::int)\<close>
+ by (simp add: power_le_zero_eq)
+
+lemma half_nonnegative_int_iff [simp]:
+ \<open>k div 2 \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
+proof (cases \<open>k \<ge> 0\<close>)
+ 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]:
+ \<open>k div 2 < 0 \<longleftrightarrow> k < 0\<close> 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:
+ \<open>- 1 div (2 :: int) ^ n = - 1\<close>
+ by (induction n) (use div_exp_eq [symmetric, of \<open>- 1 :: int\<close> 1] in \<open>simp_all add: ac_simps\<close>)
+
+lemma drop_bit_minus_one [simp]:
+ \<open>drop_bit n (- 1 :: int) = - 1\<close>
+ by (simp add: drop_bit_eq_div minus_1_div_exp_eq_int)
+
+lemma take_bit_Suc_from_most:
+ \<open>take_bit (Suc n) k = 2 ^ n * of_bool (bit k n) + take_bit n k\<close> 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:
+ \<open>take_bit n (- take_bit n k) = take_bit n (- k)\<close>
+ for k :: int
+ by (simp add: take_bit_eq_mod mod_minus_eq)
+
+lemma take_bit_diff:
+ \<open>take_bit n (take_bit n k - take_bit n l) = take_bit n (k - l)\<close>
+ for k l :: int
+ by (simp add: take_bit_eq_mod mod_diff_eq)
+
+lemma bit_imp_take_bit_positive:
+ \<open>0 < take_bit m k\<close> if \<open>n < m\<close> and \<open>bit k n\<close> for k :: int
+proof (rule ccontr)
+ assume \<open>\<not> 0 < take_bit m k\<close>
+ then have \<open>take_bit m k = 0\<close>
+ by (auto simp add: not_less intro: order_antisym)
+ then have \<open>bit (take_bit m k) n = bit 0 n\<close>
+ by simp
+ with that show False
+ by (simp add: bit_take_bit_iff)
+qed
+
+lemma take_bit_mult:
+ \<open>take_bit n (take_bit n k * take_bit n l) = take_bit n (k * l)\<close>
+ 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]:
+ \<open>of_nat (nat (take_bit n k)) = of_int (take_bit n k)\<close>
+ by simp
+
+lemma take_bit_minus_small_eq:
+ \<open>take_bit n (- k) = 2 ^ n - k\<close> if \<open>0 < k\<close> \<open>k \<le> 2 ^ n\<close> for k :: int
+proof -
+ define m where \<open>m = nat k\<close>
+ with that have \<open>k = int m\<close> and \<open>0 < m\<close> and \<open>m \<le> 2 ^ n\<close>
+ by simp_all
+ have \<open>(2 ^ n - m) mod 2 ^ n = 2 ^ n - m\<close>
+ using \<open>0 < m\<close> by simp
+ then have \<open>int ((2 ^ n - m) mod 2 ^ n) = int (2 ^ n - m)\<close>
+ by simp
+ then have \<open>(2 ^ n - int m) mod 2 ^ n = 2 ^ n - int m\<close>
+ using \<open>m \<le> 2 ^ n\<close> by (simp only: of_nat_mod of_nat_diff) simp
+ with \<open>k = int m\<close> have \<open>(2 ^ n - k) mod 2 ^ n = 2 ^ n - k\<close>
+ by simp
+ then show ?thesis
+ by (simp add: take_bit_eq_mod)
+qed
+
+lemma drop_bit_push_bit_int:
+ \<open>drop_bit m (push_bit n k) = drop_bit (m - n) (push_bit (n - m) k)\<close> for k :: int
+ by (cases \<open>m \<le> n\<close>) (auto simp add: mult.left_commute [of _ \<open>2 ^ n\<close>] mult.commute [of _ \<open>2 ^ n\<close>] 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]:
+ \<open>push_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
+ by (simp add: push_bit_eq_mult zero_le_mult_iff)
+
+lemma push_bit_negative_int_iff [simp]:
+ \<open>push_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
+ by (subst Not_eq_iff [symmetric]) (simp add: not_less)
+
+lemma drop_bit_nonnegative_int_iff [simp]:
+ \<open>drop_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
+ by (induction n) (simp_all add: drop_bit_Suc drop_bit_half)
+
+lemma drop_bit_negative_int_iff [simp]:
+ \<open>drop_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
+ by (subst Not_eq_iff [symmetric]) (simp add: not_less)
+
+
+subsection \<open>Bit operations\<close>
+
+class semiring_bit_operations = semiring_bit_shifts +
+ fixes "and" :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close> (infixr \<open>AND\<close> 64)
+ and or :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close> (infixr \<open>OR\<close> 59)
+ and xor :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close> (infixr \<open>XOR\<close> 59)
+ and mask :: \<open>nat \<Rightarrow> 'a\<close>
+ and set_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
+ and unset_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
+ and flip_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
+ assumes bit_and_iff [bit_simps]: \<open>bit (a AND b) n \<longleftrightarrow> bit a n \<and> bit b n\<close>
+ and bit_or_iff [bit_simps]: \<open>bit (a OR b) n \<longleftrightarrow> bit a n \<or> bit b n\<close>
+ and bit_xor_iff [bit_simps]: \<open>bit (a XOR b) n \<longleftrightarrow> bit a n \<noteq> bit b n\<close>
+ and mask_eq_exp_minus_1: \<open>mask n = 2 ^ n - 1\<close>
+ and set_bit_eq_or: \<open>set_bit n a = a OR push_bit n 1\<close>
+ and bit_unset_bit_iff [bit_simps]: \<open>bit (unset_bit m a) n \<longleftrightarrow> bit a n \<and> m \<noteq> n\<close>
+ and flip_bit_eq_xor: \<open>flip_bit n a = a XOR push_bit n 1\<close>
+begin
+
+text \<open>
+ We want the bitwise operations to bind slightly weaker
+ than \<open>+\<close> and \<open>-\<close>.
+ For the sake of code generation
+ the operations \<^const>\<open>and\<close>, \<^const>\<open>or\<close> and \<^const>\<open>xor\<close>
+ are specified as definitional class operations.
+\<close>
+
+sublocale "and": semilattice \<open>(AND)\<close>
+ by standard (auto simp add: bit_eq_iff bit_and_iff)
+
+sublocale or: semilattice_neutr \<open>(OR)\<close> 0
+ by standard (auto simp add: bit_eq_iff bit_or_iff)
+
+sublocale xor: comm_monoid \<open>(XOR)\<close> 0
+ by standard (auto simp add: bit_eq_iff bit_xor_iff)
+
+lemma even_and_iff:
+ \<open>even (a AND b) \<longleftrightarrow> even a \<or> even b\<close>
+ using bit_and_iff [of a b 0] by auto
+
+lemma even_or_iff:
+ \<open>even (a OR b) \<longleftrightarrow> even a \<and> even b\<close>
+ using bit_or_iff [of a b 0] by auto
+
+lemma even_xor_iff:
+ \<open>even (a XOR b) \<longleftrightarrow> (even a \<longleftrightarrow> even b)\<close>
+ using bit_xor_iff [of a b 0] by auto
+
+lemma zero_and_eq [simp]:
+ \<open>0 AND a = 0\<close>
+ by (simp add: bit_eq_iff bit_and_iff)
+
+lemma and_zero_eq [simp]:
+ \<open>a AND 0 = 0\<close>
+ by (simp add: bit_eq_iff bit_and_iff)
+
+lemma one_and_eq:
+ \<open>1 AND a = a mod 2\<close>
+ by (simp add: bit_eq_iff bit_and_iff) (auto simp add: bit_1_iff)
+
+lemma and_one_eq:
+ \<open>a AND 1 = a mod 2\<close>
+ using one_and_eq [of a] by (simp add: ac_simps)
+
+lemma one_or_eq:
+ \<open>1 OR a = a + of_bool (even a)\<close>
+ 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:
+ \<open>a OR 1 = a + of_bool (even a)\<close>
+ using one_or_eq [of a] by (simp add: ac_simps)
+
+lemma one_xor_eq:
+ \<open>1 XOR a = a + of_bool (even a) - of_bool (odd a)\<close>
+ 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:
+ \<open>a XOR 1 = a + of_bool (even a) - of_bool (odd a)\<close>
+ using one_xor_eq [of a] by (simp add: ac_simps)
+
+lemma take_bit_and [simp]:
+ \<open>take_bit n (a AND b) = take_bit n a AND take_bit n b\<close>
+ by (auto simp add: bit_eq_iff bit_take_bit_iff bit_and_iff)
+
+lemma take_bit_or [simp]:
+ \<open>take_bit n (a OR b) = take_bit n a OR take_bit n b\<close>
+ by (auto simp add: bit_eq_iff bit_take_bit_iff bit_or_iff)
+
+lemma take_bit_xor [simp]:
+ \<open>take_bit n (a XOR b) = take_bit n a XOR take_bit n b\<close>
+ by (auto simp add: bit_eq_iff bit_take_bit_iff bit_xor_iff)
+
+lemma push_bit_and [simp]:
+ \<open>push_bit n (a AND b) = push_bit n a AND push_bit n b\<close>
+ by (rule bit_eqI) (auto simp add: bit_push_bit_iff bit_and_iff)
+
+lemma push_bit_or [simp]:
+ \<open>push_bit n (a OR b) = push_bit n a OR push_bit n b\<close>
+ by (rule bit_eqI) (auto simp add: bit_push_bit_iff bit_or_iff)
+
+lemma push_bit_xor [simp]:
+ \<open>push_bit n (a XOR b) = push_bit n a XOR push_bit n b\<close>
+ by (rule bit_eqI) (auto simp add: bit_push_bit_iff bit_xor_iff)
+
+lemma drop_bit_and [simp]:
+ \<open>drop_bit n (a AND b) = drop_bit n a AND drop_bit n b\<close>
+ by (rule bit_eqI) (auto simp add: bit_drop_bit_eq bit_and_iff)
+
+lemma drop_bit_or [simp]:
+ \<open>drop_bit n (a OR b) = drop_bit n a OR drop_bit n b\<close>
+ by (rule bit_eqI) (auto simp add: bit_drop_bit_eq bit_or_iff)
+
+lemma drop_bit_xor [simp]:
+ \<open>drop_bit n (a XOR b) = drop_bit n a XOR drop_bit n b\<close>
+ by (rule bit_eqI) (auto simp add: bit_drop_bit_eq bit_xor_iff)
+
+lemma bit_mask_iff [bit_simps]:
+ \<open>bit (mask m) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n < m\<close>
+ by (simp add: mask_eq_exp_minus_1 bit_mask_iff)
+
+lemma even_mask_iff:
+ \<open>even (mask n) \<longleftrightarrow> n = 0\<close>
+ using bit_mask_iff [of n 0] by auto
+
+lemma mask_0 [simp]:
+ \<open>mask 0 = 0\<close>
+ by (simp add: mask_eq_exp_minus_1)
+
+lemma mask_Suc_0 [simp]:
+ \<open>mask (Suc 0) = 1\<close>
+ by (simp add: mask_eq_exp_minus_1 add_implies_diff sym)
+
+lemma mask_Suc_exp:
+ \<open>mask (Suc n) = 2 ^ n OR mask n\<close>
+ 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:
+ \<open>mask (Suc n) = 1 OR 2 * mask n\<close>
+proof (rule bit_eqI)
+ fix q
+ assume \<open>2 ^ q \<noteq> 0\<close>
+ show \<open>bit (mask (Suc n)) q \<longleftrightarrow> bit (1 OR 2 * mask n) q\<close>
+ 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:
+ \<open>mask (numeral n) = 1 + 2 * mask (pred_numeral n)\<close>
+ by (simp add: numeral_eq_Suc mask_Suc_double one_or_eq ac_simps)
+
+lemma take_bit_mask [simp]:
+ \<open>take_bit m (mask n) = mask (min m n)\<close>
+ by (rule bit_eqI) (simp add: bit_simps)
+
+lemma take_bit_eq_mask:
+ \<open>take_bit n a = a AND mask n\<close>
+ by (rule bit_eqI)
+ (auto simp add: bit_take_bit_iff bit_and_iff bit_mask_iff)
+
+lemma or_eq_0_iff:
+ \<open>a OR b = 0 \<longleftrightarrow> a = 0 \<and> b = 0\<close>
+ by (auto simp add: bit_eq_iff bit_or_iff)
+
+lemma disjunctive_add:
+ \<open>a + b = a OR b\<close> if \<open>\<And>n. \<not> bit a n \<or> \<not> bit b n\<close>
+ by (rule bit_eqI) (use that in \<open>simp add: bit_disjunctive_add_iff bit_or_iff\<close>)
+
+lemma bit_iff_and_drop_bit_eq_1:
+ \<open>bit a n \<longleftrightarrow> drop_bit n a AND 1 = 1\<close>
+ 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:
+ \<open>bit a n \<longleftrightarrow> a AND push_bit n 1 \<noteq> 0\<close>
+ apply (cases \<open>2 ^ n = 0\<close>)
+ 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]:
+ \<open>bit (set_bit m a) n \<longleftrightarrow> bit a n \<or> (m = n \<and> 2 ^ n \<noteq> 0)\<close>
+ by (auto simp add: set_bit_def push_bit_of_1 bit_or_iff bit_exp_iff)
+
+lemma even_set_bit_iff:
+ \<open>even (set_bit m a) \<longleftrightarrow> even a \<and> m \<noteq> 0\<close>
+ using bit_set_bit_iff [of m a 0] by auto
+
+lemma even_unset_bit_iff:
+ \<open>even (unset_bit m a) \<longleftrightarrow> even a \<or> m = 0\<close>
+ using bit_unset_bit_iff [of m a 0] by auto
+
+lemma and_exp_eq_0_iff_not_bit:
+ \<open>a AND 2 ^ n = 0 \<longleftrightarrow> \<not> bit a n\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
+proof
+ assume ?Q
+ then show ?P
+ by (auto intro: bit_eqI simp add: bit_simps)
+next
+ assume ?P
+ show ?Q
+ proof (rule notI)
+ assume \<open>bit a n\<close>
+ then have \<open>a AND 2 ^ n = 2 ^ n\<close>
+ by (auto intro: bit_eqI simp add: bit_simps)
+ with \<open>?P\<close> show False
+ using \<open>bit a n\<close> 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]:
+ \<open>bit (flip_bit m a) n \<longleftrightarrow> (m = n \<longleftrightarrow> \<not> bit a n) \<and> 2 ^ n \<noteq> 0\<close>
+ 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:
+ \<open>even (flip_bit m a) \<longleftrightarrow> \<not> (even a \<longleftrightarrow> m = 0)\<close>
+ using bit_flip_bit_iff [of m a 0] by auto
+
+lemma set_bit_0 [simp]:
+ \<open>set_bit 0 a = 1 + 2 * (a div 2)\<close>
+proof (rule bit_eqI)
+ fix m
+ assume *: \<open>2 ^ m \<noteq> 0\<close>
+ then show \<open>bit (set_bit 0 a) m = bit (1 + 2 * (a div 2)) m\<close>
+ 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:
+ \<open>set_bit (Suc n) a = a mod 2 + 2 * set_bit n (a div 2)\<close>
+proof (rule bit_eqI)
+ fix m
+ assume *: \<open>2 ^ m \<noteq> 0\<close>
+ show \<open>bit (set_bit (Suc n) a) m \<longleftrightarrow> bit (a mod 2 + 2 * set_bit n (a div 2)) m\<close>
+ proof (cases m)
+ case 0
+ then show ?thesis
+ by (simp add: even_set_bit_iff)
+ next
+ case (Suc m)
+ with * have \<open>2 ^ m \<noteq> 0\<close>
+ 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 \<open>2 ^ m \<noteq> 0\<close> bit_Suc)
+ qed
+qed
+
+lemma unset_bit_0 [simp]:
+ \<open>unset_bit 0 a = 2 * (a div 2)\<close>
+proof (rule bit_eqI)
+ fix m
+ assume *: \<open>2 ^ m \<noteq> 0\<close>
+ then show \<open>bit (unset_bit 0 a) m = bit (2 * (a div 2)) m\<close>
+ by (simp add: bit_unset_bit_iff bit_double_iff)
+ (cases m, simp_all add: bit_Suc)
+qed
+
+lemma unset_bit_Suc:
+ \<open>unset_bit (Suc n) a = a mod 2 + 2 * unset_bit n (a div 2)\<close>
+proof (rule bit_eqI)
+ fix m
+ assume *: \<open>2 ^ m \<noteq> 0\<close>
+ then show \<open>bit (unset_bit (Suc n) a) m \<longleftrightarrow> bit (a mod 2 + 2 * unset_bit n (a div 2)) m\<close>
+ 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]:
+ \<open>flip_bit 0 a = of_bool (even a) + 2 * (a div 2)\<close>
+proof (rule bit_eqI)
+ fix m
+ assume *: \<open>2 ^ m \<noteq> 0\<close>
+ then show \<open>bit (flip_bit 0 a) m = bit (of_bool (even a) + 2 * (a div 2)) m\<close>
+ 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:
+ \<open>flip_bit (Suc n) a = a mod 2 + 2 * flip_bit n (a div 2)\<close>
+proof (rule bit_eqI)
+ fix m
+ assume *: \<open>2 ^ m \<noteq> 0\<close>
+ show \<open>bit (flip_bit (Suc n) a) m \<longleftrightarrow> bit (a mod 2 + 2 * flip_bit n (a div 2)) m\<close>
+ proof (cases m)
+ case 0
+ then show ?thesis
+ by (simp add: even_flip_bit_iff)
+ next
+ case (Suc m)
+ with * have \<open>2 ^ m \<noteq> 0\<close>
+ 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 \<open>2 ^ m \<noteq> 0\<close> bit_Suc)
+ qed
+qed
+
+lemma flip_bit_eq_if:
+ \<open>flip_bit n a = (if bit a n then unset_bit else set_bit) n a\<close>
+ 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:
+ \<open>take_bit n (set_bit m a) = (if n \<le> m then take_bit n a else set_bit m (take_bit n a))\<close>
+ by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_set_bit_iff)
+
+lemma take_bit_unset_bit_eq:
+ \<open>take_bit n (unset_bit m a) = (if n \<le> m then take_bit n a else unset_bit m (take_bit n a))\<close>
+ by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_unset_bit_iff)
+
+lemma take_bit_flip_bit_eq:
+ \<open>take_bit n (flip_bit m a) = (if n \<le> m then take_bit n a else flip_bit m (take_bit n a))\<close>
+ 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 :: \<open>'a \<Rightarrow> 'a\<close> (\<open>NOT\<close>)
+ assumes bit_not_iff [bit_simps]: \<open>\<And>n. bit (NOT a) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> \<not> bit a n\<close>
+ assumes minus_eq_not_minus_1: \<open>- a = NOT (a - 1)\<close>
+begin
+
+text \<open>
+ For the sake of code generation \<^const>\<open>not\<close> is specified as
+ definitional class operation. Note that \<^const>\<open>not\<close> has no
+ sensible definition for unlimited but only positive bit strings
+ (type \<^typ>\<open>nat\<close>).
+\<close>
+
+lemma bits_minus_1_mod_2_eq [simp]:
+ \<open>(- 1) mod 2 = 1\<close>
+ by (simp add: mod_2_eq_odd)
+
+lemma not_eq_complement:
+ \<open>NOT a = - a - 1\<close>
+ using minus_eq_not_minus_1 [of \<open>a + 1\<close>] by simp
+
+lemma minus_eq_not_plus_1:
+ \<open>- a = NOT a + 1\<close>
+ using not_eq_complement [of a] by simp
+
+lemma bit_minus_iff [bit_simps]:
+ \<open>bit (- a) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> \<not> bit (a - 1) n\<close>
+ by (simp add: minus_eq_not_minus_1 bit_not_iff)
+
+lemma even_not_iff [simp]:
+ \<open>even (NOT a) \<longleftrightarrow> odd a\<close>
+ using bit_not_iff [of a 0] by auto
+
+lemma bit_not_exp_iff [bit_simps]:
+ \<open>bit (NOT (2 ^ m)) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n \<noteq> m\<close>
+ by (auto simp add: bit_not_iff bit_exp_iff)
+
+lemma bit_minus_1_iff [simp]:
+ \<open>bit (- 1) n \<longleftrightarrow> 2 ^ n \<noteq> 0\<close>
+ by (simp add: bit_minus_iff)
+
+lemma bit_minus_exp_iff [bit_simps]:
+ \<open>bit (- (2 ^ m)) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n \<ge> m\<close>
+ by (auto simp add: bit_simps simp flip: mask_eq_exp_minus_1)
+
+lemma bit_minus_2_iff [simp]:
+ \<open>bit (- 2) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n > 0\<close>
+ by (simp add: bit_minus_iff bit_1_iff)
+
+lemma not_one [simp]:
+ \<open>NOT 1 = - 2\<close>
+ by (simp add: bit_eq_iff bit_not_iff) (simp add: bit_1_iff)
+
+sublocale "and": semilattice_neutr \<open>(AND)\<close> \<open>- 1\<close>
+ by standard (rule bit_eqI, simp add: bit_and_iff)
+
+sublocale bit: boolean_algebra \<open>(AND)\<close> \<open>(OR)\<close> NOT 0 \<open>- 1\<close>
+ rewrites \<open>bit.xor = (XOR)\<close>
+proof -
+ interpret bit: boolean_algebra \<open>(AND)\<close> \<open>(OR)\<close> NOT 0 \<open>- 1\<close>
+ by standard (auto simp add: bit_and_iff bit_or_iff bit_not_iff intro: bit_eqI)
+ show \<open>boolean_algebra (AND) (OR) NOT 0 (- 1)\<close>
+ by standard
+ show \<open>boolean_algebra.xor (AND) (OR) NOT = (XOR)\<close>
+ 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:
+ \<open>a AND b = NOT (NOT a OR NOT b)\<close>
+ by simp
+
+lemma or_eq_not_not_and:
+ \<open>a OR b = NOT (NOT a AND NOT b)\<close>
+ by simp
+
+lemma not_add_distrib:
+ \<open>NOT (a + b) = NOT a - b\<close>
+ by (simp add: not_eq_complement algebra_simps)
+
+lemma not_diff_distrib:
+ \<open>NOT (a - b) = NOT a + b\<close>
+ using not_add_distrib [of a \<open>- b\<close>] by simp
+
+lemma (in ring_bit_operations) and_eq_minus_1_iff:
+ \<open>a AND b = - 1 \<longleftrightarrow> a = - 1 \<and> b = - 1\<close>
+proof
+ assume \<open>a = - 1 \<and> b = - 1\<close>
+ then show \<open>a AND b = - 1\<close>
+ by simp
+next
+ assume \<open>a AND b = - 1\<close>
+ have *: \<open>bit a n\<close> \<open>bit b n\<close> if \<open>2 ^ n \<noteq> 0\<close> for n
+ proof -
+ from \<open>a AND b = - 1\<close>
+ have \<open>bit (a AND b) n = bit (- 1) n\<close>
+ by (simp add: bit_eq_iff)
+ then show \<open>bit a n\<close> \<open>bit b n\<close>
+ using that by (simp_all add: bit_and_iff)
+ qed
+ have \<open>a = - 1\<close>
+ by (rule bit_eqI) (simp add: *)
+ moreover have \<open>b = - 1\<close>
+ by (rule bit_eqI) (simp add: *)
+ ultimately show \<open>a = - 1 \<and> b = - 1\<close>
+ by simp
+qed
+
+lemma disjunctive_diff:
+ \<open>a - b = a AND NOT b\<close> if \<open>\<And>n. bit b n \<Longrightarrow> bit a n\<close>
+proof -
+ have \<open>NOT a + b = NOT a OR b\<close>
+ by (rule disjunctive_add) (auto simp add: bit_not_iff dest: that)
+ then have \<open>NOT (NOT a + b) = NOT (NOT a OR b)\<close>
+ by simp
+ then show ?thesis
+ by (simp add: not_add_distrib)
+qed
+
+lemma push_bit_minus:
+ \<open>push_bit n (- a) = - push_bit n a\<close>
+ by (simp add: push_bit_eq_mult)
+
+lemma take_bit_not_take_bit:
+ \<open>take_bit n (NOT (take_bit n a)) = take_bit n (NOT a)\<close>
+ by (auto simp add: bit_eq_iff bit_take_bit_iff bit_not_iff)
+
+lemma take_bit_not_iff:
+ \<open>take_bit n (NOT a) = take_bit n (NOT b) \<longleftrightarrow> take_bit n a = take_bit n b\<close>
+ 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:
+ \<open>take_bit n (NOT a) = mask n - take_bit n a\<close>
+proof -
+ have \<open>take_bit n (NOT a) = take_bit n (NOT (take_bit n a))\<close>
+ by (simp add: take_bit_not_take_bit)
+ also have \<open>\<dots> = mask n AND NOT (take_bit n a)\<close>
+ by (simp add: take_bit_eq_mask ac_simps)
+ also have \<open>\<dots> = mask n - take_bit n a\<close>
+ 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:
+ \<open>mask n = take_bit n (- 1)\<close>
+ by (simp add: bit_eq_iff bit_mask_iff bit_take_bit_iff conj_commute)
+
+lemma take_bit_minus_one_eq_mask:
+ \<open>take_bit n (- 1) = mask n\<close>
+ by (simp add: mask_eq_take_bit_minus_one)
+
+lemma minus_exp_eq_not_mask:
+ \<open>- (2 ^ n) = NOT (mask n)\<close>
+ 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:
+ \<open>push_bit n (- 1) = NOT (mask n)\<close>
+ by (simp add: push_bit_eq_mult minus_exp_eq_not_mask)
+
+lemma take_bit_not_mask_eq_0:
+ \<open>take_bit m (NOT (mask n)) = 0\<close> if \<open>n \<ge> m\<close>
+ by (rule bit_eqI) (use that in \<open>simp add: bit_take_bit_iff bit_not_iff bit_mask_iff\<close>)
+
+lemma unset_bit_eq_and_not:
+ \<open>unset_bit n a = a AND NOT (push_bit n 1)\<close>
+ by (rule bit_eqI) (auto simp add: bit_simps)
+
+lemmas unset_bit_def = unset_bit_eq_and_not
+
+end
+
+
+subsection \<open>Instance \<^typ>\<open>int\<close>\<close>
+
+lemma int_bit_bound:
+ fixes k :: int
+ obtains n where \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k n\<close>
+ and \<open>n > 0 \<Longrightarrow> bit k (n - 1) \<noteq> bit k n\<close>
+proof -
+ obtain q where *: \<open>\<And>m. q \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k q\<close>
+ proof (cases \<open>k \<ge> 0\<close>)
+ case True
+ moreover from power_gt_expt [of 2 \<open>nat k\<close>]
+ have \<open>nat k < 2 ^ nat k\<close>
+ by simp
+ then have \<open>int (nat k) < int (2 ^ nat k)\<close>
+ by (simp only: of_nat_less_iff)
+ ultimately have *: \<open>k div 2 ^ nat k = 0\<close>
+ by simp
+ show thesis
+ proof (rule that [of \<open>nat k\<close>])
+ fix m
+ assume \<open>nat k \<le> m\<close>
+ then show \<open>bit k m \<longleftrightarrow> bit k (nat k)\<close>
+ 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 \<open>nat (- k)\<close>]
+ have \<open>nat (- k) < 2 ^ nat (- k)\<close>
+ by simp
+ then have \<open>int (nat (- k)) < int (2 ^ nat (- k))\<close>
+ by (simp only: of_nat_less_iff)
+ ultimately have \<open>- k div - (2 ^ nat (- k)) = - 1\<close>
+ by (subst div_pos_neg_trivial) simp_all
+ then have *: \<open>k div 2 ^ nat (- k) = - 1\<close>
+ by simp
+ show thesis
+ proof (rule that [of \<open>nat (- k)\<close>])
+ fix m
+ assume \<open>nat (- k) \<le> m\<close>
+ then show \<open>bit k m \<longleftrightarrow> bit k (nat (- k))\<close>
+ 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 \<open>\<forall>m. bit k m \<longleftrightarrow> bit k q\<close>)
+ case True
+ then have \<open>bit k 0 \<longleftrightarrow> bit k q\<close>
+ by blast
+ with True that [of 0] show thesis
+ by simp
+ next
+ case False
+ then obtain r where **: \<open>bit k r \<noteq> bit k q\<close>
+ by blast
+ have \<open>r < q\<close>
+ by (rule ccontr) (use * [of r] ** in simp)
+ define N where \<open>N = {n. n < q \<and> bit k n \<noteq> bit k q}\<close>
+ moreover have \<open>finite N\<close> \<open>r \<in> N\<close>
+ using ** N_def \<open>r < q\<close> by auto
+ moreover define n where \<open>n = Suc (Max N)\<close>
+ ultimately have \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k n\<close>
+ apply auto
+ apply (metis (full_types, lifting) "*" Max_ge_iff Suc_n_not_le_n \<open>finite N\<close> all_not_in_conv mem_Collect_eq not_le)
+ apply (metis "*" Max_ge Suc_n_not_le_n \<open>finite N\<close> linorder_not_less mem_Collect_eq)
+ apply (metis "*" Max_ge Suc_n_not_le_n \<open>finite N\<close> linorder_not_less mem_Collect_eq)
+ apply (metis (full_types, lifting) "*" Max_ge_iff Suc_n_not_le_n \<open>finite N\<close> all_not_in_conv mem_Collect_eq not_le)
+ done
+ have \<open>bit k (Max N) \<noteq> bit k n\<close>
+ by (metis (mono_tags, lifting) "*" Max_in N_def \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m = bit k n\<close> \<open>finite N\<close> \<open>r \<in> N\<close> empty_iff le_cases mem_Collect_eq)
+ show thesis apply (rule that [of n])
+ using \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m = bit k n\<close> apply blast
+ using \<open>bit k (Max N) \<noteq> bit k n\<close> n_def by auto
+ qed
+qed
+
+instantiation int :: ring_bit_operations
+begin
+
+definition not_int :: \<open>int \<Rightarrow> int\<close>
+ where \<open>not_int k = - k - 1\<close>
+
+lemma not_int_rec:
+ \<open>NOT k = of_bool (even k) + 2 * NOT (k div 2)\<close> for k :: int
+ by (auto simp add: not_int_def elim: oddE)
+
+lemma even_not_iff_int:
+ \<open>even (NOT k) \<longleftrightarrow> odd k\<close> for k :: int
+ by (simp add: not_int_def)
+
+lemma not_int_div_2:
+ \<open>NOT k div 2 = NOT (k div 2)\<close> 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]:
+ \<open>bit (NOT k) n \<longleftrightarrow> \<not> bit k n\<close>
+ for k :: int
+ by (simp add: bit_not_int_iff' not_int_def)
+
+function and_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>
+ where \<open>(k::int) AND l = (if k \<in> {0, - 1} \<and> l \<in> {0, - 1}
+ then - of_bool (odd k \<and> odd l)
+ else of_bool (odd k \<and> odd l) + 2 * ((k div 2) AND (l div 2)))\<close>
+ by auto
+
+termination proof (relation \<open>measure (\<lambda>(k, l). nat (\<bar>k\<bar> + \<bar>l\<bar>))\<close>)
+ show \<open>wf (measure (\<lambda>(k, l). nat (\<bar>k\<bar> + \<bar>l\<bar>)))\<close>
+ by simp
+ show \<open>((k div 2, l div 2), k, l) \<in> measure (\<lambda>(k, l). nat (\<bar>k\<bar> + \<bar>l\<bar>))\<close>
+ if \<open>\<not> (k \<in> {0, - 1} \<and> l \<in> {0, - 1})\<close> for k l
+ proof -
+ have less_eq: \<open>\<bar>k div 2\<bar> \<le> \<bar>k\<bar>\<close> for k :: int
+ by (cases k) (simp_all add: divide_int_def nat_add_distrib)
+ have less: \<open>\<bar>k div 2\<bar> < \<bar>k\<bar>\<close> if \<open>k \<notin> {0, - 1}\<close> 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 \<open>n \<noteq> 0\<close>
+ by simp
+ then have \<open>n div 2 < n\<close>
+ 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 *: \<open>k \<notin> {0, - 1} \<or> l \<notin> {0, - 1}\<close>
+ by simp
+ then have \<open>0 < \<bar>k\<bar> + \<bar>l\<bar>\<close>
+ by auto
+ moreover from * have \<open>\<bar>k div 2\<bar> + \<bar>l div 2\<bar> < \<bar>k\<bar> + \<bar>l\<bar>\<close>
+ proof
+ assume \<open>k \<notin> {0, - 1}\<close>
+ then have \<open>\<bar>k div 2\<bar> < \<bar>k\<bar>\<close>
+ by (rule less)
+ with less_eq [of l] show ?thesis
+ by auto
+ next
+ assume \<open>l \<notin> {0, - 1}\<close>
+ then have \<open>\<bar>l div 2\<bar> < \<bar>l\<bar>\<close>
+ 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:
+ \<open>k AND l = of_bool (odd k \<and> odd l) + 2 * ((k div 2) AND (l div 2))\<close>
+ for k l :: int
+proof (cases \<open>k \<in> {0, - 1} \<and> l \<in> {0, - 1}\<close>)
+ 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:
+ \<open>bit (k AND l) n \<longleftrightarrow> bit k n \<and> bit l n\<close> 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:
+ \<open>even (k AND l) \<longleftrightarrow> even k \<or> even l\<close> for k l :: int
+ using bit_and_int_iff [of k l 0] by auto
+
+definition or_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>
+ where \<open>k OR l = NOT (NOT k AND NOT l)\<close> for k l :: int
+
+lemma or_int_rec:
+ \<open>k OR l = of_bool (odd k \<or> odd l) + 2 * ((k div 2) OR (l div 2))\<close>
+ for k l :: int
+ using and_int_rec [of \<open>NOT k\<close> \<open>NOT l\<close>]
+ 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:
+ \<open>bit (k OR l) n \<longleftrightarrow> bit k n \<or> bit l n\<close> for k l :: int
+ by (simp add: or_int_def bit_not_int_iff bit_and_int_iff)
+
+definition xor_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>
+ where \<open>k XOR l = k AND NOT l OR NOT k AND l\<close> for k l :: int
+
+lemma xor_int_rec:
+ \<open>k XOR l = of_bool (odd k \<noteq> odd l) + 2 * ((k div 2) XOR (l div 2))\<close>
+ for k l :: int
+ by (simp add: xor_int_def or_int_rec [of \<open>k AND NOT l\<close> \<open>NOT k AND l\<close>] even_and_iff_int even_not_iff_int)
+ (simp add: and_int_rec [of \<open>NOT k\<close> \<open>l\<close>] and_int_rec [of \<open>k\<close> \<open>NOT l\<close>] not_int_div_2)
+
+lemma bit_xor_int_iff:
+ \<open>bit (k XOR l) n \<longleftrightarrow> bit k n \<noteq> bit l n\<close> 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 :: \<open>nat \<Rightarrow> int\<close>
+ where \<open>mask n = (2 :: int) ^ n - 1\<close>
+
+definition set_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
+ where \<open>set_bit n k = k OR push_bit n 1\<close> for k :: int
+
+definition unset_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
+ where \<open>unset_bit n k = k AND NOT (push_bit n 1)\<close> for k :: int
+
+definition flip_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
+ where \<open>flip_bit n k = k XOR push_bit n 1\<close> for k :: int
+
+instance proof
+ fix k l :: int and m n :: nat
+ show \<open>- k = NOT (k - 1)\<close>
+ by (simp add: not_int_def)
+ show \<open>bit (k AND l) n \<longleftrightarrow> bit k n \<and> bit l n\<close>
+ by (fact bit_and_int_iff)
+ show \<open>bit (k OR l) n \<longleftrightarrow> bit k n \<or> bit l n\<close>
+ by (fact bit_or_int_iff)
+ show \<open>bit (k XOR l) n \<longleftrightarrow> bit k n \<noteq> bit l n\<close>
+ by (fact bit_xor_int_iff)
+ show \<open>bit (unset_bit m k) n \<longleftrightarrow> bit k n \<and> m \<noteq> n\<close>
+ proof -
+ have \<open>unset_bit m k = k AND NOT (push_bit m 1)\<close>
+ by (simp add: unset_bit_int_def)
+ also have \<open>NOT (push_bit m 1 :: int) = - (push_bit m 1 + 1)\<close>
+ 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:
+ \<open>mask n div 2 = (mask (n - 1) :: int)\<close>
+ by (cases n) (simp_all add: mask_eq_exp_minus_1 algebra_simps)
+
+lemma mask_nonnegative_int [simp]:
+ \<open>mask n \<ge> (0::int)\<close>
+ by (simp add: mask_eq_exp_minus_1)
+
+lemma not_mask_negative_int [simp]:
+ \<open>\<not> mask n < (0::int)\<close>
+ by (simp add: not_less)
+
+lemma not_nonnegative_int_iff [simp]:
+ \<open>NOT k \<ge> 0 \<longleftrightarrow> k < 0\<close> for k :: int
+ by (simp add: not_int_def)
+
+lemma not_negative_int_iff [simp]:
+ \<open>NOT k < 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
+ by (subst Not_eq_iff [symmetric]) (simp add: not_less not_le)
+
+lemma and_nonnegative_int_iff [simp]:
+ \<open>k AND l \<ge> 0 \<longleftrightarrow> k \<ge> 0 \<or> l \<ge> 0\<close> 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 \<open>k * 2\<close> l]
+ by (simp add: pos_imp_zdiv_nonneg_iff zero_le_mult_iff)
+next
+ case (odd k)
+ from odd have \<open>0 \<le> k AND l div 2 \<longleftrightarrow> 0 \<le> k \<or> 0 \<le> l div 2\<close>
+ by simp
+ then have \<open>0 \<le> (1 + k * 2) div 2 AND l div 2 \<longleftrightarrow> 0 \<le> (1 + k * 2) div 2 \<or> 0 \<le> l div 2\<close>
+ by simp
+ with and_int_rec [of \<open>1 + k * 2\<close> l]
+ show ?case
+ by (auto simp add: zero_le_mult_iff not_le)
+qed
+
+lemma and_negative_int_iff [simp]:
+ \<open>k AND l < 0 \<longleftrightarrow> k < 0 \<and> l < 0\<close> for k l :: int
+ by (subst Not_eq_iff [symmetric]) (simp add: not_less)
+
+lemma and_less_eq:
+ \<open>k AND l \<le> k\<close> if \<open>l < 0\<close> 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 \<open>l div 2\<close>] even.hyps even.prems
+ show ?case
+ by (simp add: and_int_rec [of _ l])
+next
+ case (odd k)
+ from odd.IH [of \<open>l div 2\<close>] odd.hyps odd.prems
+ show ?case
+ by (simp add: and_int_rec [of _ l])
+qed
+
+lemma or_nonnegative_int_iff [simp]:
+ \<open>k OR l \<ge> 0 \<longleftrightarrow> k \<ge> 0 \<and> l \<ge> 0\<close> for k l :: int
+ by (simp only: or_eq_not_not_and not_nonnegative_int_iff) simp
+
+lemma or_negative_int_iff [simp]:
+ \<open>k OR l < 0 \<longleftrightarrow> k < 0 \<or> l < 0\<close> for k l :: int
+ by (subst Not_eq_iff [symmetric]) (simp add: not_less)
+
+lemma or_greater_eq:
+ \<open>k OR l \<ge> k\<close> if \<open>l \<ge> 0\<close> 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 \<open>l div 2\<close>] even.hyps even.prems
+ show ?case
+ by (simp add: or_int_rec [of _ l])
+next
+ case (odd k)
+ from odd.IH [of \<open>l div 2\<close>] odd.hyps odd.prems
+ show ?case
+ by (simp add: or_int_rec [of _ l])
+qed
+
+lemma xor_nonnegative_int_iff [simp]:
+ \<open>k XOR l \<ge> 0 \<longleftrightarrow> (k \<ge> 0 \<longleftrightarrow> l \<ge> 0)\<close> for k l :: int
+ by (simp only: bit.xor_def or_nonnegative_int_iff) auto
+
+lemma xor_negative_int_iff [simp]:
+ \<open>k XOR l < 0 \<longleftrightarrow> (k < 0) \<noteq> (l < 0)\<close> for k l :: int
+ by (subst Not_eq_iff [symmetric]) (auto simp add: not_less)
+
+lemma OR_upper: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
+ fixes x y :: int
+ assumes \<open>0 \<le> x\<close> \<open>x < 2 ^ n\<close> \<open>y < 2 ^ n\<close>
+ shows \<open>x OR y < 2 ^ n\<close>
+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 \<open>n - 1\<close> \<open>y div 2\<close>] even.prems even.hyps
+ show ?case
+ by (cases n) (auto simp add: or_int_rec [of \<open>_ * 2\<close>] elim: oddE)
+next
+ case (odd x)
+ from odd.IH [of \<open>n - 1\<close> \<open>y div 2\<close>] odd.prems odd.hyps
+ show ?case
+ by (cases n) (auto simp add: or_int_rec [of \<open>1 + _ * 2\<close>], linarith)
+qed
+
+lemma XOR_upper: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
+ fixes x y :: int
+ assumes \<open>0 \<le> x\<close> \<open>x < 2 ^ n\<close> \<open>y < 2 ^ n\<close>
+ shows \<open>x XOR y < 2 ^ n\<close>
+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 \<open>n - 1\<close> \<open>y div 2\<close>] even.prems even.hyps
+ show ?case
+ by (cases n) (auto simp add: xor_int_rec [of \<open>_ * 2\<close>] elim: oddE)
+next
+ case (odd x)
+ from odd.IH [of \<open>n - 1\<close> \<open>y div 2\<close>] odd.prems odd.hyps
+ show ?case
+ by (cases n) (auto simp add: xor_int_rec [of \<open>1 + _ * 2\<close>])
+qed
+
+lemma AND_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
+ fixes x y :: int
+ assumes \<open>0 \<le> x\<close>
+ shows \<open>0 \<le> x AND y\<close>
+ using assms by simp
+
+lemma OR_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
+ fixes x y :: int
+ assumes \<open>0 \<le> x\<close> \<open>0 \<le> y\<close>
+ shows \<open>0 \<le> x OR y\<close>
+ using assms by simp
+
+lemma XOR_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
+ fixes x y :: int
+ assumes \<open>0 \<le> x\<close> \<open>0 \<le> y\<close>
+ shows \<open>0 \<le> x XOR y\<close>
+ using assms by simp
+
+lemma AND_upper1 [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
+ fixes x y :: int
+ assumes \<open>0 \<le> x\<close>
+ shows \<open>x AND y \<le> x\<close>
+using assms proof (induction x arbitrary: y rule: int_bit_induct)
+ case (odd k)
+ then have \<open>k AND y div 2 \<le> k\<close>
+ by simp
+ then show ?case
+ by (simp add: and_int_rec [of \<open>1 + _ * 2\<close>])
+qed (simp_all add: and_int_rec [of \<open>_ * 2\<close>])
+
+lemmas AND_upper1' [simp] = order_trans [OF AND_upper1] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
+lemmas AND_upper1'' [simp] = order_le_less_trans [OF AND_upper1] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
+
+lemma AND_upper2 [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
+ fixes x y :: int
+ assumes \<open>0 \<le> y\<close>
+ shows \<open>x AND y \<le> y\<close>
+ using assms AND_upper1 [of y x] by (simp add: ac_simps)
+
+lemmas AND_upper2' [simp] = order_trans [OF AND_upper2] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
+lemmas AND_upper2'' [simp] = order_le_less_trans [OF AND_upper2] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
+
+lemma plus_and_or: \<open>(x AND y) + (x OR y) = x + y\<close> 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 \<open>y div 2\<close>]
+ 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 \<open>y div 2\<close>]
+ 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]:
+ \<open>set_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
+ by (simp add: set_bit_def)
+
+lemma set_bit_negative_int_iff [simp]:
+ \<open>set_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
+ by (simp add: set_bit_def)
+
+lemma unset_bit_nonnegative_int_iff [simp]:
+ \<open>unset_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
+ by (simp add: unset_bit_def)
+
+lemma unset_bit_negative_int_iff [simp]:
+ \<open>unset_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
+ by (simp add: unset_bit_def)
+
+lemma flip_bit_nonnegative_int_iff [simp]:
+ \<open>flip_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
+ by (simp add: flip_bit_def)
+
+lemma flip_bit_negative_int_iff [simp]:
+ \<open>flip_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
+ by (simp add: flip_bit_def)
+
+lemma set_bit_greater_eq:
+ \<open>set_bit n k \<ge> k\<close> for k :: int
+ by (simp add: set_bit_def or_greater_eq)
+
+lemma unset_bit_less_eq:
+ \<open>unset_bit n k \<le> k\<close> for k :: int
+ by (simp add: unset_bit_def and_less_eq)
+
+lemma set_bit_eq:
+ \<open>set_bit n k = k + of_bool (\<not> bit k n) * 2 ^ n\<close> for k :: int
+proof (rule bit_eqI)
+ fix m
+ show \<open>bit (set_bit n k) m \<longleftrightarrow> bit (k + of_bool (\<not> bit k n) * 2 ^ n) m\<close>
+ proof (cases \<open>m = n\<close>)
+ 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:
+ \<open>unset_bit n k = k - of_bool (bit k n) * 2 ^ n\<close> for k :: int
+proof (rule bit_eqI)
+ fix m
+ show \<open>bit (unset_bit n k) m \<longleftrightarrow> bit (k - of_bool (bit k n) * 2 ^ n) m\<close>
+ proof (cases \<open>m = n\<close>)
+ 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 \<open>2 ^ n\<close> \<open>- (2 ^ n)\<close> 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:
+ \<open>take_bit n k = mask n \<longleftrightarrow> take_bit n (k + 1) = 0\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
+ for k :: int
+proof
+ assume ?P
+ then have \<open>take_bit n (take_bit n k + take_bit n 1) = 0\<close>
+ by (simp add: mask_eq_exp_minus_1)
+ then show ?Q
+ by (simp only: take_bit_add)
+next
+ assume ?Q
+ then have \<open>take_bit n (k + 1) - 1 = - 1\<close>
+ by simp
+ then have \<open>take_bit n (take_bit n (k + 1) - 1) = take_bit n (- 1)\<close>
+ by simp
+ moreover have \<open>take_bit n (take_bit n (k + 1) - 1) = take_bit n k\<close>
+ 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:
+ \<open>take_bit n k = mask n \<longleftrightarrow> 2 ^ n dvd k + 1\<close>
+ 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:
+ \<open>even (of_int k) \<longleftrightarrow> even k\<close>
+ by (induction k rule: int_bit_induct) simp_all
+
+lemma bit_of_int_iff [bit_simps]:
+ \<open>bit (of_int k) n \<longleftrightarrow> (2::'a) ^ n \<noteq> 0 \<and> bit k n\<close>
+proof (cases \<open>(2::'a) ^ n = 0\<close>)
+ case True
+ then show ?thesis
+ by (simp add: exp_eq_0_imp_not_bit)
+next
+ case False
+ then have \<open>bit (of_int k) n \<longleftrightarrow> bit k n\<close>
+ 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 \<open>of_int k\<close> 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 \<open>of_int k\<close> 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:
+ \<open>push_bit n (of_int k) = of_int (push_bit n k)\<close>
+ by (simp add: push_bit_eq_mult semiring_bit_shifts_class.push_bit_eq_mult)
+
+lemma of_int_push_bit:
+ \<open>of_int (push_bit n k) = push_bit n (of_int k)\<close>
+ by (simp add: push_bit_eq_mult semiring_bit_shifts_class.push_bit_eq_mult)
+
+lemma take_bit_of_int:
+ \<open>take_bit n (of_int k) = of_int (take_bit n k)\<close>
+ 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:
+ \<open>of_int (take_bit n k) = take_bit n (of_int k)\<close>
+ 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:
+ \<open>of_int (NOT k) = NOT (of_int k)\<close>
+ by (rule bit_eqI) (simp add: bit_not_iff Bit_Operations.bit_not_iff bit_of_int_iff)
+
+lemma of_int_and_eq:
+ \<open>of_int (k AND l) = of_int k AND of_int l\<close>
+ by (rule bit_eqI) (simp add: bit_of_int_iff bit_and_iff Bit_Operations.bit_and_iff)
+
+lemma of_int_or_eq:
+ \<open>of_int (k OR l) = of_int k OR of_int l\<close>
+ by (rule bit_eqI) (simp add: bit_of_int_iff bit_or_iff Bit_Operations.bit_or_iff)
+
+lemma of_int_xor_eq:
+ \<open>of_int (k XOR l) = of_int k XOR of_int l\<close>
+ by (rule bit_eqI) (simp add: bit_of_int_iff bit_xor_iff Bit_Operations.bit_xor_iff)
+
+lemma of_int_mask_eq:
+ \<open>of_int (mask n) = mask n\<close>
+ by (induction n) (simp_all add: mask_Suc_double Bit_Operations.mask_Suc_double of_int_or_eq)
+
+end
+
+lemma take_bit_incr_eq:
+ \<open>take_bit n (k + 1) = 1 + take_bit n k\<close> if \<open>take_bit n k \<noteq> 2 ^ n - 1\<close>
+ for k :: int
+proof -
+ from that have \<open>2 ^ n \<noteq> k mod 2 ^ n + 1\<close>
+ by (simp add: take_bit_eq_mod)
+ moreover have \<open>k mod 2 ^ n < 2 ^ n\<close>
+ by simp
+ ultimately have *: \<open>k mod 2 ^ n + 1 < 2 ^ n\<close>
+ by linarith
+ have \<open>(k + 1) mod 2 ^ n = (k mod 2 ^ n + 1) mod 2 ^ n\<close>
+ by (simp add: mod_simps)
+ also have \<open>\<dots> = k mod 2 ^ n + 1\<close>
+ using * by (simp add: zmod_trivial_iff)
+ finally have \<open>(k + 1) mod 2 ^ n = k mod 2 ^ n + 1\<close> .
+ then show ?thesis
+ by (simp add: take_bit_eq_mod)
+qed
+
+lemma take_bit_decr_eq:
+ \<open>take_bit n (k - 1) = take_bit n k - 1\<close> if \<open>take_bit n k \<noteq> 0\<close>
+ for k :: int
+proof -
+ from that have \<open>k mod 2 ^ n \<noteq> 0\<close>
+ by (simp add: take_bit_eq_mod)
+ moreover have \<open>k mod 2 ^ n \<ge> 0\<close> \<open>k mod 2 ^ n < 2 ^ n\<close>
+ by simp_all
+ ultimately have *: \<open>k mod 2 ^ n > 0\<close>
+ by linarith
+ have \<open>(k - 1) mod 2 ^ n = (k mod 2 ^ n - 1) mod 2 ^ n\<close>
+ by (simp add: mod_simps)
+ also have \<open>\<dots> = k mod 2 ^ n - 1\<close>
+ by (simp add: zmod_trivial_iff)
+ (use \<open>k mod 2 ^ n < 2 ^ n\<close> * in linarith)
+ finally have \<open>(k - 1) mod 2 ^ n = k mod 2 ^ n - 1\<close> .
+ then show ?thesis
+ by (simp add: take_bit_eq_mod)
+qed
+
+lemma take_bit_int_greater_eq:
+ \<open>k + 2 ^ n \<le> take_bit n k\<close> if \<open>k < 0\<close> for k :: int
+proof -
+ have \<open>k + 2 ^ n \<le> take_bit n (k + 2 ^ n)\<close>
+ proof (cases \<open>k > - (2 ^ n)\<close>)
+ case False
+ then have \<open>k + 2 ^ n \<le> 0\<close>
+ by simp
+ also note take_bit_nonnegative
+ finally show ?thesis .
+ next
+ case True
+ with that have \<open>0 \<le> k + 2 ^ n\<close> and \<open>k + 2 ^ n < 2 ^ n\<close>
+ 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:
+ \<open>take_bit n k \<le> k - 2 ^ n\<close> if \<open>2 ^ n \<le> k\<close> and \<open>n > 0\<close> for k :: int
+ using that zmod_le_nonneg_dividend [of \<open>k - 2 ^ n\<close> \<open>2 ^ n\<close>]
+ by (simp add: take_bit_eq_mod)
+
+lemma take_bit_int_less_eq_self_iff:
+ \<open>take_bit n k \<le> k \<longleftrightarrow> 0 \<le> k\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
+ for k :: int
+proof
+ assume ?P
+ show ?Q
+ proof (rule ccontr)
+ assume \<open>\<not> 0 \<le> k\<close>
+ then have \<open>k < 0\<close>
+ by simp
+ with \<open>?P\<close>
+ have \<open>take_bit n k < 0\<close>
+ 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:
+ \<open>take_bit n k < k \<longleftrightarrow> 2 ^ n \<le> k\<close>
+ 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 \<open>2 ^ n\<close> k])
+
+lemma take_bit_int_greater_self_iff:
+ \<open>k < take_bit n k \<longleftrightarrow> k < 0\<close>
+ for k :: int
+ using take_bit_int_less_eq_self_iff [of n k] by auto
+
+lemma take_bit_int_greater_eq_self_iff:
+ \<open>k \<le> take_bit n k \<longleftrightarrow> k < 2 ^ n\<close>
+ 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 \<open>2 ^ n\<close>])
+
+lemma minus_numeral_inc_eq:
+ \<open>- numeral (Num.inc n) = NOT (numeral n :: int)\<close>
+ by (simp add: not_int_def sub_inc_One_eq add_One)
+
+lemma sub_one_eq_not_neg:
+ \<open>Num.sub n num.One = NOT (- numeral n :: int)\<close>
+ by (simp add: not_int_def)
+
+lemma int_not_numerals [simp]:
+ \<open>NOT (numeral (Num.Bit0 n) :: int) = - numeral (Num.Bit1 n)\<close>
+ \<open>NOT (numeral (Num.Bit1 n) :: int) = - numeral (Num.inc (num.Bit1 n))\<close>
+ \<open>NOT (numeral (Num.BitM n) :: int) = - numeral (num.Bit0 n)\<close>
+ \<open>NOT (- numeral (Num.Bit0 n) :: int) = numeral (Num.BitM n)\<close>
+ \<open>NOT (- numeral (Num.Bit1 n) :: int) = numeral (Num.Bit0 n)\<close>
+ by (simp_all add: not_int_def add_One inc_BitM_eq)
+
+text \<open>FIXME: The rule sets below are very large (24 rules for each
+ operator). Is there a simpler way to do this?\<close>
+
+context
+begin
+
+private lemma eqI:
+ \<open>k = l\<close>
+ if num: \<open>\<And>n. bit k (numeral n) \<longleftrightarrow> bit l (numeral n)\<close>
+ and even: \<open>even k \<longleftrightarrow> even l\<close>
+ for k l :: int
+proof (rule bit_eqI)
+ fix n
+ show \<open>bit k n \<longleftrightarrow> bit l n\<close>
+ proof (cases n)
+ case 0
+ with even show ?thesis
+ by simp
+ next
+ case (Suc n)
+ with num [of \<open>num_of_nat (Suc n)\<close>] show ?thesis
+ by (simp only: numeral_num_of_nat)
+ qed
+qed
+
+lemma int_and_numerals [simp]:
+ \<open>numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND numeral y)\<close>
+ \<open>numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (2 :: int) * (numeral x AND numeral y)\<close>
+ \<open>numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND numeral y)\<close>
+ \<open>numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x AND numeral y)\<close>
+ \<open>numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND - numeral y)\<close>
+ \<open>numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (2 :: int) * (numeral x AND - numeral (y + Num.One))\<close>
+ \<open>numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND - numeral y)\<close>
+ \<open>numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x AND - numeral (y + Num.One))\<close>
+ \<open>- numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (2 :: int) * (- numeral x AND numeral y)\<close>
+ \<open>- numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (2 :: int) * (- numeral x AND numeral y)\<close>
+ \<open>- numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (2 :: int) * (- numeral (x + Num.One) AND numeral y)\<close>
+ \<open>- numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) AND numeral y)\<close>
+ \<open>- numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x AND - numeral y)\<close>
+ \<open>- numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (2 :: int) * (- numeral x AND - numeral (y + Num.One))\<close>
+ \<open>- numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (- numeral (x + Num.One) AND - numeral y)\<close>
+ \<open>- numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) AND - numeral (y + Num.One))\<close>
+ \<open>(1::int) AND numeral (Num.Bit0 y) = 0\<close>
+ \<open>(1::int) AND numeral (Num.Bit1 y) = 1\<close>
+ \<open>(1::int) AND - numeral (Num.Bit0 y) = 0\<close>
+ \<open>(1::int) AND - numeral (Num.Bit1 y) = 1\<close>
+ \<open>numeral (Num.Bit0 x) AND (1::int) = 0\<close>
+ \<open>numeral (Num.Bit1 x) AND (1::int) = 1\<close>
+ \<open>- numeral (Num.Bit0 x) AND (1::int) = 0\<close>
+ \<open>- numeral (Num.Bit1 x) AND (1::int) = 1\<close>
+ 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]:
+ \<open>numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (2 :: int) * (numeral x OR numeral y)\<close>
+ \<open>numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR numeral y)\<close>
+ \<open>numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x OR numeral y)\<close>
+ \<open>numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR numeral y)\<close>
+ \<open>numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (2 :: int) * (numeral x OR - numeral y)\<close>
+ \<open>numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR - numeral (y + Num.One))\<close>
+ \<open>numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x OR - numeral y)\<close>
+ \<open>numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR - numeral (y + Num.One))\<close>
+ \<open>- numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (2 :: int) * (- numeral x OR numeral y)\<close>
+ \<open>- numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x OR numeral y)\<close>
+ \<open>- numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR numeral y)\<close>
+ \<open>- numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR numeral y)\<close>
+ \<open>- numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x OR - numeral y)\<close>
+ \<open>- numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x OR - numeral (y + Num.One))\<close>
+ \<open>- numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR - numeral y)\<close>
+ \<open>- numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR - numeral (y + Num.One))\<close>
+ \<open>(1::int) OR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)\<close>
+ \<open>(1::int) OR numeral (Num.Bit1 y) = numeral (Num.Bit1 y)\<close>
+ \<open>(1::int) OR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)\<close>
+ \<open>(1::int) OR - numeral (Num.Bit1 y) = - numeral (Num.Bit1 y)\<close>
+ \<open>numeral (Num.Bit0 x) OR (1::int) = numeral (Num.Bit1 x)\<close>
+ \<open>numeral (Num.Bit1 x) OR (1::int) = numeral (Num.Bit1 x)\<close>
+ \<open>- numeral (Num.Bit0 x) OR (1::int) = - numeral (Num.BitM x)\<close>
+ \<open>- numeral (Num.Bit1 x) OR (1::int) = - numeral (Num.Bit1 x)\<close>
+ 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]:
+ \<open>numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (2 :: int) * (numeral x XOR numeral y)\<close>
+ \<open>numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x XOR numeral y)\<close>
+ \<open>numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x XOR numeral y)\<close>
+ \<open>numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (2 :: int) * (numeral x XOR numeral y)\<close>
+ \<open>numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (2 :: int) * (numeral x XOR - numeral y)\<close>
+ \<open>numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x XOR - numeral (y + Num.One))\<close>
+ \<open>numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x XOR - numeral y)\<close>
+ \<open>numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (2 :: int) * (numeral x XOR - numeral (y + Num.One))\<close>
+ \<open>- numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (2 :: int) * (- numeral x XOR numeral y)\<close>
+ \<open>- numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x XOR numeral y)\<close>
+ \<open>- numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) XOR numeral y)\<close>
+ \<open>- numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (2 :: int) * (- numeral (x + Num.One) XOR numeral y)\<close>
+ \<open>- numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x XOR - numeral y)\<close>
+ \<open>- numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x XOR - numeral (y + Num.One))\<close>
+ \<open>- numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) XOR - numeral y)\<close>
+ \<open>- numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (2 :: int) * (- numeral (x + Num.One) XOR - numeral (y + Num.One))\<close>
+ \<open>(1::int) XOR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)\<close>
+ \<open>(1::int) XOR numeral (Num.Bit1 y) = numeral (Num.Bit0 y)\<close>
+ \<open>(1::int) XOR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)\<close>
+ \<open>(1::int) XOR - numeral (Num.Bit1 y) = - numeral (Num.Bit0 (y + Num.One))\<close>
+ \<open>numeral (Num.Bit0 x) XOR (1::int) = numeral (Num.Bit1 x)\<close>
+ \<open>numeral (Num.Bit1 x) XOR (1::int) = numeral (Num.Bit0 x)\<close>
+ \<open>- numeral (Num.Bit0 x) XOR (1::int) = - numeral (Num.BitM x)\<close>
+ \<open>- numeral (Num.Bit1 x) XOR (1::int) = - numeral (Num.Bit0 (x + Num.One))\<close>
+ 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 \<open>Bit concatenation\<close>
+
+definition concat_bit :: \<open>nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> int\<close>
+ where \<open>concat_bit n k l = take_bit n k OR push_bit n l\<close>
+
+lemma bit_concat_bit_iff [bit_simps]:
+ \<open>bit (concat_bit m k l) n \<longleftrightarrow> n < m \<and> bit k n \<or> m \<le> n \<and> bit l (n - m)\<close>
+ 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:
+ \<open>concat_bit n k l = take_bit n k + push_bit n l\<close>
+ 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]:
+ \<open>concat_bit 0 k l = l\<close>
+ by (simp add: concat_bit_def)
+
+lemma concat_bit_Suc:
+ \<open>concat_bit (Suc n) k l = k mod 2 + 2 * concat_bit n (k div 2) l\<close>
+ by (simp add: concat_bit_eq take_bit_Suc push_bit_double)
+
+lemma concat_bit_of_zero_1 [simp]:
+ \<open>concat_bit n 0 l = push_bit n l\<close>
+ by (simp add: concat_bit_def)
+
+lemma concat_bit_of_zero_2 [simp]:
+ \<open>concat_bit n k 0 = take_bit n k\<close>
+ by (simp add: concat_bit_def take_bit_eq_mask)
+
+lemma concat_bit_nonnegative_iff [simp]:
+ \<open>concat_bit n k l \<ge> 0 \<longleftrightarrow> l \<ge> 0\<close>
+ by (simp add: concat_bit_def)
+
+lemma concat_bit_negative_iff [simp]:
+ \<open>concat_bit n k l < 0 \<longleftrightarrow> l < 0\<close>
+ by (simp add: concat_bit_def)
+
+lemma concat_bit_assoc:
+ \<open>concat_bit n k (concat_bit m l r) = concat_bit (m + n) (concat_bit n k l) r\<close>
+ by (rule bit_eqI) (auto simp add: bit_concat_bit_iff ac_simps)
+
+lemma concat_bit_assoc_sym:
+ \<open>concat_bit m (concat_bit n k l) r = concat_bit (min m n) k (concat_bit (m - n) l r)\<close>
+ by (rule bit_eqI) (auto simp add: bit_concat_bit_iff ac_simps min_def)
+
+lemma concat_bit_eq_iff:
+ \<open>concat_bit n k l = concat_bit n r s
+ \<longleftrightarrow> take_bit n k = take_bit n r \<and> l = s\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
+proof
+ assume ?Q
+ then show ?P
+ by (simp add: concat_bit_def)
+next
+ assume ?P
+ then have *: \<open>bit (concat_bit n k l) m = bit (concat_bit n r s) m\<close> for m
+ by (simp add: bit_eq_iff)
+ have \<open>take_bit n k = take_bit n r\<close>
+ proof (rule bit_eqI)
+ fix m
+ from * [of m]
+ show \<open>bit (take_bit n k) m \<longleftrightarrow> bit (take_bit n r) m\<close>
+ by (auto simp add: bit_take_bit_iff bit_concat_bit_iff)
+ qed
+ moreover have \<open>push_bit n l = push_bit n s\<close>
+ proof (rule bit_eqI)
+ fix m
+ from * [of m]
+ show \<open>bit (push_bit n l) m \<longleftrightarrow> bit (push_bit n s) m\<close>
+ by (auto simp add: bit_push_bit_iff bit_concat_bit_iff)
+ qed
+ then have \<open>l = s\<close>
+ by (simp add: push_bit_eq_mult)
+ ultimately show ?Q
+ by (simp add: concat_bit_def)
+qed
+
+lemma take_bit_concat_bit_eq:
+ \<open>take_bit m (concat_bit n k l) = concat_bit (min m n) k (take_bit (m - n) l)\<close>
+ by (rule bit_eqI)
+ (auto simp add: bit_take_bit_iff bit_concat_bit_iff min_def)
+
+lemma concat_bit_take_bit_eq:
+ \<open>concat_bit n (take_bit n b) = concat_bit n b\<close>
+ by (simp add: concat_bit_def [abs_def])
+
+
+subsection \<open>Taking bits with sign propagation\<close>
+
+context ring_bit_operations
+begin
+
+definition signed_take_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
+ where \<open>signed_take_bit n a = take_bit n a OR (of_bool (bit a n) * NOT (mask n))\<close>
+
+lemma signed_take_bit_eq_if_positive:
+ \<open>signed_take_bit n a = take_bit n a\<close> if \<open>\<not> bit a n\<close>
+ using that by (simp add: signed_take_bit_def)
+
+lemma signed_take_bit_eq_if_negative:
+ \<open>signed_take_bit n a = take_bit n a OR NOT (mask n)\<close> if \<open>bit a n\<close>
+ using that by (simp add: signed_take_bit_def)
+
+lemma even_signed_take_bit_iff:
+ \<open>even (signed_take_bit m a) \<longleftrightarrow> even a\<close>
+ 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]:
+ \<open>bit (signed_take_bit m a) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> bit a (min m n)\<close>
+ 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]:
+ \<open>signed_take_bit 0 a = - (a mod 2)\<close>
+ by (simp add: signed_take_bit_def odd_iff_mod_2_eq_one)
+
+lemma signed_take_bit_Suc:
+ \<open>signed_take_bit (Suc n) a = a mod 2 + 2 * signed_take_bit n (a div 2)\<close>
+proof (rule bit_eqI)
+ fix m
+ assume *: \<open>2 ^ m \<noteq> 0\<close>
+ show \<open>bit (signed_take_bit (Suc n) a) m \<longleftrightarrow>
+ bit (a mod 2 + 2 * signed_take_bit n (a div 2)) m\<close>
+ proof (cases m)
+ case 0
+ then show ?thesis
+ by (simp add: even_signed_take_bit_iff)
+ next
+ case (Suc m)
+ with * have \<open>2 ^ m \<noteq> 0\<close>
+ 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]:
+ \<open>signed_take_bit n 0 = 0\<close>
+ by (simp add: signed_take_bit_def)
+
+lemma signed_take_bit_of_minus_1 [simp]:
+ \<open>signed_take_bit n (- 1) = - 1\<close>
+ 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]:
+ \<open>signed_take_bit (Suc n) 1 = 1\<close>
+ by (simp add: signed_take_bit_Suc)
+
+lemma signed_take_bit_rec:
+ \<open>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))\<close>
+ by (cases n) (simp_all add: signed_take_bit_Suc)
+
+lemma signed_take_bit_eq_iff_take_bit_eq:
+ \<open>signed_take_bit n a = signed_take_bit n b \<longleftrightarrow> take_bit (Suc n) a = take_bit (Suc n) b\<close>
+proof -
+ have \<open>bit (signed_take_bit n a) = bit (signed_take_bit n b) \<longleftrightarrow> bit (take_bit (Suc n) a) = bit (take_bit (Suc n) b)\<close>
+ 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]:
+ \<open>signed_take_bit m (signed_take_bit n a) = signed_take_bit (min m n) a\<close>
+proof (rule bit_eqI)
+ fix q
+ show \<open>bit (signed_take_bit m (signed_take_bit n a)) q \<longleftrightarrow>
+ bit (signed_take_bit (min m n) a) q\<close>
+ 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:
+ \<open>signed_take_bit m (take_bit n a) = (if n \<le> m then take_bit n else signed_take_bit m) a\<close>
+ by (rule bit_eqI) (auto simp add: bit_signed_take_bit_iff min_def bit_take_bit_iff)
+
+lemma take_bit_signed_take_bit:
+ \<open>take_bit m (signed_take_bit n a) = take_bit m a\<close> if \<open>m \<le> Suc n\<close>
+ 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 \<open>Modulus centered around 0\<close>
+
+lemma signed_take_bit_eq_concat_bit:
+ \<open>signed_take_bit n k = concat_bit n k (- of_bool (bit k n))\<close>
+ by (simp add: concat_bit_def signed_take_bit_def push_bit_minus_one_eq_not_mask)
+
+lemma signed_take_bit_add:
+ \<open>signed_take_bit n (signed_take_bit n k + signed_take_bit n l) = signed_take_bit n (k + l)\<close>
+ for k l :: int
+proof -
+ have \<open>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)\<close>
+ 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:
+ \<open>signed_take_bit n (signed_take_bit n k - signed_take_bit n l) = signed_take_bit n (k - l)\<close>
+ for k l :: int
+proof -
+ have \<open>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)\<close>
+ 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:
+ \<open>signed_take_bit n (- signed_take_bit n k) = signed_take_bit n (- k)\<close>
+ for k :: int
+proof -
+ have \<open>take_bit (Suc n)
+ (- take_bit (Suc n) (signed_take_bit n k)) =
+ take_bit (Suc n) (- k)\<close>
+ 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:
+ \<open>signed_take_bit n (signed_take_bit n k * signed_take_bit n l) = signed_take_bit n (k * l)\<close>
+ for k l :: int
+proof -
+ have \<open>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)\<close>
+ 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:
+ \<open>signed_take_bit n k = take_bit (Suc n) k - 2 ^ Suc n * of_bool (bit k n)\<close>
+ for k :: int
+proof (cases \<open>bit k n\<close>)
+ case True
+ have \<open>signed_take_bit n k = take_bit (Suc n) k OR NOT (mask (Suc n))\<close>
+ 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 \<open>signed_take_bit n k = take_bit (Suc n) k + NOT (mask (Suc n))\<close>
+ 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:
+ \<open>signed_take_bit n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n\<close>
+ for k :: int
+proof -
+ have *: \<open>take_bit n k OR 2 ^ n = take_bit n k + 2 ^ n\<close>
+ by (simp add: disjunctive_add bit_exp_iff bit_take_bit_iff)
+ have \<open>take_bit n k - 2 ^ n = take_bit n k + NOT (mask n)\<close>
+ by (simp add: minus_exp_eq_not_mask)
+ also have \<open>\<dots> = take_bit n k OR NOT (mask n)\<close>
+ by (rule disjunctive_add)
+ (simp add: bit_exp_iff bit_take_bit_iff bit_not_iff bit_mask_iff)
+ finally have **: \<open>take_bit n k - 2 ^ n = take_bit n k OR NOT (mask n)\<close> .
+ have \<open>take_bit (Suc n) (k + 2 ^ n) = take_bit (Suc n) (take_bit (Suc n) k + take_bit (Suc n) (2 ^ n))\<close>
+ by (simp only: take_bit_add)
+ also have \<open>take_bit (Suc n) k = 2 ^ n * of_bool (bit k n) + take_bit n k\<close>
+ by (simp add: take_bit_Suc_from_most)
+ finally have \<open>take_bit (Suc n) (k + 2 ^ n) = take_bit (Suc n) (2 ^ (n + of_bool (bit k n)) + take_bit n k)\<close>
+ by (simp add: ac_simps)
+ also have \<open>2 ^ (n + of_bool (bit k n)) + take_bit n k = 2 ^ (n + of_bool (bit k n)) OR take_bit n k\<close>
+ 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]:
+ \<open>0 \<le> signed_take_bit n k \<longleftrightarrow> \<not> bit k n\<close>
+ for k :: int
+ by (simp add: signed_take_bit_def not_less concat_bit_def)
+
+lemma signed_take_bit_negative_iff [simp]:
+ \<open>signed_take_bit n k < 0 \<longleftrightarrow> bit k n\<close>
+ 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]:
+ \<open>- (2 ^ n) \<le> signed_take_bit n k\<close>
+ for k :: int
+ by (simp add: signed_take_bit_eq_take_bit_shift)
+
+lemma signed_take_bit_int_less_exp [simp]:
+ \<open>signed_take_bit n k < 2 ^ n\<close>
+ for k :: int
+ using take_bit_int_less_exp [of \<open>Suc n\<close>]
+ by (simp add: signed_take_bit_eq_take_bit_shift)
+
+lemma signed_take_bit_int_eq_self_iff:
+ \<open>signed_take_bit n k = k \<longleftrightarrow> - (2 ^ n) \<le> k \<and> k < 2 ^ n\<close>
+ 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:
+ \<open>signed_take_bit n k = k\<close> if \<open>- (2 ^ n) \<le> k\<close> \<open>k < 2 ^ n\<close>
+ for k :: int
+ using that by (simp add: signed_take_bit_int_eq_self_iff)
+
+lemma signed_take_bit_int_less_eq_self_iff:
+ \<open>signed_take_bit n k \<le> k \<longleftrightarrow> - (2 ^ n) \<le> k\<close>
+ 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:
+ \<open>signed_take_bit n k < k \<longleftrightarrow> 2 ^ n \<le> k\<close>
+ 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:
+ \<open>k < signed_take_bit n k \<longleftrightarrow> k < - (2 ^ n)\<close>
+ 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:
+ \<open>k \<le> signed_take_bit n k \<longleftrightarrow> k < 2 ^ n\<close>
+ 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:
+ \<open>k + 2 ^ Suc n \<le> signed_take_bit n k\<close> if \<open>k < - (2 ^ n)\<close>
+ for k :: int
+ using that take_bit_int_greater_eq [of \<open>k + 2 ^ n\<close> \<open>Suc n\<close>]
+ by (simp add: signed_take_bit_eq_take_bit_shift)
+
+lemma signed_take_bit_int_less_eq:
+ \<open>signed_take_bit n k \<le> k - 2 ^ Suc n\<close> if \<open>k \<ge> 2 ^ n\<close>
+ for k :: int
+ using that take_bit_int_less_eq [of \<open>Suc n\<close> \<open>k + 2 ^ n\<close>]
+ by (simp add: signed_take_bit_eq_take_bit_shift)
+
+lemma signed_take_bit_Suc_bit0 [simp]:
+ \<open>signed_take_bit (Suc n) (numeral (Num.Bit0 k)) = signed_take_bit n (numeral k) * (2 :: int)\<close>
+ by (simp add: signed_take_bit_Suc)
+
+lemma signed_take_bit_Suc_bit1 [simp]:
+ \<open>signed_take_bit (Suc n) (numeral (Num.Bit1 k)) = signed_take_bit n (numeral k) * 2 + (1 :: int)\<close>
+ by (simp add: signed_take_bit_Suc)
+
+lemma signed_take_bit_Suc_minus_bit0 [simp]:
+ \<open>signed_take_bit (Suc n) (- numeral (Num.Bit0 k)) = signed_take_bit n (- numeral k) * (2 :: int)\<close>
+ by (simp add: signed_take_bit_Suc)
+
+lemma signed_take_bit_Suc_minus_bit1 [simp]:
+ \<open>signed_take_bit (Suc n) (- numeral (Num.Bit1 k)) = signed_take_bit n (- numeral k - 1) * 2 + (1 :: int)\<close>
+ by (simp add: signed_take_bit_Suc)
+
+lemma signed_take_bit_numeral_bit0 [simp]:
+ \<open>signed_take_bit (numeral l) (numeral (Num.Bit0 k)) = signed_take_bit (pred_numeral l) (numeral k) * (2 :: int)\<close>
+ by (simp add: signed_take_bit_rec)
+
+lemma signed_take_bit_numeral_bit1 [simp]:
+ \<open>signed_take_bit (numeral l) (numeral (Num.Bit1 k)) = signed_take_bit (pred_numeral l) (numeral k) * 2 + (1 :: int)\<close>
+ by (simp add: signed_take_bit_rec)
+
+lemma signed_take_bit_numeral_minus_bit0 [simp]:
+ \<open>signed_take_bit (numeral l) (- numeral (Num.Bit0 k)) = signed_take_bit (pred_numeral l) (- numeral k) * (2 :: int)\<close>
+ by (simp add: signed_take_bit_rec)
+
+lemma signed_take_bit_numeral_minus_bit1 [simp]:
+ \<open>signed_take_bit (numeral l) (- numeral (Num.Bit1 k)) = signed_take_bit (pred_numeral l) (- numeral k - 1) * 2 + (1 :: int)\<close>
+ by (simp add: signed_take_bit_rec)
+
+lemma signed_take_bit_code [code]:
+ \<open>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)\<close>
+proof -
+ have *: \<open>take_bit (Suc n) a + push_bit n (- 2) =
+ take_bit (Suc n) a OR NOT (mask (Suc n))\<close>
+ 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 \<open>Instance \<^typ>\<open>nat\<close>\<close>
+
+instantiation nat :: semiring_bit_operations
+begin
+
+definition and_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
+ where \<open>m AND n = nat (int m AND int n)\<close> for m n :: nat
+
+definition or_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
+ where \<open>m OR n = nat (int m OR int n)\<close> for m n :: nat
+
+definition xor_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
+ where \<open>m XOR n = nat (int m XOR int n)\<close> for m n :: nat
+
+definition mask_nat :: \<open>nat \<Rightarrow> nat\<close>
+ where \<open>mask n = (2 :: nat) ^ n - 1\<close>
+
+definition set_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
+ where \<open>set_bit m n = n OR push_bit m 1\<close> for m n :: nat
+
+definition unset_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
+ where \<open>unset_bit m n = (if bit n m then n - push_bit m 1 else n)\<close> for m n :: nat
+
+definition flip_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
+ where \<open>flip_bit m n = n XOR push_bit m 1\<close> for m n :: nat
+
+instance proof
+ fix m n q :: nat
+ show \<open>bit (m AND n) q \<longleftrightarrow> bit m q \<and> bit n q\<close>
+ by (simp add: and_nat_def bit_simps)
+ show \<open>bit (m OR n) q \<longleftrightarrow> bit m q \<or> bit n q\<close>
+ by (simp add: or_nat_def bit_simps)
+ show \<open>bit (m XOR n) q \<longleftrightarrow> bit m q \<noteq> bit n q\<close>
+ by (simp add: xor_nat_def bit_simps)
+ show \<open>bit (unset_bit m n) q \<longleftrightarrow> bit n q \<and> m \<noteq> q\<close>
+ proof (cases \<open>bit n m\<close>)
+ case False
+ then show ?thesis by (auto simp add: unset_bit_nat_def)
+ next
+ case True
+ have \<open>push_bit m (drop_bit m n) + take_bit m n = n\<close>
+ by (fact bits_ident)
+ also from \<open>bit n m\<close> have \<open>drop_bit m n = 2 * drop_bit (Suc m) n + 1\<close>
+ by (simp add: drop_bit_Suc drop_bit_half even_drop_bit_iff_not_bit ac_simps)
+ finally have \<open>push_bit m (2 * drop_bit (Suc m) n) + take_bit m n + push_bit m 1 = n\<close>
+ by (simp only: push_bit_add ac_simps)
+ then have \<open>n - push_bit m 1 = push_bit m (2 * drop_bit (Suc m) n) + take_bit m n\<close>
+ by simp
+ then have \<open>n - push_bit m 1 = push_bit m (2 * drop_bit (Suc m) n) OR take_bit m n\<close>
+ by (simp add: or_nat_def bit_simps flip: disjunctive_add)
+ with \<open>bit n m\<close> 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:
+ \<open>m AND n = of_bool (odd m \<and> odd n) + 2 * ((m div 2) AND (n div 2))\<close> for m n :: nat
+ by (simp add: and_nat_def and_int_rec [of \<open>int m\<close> \<open>int n\<close>] zdiv_int nat_add_distrib nat_mult_distrib)
+
+lemma or_nat_rec:
+ \<open>m OR n = of_bool (odd m \<or> odd n) + 2 * ((m div 2) OR (n div 2))\<close> for m n :: nat
+ by (simp add: or_nat_def or_int_rec [of \<open>int m\<close> \<open>int n\<close>] zdiv_int nat_add_distrib nat_mult_distrib)
+
+lemma xor_nat_rec:
+ \<open>m XOR n = of_bool (odd m \<noteq> odd n) + 2 * ((m div 2) XOR (n div 2))\<close> for m n :: nat
+ by (simp add: xor_nat_def xor_int_rec [of \<open>int m\<close> \<open>int n\<close>] zdiv_int nat_add_distrib nat_mult_distrib)
+
+lemma Suc_0_and_eq [simp]:
+ \<open>Suc 0 AND n = n mod 2\<close>
+ using one_and_eq [of n] by simp
+
+lemma and_Suc_0_eq [simp]:
+ \<open>n AND Suc 0 = n mod 2\<close>
+ using and_one_eq [of n] by simp
+
+lemma Suc_0_or_eq:
+ \<open>Suc 0 OR n = n + of_bool (even n)\<close>
+ using one_or_eq [of n] by simp
+
+lemma or_Suc_0_eq:
+ \<open>n OR Suc 0 = n + of_bool (even n)\<close>
+ using or_one_eq [of n] by simp
+
+lemma Suc_0_xor_eq:
+ \<open>Suc 0 XOR n = n + of_bool (even n) - of_bool (odd n)\<close>
+ using one_xor_eq [of n] by simp
+
+lemma xor_Suc_0_eq:
+ \<open>n XOR Suc 0 = n + of_bool (even n) - of_bool (odd n)\<close>
+ using xor_one_eq [of n] by simp
+
+context semiring_bit_operations
+begin
+
+lemma of_nat_and_eq:
+ \<open>of_nat (m AND n) = of_nat m AND of_nat n\<close>
+ by (rule bit_eqI) (simp add: bit_of_nat_iff bit_and_iff Bit_Operations.bit_and_iff)
+
+lemma of_nat_or_eq:
+ \<open>of_nat (m OR n) = of_nat m OR of_nat n\<close>
+ by (rule bit_eqI) (simp add: bit_of_nat_iff bit_or_iff Bit_Operations.bit_or_iff)
+
+lemma of_nat_xor_eq:
+ \<open>of_nat (m XOR n) = of_nat m XOR of_nat n\<close>
+ 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:
+ \<open>of_nat (mask n) = mask n\<close>
+ by (induction n) (simp_all add: mask_Suc_double Bit_Operations.mask_Suc_double of_nat_or_eq)
+
+end
+
+lemma Suc_mask_eq_exp:
+ \<open>Suc (mask n) = 2 ^ n\<close>
+ by (simp add: mask_eq_exp_minus_1)
+
+lemma less_eq_mask:
+ \<open>n \<le> mask n\<close>
+ 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:
+ \<open>n < mask n\<close> if \<open>Suc 0 < n\<close>
+proof -
+ define m where \<open>m = n - 2\<close>
+ with that have *: \<open>n = m + 2\<close>
+ by simp
+ have \<open>Suc (Suc (Suc m)) < 4 * 2 ^ m\<close>
+ by (induction m) simp_all
+ then have \<open>Suc (m + 2) < Suc (mask (m + 2))\<close>
+ by (simp add: Suc_mask_eq_exp)
+ then have \<open>m + 2 < mask (m + 2)\<close>
+ by (simp add: less_le)
+ with * show ?thesis
+ by simp
+qed
+
+
+subsection \<open>Horner sums\<close>
+
+context semiring_bit_shifts
+begin
+
+lemma horner_sum_bit_eq_take_bit:
+ \<open>horner_sum of_bool 2 (map (bit a) [0..<n]) = take_bit n a\<close>
+proof (induction a arbitrary: n rule: bits_induct)
+ case (stable a)
+ moreover have \<open>bit a = (\<lambda>_. odd a)\<close>
+ using stable by (simp add: stable_imp_bit_iff_odd fun_eq_iff)
+ moreover have \<open>{q. q < n} = {0..<n}\<close>
+ 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 \<open>map (bit (of_bool b + 2 * a)) [0..<Suc m] = b # map (bit (of_bool b + 2 * a)) [Suc 0..<Suc m]\<close>
+ by (simp only: upt_conv_Cons) simp
+ also have \<open>\<dots> = b # map (bit a) [0..<m]\<close>
+ 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]:
+ \<open>bit (horner_sum of_bool 2 bs) n \<longleftrightarrow> n < length bs \<and> bs ! n\<close>
+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:
+ \<open>take_bit n (horner_sum of_bool 2 bs) = horner_sum of_bool 2 (take n bs)\<close>
+ by (auto simp add: bit_eq_iff bit_take_bit_iff bit_horner_sum_bit_iff)
+
+end
+
+lemma horner_sum_of_bool_2_less:
+ \<open>(horner_sum of_bool 2 bs :: int) < 2 ^ length bs\<close>
+proof -
+ have \<open>(\<Sum>n = 0..<length bs. of_bool (bs ! n) * (2::int) ^ n) \<le> (\<Sum>n = 0..<length bs. 2 ^ n)\<close>
+ by (rule sum_mono) simp
+ also have \<open>\<dots> = 2 ^ length bs - 1\<close>
+ by (induction bs) simp_all
+ finally show ?thesis
+ by (simp add: horner_sum_eq_sum)
+qed
+
+
+subsection \<open>Symbolic computations on numeral expressions\<close> \<^marker>\<open>contributor \<open>Andreas Lochbihler\<close>\<close>
+
+fun and_num :: \<open>num \<Rightarrow> num \<Rightarrow> num option\<close>
+where
+ \<open>and_num num.One num.One = Some num.One\<close>
+| \<open>and_num num.One (num.Bit0 n) = None\<close>
+| \<open>and_num num.One (num.Bit1 n) = Some num.One\<close>
+| \<open>and_num (num.Bit0 m) num.One = None\<close>
+| \<open>and_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (and_num m n)\<close>
+| \<open>and_num (num.Bit0 m) (num.Bit1 n) = map_option num.Bit0 (and_num m n)\<close>
+| \<open>and_num (num.Bit1 m) num.One = Some num.One\<close>
+| \<open>and_num (num.Bit1 m) (num.Bit0 n) = map_option num.Bit0 (and_num m n)\<close>
+| \<open>and_num (num.Bit1 m) (num.Bit1 n) = (case and_num m n of None \<Rightarrow> Some num.One | Some n' \<Rightarrow> Some (num.Bit1 n'))\<close>
+
+fun and_not_num :: \<open>num \<Rightarrow> num \<Rightarrow> num option\<close>
+where
+ \<open>and_not_num num.One num.One = None\<close>
+| \<open>and_not_num num.One (num.Bit0 n) = Some num.One\<close>
+| \<open>and_not_num num.One (num.Bit1 n) = None\<close>
+| \<open>and_not_num (num.Bit0 m) num.One = Some (num.Bit0 m)\<close>
+| \<open>and_not_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (and_not_num m n)\<close>
+| \<open>and_not_num (num.Bit0 m) (num.Bit1 n) = map_option num.Bit0 (and_not_num m n)\<close>
+| \<open>and_not_num (num.Bit1 m) num.One = Some (num.Bit0 m)\<close>
+| \<open>and_not_num (num.Bit1 m) (num.Bit0 n) = (case and_not_num m n of None \<Rightarrow> Some num.One | Some n' \<Rightarrow> Some (num.Bit1 n'))\<close>
+| \<open>and_not_num (num.Bit1 m) (num.Bit1 n) = map_option num.Bit0 (and_not_num m n)\<close>
+
+fun or_num :: \<open>num \<Rightarrow> num \<Rightarrow> num\<close>
+where
+ \<open>or_num num.One num.One = num.One\<close>
+| \<open>or_num num.One (num.Bit0 n) = num.Bit1 n\<close>
+| \<open>or_num num.One (num.Bit1 n) = num.Bit1 n\<close>
+| \<open>or_num (num.Bit0 m) num.One = num.Bit1 m\<close>
+| \<open>or_num (num.Bit0 m) (num.Bit0 n) = num.Bit0 (or_num m n)\<close>
+| \<open>or_num (num.Bit0 m) (num.Bit1 n) = num.Bit1 (or_num m n)\<close>
+| \<open>or_num (num.Bit1 m) num.One = num.Bit1 m\<close>
+| \<open>or_num (num.Bit1 m) (num.Bit0 n) = num.Bit1 (or_num m n)\<close>
+| \<open>or_num (num.Bit1 m) (num.Bit1 n) = num.Bit1 (or_num m n)\<close>
+
+fun or_not_num_neg :: \<open>num \<Rightarrow> num \<Rightarrow> num\<close>
+where
+ \<open>or_not_num_neg num.One num.One = num.One\<close>
+| \<open>or_not_num_neg num.One (num.Bit0 m) = num.Bit1 m\<close>
+| \<open>or_not_num_neg num.One (num.Bit1 m) = num.Bit1 m\<close>
+| \<open>or_not_num_neg (num.Bit0 n) num.One = num.Bit0 num.One\<close>
+| \<open>or_not_num_neg (num.Bit0 n) (num.Bit0 m) = Num.BitM (or_not_num_neg n m)\<close>
+| \<open>or_not_num_neg (num.Bit0 n) (num.Bit1 m) = num.Bit0 (or_not_num_neg n m)\<close>
+| \<open>or_not_num_neg (num.Bit1 n) num.One = num.One\<close>
+| \<open>or_not_num_neg (num.Bit1 n) (num.Bit0 m) = Num.BitM (or_not_num_neg n m)\<close>
+| \<open>or_not_num_neg (num.Bit1 n) (num.Bit1 m) = Num.BitM (or_not_num_neg n m)\<close>
+
+fun xor_num :: \<open>num \<Rightarrow> num \<Rightarrow> num option\<close>
+where
+ \<open>xor_num num.One num.One = None\<close>
+| \<open>xor_num num.One (num.Bit0 n) = Some (num.Bit1 n)\<close>
+| \<open>xor_num num.One (num.Bit1 n) = Some (num.Bit0 n)\<close>
+| \<open>xor_num (num.Bit0 m) num.One = Some (num.Bit1 m)\<close>
+| \<open>xor_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (xor_num m n)\<close>
+| \<open>xor_num (num.Bit0 m) (num.Bit1 n) = Some (case xor_num m n of None \<Rightarrow> num.One | Some n' \<Rightarrow> num.Bit1 n')\<close>
+| \<open>xor_num (num.Bit1 m) num.One = Some (num.Bit0 m)\<close>
+| \<open>xor_num (num.Bit1 m) (num.Bit0 n) = Some (case xor_num m n of None \<Rightarrow> num.One | Some n' \<Rightarrow> num.Bit1 n')\<close>
+| \<open>xor_num (num.Bit1 m) (num.Bit1 n) = map_option num.Bit0 (xor_num m n)\<close>
+
+lemma int_numeral_and_num:
+ \<open>numeral m AND numeral n = (case and_num m n of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close>
+ by (induction m n rule: and_num.induct) (simp_all split: option.split)
+
+lemma and_num_eq_None_iff:
+ \<open>and_num m n = None \<longleftrightarrow> numeral m AND numeral n = (0::int)\<close>
+ by (simp add: int_numeral_and_num split: option.split)
+
+lemma and_num_eq_Some_iff:
+ \<open>and_num m n = Some q \<longleftrightarrow> numeral m AND numeral n = (numeral q :: int)\<close>
+ by (simp add: int_numeral_and_num split: option.split)
+
+lemma int_numeral_and_not_num:
+ \<open>numeral m AND NOT (numeral n) = (case and_not_num m n of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close>
+ 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:
+ \<open>NOT (numeral m) AND numeral n = (case and_not_num n m of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close>
+ using int_numeral_and_not_num [of n m] by (simp add: ac_simps)
+
+lemma and_not_num_eq_None_iff:
+ \<open>and_not_num m n = None \<longleftrightarrow> numeral m AND NOT (numeral n) = (0::int)\<close>
+ by (simp add: int_numeral_and_not_num split: option.split)
+
+lemma and_not_num_eq_Some_iff:
+ \<open>and_not_num m n = Some q \<longleftrightarrow> numeral m AND NOT (numeral n) = (numeral q :: int)\<close>
+ by (simp add: int_numeral_and_not_num split: option.split)
+
+lemma int_numeral_or_num:
+ \<open>numeral m OR numeral n = (numeral (or_num m n) :: int)\<close>
+ by (induction m n rule: or_num.induct) simp_all
+
+lemma numeral_or_num_eq:
+ \<open>numeral (or_num m n) = (numeral m OR numeral n :: int)\<close>
+ by (simp add: int_numeral_or_num)
+
+lemma int_numeral_or_not_num_neg:
+ \<open>numeral m OR NOT (numeral n :: int) = - numeral (or_not_num_neg m n)\<close>
+ 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:
+ \<open>NOT (numeral m) OR (numeral n :: int) = - numeral (or_not_num_neg n m)\<close>
+ using int_numeral_or_not_num_neg [of n m] by (simp add: ac_simps)
+
+lemma numeral_or_not_num_eq:
+ \<open>numeral (or_not_num_neg m n) = - (numeral m OR NOT (numeral n :: int))\<close>
+ using int_numeral_or_not_num_neg [of m n] by simp
+
+lemma int_numeral_xor_num:
+ \<open>numeral m XOR numeral n = (case xor_num m n of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close>
+ by (induction m n rule: xor_num.induct) (simp_all split: option.split)
+
+lemma xor_num_eq_None_iff:
+ \<open>xor_num m n = None \<longleftrightarrow> numeral m XOR numeral n = (0::int)\<close>
+ by (simp add: int_numeral_xor_num split: option.split)
+
+lemma xor_num_eq_Some_iff:
+ \<open>xor_num m n = Some q \<longleftrightarrow> numeral m XOR numeral n = (numeral q :: int)\<close>
+ by (simp add: int_numeral_xor_num split: option.split)
+
+
+subsection \<open>Key ideas of bit operations\<close>
+
+text \<open>
+ 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>\<open>int\<close>,
+ bounded bit values by quotient types over \<^typ>\<open>int\<close>.
+
+ \<^item> (A special case are idealized unbounded bit values ending
+ in @{term [source] 0} which can be represented by type \<^typ>\<open>nat\<close> but
+ only support a restricted set of operations).
+
+ \<^item> From this idea follows that
+
+ \<^item> multiplication by \<^term>\<open>2 :: int\<close> is a bit shift to the left and
+
+ \<^item> division by \<^term>\<open>2 :: int\<close> 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>\<open>(2 :: int) ^ n \<noteq> 0\<close>
+ represents that \<^term>\<open>n\<close> is \<^emph>\<open>not\<close> 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>\<open>n\<close>th bit: \<^term>\<open>(2 :: int) ^ n\<close>
+
+ \<^item> Bit mask upto bit \<^term>\<open>n\<close>: @{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>\<open>0::int\<close>: @{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]}
+\<close>
+
+no_notation
+ "and" (infixr \<open>AND\<close> 64)
+ and or (infixr \<open>OR\<close> 59)
+ and xor (infixr \<open>XOR\<close> 59)
+
+bundle bit_operations_syntax
+begin
+
+notation
+ "and" (infixr \<open>AND\<close> 64)
+ and or (infixr \<open>OR\<close> 59)
+ and xor (infixr \<open>XOR\<close> 59)
+
+end
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Boolean_Algebra.thy Mon Aug 02 10:01:06 2021 +0000
@@ -0,0 +1,296 @@
+(* Title: HOL/Boolean_Algebra.thy
+ Author: Brian Huffman
+*)
+
+section \<open>Abstract boolean Algebras\<close>
+
+theory Boolean_Algebra
+ imports Lattices
+begin
+
+locale boolean_algebra = conj: abel_semigroup "(\<^bold>\<sqinter>)" + disj: abel_semigroup "(\<^bold>\<squnion>)"
+ for conj :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<^bold>\<sqinter>" 70)
+ and disj :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<^bold>\<squnion>" 65) +
+ fixes compl :: "'a \<Rightarrow> 'a" ("\<^bold>- _" [81] 80)
+ and zero :: "'a" ("\<^bold>0")
+ and one :: "'a" ("\<^bold>1")
+ assumes conj_disj_distrib: "x \<^bold>\<sqinter> (y \<^bold>\<squnion> z) = (x \<^bold>\<sqinter> y) \<^bold>\<squnion> (x \<^bold>\<sqinter> z)"
+ and disj_conj_distrib: "x \<^bold>\<squnion> (y \<^bold>\<sqinter> z) = (x \<^bold>\<squnion> y) \<^bold>\<sqinter> (x \<^bold>\<squnion> z)"
+ and conj_one_right: "x \<^bold>\<sqinter> \<^bold>1 = x"
+ and disj_zero_right: "x \<^bold>\<squnion> \<^bold>0 = x"
+ and conj_cancel_right [simp]: "x \<^bold>\<sqinter> \<^bold>- x = \<^bold>0"
+ and disj_cancel_right [simp]: "x \<^bold>\<squnion> \<^bold>- x = \<^bold>1"
+begin
+
+sublocale conj: semilattice_neutr "(\<^bold>\<sqinter>)" "\<^bold>1"
+proof
+ show "x \<^bold>\<sqinter> \<^bold>1 = x" for x
+ by (fact conj_one_right)
+ show "x \<^bold>\<sqinter> x = x" for x
+ proof -
+ have "x \<^bold>\<sqinter> x = (x \<^bold>\<sqinter> x) \<^bold>\<squnion> \<^bold>0"
+ by (simp add: disj_zero_right)
+ also have "\<dots> = (x \<^bold>\<sqinter> x) \<^bold>\<squnion> (x \<^bold>\<sqinter> \<^bold>- x)"
+ by simp
+ also have "\<dots> = x \<^bold>\<sqinter> (x \<^bold>\<squnion> \<^bold>- x)"
+ by (simp only: conj_disj_distrib)
+ also have "\<dots> = x \<^bold>\<sqinter> \<^bold>1"
+ by simp
+ also have "\<dots> = x"
+ by (simp add: conj_one_right)
+ finally show ?thesis .
+ qed
+qed
+
+sublocale disj: semilattice_neutr "(\<^bold>\<squnion>)" "\<^bold>0"
+proof
+ show "x \<^bold>\<squnion> \<^bold>0 = x" for x
+ by (fact disj_zero_right)
+ show "x \<^bold>\<squnion> x = x" for x
+ proof -
+ have "x \<^bold>\<squnion> x = (x \<^bold>\<squnion> x) \<^bold>\<sqinter> \<^bold>1"
+ by simp
+ also have "\<dots> = (x \<^bold>\<squnion> x) \<^bold>\<sqinter> (x \<^bold>\<squnion> \<^bold>- x)"
+ by simp
+ also have "\<dots> = x \<^bold>\<squnion> (x \<^bold>\<sqinter> \<^bold>- x)"
+ by (simp only: disj_conj_distrib)
+ also have "\<dots> = x \<^bold>\<squnion> \<^bold>0"
+ by simp
+ also have "\<dots> = x"
+ by (simp add: disj_zero_right)
+ finally show ?thesis .
+ qed
+qed
+
+
+subsection \<open>Complement\<close>
+
+lemma complement_unique:
+ assumes 1: "a \<^bold>\<sqinter> x = \<^bold>0"
+ assumes 2: "a \<^bold>\<squnion> x = \<^bold>1"
+ assumes 3: "a \<^bold>\<sqinter> y = \<^bold>0"
+ assumes 4: "a \<^bold>\<squnion> y = \<^bold>1"
+ shows "x = y"
+proof -
+ from 1 3 have "(a \<^bold>\<sqinter> x) \<^bold>\<squnion> (x \<^bold>\<sqinter> y) = (a \<^bold>\<sqinter> y) \<^bold>\<squnion> (x \<^bold>\<sqinter> y)"
+ by simp
+ then have "(x \<^bold>\<sqinter> a) \<^bold>\<squnion> (x \<^bold>\<sqinter> y) = (y \<^bold>\<sqinter> a) \<^bold>\<squnion> (y \<^bold>\<sqinter> x)"
+ by (simp add: ac_simps)
+ then have "x \<^bold>\<sqinter> (a \<^bold>\<squnion> y) = y \<^bold>\<sqinter> (a \<^bold>\<squnion> x)"
+ by (simp add: conj_disj_distrib)
+ with 2 4 have "x \<^bold>\<sqinter> \<^bold>1 = y \<^bold>\<sqinter> \<^bold>1"
+ by simp
+ then show "x = y"
+ by simp
+qed
+
+lemma compl_unique: "x \<^bold>\<sqinter> y = \<^bold>0 \<Longrightarrow> x \<^bold>\<squnion> y = \<^bold>1 \<Longrightarrow> \<^bold>- x = y"
+ by (rule complement_unique [OF conj_cancel_right disj_cancel_right])
+
+lemma double_compl [simp]: "\<^bold>- (\<^bold>- x) = x"
+proof (rule compl_unique)
+ show "\<^bold>- x \<^bold>\<sqinter> x = \<^bold>0"
+ by (simp only: conj_cancel_right conj.commute)
+ show "\<^bold>- x \<^bold>\<squnion> x = \<^bold>1"
+ by (simp only: disj_cancel_right disj.commute)
+qed
+
+lemma compl_eq_compl_iff [simp]:
+ \<open>\<^bold>- x = \<^bold>- y \<longleftrightarrow> x = y\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
+proof
+ assume \<open>?Q\<close>
+ then show ?P by simp
+next
+ assume \<open>?P\<close>
+ then have \<open>\<^bold>- (\<^bold>- x) = \<^bold>- (\<^bold>- y)\<close>
+ by simp
+ then show ?Q
+ by simp
+qed
+
+
+subsection \<open>Conjunction\<close>
+
+lemma conj_zero_right [simp]: "x \<^bold>\<sqinter> \<^bold>0 = \<^bold>0"
+ using conj.left_idem conj_cancel_right by fastforce
+
+lemma compl_one [simp]: "\<^bold>- \<^bold>1 = \<^bold>0"
+ by (rule compl_unique [OF conj_zero_right disj_zero_right])
+
+lemma conj_zero_left [simp]: "\<^bold>0 \<^bold>\<sqinter> x = \<^bold>0"
+ by (subst conj.commute) (rule conj_zero_right)
+
+lemma conj_cancel_left [simp]: "\<^bold>- x \<^bold>\<sqinter> x = \<^bold>0"
+ by (subst conj.commute) (rule conj_cancel_right)
+
+lemma conj_disj_distrib2: "(y \<^bold>\<squnion> z) \<^bold>\<sqinter> x = (y \<^bold>\<sqinter> x) \<^bold>\<squnion> (z \<^bold>\<sqinter> x)"
+ by (simp only: conj.commute conj_disj_distrib)
+
+lemmas conj_disj_distribs = conj_disj_distrib conj_disj_distrib2
+
+lemma conj_assoc: "(x \<^bold>\<sqinter> y) \<^bold>\<sqinter> z = x \<^bold>\<sqinter> (y \<^bold>\<sqinter> z)"
+ by (fact ac_simps)
+
+lemma conj_commute: "x \<^bold>\<sqinter> y = y \<^bold>\<sqinter> x"
+ by (fact ac_simps)
+
+lemmas conj_left_commute = conj.left_commute
+lemmas conj_ac = conj.assoc conj.commute conj.left_commute
+
+lemma conj_one_left: "\<^bold>1 \<^bold>\<sqinter> x = x"
+ by (fact conj.left_neutral)
+
+lemma conj_left_absorb: "x \<^bold>\<sqinter> (x \<^bold>\<sqinter> y) = x \<^bold>\<sqinter> y"
+ by (fact conj.left_idem)
+
+lemma conj_absorb: "x \<^bold>\<sqinter> x = x"
+ by (fact conj.idem)
+
+
+subsection \<open>Disjunction\<close>
+
+interpretation dual: boolean_algebra "(\<^bold>\<squnion>)" "(\<^bold>\<sqinter>)" compl \<open>\<^bold>1\<close> \<open>\<^bold>0\<close>
+ apply standard
+ apply (rule disj_conj_distrib)
+ apply (rule conj_disj_distrib)
+ apply simp_all
+ done
+
+lemma compl_zero [simp]: "\<^bold>- \<^bold>0 = \<^bold>1"
+ by (fact dual.compl_one)
+
+lemma disj_one_right [simp]: "x \<^bold>\<squnion> \<^bold>1 = \<^bold>1"
+ by (fact dual.conj_zero_right)
+
+lemma disj_one_left [simp]: "\<^bold>1 \<^bold>\<squnion> x = \<^bold>1"
+ by (fact dual.conj_zero_left)
+
+lemma disj_cancel_left [simp]: "\<^bold>- x \<^bold>\<squnion> x = \<^bold>1"
+ by (rule dual.conj_cancel_left)
+
+lemma disj_conj_distrib2: "(y \<^bold>\<sqinter> z) \<^bold>\<squnion> x = (y \<^bold>\<squnion> x) \<^bold>\<sqinter> (z \<^bold>\<squnion> x)"
+ by (rule dual.conj_disj_distrib2)
+
+lemmas disj_conj_distribs = disj_conj_distrib disj_conj_distrib2
+
+lemma disj_assoc: "(x \<^bold>\<squnion> y) \<^bold>\<squnion> z = x \<^bold>\<squnion> (y \<^bold>\<squnion> z)"
+ by (fact ac_simps)
+
+lemma disj_commute: "x \<^bold>\<squnion> y = y \<^bold>\<squnion> x"
+ by (fact ac_simps)
+
+lemmas disj_left_commute = disj.left_commute
+
+lemmas disj_ac = disj.assoc disj.commute disj.left_commute
+
+lemma disj_zero_left: "\<^bold>0 \<^bold>\<squnion> x = x"
+ by (fact disj.left_neutral)
+
+lemma disj_left_absorb: "x \<^bold>\<squnion> (x \<^bold>\<squnion> y) = x \<^bold>\<squnion> y"
+ by (fact disj.left_idem)
+
+lemma disj_absorb: "x \<^bold>\<squnion> x = x"
+ by (fact disj.idem)
+
+
+subsection \<open>De Morgan's Laws\<close>
+
+lemma de_Morgan_conj [simp]: "\<^bold>- (x \<^bold>\<sqinter> y) = \<^bold>- x \<^bold>\<squnion> \<^bold>- y"
+proof (rule compl_unique)
+ have "(x \<^bold>\<sqinter> y) \<^bold>\<sqinter> (\<^bold>- x \<^bold>\<squnion> \<^bold>- y) = ((x \<^bold>\<sqinter> y) \<^bold>\<sqinter> \<^bold>- x) \<^bold>\<squnion> ((x \<^bold>\<sqinter> y) \<^bold>\<sqinter> \<^bold>- y)"
+ by (rule conj_disj_distrib)
+ also have "\<dots> = (y \<^bold>\<sqinter> (x \<^bold>\<sqinter> \<^bold>- x)) \<^bold>\<squnion> (x \<^bold>\<sqinter> (y \<^bold>\<sqinter> \<^bold>- y))"
+ by (simp only: conj_ac)
+ finally show "(x \<^bold>\<sqinter> y) \<^bold>\<sqinter> (\<^bold>- x \<^bold>\<squnion> \<^bold>- y) = \<^bold>0"
+ by (simp only: conj_cancel_right conj_zero_right disj_zero_right)
+next
+ have "(x \<^bold>\<sqinter> y) \<^bold>\<squnion> (\<^bold>- x \<^bold>\<squnion> \<^bold>- y) = (x \<^bold>\<squnion> (\<^bold>- x \<^bold>\<squnion> \<^bold>- y)) \<^bold>\<sqinter> (y \<^bold>\<squnion> (\<^bold>- x \<^bold>\<squnion> \<^bold>- y))"
+ by (rule disj_conj_distrib2)
+ also have "\<dots> = (\<^bold>- y \<^bold>\<squnion> (x \<^bold>\<squnion> \<^bold>- x)) \<^bold>\<sqinter> (\<^bold>- x \<^bold>\<squnion> (y \<^bold>\<squnion> \<^bold>- y))"
+ by (simp only: disj_ac)
+ finally show "(x \<^bold>\<sqinter> y) \<^bold>\<squnion> (\<^bold>- x \<^bold>\<squnion> \<^bold>- y) = \<^bold>1"
+ by (simp only: disj_cancel_right disj_one_right conj_one_right)
+qed
+
+lemma de_Morgan_disj [simp]: "\<^bold>- (x \<^bold>\<squnion> y) = \<^bold>- x \<^bold>\<sqinter> \<^bold>- y"
+ using dual.boolean_algebra_axioms by (rule boolean_algebra.de_Morgan_conj)
+
+
+subsection \<open>Symmetric Difference\<close>
+
+definition xor :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<^bold>\<ominus>" 65)
+ where "x \<^bold>\<ominus> y = (x \<^bold>\<sqinter> \<^bold>- y) \<^bold>\<squnion> (\<^bold>- x \<^bold>\<sqinter> y)"
+
+sublocale xor: comm_monoid xor \<open>\<^bold>0\<close>
+proof
+ fix x y z :: 'a
+ let ?t = "(x \<^bold>\<sqinter> y \<^bold>\<sqinter> z) \<^bold>\<squnion> (x \<^bold>\<sqinter> \<^bold>- y \<^bold>\<sqinter> \<^bold>- z) \<^bold>\<squnion> (\<^bold>- x \<^bold>\<sqinter> y \<^bold>\<sqinter> \<^bold>- z) \<^bold>\<squnion> (\<^bold>- x \<^bold>\<sqinter> \<^bold>- y \<^bold>\<sqinter> z)"
+ have "?t \<^bold>\<squnion> (z \<^bold>\<sqinter> x \<^bold>\<sqinter> \<^bold>- x) \<^bold>\<squnion> (z \<^bold>\<sqinter> y \<^bold>\<sqinter> \<^bold>- y) = ?t \<^bold>\<squnion> (x \<^bold>\<sqinter> y \<^bold>\<sqinter> \<^bold>- y) \<^bold>\<squnion> (x \<^bold>\<sqinter> z \<^bold>\<sqinter> \<^bold>- z)"
+ by (simp only: conj_cancel_right conj_zero_right)
+ then show "(x \<^bold>\<ominus> y) \<^bold>\<ominus> z = x \<^bold>\<ominus> (y \<^bold>\<ominus> z)"
+ by (simp only: xor_def de_Morgan_disj de_Morgan_conj double_compl)
+ (simp only: conj_disj_distribs conj_ac disj_ac)
+ show "x \<^bold>\<ominus> y = y \<^bold>\<ominus> x"
+ by (simp only: xor_def conj_commute disj_commute)
+ show "x \<^bold>\<ominus> \<^bold>0 = x"
+ by (simp add: xor_def)
+qed
+
+lemmas xor_assoc = xor.assoc
+lemmas xor_commute = xor.commute
+lemmas xor_left_commute = xor.left_commute
+
+lemmas xor_ac = xor.assoc xor.commute xor.left_commute
+
+lemma xor_def2: "x \<^bold>\<ominus> y = (x \<^bold>\<squnion> y) \<^bold>\<sqinter> (\<^bold>- x \<^bold>\<squnion> \<^bold>- y)"
+ using conj.commute conj_disj_distrib2 disj.commute xor_def by auto
+
+lemma xor_zero_right: "x \<^bold>\<ominus> \<^bold>0 = x"
+ by (fact xor.comm_neutral)
+
+lemma xor_zero_left: "\<^bold>0 \<^bold>\<ominus> x = x"
+ by (fact xor.left_neutral)
+
+lemma xor_one_right [simp]: "x \<^bold>\<ominus> \<^bold>1 = \<^bold>- x"
+ by (simp only: xor_def compl_one conj_zero_right conj_one_right disj_zero_left)
+
+lemma xor_one_left [simp]: "\<^bold>1 \<^bold>\<ominus> x = \<^bold>- x"
+ by (subst xor_commute) (rule xor_one_right)
+
+lemma xor_self [simp]: "x \<^bold>\<ominus> x = \<^bold>0"
+ by (simp only: xor_def conj_cancel_right conj_cancel_left disj_zero_right)
+
+lemma xor_left_self [simp]: "x \<^bold>\<ominus> (x \<^bold>\<ominus> y) = y"
+ by (simp only: xor_assoc [symmetric] xor_self xor_zero_left)
+
+lemma xor_compl_left [simp]: "\<^bold>- x \<^bold>\<ominus> y = \<^bold>- (x \<^bold>\<ominus> y)"
+ by (simp add: ac_simps flip: xor_one_left)
+
+lemma xor_compl_right [simp]: "x \<^bold>\<ominus> \<^bold>- y = \<^bold>- (x \<^bold>\<ominus> y)"
+ using xor_commute xor_compl_left by auto
+
+lemma xor_cancel_right: "x \<^bold>\<ominus> \<^bold>- x = \<^bold>1"
+ by (simp only: xor_compl_right xor_self compl_zero)
+
+lemma xor_cancel_left: "\<^bold>- x \<^bold>\<ominus> x = \<^bold>1"
+ by (simp only: xor_compl_left xor_self compl_zero)
+
+lemma conj_xor_distrib: "x \<^bold>\<sqinter> (y \<^bold>\<ominus> z) = (x \<^bold>\<sqinter> y) \<^bold>\<ominus> (x \<^bold>\<sqinter> z)"
+proof -
+ have *: "(x \<^bold>\<sqinter> y \<^bold>\<sqinter> \<^bold>- z) \<^bold>\<squnion> (x \<^bold>\<sqinter> \<^bold>- y \<^bold>\<sqinter> z) =
+ (y \<^bold>\<sqinter> x \<^bold>\<sqinter> \<^bold>- x) \<^bold>\<squnion> (z \<^bold>\<sqinter> x \<^bold>\<sqinter> \<^bold>- x) \<^bold>\<squnion> (x \<^bold>\<sqinter> y \<^bold>\<sqinter> \<^bold>- z) \<^bold>\<squnion> (x \<^bold>\<sqinter> \<^bold>- y \<^bold>\<sqinter> z)"
+ by (simp only: conj_cancel_right conj_zero_right disj_zero_left)
+ then show "x \<^bold>\<sqinter> (y \<^bold>\<ominus> z) = (x \<^bold>\<sqinter> y) \<^bold>\<ominus> (x \<^bold>\<sqinter> z)"
+ by (simp (no_asm_use) only:
+ xor_def de_Morgan_disj de_Morgan_conj double_compl
+ conj_disj_distribs conj_ac disj_ac)
+qed
+
+lemma conj_xor_distrib2: "(y \<^bold>\<ominus> z) \<^bold>\<sqinter> x = (y \<^bold>\<sqinter> x) \<^bold>\<ominus> (z \<^bold>\<sqinter> x)"
+ by (simp add: conj.commute conj_xor_distrib)
+
+lemmas conj_xor_distribs = conj_xor_distrib conj_xor_distrib2
+
+end
+
+end
--- a/src/HOL/Code_Numeral.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/Code_Numeral.thy Mon Aug 02 10:01:06 2021 +0000
@@ -5,7 +5,7 @@
section \<open>Numeric types for code generation onto target language numerals only\<close>
theory Code_Numeral
-imports Divides Lifting
+imports Divides Lifting Bit_Operations
begin
subsection \<open>Type of target language integers\<close>
@@ -1203,12 +1203,6 @@
hide_const (open) Nat
-lifting_update integer.lifting
-lifting_forget integer.lifting
-
-lifting_update natural.lifting
-lifting_forget natural.lifting
-
code_reflect Code_Numeral
datatypes natural
functions "Code_Numeral.Suc" "0 :: natural" "1 :: natural"
@@ -1217,4 +1211,87 @@
"modulo :: natural \<Rightarrow> _"
integer_of_natural natural_of_integer
+
+subsection \<open>Bit operations\<close>
+
+instantiation integer :: ring_bit_operations
+begin
+
+lift_definition not_integer :: \<open>integer \<Rightarrow> integer\<close>
+ is not .
+
+lift_definition and_integer :: \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close>
+ is \<open>and\<close> .
+
+lift_definition or_integer :: \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close>
+ is or .
+
+lift_definition xor_integer :: \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close>
+ is xor .
+
+lift_definition mask_integer :: \<open>nat \<Rightarrow> integer\<close>
+ is mask .
+
+lift_definition set_bit_integer :: \<open>nat \<Rightarrow> integer \<Rightarrow> integer\<close>
+ is set_bit .
+
+lift_definition unset_bit_integer :: \<open>nat \<Rightarrow> integer \<Rightarrow> integer\<close>
+ is unset_bit .
+
+lift_definition flip_bit_integer :: \<open>nat \<Rightarrow> integer \<Rightarrow> integer\<close>
+ is flip_bit .
+
+instance by (standard; transfer)
+ (simp_all add: minus_eq_not_minus_1 mask_eq_exp_minus_1
+ bit_not_iff bit_and_iff bit_or_iff bit_xor_iff
+ set_bit_def bit_unset_bit_iff flip_bit_def)
+
end
+
+lemma [code]:
+ \<open>mask n = 2 ^ n - (1::integer)\<close>
+ by (simp add: mask_eq_exp_minus_1)
+
+instantiation natural :: semiring_bit_operations
+begin
+
+lift_definition and_natural :: \<open>natural \<Rightarrow> natural \<Rightarrow> natural\<close>
+ is \<open>and\<close> .
+
+lift_definition or_natural :: \<open>natural \<Rightarrow> natural \<Rightarrow> natural\<close>
+ is or .
+
+lift_definition xor_natural :: \<open>natural \<Rightarrow> natural \<Rightarrow> natural\<close>
+ is xor .
+
+lift_definition mask_natural :: \<open>nat \<Rightarrow> natural\<close>
+ is mask .
+
+lift_definition set_bit_natural :: \<open>nat \<Rightarrow> natural \<Rightarrow> natural\<close>
+ is set_bit .
+
+lift_definition unset_bit_natural :: \<open>nat \<Rightarrow> natural \<Rightarrow> natural\<close>
+ is unset_bit .
+
+lift_definition flip_bit_natural :: \<open>nat \<Rightarrow> natural \<Rightarrow> natural\<close>
+ is flip_bit .
+
+instance by (standard; transfer)
+ (simp_all add: mask_eq_exp_minus_1
+ bit_and_iff bit_or_iff bit_xor_iff
+ set_bit_def bit_unset_bit_iff flip_bit_def)
+
+end
+
+lemma [code]:
+ \<open>integer_of_natural (mask n) = mask n\<close>
+ by transfer (simp add: mask_eq_exp_minus_1 of_nat_diff)
+
+
+lifting_update integer.lifting
+lifting_forget integer.lifting
+
+lifting_update natural.lifting
+lifting_forget natural.lifting
+
+end
--- a/src/HOL/Decision_Procs/Cooper.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/Decision_Procs/Cooper.thy Mon Aug 02 10:01:06 2021 +0000
@@ -46,7 +46,7 @@
datatype (plugins del: size) fm = T | F
| Lt num | Le num | Gt num | Ge num | Eq num | NEq num
| Dvd int num | NDvd int num
- | NOT fm | And fm fm | Or fm fm | Imp fm fm | Iff fm fm
+ | Not fm | And fm fm | Or fm fm | Imp fm fm | Iff fm fm
| E fm | A fm | Closed nat | NClosed nat
instantiation fm :: size
@@ -54,7 +54,7 @@
primrec size_fm :: "fm \<Rightarrow> nat"
where
- "size_fm (NOT p) = 1 + size_fm p"
+ "size_fm (Not p) = 1 + size_fm p"
| "size_fm (And p q) = 1 + size_fm p + size_fm q"
| "size_fm (Or p q) = 1 + size_fm p + size_fm q"
| "size_fm (Imp p q) = 3 + size_fm p + size_fm q"
@@ -94,7 +94,7 @@
| "Ifm bbs bs (NEq a) \<longleftrightarrow> Inum bs a \<noteq> 0"
| "Ifm bbs bs (Dvd i b) \<longleftrightarrow> i dvd Inum bs b"
| "Ifm bbs bs (NDvd i b) \<longleftrightarrow> \<not> i dvd Inum bs b"
- | "Ifm bbs bs (NOT p) \<longleftrightarrow> \<not> Ifm bbs bs p"
+ | "Ifm bbs bs (Not p) \<longleftrightarrow> \<not> Ifm bbs bs p"
| "Ifm bbs bs (And p q) \<longleftrightarrow> Ifm bbs bs p \<and> Ifm bbs bs q"
| "Ifm bbs bs (Or p q) \<longleftrightarrow> Ifm bbs bs p \<or> Ifm bbs bs q"
| "Ifm bbs bs (Imp p q) \<longleftrightarrow> (Ifm bbs bs p \<longrightarrow> Ifm bbs bs q)"
@@ -109,25 +109,25 @@
"prep (E T) = T"
| "prep (E F) = F"
| "prep (E (Or p q)) = Or (prep (E p)) (prep (E q))"
- | "prep (E (Imp p q)) = Or (prep (E (NOT p))) (prep (E q))"
- | "prep (E (Iff p q)) = Or (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))"
- | "prep (E (NOT (And p q))) = Or (prep (E (NOT p))) (prep (E(NOT q)))"
- | "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))"
- | "prep (E (NOT (Iff p q))) = Or (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))"
+ | "prep (E (Imp p q)) = Or (prep (E (Not p))) (prep (E q))"
+ | "prep (E (Iff p q)) = Or (prep (E (And p q))) (prep (E (And (Not p) (Not q))))"
+ | "prep (E (Not (And p q))) = Or (prep (E (Not p))) (prep (E(Not q)))"
+ | "prep (E (Not (Imp p q))) = prep (E (And p (Not q)))"
+ | "prep (E (Not (Iff p q))) = Or (prep (E (And p (Not q)))) (prep (E(And (Not p) q)))"
| "prep (E p) = E (prep p)"
| "prep (A (And p q)) = And (prep (A p)) (prep (A q))"
- | "prep (A p) = prep (NOT (E (NOT p)))"
- | "prep (NOT (NOT p)) = prep p"
- | "prep (NOT (And p q)) = Or (prep (NOT p)) (prep (NOT q))"
- | "prep (NOT (A p)) = prep (E (NOT p))"
- | "prep (NOT (Or p q)) = And (prep (NOT p)) (prep (NOT q))"
- | "prep (NOT (Imp p q)) = And (prep p) (prep (NOT q))"
- | "prep (NOT (Iff p q)) = Or (prep (And p (NOT q))) (prep (And (NOT p) q))"
- | "prep (NOT p) = NOT (prep p)"
+ | "prep (A p) = prep (Not (E (Not p)))"
+ | "prep (Not (Not p)) = prep p"
+ | "prep (Not (And p q)) = Or (prep (Not p)) (prep (Not q))"
+ | "prep (Not (A p)) = prep (E (Not p))"
+ | "prep (Not (Or p q)) = And (prep (Not p)) (prep (Not q))"
+ | "prep (Not (Imp p q)) = And (prep p) (prep (Not q))"
+ | "prep (Not (Iff p q)) = Or (prep (And p (Not q))) (prep (And (Not p) q))"
+ | "prep (Not p) = Not (prep p)"
| "prep (Or p q) = Or (prep p) (prep q)"
| "prep (And p q) = And (prep p) (prep q)"
- | "prep (Imp p q) = prep (Or (NOT p) q)"
- | "prep (Iff p q) = Or (prep (And p q)) (prep (And (NOT p) (NOT q)))"
+ | "prep (Imp p q) = prep (Or (Not p) q)"
+ | "prep (Iff p q) = Or (prep (And p q)) (prep (And (Not p) (Not q)))"
| "prep p = p"
lemma prep: "Ifm bbs bs (prep p) = Ifm bbs bs p"
@@ -138,7 +138,7 @@
where
"qfree (E p) \<longleftrightarrow> False"
| "qfree (A p) \<longleftrightarrow> False"
- | "qfree (NOT p) \<longleftrightarrow> qfree p"
+ | "qfree (Not p) \<longleftrightarrow> qfree p"
| "qfree (And p q) \<longleftrightarrow> qfree p \<and> qfree q"
| "qfree (Or p q) \<longleftrightarrow> qfree p \<and> qfree q"
| "qfree (Imp p q) \<longleftrightarrow> qfree p \<and> qfree q"
@@ -175,7 +175,7 @@
| "bound0 (NEq a) \<longleftrightarrow> numbound0 a"
| "bound0 (Dvd i a) \<longleftrightarrow> numbound0 a"
| "bound0 (NDvd i a) \<longleftrightarrow> numbound0 a"
- | "bound0 (NOT p) \<longleftrightarrow> bound0 p"
+ | "bound0 (Not p) \<longleftrightarrow> bound0 p"
| "bound0 (And p q) \<longleftrightarrow> bound0 p \<and> bound0 q"
| "bound0 (Or p q) \<longleftrightarrow> bound0 p \<and> bound0 q"
| "bound0 (Imp p q) \<longleftrightarrow> bound0 p \<and> bound0 q"
@@ -220,7 +220,7 @@
| "subst0 t (NEq a) = NEq (numsubst0 t a)"
| "subst0 t (Dvd i a) = Dvd i (numsubst0 t a)"
| "subst0 t (NDvd i a) = NDvd i (numsubst0 t a)"
- | "subst0 t (NOT p) = NOT (subst0 t p)"
+ | "subst0 t (Not p) = Not (subst0 t p)"
| "subst0 t (And p q) = And (subst0 t p) (subst0 t q)"
| "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)"
| "subst0 t (Imp p q) = Imp (subst0 t p) (subst0 t q)"
@@ -254,7 +254,7 @@
| "decr (NEq a) = NEq (decrnum a)"
| "decr (Dvd i a) = Dvd i (decrnum a)"
| "decr (NDvd i a) = NDvd i (decrnum a)"
- | "decr (NOT p) = NOT (decr p)"
+ | "decr (Not p) = Not (decr p)"
| "decr (And p q) = And (decr p) (decr q)"
| "decr (Or p q) = Or (decr p) (decr q)"
| "decr (Imp p q) = Imp (decr p) (decr q)"
@@ -508,12 +508,12 @@
fun not :: "fm \<Rightarrow> fm"
where
- "not (NOT p) = p"
+ "not (Not p) = p"
| "not T = F"
| "not F = T"
- | "not p = NOT p"
+ | "not p = Not p"
-lemma not: "Ifm bbs bs (not p) = Ifm bbs bs (NOT p)"
+lemma not: "Ifm bbs bs (not p) = Ifm bbs bs (Not p)"
by (cases p) auto
lemma not_qf: "qfree p \<Longrightarrow> qfree (not p)"
@@ -596,7 +596,7 @@
| "simpfm (Or p q) = disj (simpfm p) (simpfm q)"
| "simpfm (Imp p q) = imp (simpfm p) (simpfm q)"
| "simpfm (Iff p q) = iff (simpfm p) (simpfm q)"
- | "simpfm (NOT p) = not (simpfm p)"
+ | "simpfm (Not p) = not (simpfm p)"
| "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v < 0 then T else F | _ \<Rightarrow> Lt a')"
| "simpfm (Le a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v \<le> 0 then T else F | _ \<Rightarrow> Le a')"
| "simpfm (Gt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v > 0 then T else F | _ \<Rightarrow> Gt a')"
@@ -825,8 +825,8 @@
fun qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"
where
"qelim (E p) = (\<lambda>qe. DJ qe (qelim p qe))"
- | "qelim (A p) = (\<lambda>qe. not (qe ((qelim (NOT p) qe))))"
- | "qelim (NOT p) = (\<lambda>qe. not (qelim p qe))"
+ | "qelim (A p) = (\<lambda>qe. not (qe ((qelim (Not p) qe))))"
+ | "qelim (Not p) = (\<lambda>qe. not (qelim p qe))"
| "qelim (And p q) = (\<lambda>qe. conj (qelim p qe) (qelim q qe))"
| "qelim (Or p q) = (\<lambda>qe. disj (qelim p qe) (qelim q qe))"
| "qelim (Imp p q) = (\<lambda>qe. imp (qelim p qe) (qelim q qe))"
@@ -995,8 +995,8 @@
where
"zlfm (And p q) = And (zlfm p) (zlfm q)"
| "zlfm (Or p q) = Or (zlfm p) (zlfm q)"
- | "zlfm (Imp p q) = Or (zlfm (NOT p)) (zlfm q)"
- | "zlfm (Iff p q) = Or (And (zlfm p) (zlfm q)) (And (zlfm (NOT p)) (zlfm (NOT q)))"
+ | "zlfm (Imp p q) = Or (zlfm (Not p)) (zlfm q)"
+ | "zlfm (Iff p q) = Or (And (zlfm p) (zlfm q)) (And (zlfm (Not p)) (zlfm (Not q)))"
| "zlfm (Lt a) =
(let (c, r) = zsplit0 a in
if c = 0 then Lt r else
@@ -1041,23 +1041,23 @@
if c = 0 then NDvd \<bar>i\<bar> r
else if c > 0 then NDvd \<bar>i\<bar> (CN 0 c r)
else NDvd \<bar>i\<bar> (CN 0 (- c) (Neg r)))"
- | "zlfm (NOT (And p q)) = Or (zlfm (NOT p)) (zlfm (NOT q))"
- | "zlfm (NOT (Or p q)) = And (zlfm (NOT p)) (zlfm (NOT q))"
- | "zlfm (NOT (Imp p q)) = And (zlfm p) (zlfm (NOT q))"
- | "zlfm (NOT (Iff p q)) = Or (And(zlfm p) (zlfm(NOT q))) (And (zlfm(NOT p)) (zlfm q))"
- | "zlfm (NOT (NOT p)) = zlfm p"
- | "zlfm (NOT T) = F"
- | "zlfm (NOT F) = T"
- | "zlfm (NOT (Lt a)) = zlfm (Ge a)"
- | "zlfm (NOT (Le a)) = zlfm (Gt a)"
- | "zlfm (NOT (Gt a)) = zlfm (Le a)"
- | "zlfm (NOT (Ge a)) = zlfm (Lt a)"
- | "zlfm (NOT (Eq a)) = zlfm (NEq a)"
- | "zlfm (NOT (NEq a)) = zlfm (Eq a)"
- | "zlfm (NOT (Dvd i a)) = zlfm (NDvd i a)"
- | "zlfm (NOT (NDvd i a)) = zlfm (Dvd i a)"
- | "zlfm (NOT (Closed P)) = NClosed P"
- | "zlfm (NOT (NClosed P)) = Closed P"
+ | "zlfm (Not (And p q)) = Or (zlfm (Not p)) (zlfm (Not q))"
+ | "zlfm (Not (Or p q)) = And (zlfm (Not p)) (zlfm (Not q))"
+ | "zlfm (Not (Imp p q)) = And (zlfm p) (zlfm (Not q))"
+ | "zlfm (Not (Iff p q)) = Or (And(zlfm p) (zlfm(Not q))) (And (zlfm(Not p)) (zlfm q))"
+ | "zlfm (Not (Not p)) = zlfm p"
+ | "zlfm (Not T) = F"
+ | "zlfm (Not F) = T"
+ | "zlfm (Not (Lt a)) = zlfm (Ge a)"
+ | "zlfm (Not (Le a)) = zlfm (Gt a)"
+ | "zlfm (Not (Gt a)) = zlfm (Le a)"
+ | "zlfm (Not (Ge a)) = zlfm (Lt a)"
+ | "zlfm (Not (Eq a)) = zlfm (NEq a)"
+ | "zlfm (Not (NEq a)) = zlfm (Eq a)"
+ | "zlfm (Not (Dvd i a)) = zlfm (NDvd i a)"
+ | "zlfm (Not (NDvd i a)) = zlfm (Dvd i a)"
+ | "zlfm (Not (Closed P)) = NClosed P"
+ | "zlfm (Not (NClosed P)) = Closed P"
| "zlfm p = p"
lemma zlfm_I:
@@ -2425,8 +2425,8 @@
@{code Or} (fm_of_term ps vs t1, fm_of_term ps vs t2)
| fm_of_term ps vs (\<^term>\<open>HOL.implies\<close> $ t1 $ t2) =
@{code Imp} (fm_of_term ps vs t1, fm_of_term ps vs t2)
- | fm_of_term ps vs (\<^term>\<open>Not\<close> $ t') =
- @{code NOT} (fm_of_term ps vs t')
+ | fm_of_term ps vs (\<^term>\<open>HOL.Not\<close> $ t') =
+ @{code Not} (fm_of_term ps vs t')
| fm_of_term ps vs (Const (\<^const_name>\<open>Ex\<close>, _) $ Abs (xn, xT, p)) =
let
val (xn', p') = Syntax_Trans.variant_abs (xn, xT, p); (* FIXME !? *)
@@ -2467,12 +2467,12 @@
| term_of_fm ps vs (@{code Eq} t) =
\<^term>\<open>(=) :: int \<Rightarrow> int \<Rightarrow> bool\<close> $ term_of_num vs t $ \<^term>\<open>0::int\<close>
| term_of_fm ps vs (@{code NEq} t) =
- term_of_fm ps vs (@{code NOT} (@{code Eq} t))
+ term_of_fm ps vs (@{code Not} (@{code Eq} t))
| term_of_fm ps vs (@{code Dvd} (i, t)) =
\<^term>\<open>(dvd) :: int \<Rightarrow> int \<Rightarrow> bool\<close> $ term_of_num vs (@{code C} i) $ term_of_num vs t
| term_of_fm ps vs (@{code NDvd} (i, t)) =
- term_of_fm ps vs (@{code NOT} (@{code Dvd} (i, t)))
- | term_of_fm ps vs (@{code NOT} t') =
+ term_of_fm ps vs (@{code Not} (@{code Dvd} (i, t)))
+ | term_of_fm ps vs (@{code Not} t') =
HOLogic.Not $ term_of_fm ps vs t'
| term_of_fm ps vs (@{code And} (t1, t2)) =
HOLogic.conj $ term_of_fm ps vs t1 $ term_of_fm ps vs t2
@@ -2486,7 +2486,7 @@
let
val q = @{code integer_of_nat} n
in (fst o the) (find_first (fn (_, m) => m = q) ps) end
- | term_of_fm ps vs (@{code NClosed} n) = term_of_fm ps vs (@{code NOT} (@{code Closed} n));
+ | term_of_fm ps vs (@{code NClosed} n) = term_of_fm ps vs (@{code Not} (@{code Closed} n));
fun term_bools acc t =
let
@@ -2494,7 +2494,7 @@
member (=) [\<^term>\<open>HOL.conj\<close>, \<^term>\<open>HOL.disj\<close>, \<^term>\<open>HOL.implies\<close>,
\<^term>\<open>(=) :: bool \<Rightarrow> _\<close>,
\<^term>\<open>(=) :: int \<Rightarrow> _\<close>, \<^term>\<open>(<) :: int \<Rightarrow> _\<close>,
- \<^term>\<open>(\<le>) :: int \<Rightarrow> _\<close>, \<^term>\<open>Not\<close>, \<^term>\<open>All :: (int \<Rightarrow> _) \<Rightarrow> _\<close>,
+ \<^term>\<open>(\<le>) :: int \<Rightarrow> _\<close>, \<^term>\<open>HOL.Not\<close>, \<^term>\<open>All :: (int \<Rightarrow> _) \<Rightarrow> _\<close>,
\<^term>\<open>Ex :: (int \<Rightarrow> _) \<Rightarrow> _\<close>, \<^term>\<open>True\<close>, \<^term>\<open>False\<close>]
fun is_ty t = not (fastype_of t = HOLogic.boolT)
in
--- a/src/HOL/Decision_Procs/Ferrack.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/Decision_Procs/Ferrack.thy Mon Aug 02 10:01:06 2021 +0000
@@ -46,14 +46,14 @@
(* FORMULAE *)
datatype (plugins del: size) fm =
T| F| Lt num| Le num| Gt num| Ge num| Eq num| NEq num|
- NOT fm| And fm fm| Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm
+ Not fm| And fm fm| Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm
instantiation fm :: size
begin
primrec size_fm :: "fm \<Rightarrow> nat"
where
- "size_fm (NOT p) = 1 + size_fm p"
+ "size_fm (Not p) = 1 + size_fm p"
| "size_fm (And p q) = 1 + size_fm p + size_fm q"
| "size_fm (Or p q) = 1 + size_fm p + size_fm q"
| "size_fm (Imp p q) = 3 + size_fm p + size_fm q"
@@ -87,7 +87,7 @@
| "Ifm bs (Ge a) = (Inum bs a \<ge> 0)"
| "Ifm bs (Eq a) = (Inum bs a = 0)"
| "Ifm bs (NEq a) = (Inum bs a \<noteq> 0)"
-| "Ifm bs (NOT p) = (\<not> (Ifm bs p))"
+| "Ifm bs (Not p) = (\<not> (Ifm bs p))"
| "Ifm bs (And p q) = (Ifm bs p \<and> Ifm bs q)"
| "Ifm bs (Or p q) = (Ifm bs p \<or> Ifm bs q)"
| "Ifm bs (Imp p q) = ((Ifm bs p) \<longrightarrow> (Ifm bs q))"
@@ -104,7 +104,7 @@
lemma IfmEqSub: "\<lbrakk> Inum bs s = s' ; Inum bs t = t' \<rbrakk> \<Longrightarrow> Ifm bs (Eq (Sub s t)) = (s' = t')"
by simp
-lemma IfmNOT: " (Ifm bs p = P) \<Longrightarrow> (Ifm bs (NOT p) = (\<not>P))"
+lemma IfmNot: " (Ifm bs p = P) \<Longrightarrow> (Ifm bs (Not p) = (\<not>P))"
by simp
lemma IfmAnd: " \<lbrakk> Ifm bs p = P ; Ifm bs q = Q\<rbrakk> \<Longrightarrow> (Ifm bs (And p q) = (P \<and> Q))"
@@ -127,12 +127,12 @@
fun not:: "fm \<Rightarrow> fm"
where
- "not (NOT p) = p"
+ "not (Not p) = p"
| "not T = F"
| "not F = T"
-| "not p = NOT p"
+| "not p = Not p"
-lemma not[simp]: "Ifm bs (not p) = Ifm bs (NOT p)"
+lemma not[simp]: "Ifm bs (not p) = Ifm bs (Not p)"
by (cases p) auto
definition conj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
@@ -172,7 +172,7 @@
where
"iff p q =
(if p = q then T
- else if p = NOT q \<or> NOT p = q then F
+ else if p = Not q \<or> Not p = q then F
else if p = F then not q
else if q = F then not p
else if p = T then q
@@ -180,7 +180,7 @@
else Iff p q)"
lemma iff[simp]: "Ifm bs (iff p q) = Ifm bs (Iff p q)"
- by (unfold iff_def, cases "p = q", simp, cases "p = NOT q", simp) (cases "NOT p = q", auto)
+ by (unfold iff_def, cases "p = q", simp, cases "p = Not q", simp) (cases "Not p = q", auto)
lemma conj_simps:
"conj F Q = F"
@@ -209,19 +209,19 @@
"P \<noteq> T \<Longrightarrow> P \<noteq> F \<Longrightarrow> P \<noteq> Q \<Longrightarrow> Q \<noteq> T \<Longrightarrow> Q \<noteq> F \<Longrightarrow> imp P Q = Imp P Q"
by (simp_all add: imp_def)
-lemma trivNOT: "p \<noteq> NOT p" "NOT p \<noteq> p"
+lemma trivNot: "p \<noteq> Not p" "Not p \<noteq> p"
by (induct p) auto
lemma iff_simps:
"iff p p = T"
- "iff p (NOT p) = F"
- "iff (NOT p) p = F"
+ "iff p (Not p) = F"
+ "iff (Not p) p = F"
"iff p F = not p"
"iff F p = not p"
- "p \<noteq> NOT T \<Longrightarrow> iff T p = p"
- "p\<noteq> NOT T \<Longrightarrow> iff p T = p"
- "p\<noteq>q \<Longrightarrow> p\<noteq> NOT q \<Longrightarrow> q\<noteq> NOT p \<Longrightarrow> p\<noteq> F \<Longrightarrow> q\<noteq> F \<Longrightarrow> p \<noteq> T \<Longrightarrow> q \<noteq> T \<Longrightarrow> iff p q = Iff p q"
- using trivNOT
+ "p \<noteq> Not T \<Longrightarrow> iff T p = p"
+ "p\<noteq> Not T \<Longrightarrow> iff p T = p"
+ "p\<noteq>q \<Longrightarrow> p\<noteq> Not q \<Longrightarrow> q\<noteq> Not p \<Longrightarrow> p\<noteq> F \<Longrightarrow> q\<noteq> F \<Longrightarrow> p \<noteq> T \<Longrightarrow> q \<noteq> T \<Longrightarrow> iff p q = Iff p q"
+ using trivNot
by (simp_all add: iff_def, cases p, auto)
(* Quantifier freeness *)
@@ -229,7 +229,7 @@
where
"qfree (E p) = False"
| "qfree (A p) = False"
-| "qfree (NOT p) = qfree p"
+| "qfree (Not p) = qfree p"
| "qfree (And p q) = (qfree p \<and> qfree q)"
| "qfree (Or p q) = (qfree p \<and> qfree q)"
| "qfree (Imp p q) = (qfree p \<and> qfree q)"
@@ -262,7 +262,7 @@
| "bound0 (Ge a) = numbound0 a"
| "bound0 (Eq a) = numbound0 a"
| "bound0 (NEq a) = numbound0 a"
-| "bound0 (NOT p) = bound0 p"
+| "bound0 (Not p) = bound0 p"
| "bound0 (And p q) = (bound0 p \<and> bound0 q)"
| "bound0 (Or p q) = (bound0 p \<and> bound0 q)"
| "bound0 (Imp p q) = ((bound0 p) \<and> (bound0 q))"
@@ -321,7 +321,7 @@
| "decr (Ge a) = Ge (decrnum a)"
| "decr (Eq a) = Eq (decrnum a)"
| "decr (NEq a) = NEq (decrnum a)"
-| "decr (NOT p) = NOT (decr p)"
+| "decr (Not p) = Not (decr p)"
| "decr (And p q) = conj (decr p) (decr q)"
| "decr (Or p q) = disj (decr p) (decr q)"
| "decr (Imp p q) = imp (decr p) (decr q)"
@@ -890,7 +890,7 @@
| "simpfm (Or p q) = disj (simpfm p) (simpfm q)"
| "simpfm (Imp p q) = imp (simpfm p) (simpfm q)"
| "simpfm (Iff p q) = iff (simpfm p) (simpfm q)"
-| "simpfm (NOT p) = not (simpfm p)"
+| "simpfm (Not p) = not (simpfm p)"
| "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v < 0) then T else F | _ \<Rightarrow> Lt a')"
| "simpfm (Le a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v \<le> 0) then T else F | _ \<Rightarrow> Le a')"
| "simpfm (Gt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v > 0) then T else F | _ \<Rightarrow> Gt a')"
@@ -1031,25 +1031,25 @@
"prep (E T) = T"
| "prep (E F) = F"
| "prep (E (Or p q)) = disj (prep (E p)) (prep (E q))"
-| "prep (E (Imp p q)) = disj (prep (E (NOT p))) (prep (E q))"
-| "prep (E (Iff p q)) = disj (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))"
-| "prep (E (NOT (And p q))) = disj (prep (E (NOT p))) (prep (E(NOT q)))"
-| "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))"
-| "prep (E (NOT (Iff p q))) = disj (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))"
+| "prep (E (Imp p q)) = disj (prep (E (Not p))) (prep (E q))"
+| "prep (E (Iff p q)) = disj (prep (E (And p q))) (prep (E (And (Not p) (Not q))))"
+| "prep (E (Not (And p q))) = disj (prep (E (Not p))) (prep (E(Not q)))"
+| "prep (E (Not (Imp p q))) = prep (E (And p (Not q)))"
+| "prep (E (Not (Iff p q))) = disj (prep (E (And p (Not q)))) (prep (E(And (Not p) q)))"
| "prep (E p) = E (prep p)"
| "prep (A (And p q)) = conj (prep (A p)) (prep (A q))"
-| "prep (A p) = prep (NOT (E (NOT p)))"
-| "prep (NOT (NOT p)) = prep p"
-| "prep (NOT (And p q)) = disj (prep (NOT p)) (prep (NOT q))"
-| "prep (NOT (A p)) = prep (E (NOT p))"
-| "prep (NOT (Or p q)) = conj (prep (NOT p)) (prep (NOT q))"
-| "prep (NOT (Imp p q)) = conj (prep p) (prep (NOT q))"
-| "prep (NOT (Iff p q)) = disj (prep (And p (NOT q))) (prep (And (NOT p) q))"
-| "prep (NOT p) = not (prep p)"
+| "prep (A p) = prep (Not (E (Not p)))"
+| "prep (Not (Not p)) = prep p"
+| "prep (Not (And p q)) = disj (prep (Not p)) (prep (Not q))"
+| "prep (Not (A p)) = prep (E (Not p))"
+| "prep (Not (Or p q)) = conj (prep (Not p)) (prep (Not q))"
+| "prep (Not (Imp p q)) = conj (prep p) (prep (Not q))"
+| "prep (Not (Iff p q)) = disj (prep (And p (Not q))) (prep (And (Not p) q))"
+| "prep (Not p) = not (prep p)"
| "prep (Or p q) = disj (prep p) (prep q)"
| "prep (And p q) = conj (prep p) (prep q)"
-| "prep (Imp p q) = prep (Or (NOT p) q)"
-| "prep (Iff p q) = disj (prep (And p q)) (prep (And (NOT p) (NOT q)))"
+| "prep (Imp p q) = prep (Or (Not p) q)"
+| "prep (Iff p q) = disj (prep (And p q)) (prep (And (Not p) (Not q)))"
| "prep p = p"
lemma prep: "\<And>bs. Ifm bs (prep p) = Ifm bs p"
@@ -1059,8 +1059,8 @@
fun qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"
where
"qelim (E p) = (\<lambda>qe. DJ qe (qelim p qe))"
-| "qelim (A p) = (\<lambda>qe. not (qe ((qelim (NOT p) qe))))"
-| "qelim (NOT p) = (\<lambda>qe. not (qelim p qe))"
+| "qelim (A p) = (\<lambda>qe. not (qe ((qelim (Not p) qe))))"
+| "qelim (Not p) = (\<lambda>qe. not (qelim p qe))"
| "qelim (And p q) = (\<lambda>qe. conj (qelim p qe) (qelim q qe))"
| "qelim (Or p q) = (\<lambda>qe. disj (qelim p qe) (qelim q qe))"
| "qelim (Imp p q) = (\<lambda>qe. imp (qelim p qe) (qelim q qe))"
@@ -1219,27 +1219,27 @@
where
"rlfm (And p q) = conj (rlfm p) (rlfm q)"
| "rlfm (Or p q) = disj (rlfm p) (rlfm q)"
-| "rlfm (Imp p q) = disj (rlfm (NOT p)) (rlfm q)"
-| "rlfm (Iff p q) = disj (conj (rlfm p) (rlfm q)) (conj (rlfm (NOT p)) (rlfm (NOT q)))"
+| "rlfm (Imp p q) = disj (rlfm (Not p)) (rlfm q)"
+| "rlfm (Iff p q) = disj (conj (rlfm p) (rlfm q)) (conj (rlfm (Not p)) (rlfm (Not q)))"
| "rlfm (Lt a) = case_prod lt (rsplit0 a)"
| "rlfm (Le a) = case_prod le (rsplit0 a)"
| "rlfm (Gt a) = case_prod gt (rsplit0 a)"
| "rlfm (Ge a) = case_prod ge (rsplit0 a)"
| "rlfm (Eq a) = case_prod eq (rsplit0 a)"
| "rlfm (NEq a) = case_prod neq (rsplit0 a)"
-| "rlfm (NOT (And p q)) = disj (rlfm (NOT p)) (rlfm (NOT q))"
-| "rlfm (NOT (Or p q)) = conj (rlfm (NOT p)) (rlfm (NOT q))"
-| "rlfm (NOT (Imp p q)) = conj (rlfm p) (rlfm (NOT q))"
-| "rlfm (NOT (Iff p q)) = disj (conj(rlfm p) (rlfm(NOT q))) (conj(rlfm(NOT p)) (rlfm q))"
-| "rlfm (NOT (NOT p)) = rlfm p"
-| "rlfm (NOT T) = F"
-| "rlfm (NOT F) = T"
-| "rlfm (NOT (Lt a)) = rlfm (Ge a)"
-| "rlfm (NOT (Le a)) = rlfm (Gt a)"
-| "rlfm (NOT (Gt a)) = rlfm (Le a)"
-| "rlfm (NOT (Ge a)) = rlfm (Lt a)"
-| "rlfm (NOT (Eq a)) = rlfm (NEq a)"
-| "rlfm (NOT (NEq a)) = rlfm (Eq a)"
+| "rlfm (Not (And p q)) = disj (rlfm (Not p)) (rlfm (Not q))"
+| "rlfm (Not (Or p q)) = conj (rlfm (Not p)) (rlfm (Not q))"
+| "rlfm (Not (Imp p q)) = conj (rlfm p) (rlfm (Not q))"
+| "rlfm (Not (Iff p q)) = disj (conj(rlfm p) (rlfm(Not q))) (conj(rlfm(Not p)) (rlfm q))"
+| "rlfm (Not (Not p)) = rlfm p"
+| "rlfm (Not T) = F"
+| "rlfm (Not F) = T"
+| "rlfm (Not (Lt a)) = rlfm (Ge a)"
+| "rlfm (Not (Le a)) = rlfm (Gt a)"
+| "rlfm (Not (Gt a)) = rlfm (Le a)"
+| "rlfm (Not (Ge a)) = rlfm (Lt a)"
+| "rlfm (Not (Eq a)) = rlfm (NEq a)"
+| "rlfm (Not (NEq a)) = rlfm (Eq a)"
| "rlfm p = p"
lemma rlfm_I:
@@ -2489,7 +2489,7 @@
| fm_of_term vs (\<^term>\<open>HOL.conj\<close> $ t1 $ t2) = @{code And} (fm_of_term vs t1, fm_of_term vs t2)
| fm_of_term vs (\<^term>\<open>HOL.disj\<close> $ t1 $ t2) = @{code Or} (fm_of_term vs t1, fm_of_term vs t2)
| fm_of_term vs (\<^term>\<open>HOL.implies\<close> $ t1 $ t2) = @{code Imp} (fm_of_term vs t1, fm_of_term vs t2)
- | fm_of_term vs (\<^term>\<open>Not\<close> $ t') = @{code NOT} (fm_of_term vs t')
+ | fm_of_term vs (\<^term>\<open>HOL.Not\<close> $ t') = @{code Not} (fm_of_term vs t')
| fm_of_term vs (Const (\<^const_name>\<open>Ex\<close>, _) $ Abs (xn, xT, p)) =
@{code E} (fm_of_term (("", dummyT) :: vs) p)
| fm_of_term vs (Const (\<^const_name>\<open>All\<close>, _) $ Abs (xn, xT, p)) =
@@ -2520,8 +2520,8 @@
\<^term>\<open>0::real\<close> $ term_of_num vs t
| term_of_fm vs (@{code Eq} t) = \<^term>\<open>(=) :: real \<Rightarrow> real \<Rightarrow> bool\<close> $
term_of_num vs t $ \<^term>\<open>0::real\<close>
- | term_of_fm vs (@{code NEq} t) = term_of_fm vs (@{code NOT} (@{code Eq} t))
- | term_of_fm vs (@{code NOT} t') = HOLogic.Not $ term_of_fm vs t'
+ | term_of_fm vs (@{code NEq} t) = term_of_fm vs (@{code Not} (@{code Eq} t))
+ | term_of_fm vs (@{code Not} t') = HOLogic.Not $ term_of_fm vs t'
| term_of_fm vs (@{code And} (t1, t2)) = HOLogic.conj $ term_of_fm vs t1 $ term_of_fm vs t2
| term_of_fm vs (@{code Or} (t1, t2)) = HOLogic.disj $ term_of_fm vs t1 $ term_of_fm vs t2
| term_of_fm vs (@{code Imp} (t1, t2)) = HOLogic.imp $ term_of_fm vs t1 $ term_of_fm vs t2
--- a/src/HOL/Decision_Procs/MIR.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/Decision_Procs/MIR.thy Mon Aug 02 10:01:06 2021 +0000
@@ -187,14 +187,14 @@
datatype (plugins del: size) fm =
T | F | Lt num | Le num | Gt num | Ge num | Eq num | NEq num |
Dvd int num | NDvd int num |
- NOT fm | And fm fm | Or fm fm | Imp fm fm | Iff fm fm | E fm | A fm
+ Not fm | And fm fm | Or fm fm | Imp fm fm | Iff fm fm | E fm | A fm
instantiation fm :: size
begin
primrec size_fm :: "fm \<Rightarrow> nat"
where
- "size_fm (NOT p) = 1 + size_fm p"
+ "size_fm (Not p) = 1 + size_fm p"
| "size_fm (And p q) = 1 + size_fm p + size_fm q"
| "size_fm (Or p q) = 1 + size_fm p + size_fm q"
| "size_fm (Imp p q) = 3 + size_fm p + size_fm q"
@@ -232,7 +232,7 @@
| "Ifm bs (NEq a) \<longleftrightarrow> Inum bs a \<noteq> 0"
| "Ifm bs (Dvd i b) \<longleftrightarrow> real_of_int i rdvd Inum bs b"
| "Ifm bs (NDvd i b) \<longleftrightarrow> \<not> (real_of_int i rdvd Inum bs b)"
-| "Ifm bs (NOT p) \<longleftrightarrow> \<not> (Ifm bs p)"
+| "Ifm bs (Not p) \<longleftrightarrow> \<not> (Ifm bs p)"
| "Ifm bs (And p q) \<longleftrightarrow> Ifm bs p \<and> Ifm bs q"
| "Ifm bs (Or p q) \<longleftrightarrow> Ifm bs p \<or> Ifm bs q"
| "Ifm bs (Imp p q) \<longleftrightarrow> (Ifm bs p \<longrightarrow> Ifm bs q)"
@@ -245,25 +245,25 @@
"prep (E T) = T"
| "prep (E F) = F"
| "prep (E (Or p q)) = Or (prep (E p)) (prep (E q))"
-| "prep (E (Imp p q)) = Or (prep (E (NOT p))) (prep (E q))"
-| "prep (E (Iff p q)) = Or (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))"
-| "prep (E (NOT (And p q))) = Or (prep (E (NOT p))) (prep (E(NOT q)))"
-| "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))"
-| "prep (E (NOT (Iff p q))) = Or (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))"
+| "prep (E (Imp p q)) = Or (prep (E (Not p))) (prep (E q))"
+| "prep (E (Iff p q)) = Or (prep (E (And p q))) (prep (E (And (Not p) (Not q))))"
+| "prep (E (Not (And p q))) = Or (prep (E (Not p))) (prep (E(Not q)))"
+| "prep (E (Not (Imp p q))) = prep (E (And p (Not q)))"
+| "prep (E (Not (Iff p q))) = Or (prep (E (And p (Not q)))) (prep (E(And (Not p) q)))"
| "prep (E p) = E (prep p)"
| "prep (A (And p q)) = And (prep (A p)) (prep (A q))"
-| "prep (A p) = prep (NOT (E (NOT p)))"
-| "prep (NOT (NOT p)) = prep p"
-| "prep (NOT (And p q)) = Or (prep (NOT p)) (prep (NOT q))"
-| "prep (NOT (A p)) = prep (E (NOT p))"
-| "prep (NOT (Or p q)) = And (prep (NOT p)) (prep (NOT q))"
-| "prep (NOT (Imp p q)) = And (prep p) (prep (NOT q))"
-| "prep (NOT (Iff p q)) = Or (prep (And p (NOT q))) (prep (And (NOT p) q))"
-| "prep (NOT p) = NOT (prep p)"
+| "prep (A p) = prep (Not (E (Not p)))"
+| "prep (Not (Not p)) = prep p"
+| "prep (Not (And p q)) = Or (prep (Not p)) (prep (Not q))"
+| "prep (Not (A p)) = prep (E (Not p))"
+| "prep (Not (Or p q)) = And (prep (Not p)) (prep (Not q))"
+| "prep (Not (Imp p q)) = And (prep p) (prep (Not q))"
+| "prep (Not (Iff p q)) = Or (prep (And p (Not q))) (prep (And (Not p) q))"
+| "prep (Not p) = Not (prep p)"
| "prep (Or p q) = Or (prep p) (prep q)"
| "prep (And p q) = And (prep p) (prep q)"
-| "prep (Imp p q) = prep (Or (NOT p) q)"
-| "prep (Iff p q) = Or (prep (And p q)) (prep (And (NOT p) (NOT q)))"
+| "prep (Imp p q) = prep (Or (Not p) q)"
+| "prep (Iff p q) = Or (prep (And p q)) (prep (And (Not p) (Not q)))"
| "prep p = p"
lemma prep: "\<And> bs. Ifm bs (prep p) = Ifm bs p"
@@ -275,7 +275,7 @@
where
"qfree (E p) = False"
| "qfree (A p) = False"
-| "qfree (NOT p) = qfree p"
+| "qfree (Not p) = qfree p"
| "qfree (And p q) = (qfree p \<and> qfree q)"
| "qfree (Or p q) = (qfree p \<and> qfree q)"
| "qfree (Imp p q) = (qfree p \<and> qfree q)"
@@ -323,7 +323,7 @@
| "bound0 (NEq a) = numbound0 a"
| "bound0 (Dvd i a) = numbound0 a"
| "bound0 (NDvd i a) = numbound0 a"
-| "bound0 (NOT p) = bound0 p"
+| "bound0 (Not p) = bound0 p"
| "bound0 (And p q) = (bound0 p \<and> bound0 q)"
| "bound0 (Or p q) = (bound0 p \<and> bound0 q)"
| "bound0 (Imp p q) = ((bound0 p) \<and> (bound0 q))"
@@ -365,7 +365,7 @@
| "subst0 t (NEq a) = NEq (numsubst0 t a)"
| "subst0 t (Dvd i a) = Dvd i (numsubst0 t a)"
| "subst0 t (NDvd i a) = NDvd i (numsubst0 t a)"
-| "subst0 t (NOT p) = NOT (subst0 t p)"
+| "subst0 t (Not p) = Not (subst0 t p)"
| "subst0 t (And p q) = And (subst0 t p) (subst0 t q)"
| "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)"
| "subst0 t (Imp p q) = Imp (subst0 t p) (subst0 t q)"
@@ -398,7 +398,7 @@
| "decr (NEq a) = NEq (decrnum a)"
| "decr (Dvd i a) = Dvd i (decrnum a)"
| "decr (NDvd i a) = NDvd i (decrnum a)"
-| "decr (NOT p) = NOT (decr p)"
+| "decr (Not p) = Not (decr p)"
| "decr (And p q) = And (decr p) (decr q)"
| "decr (Or p q) = Or (decr p) (decr q)"
| "decr (Imp p q) = Imp (decr p) (decr q)"
@@ -1044,7 +1044,7 @@
fun not:: "fm \<Rightarrow> fm"
where
- "not (NOT p) = p"
+ "not (Not p) = p"
| "not T = F"
| "not F = T"
| "not (Lt t) = Ge t"
@@ -1057,8 +1057,8 @@
| "not (NDvd i t) = Dvd i t"
| "not (And p q) = Or (not p) (not q)"
| "not (Or p q) = And (not p) (not q)"
-| "not p = NOT p"
-lemma not[simp]: "Ifm bs (not p) = Ifm bs (NOT p)"
+| "not p = Not p"
+lemma not[simp]: "Ifm bs (not p) = Ifm bs (Not p)"
by (induct p) auto
lemma not_qf[simp]: "qfree p \<Longrightarrow> qfree (not p)"
by (induct p) auto
@@ -1188,7 +1188,7 @@
| "simpfm (Or p q) = disj (simpfm p) (simpfm q)"
| "simpfm (Imp p q) = imp (simpfm p) (simpfm q)"
| "simpfm (Iff p q) = iff (simpfm p) (simpfm q)"
-| "simpfm (NOT p) = not (simpfm p)"
+| "simpfm (Not p) = not (simpfm p)"
| "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v < 0) then T else F
| _ \<Rightarrow> Lt (reducecoeff a'))"
| "simpfm (Le a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v \<le> 0) then T else F | _ \<Rightarrow> Le (reducecoeff a'))"
@@ -1459,11 +1459,11 @@
fun qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"
where
"qelim (E p) = (\<lambda> qe. DJ (CJNB qe) (qelim p qe))"
-| "qelim (A p) = (\<lambda> qe. not (qe ((qelim (NOT p) qe))))"
-| "qelim (NOT p) = (\<lambda> qe. not (qelim p qe))"
+| "qelim (A p) = (\<lambda> qe. not (qe ((qelim (Not p) qe))))"
+| "qelim (Not p) = (\<lambda> qe. not (qelim p qe))"
| "qelim (And p q) = (\<lambda> qe. conj (qelim p qe) (qelim q qe))"
| "qelim (Or p q) = (\<lambda> qe. disj (qelim p qe) (qelim q qe))"
-| "qelim (Imp p q) = (\<lambda> qe. disj (qelim (NOT p) qe) (qelim q qe))"
+| "qelim (Imp p q) = (\<lambda> qe. disj (qelim (Not p) qe) (qelim q qe))"
| "qelim (Iff p q) = (\<lambda> qe. iff (qelim p qe) (qelim q qe))"
| "qelim p = (\<lambda> y. simpfm p)"
@@ -1611,8 +1611,8 @@
where
"zlfm (And p q) = conj (zlfm p) (zlfm q)"
| "zlfm (Or p q) = disj (zlfm p) (zlfm q)"
-| "zlfm (Imp p q) = disj (zlfm (NOT p)) (zlfm q)"
-| "zlfm (Iff p q) = disj (conj (zlfm p) (zlfm q)) (conj (zlfm (NOT p)) (zlfm (NOT q)))"
+| "zlfm (Imp p q) = disj (zlfm (Not p)) (zlfm q)"
+| "zlfm (Iff p q) = disj (conj (zlfm p) (zlfm q)) (conj (zlfm (Not p)) (zlfm (Not q)))"
| "zlfm (Lt a) = (let (c,r) = zsplit0 a in
if c=0 then Lt r else
if c>0 then Or (Lt (CN 0 c (Neg (Floor (Neg r))))) (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Lt (Add (Floor (Neg r)) r)))
@@ -1647,21 +1647,21 @@
if c=0 then NDvd \<bar>i\<bar> r else
if c>0 then Or (NEq (Sub (Floor r) r)) (NDvd \<bar>i\<bar> (CN 0 c (Floor r)))
else Or (NEq (Sub (Floor r) r)) (NDvd \<bar>i\<bar> (CN 0 (-c) (Neg (Floor r))))))"
-| "zlfm (NOT (And p q)) = disj (zlfm (NOT p)) (zlfm (NOT q))"
-| "zlfm (NOT (Or p q)) = conj (zlfm (NOT p)) (zlfm (NOT q))"
-| "zlfm (NOT (Imp p q)) = conj (zlfm p) (zlfm (NOT q))"
-| "zlfm (NOT (Iff p q)) = disj (conj(zlfm p) (zlfm(NOT q))) (conj (zlfm(NOT p)) (zlfm q))"
-| "zlfm (NOT (NOT p)) = zlfm p"
-| "zlfm (NOT T) = F"
-| "zlfm (NOT F) = T"
-| "zlfm (NOT (Lt a)) = zlfm (Ge a)"
-| "zlfm (NOT (Le a)) = zlfm (Gt a)"
-| "zlfm (NOT (Gt a)) = zlfm (Le a)"
-| "zlfm (NOT (Ge a)) = zlfm (Lt a)"
-| "zlfm (NOT (Eq a)) = zlfm (NEq a)"
-| "zlfm (NOT (NEq a)) = zlfm (Eq a)"
-| "zlfm (NOT (Dvd i a)) = zlfm (NDvd i a)"
-| "zlfm (NOT (NDvd i a)) = zlfm (Dvd i a)"
+| "zlfm (Not (And p q)) = disj (zlfm (Not p)) (zlfm (Not q))"
+| "zlfm (Not (Or p q)) = conj (zlfm (Not p)) (zlfm (Not q))"
+| "zlfm (Not (Imp p q)) = conj (zlfm p) (zlfm (Not q))"
+| "zlfm (Not (Iff p q)) = disj (conj(zlfm p) (zlfm(Not q))) (conj (zlfm(Not p)) (zlfm q))"
+| "zlfm (Not (Not p)) = zlfm p"
+| "zlfm (Not T) = F"
+| "zlfm (Not F) = T"
+| "zlfm (Not (Lt a)) = zlfm (Ge a)"
+| "zlfm (Not (Le a)) = zlfm (Gt a)"
+| "zlfm (Not (Gt a)) = zlfm (Le a)"
+| "zlfm (Not (Ge a)) = zlfm (Lt a)"
+| "zlfm (Not (Eq a)) = zlfm (NEq a)"
+| "zlfm (Not (NEq a)) = zlfm (Eq a)"
+| "zlfm (Not (Dvd i a)) = zlfm (NDvd i a)"
+| "zlfm (Not (NDvd i a)) = zlfm (Dvd i a)"
| "zlfm p = p"
lemma split_int_less_real:
@@ -3210,14 +3210,14 @@
by (simp add: isint_iff)
from 10 have id: "j dvd d" by simp
from ie[symmetric] have "?p i = (\<not> (real_of_int j rdvd real_of_int (c*i + \<lfloor>?e\<rfloor>)))" by simp
- also have "\<dots> = Not (j dvd c*i + \<lfloor>?e\<rfloor>)"
+ also have "\<dots> \<longleftrightarrow> \<not> (j dvd c*i + \<lfloor>?e\<rfloor>)"
using int_rdvd_iff [where i="j" and t="c*i + \<lfloor>?e\<rfloor>"] by simp
- also have "\<dots> = Not (j dvd c*i - c*d + \<lfloor>?e\<rfloor>)"
+ also have "\<dots> \<longleftrightarrow> \<not> (j dvd c*i - c*d + \<lfloor>?e\<rfloor>)"
using dvd_period[OF id, where x="c*i" and c="-c" and t="\<lfloor>?e\<rfloor>"] by simp
- also have "\<dots> = Not (real_of_int j rdvd real_of_int (c*i - c*d + \<lfloor>?e\<rfloor>))"
+ also have "\<dots> \<longleftrightarrow> \<not> (real_of_int j rdvd real_of_int (c*i - c*d + \<lfloor>?e\<rfloor>))"
using int_rdvd_iff[where i="j" and t="(c*i - c*d + \<lfloor>?e\<rfloor>)",symmetric, simplified]
ie by simp
- also have "\<dots> = Not (real_of_int j rdvd real_of_int (c*(i - d)) + ?e)"
+ also have "\<dots> \<longleftrightarrow> \<not> (real_of_int j rdvd real_of_int (c*(i - d)) + ?e)"
using ie by (simp add:algebra_simps)
finally show ?case
using numbound0_I[OF bn,where b="real_of_int i - real_of_int d" and b'="real_of_int i" and bs="bs"] p
@@ -3960,8 +3960,8 @@
where
"rlfm (And p q) = conj (rlfm p) (rlfm q)"
| "rlfm (Or p q) = disj (rlfm p) (rlfm q)"
-| "rlfm (Imp p q) = disj (rlfm (NOT p)) (rlfm q)"
-| "rlfm (Iff p q) = disj (conj(rlfm p) (rlfm q)) (conj(rlfm (NOT p)) (rlfm (NOT q)))"
+| "rlfm (Imp p q) = disj (rlfm (Not p)) (rlfm q)"
+| "rlfm (Iff p q) = disj (conj(rlfm p) (rlfm q)) (conj(rlfm (Not p)) (rlfm (Not q)))"
| "rlfm (Lt a) = rsplit lt a"
| "rlfm (Le a) = rsplit le a"
| "rlfm (Gt a) = rsplit gt a"
@@ -3970,21 +3970,21 @@
| "rlfm (NEq a) = rsplit neq a"
| "rlfm (Dvd i a) = rsplit (\<lambda> t. DVD i t) a"
| "rlfm (NDvd i a) = rsplit (\<lambda> t. NDVD i t) a"
-| "rlfm (NOT (And p q)) = disj (rlfm (NOT p)) (rlfm (NOT q))"
-| "rlfm (NOT (Or p q)) = conj (rlfm (NOT p)) (rlfm (NOT q))"
-| "rlfm (NOT (Imp p q)) = conj (rlfm p) (rlfm (NOT q))"
-| "rlfm (NOT (Iff p q)) = disj (conj(rlfm p) (rlfm(NOT q))) (conj(rlfm(NOT p)) (rlfm q))"
-| "rlfm (NOT (NOT p)) = rlfm p"
-| "rlfm (NOT T) = F"
-| "rlfm (NOT F) = T"
-| "rlfm (NOT (Lt a)) = simpfm (rlfm (Ge a))"
-| "rlfm (NOT (Le a)) = simpfm (rlfm (Gt a))"
-| "rlfm (NOT (Gt a)) = simpfm (rlfm (Le a))"
-| "rlfm (NOT (Ge a)) = simpfm (rlfm (Lt a))"
-| "rlfm (NOT (Eq a)) = simpfm (rlfm (NEq a))"
-| "rlfm (NOT (NEq a)) = simpfm (rlfm (Eq a))"
-| "rlfm (NOT (Dvd i a)) = simpfm (rlfm (NDvd i a))"
-| "rlfm (NOT (NDvd i a)) = simpfm (rlfm (Dvd i a))"
+| "rlfm (Not (And p q)) = disj (rlfm (Not p)) (rlfm (Not q))"
+| "rlfm (Not (Or p q)) = conj (rlfm (Not p)) (rlfm (Not q))"
+| "rlfm (Not (Imp p q)) = conj (rlfm p) (rlfm (Not q))"
+| "rlfm (Not (Iff p q)) = disj (conj(rlfm p) (rlfm(Not q))) (conj(rlfm(Not p)) (rlfm q))"
+| "rlfm (Not (Not p)) = rlfm p"
+| "rlfm (Not T) = F"
+| "rlfm (Not F) = T"
+| "rlfm (Not (Lt a)) = simpfm (rlfm (Ge a))"
+| "rlfm (Not (Le a)) = simpfm (rlfm (Gt a))"
+| "rlfm (Not (Gt a)) = simpfm (rlfm (Le a))"
+| "rlfm (Not (Ge a)) = simpfm (rlfm (Lt a))"
+| "rlfm (Not (Eq a)) = simpfm (rlfm (NEq a))"
+| "rlfm (Not (NEq a)) = simpfm (rlfm (Eq a))"
+| "rlfm (Not (Dvd i a)) = simpfm (rlfm (NDvd i a))"
+| "rlfm (Not (NDvd i a)) = simpfm (rlfm (Dvd i a))"
| "rlfm p = p"
lemma bound0at_l : "\<lbrakk>isatom p ; bound0 p\<rbrakk> \<Longrightarrow> isrlfm p"
@@ -4834,7 +4834,7 @@
| "exsplit (Or p q) = Or (exsplit p) (exsplit q)"
| "exsplit (Imp p q) = Imp (exsplit p) (exsplit q)"
| "exsplit (Iff p q) = Iff (exsplit p) (exsplit q)"
-| "exsplit (NOT p) = NOT (exsplit p)"
+| "exsplit (Not p) = Not (exsplit p)"
| "exsplit p = p"
lemma exsplitnum:
@@ -5624,8 +5624,8 @@
@{code Or} (fm_of_term vs t1, fm_of_term vs t2)
| fm_of_term vs (\<^term>\<open>HOL.implies\<close> $ t1 $ t2) =
@{code Imp} (fm_of_term vs t1, fm_of_term vs t2)
- | fm_of_term vs (\<^term>\<open>Not\<close> $ t') =
- @{code NOT} (fm_of_term vs t')
+ | fm_of_term vs (\<^term>\<open>HOL.Not\<close> $ t') =
+ @{code Not} (fm_of_term vs t')
| fm_of_term vs (Const (\<^const_name>\<open>Ex\<close>, _) $ Abs (xn, xT, p)) =
@{code E} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p)
| fm_of_term vs (Const (\<^const_name>\<open>All\<close>, _) $ Abs (xn, xT, p)) =
@@ -5664,12 +5664,12 @@
| term_of_fm vs (@{code Eq} t) =
\<^term>\<open>(=) :: real \<Rightarrow> real \<Rightarrow> bool\<close> $ term_of_num vs t $ \<^term>\<open>0::real\<close>
| term_of_fm vs (@{code NEq} t) =
- term_of_fm vs (@{code NOT} (@{code Eq} t))
+ term_of_fm vs (@{code Not} (@{code Eq} t))
| term_of_fm vs (@{code Dvd} (i, t)) =
\<^term>\<open>(rdvd)\<close> $ term_of_num vs (@{code C} i) $ term_of_num vs t
| term_of_fm vs (@{code NDvd} (i, t)) =
- term_of_fm vs (@{code NOT} (@{code Dvd} (i, t)))
- | term_of_fm vs (@{code NOT} t') =
+ term_of_fm vs (@{code Not} (@{code Dvd} (i, t)))
+ | term_of_fm vs (@{code Not} t') =
HOLogic.Not $ term_of_fm vs t'
| term_of_fm vs (@{code And} (t1, t2)) =
HOLogic.conj $ term_of_fm vs t1 $ term_of_fm vs t2
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Mon Aug 02 10:01:06 2021 +0000
@@ -501,14 +501,14 @@
subsection \<open>Formulae\<close>
datatype (plugins del: size) fm = T | F | Le tm | Lt tm | Eq tm | NEq tm |
- NOT fm | And fm fm | Or fm fm | Imp fm fm | Iff fm fm | E fm | A fm
+ Not fm | And fm fm | Or fm fm | Imp fm fm | Iff fm fm | E fm | A fm
instantiation fm :: size
begin
primrec size_fm :: "fm \<Rightarrow> nat"
where
- "size_fm (NOT p) = 1 + size_fm p"
+ "size_fm (Not p) = 1 + size_fm p"
| "size_fm (And p q) = 1 + size_fm p + size_fm q"
| "size_fm (Or p q) = 1 + size_fm p + size_fm q"
| "size_fm (Imp p q) = 3 + size_fm p + size_fm q"
@@ -538,7 +538,7 @@
| "Ifm vs bs (Le a) = (Itm vs bs a \<le> 0)"
| "Ifm vs bs (Eq a) = (Itm vs bs a = 0)"
| "Ifm vs bs (NEq a) = (Itm vs bs a \<noteq> 0)"
- | "Ifm vs bs (NOT p) = (\<not> (Ifm vs bs p))"
+ | "Ifm vs bs (Not p) = (\<not> (Ifm vs bs p))"
| "Ifm vs bs (And p q) = (Ifm vs bs p \<and> Ifm vs bs q)"
| "Ifm vs bs (Or p q) = (Ifm vs bs p \<or> Ifm vs bs q)"
| "Ifm vs bs (Imp p q) = ((Ifm vs bs p) \<longrightarrow> (Ifm vs bs q))"
@@ -548,17 +548,17 @@
fun not:: "fm \<Rightarrow> fm"
where
- "not (NOT (NOT p)) = not p"
- | "not (NOT p) = p"
+ "not (Not (Not p)) = not p"
+ | "not (Not p) = p"
| "not T = F"
| "not F = T"
| "not (Lt t) = Le (tmneg t)"
| "not (Le t) = Lt (tmneg t)"
| "not (Eq t) = NEq t"
| "not (NEq t) = Eq t"
- | "not p = NOT p"
-
-lemma not[simp]: "Ifm vs bs (not p) = Ifm vs bs (NOT p)"
+ | "not p = Not p"
+
+lemma not[simp]: "Ifm vs bs (not p) = Ifm vs bs (Not p)"
by (induct p rule: not.induct) auto
definition conj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
@@ -596,7 +596,7 @@
definition iff :: "fm \<Rightarrow> fm \<Rightarrow> fm"
where "iff p q \<equiv>
(if p = q then T
- else if p = NOT q \<or> NOT p = q then F
+ else if p = Not q \<or> Not p = q then F
else if p = F then not q
else if q = F then not p
else if p = T then q
@@ -604,14 +604,14 @@
else Iff p q)"
lemma iff[simp]: "Ifm vs bs (iff p q) = Ifm vs bs (Iff p q)"
- by (unfold iff_def, cases "p = q", simp, cases "p = NOT q", simp) (cases "NOT p= q", auto)
+ by (unfold iff_def, cases "p = q", simp, cases "p = Not q", simp) (cases "Not p= q", auto)
text \<open>Quantifier freeness.\<close>
fun qfree:: "fm \<Rightarrow> bool"
where
"qfree (E p) = False"
| "qfree (A p) = False"
- | "qfree (NOT p) = qfree p"
+ | "qfree (Not p) = qfree p"
| "qfree (And p q) = (qfree p \<and> qfree q)"
| "qfree (Or p q) = (qfree p \<and> qfree q)"
| "qfree (Imp p q) = (qfree p \<and> qfree q)"
@@ -627,7 +627,7 @@
| "boundslt n (Le t) = tmboundslt n t"
| "boundslt n (Eq t) = tmboundslt n t"
| "boundslt n (NEq t) = tmboundslt n t"
- | "boundslt n (NOT p) = boundslt n p"
+ | "boundslt n (Not p) = boundslt n p"
| "boundslt n (And p q) = (boundslt n p \<and> boundslt n q)"
| "boundslt n (Or p q) = (boundslt n p \<and> boundslt n q)"
| "boundslt n (Imp p q) = ((boundslt n p) \<and> (boundslt n q))"
@@ -643,7 +643,7 @@
| "bound0 (Le a) = tmbound0 a"
| "bound0 (Eq a) = tmbound0 a"
| "bound0 (NEq a) = tmbound0 a"
- | "bound0 (NOT p) = bound0 p"
+ | "bound0 (Not p) = bound0 p"
| "bound0 (And p q) = (bound0 p \<and> bound0 q)"
| "bound0 (Or p q) = (bound0 p \<and> bound0 q)"
| "bound0 (Imp p q) = ((bound0 p) \<and> (bound0 q))"
@@ -664,7 +664,7 @@
| "bound m (Le t) = tmbound m t"
| "bound m (Eq t) = tmbound m t"
| "bound m (NEq t) = tmbound m t"
- | "bound m (NOT p) = bound m p"
+ | "bound m (Not p) = bound m p"
| "bound m (And p q) = (bound m p \<and> bound m q)"
| "bound m (Or p q) = (bound m p \<and> bound m q)"
| "bound m (Imp p q) = ((bound m p) \<and> (bound m q))"
@@ -706,7 +706,7 @@
| "decr0 (Le a) = Le (decrtm0 a)"
| "decr0 (Eq a) = Eq (decrtm0 a)"
| "decr0 (NEq a) = NEq (decrtm0 a)"
- | "decr0 (NOT p) = NOT (decr0 p)"
+ | "decr0 (Not p) = Not (decr0 p)"
| "decr0 (And p q) = conj (decr0 p) (decr0 q)"
| "decr0 (Or p q) = disj (decr0 p) (decr0 q)"
| "decr0 (Imp p q) = imp (decr0 p) (decr0 q)"
@@ -726,7 +726,7 @@
| "decr m (Le t) = (Le (decrtm m t))"
| "decr m (Eq t) = (Eq (decrtm m t))"
| "decr m (NEq t) = (NEq (decrtm m t))"
- | "decr m (NOT p) = NOT (decr m p)"
+ | "decr m (Not p) = Not (decr m p)"
| "decr m (And p q) = conj (decr m p) (decr m q)"
| "decr m (Or p q) = disj (decr m p) (decr m q)"
| "decr m (Imp p q) = imp (decr m p) (decr m q)"
@@ -774,7 +774,7 @@
| "subst0 t (Le a) = Le (tmsubst0 t a)"
| "subst0 t (Eq a) = Eq (tmsubst0 t a)"
| "subst0 t (NEq a) = NEq (tmsubst0 t a)"
- | "subst0 t (NOT p) = NOT (subst0 t p)"
+ | "subst0 t (Not p) = Not (subst0 t p)"
| "subst0 t (And p q) = And (subst0 t p) (subst0 t q)"
| "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)"
| "subst0 t (Imp p q) = Imp (subst0 t p) (subst0 t q)"
@@ -802,7 +802,7 @@
| "subst n t (Le a) = Le (tmsubst n t a)"
| "subst n t (Eq a) = Eq (tmsubst n t a)"
| "subst n t (NEq a) = NEq (tmsubst n t a)"
- | "subst n t (NOT p) = NOT (subst n t p)"
+ | "subst n t (Not p) = Not (subst n t p)"
| "subst n t (And p q) = And (subst n t p) (subst n t q)"
| "subst n t (Or p q) = Or (subst n t p) (subst n t q)"
| "subst n t (Imp p q) = Imp (subst n t p) (subst n t q)"
@@ -1081,7 +1081,7 @@
| "islin (NEq (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
| "islin (Lt (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
| "islin (Le (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
- | "islin (NOT p) = False"
+ | "islin (Not p) = False"
| "islin (Imp p q) = False"
| "islin (Iff p q) = False"
| "islin p = bound0 p"
@@ -1520,21 +1520,21 @@
| "simpfm (NEq t) = simpneq(simptm t)"
| "simpfm (And p q) = conj (simpfm p) (simpfm q)"
| "simpfm (Or p q) = disj (simpfm p) (simpfm q)"
- | "simpfm (Imp p q) = disj (simpfm (NOT p)) (simpfm q)"
+ | "simpfm (Imp p q) = disj (simpfm (Not p)) (simpfm q)"
| "simpfm (Iff p q) =
- disj (conj (simpfm p) (simpfm q)) (conj (simpfm (NOT p)) (simpfm (NOT q)))"
- | "simpfm (NOT (And p q)) = disj (simpfm (NOT p)) (simpfm (NOT q))"
- | "simpfm (NOT (Or p q)) = conj (simpfm (NOT p)) (simpfm (NOT q))"
- | "simpfm (NOT (Imp p q)) = conj (simpfm p) (simpfm (NOT q))"
- | "simpfm (NOT (Iff p q)) =
- disj (conj (simpfm p) (simpfm (NOT q))) (conj (simpfm (NOT p)) (simpfm q))"
- | "simpfm (NOT (Eq t)) = simpneq t"
- | "simpfm (NOT (NEq t)) = simpeq t"
- | "simpfm (NOT (Le t)) = simplt (Neg t)"
- | "simpfm (NOT (Lt t)) = simple (Neg t)"
- | "simpfm (NOT (NOT p)) = simpfm p"
- | "simpfm (NOT T) = F"
- | "simpfm (NOT F) = T"
+ disj (conj (simpfm p) (simpfm q)) (conj (simpfm (Not p)) (simpfm (Not q)))"
+ | "simpfm (Not (And p q)) = disj (simpfm (Not p)) (simpfm (Not q))"
+ | "simpfm (Not (Or p q)) = conj (simpfm (Not p)) (simpfm (Not q))"
+ | "simpfm (Not (Imp p q)) = conj (simpfm p) (simpfm (Not q))"
+ | "simpfm (Not (Iff p q)) =
+ disj (conj (simpfm p) (simpfm (Not q))) (conj (simpfm (Not p)) (simpfm q))"
+ | "simpfm (Not (Eq t)) = simpneq t"
+ | "simpfm (Not (NEq t)) = simpeq t"
+ | "simpfm (Not (Le t)) = simplt (Neg t)"
+ | "simpfm (Not (Lt t)) = simple (Neg t)"
+ | "simpfm (Not (Not p)) = simpfm p"
+ | "simpfm (Not T) = F"
+ | "simpfm (Not F) = T"
| "simpfm p = p"
lemma simpfm[simp]: "Ifm vs bs (simpfm p) = Ifm vs bs p"
@@ -1600,25 +1600,25 @@
"prep (E T) = T"
| "prep (E F) = F"
| "prep (E (Or p q)) = disj (prep (E p)) (prep (E q))"
- | "prep (E (Imp p q)) = disj (prep (E (NOT p))) (prep (E q))"
- | "prep (E (Iff p q)) = disj (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))"
- | "prep (E (NOT (And p q))) = disj (prep (E (NOT p))) (prep (E(NOT q)))"
- | "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))"
- | "prep (E (NOT (Iff p q))) = disj (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))"
+ | "prep (E (Imp p q)) = disj (prep (E (Not p))) (prep (E q))"
+ | "prep (E (Iff p q)) = disj (prep (E (And p q))) (prep (E (And (Not p) (Not q))))"
+ | "prep (E (Not (And p q))) = disj (prep (E (Not p))) (prep (E(Not q)))"
+ | "prep (E (Not (Imp p q))) = prep (E (And p (Not q)))"
+ | "prep (E (Not (Iff p q))) = disj (prep (E (And p (Not q)))) (prep (E(And (Not p) q)))"
| "prep (E p) = E (prep p)"
| "prep (A (And p q)) = conj (prep (A p)) (prep (A q))"
- | "prep (A p) = prep (NOT (E (NOT p)))"
- | "prep (NOT (NOT p)) = prep p"
- | "prep (NOT (And p q)) = disj (prep (NOT p)) (prep (NOT q))"
- | "prep (NOT (A p)) = prep (E (NOT p))"
- | "prep (NOT (Or p q)) = conj (prep (NOT p)) (prep (NOT q))"
- | "prep (NOT (Imp p q)) = conj (prep p) (prep (NOT q))"
- | "prep (NOT (Iff p q)) = disj (prep (And p (NOT q))) (prep (And (NOT p) q))"
- | "prep (NOT p) = not (prep p)"
+ | "prep (A p) = prep (Not (E (Not p)))"
+ | "prep (Not (Not p)) = prep p"
+ | "prep (Not (And p q)) = disj (prep (Not p)) (prep (Not q))"
+ | "prep (Not (A p)) = prep (E (Not p))"
+ | "prep (Not (Or p q)) = conj (prep (Not p)) (prep (Not q))"
+ | "prep (Not (Imp p q)) = conj (prep p) (prep (Not q))"
+ | "prep (Not (Iff p q)) = disj (prep (And p (Not q))) (prep (And (Not p) q))"
+ | "prep (Not p) = not (prep p)"
| "prep (Or p q) = disj (prep p) (prep q)"
| "prep (And p q) = conj (prep p) (prep q)"
- | "prep (Imp p q) = prep (Or (NOT p) q)"
- | "prep (Iff p q) = disj (prep (And p q)) (prep (And (NOT p) (NOT q)))"
+ | "prep (Imp p q) = prep (Or (Not p) q)"
+ | "prep (Iff p q) = disj (prep (And p q)) (prep (And (Not p) (Not q)))"
| "prep p = p"
lemma prep: "Ifm vs bs (prep p) = Ifm vs bs p"
@@ -1629,8 +1629,8 @@
fun qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"
where
"qelim (E p) = (\<lambda>qe. DJ (CJNB qe) (qelim p qe))"
- | "qelim (A p) = (\<lambda>qe. not (qe ((qelim (NOT p) qe))))"
- | "qelim (NOT p) = (\<lambda>qe. not (qelim p qe))"
+ | "qelim (A p) = (\<lambda>qe. not (qe ((qelim (Not p) qe))))"
+ | "qelim (Not p) = (\<lambda>qe. not (qelim p qe))"
| "qelim (And p q) = (\<lambda>qe. conj (qelim p qe) (qelim q qe))"
| "qelim (Or p q) = (\<lambda>qe. disj (qelim p qe) (qelim q qe))"
| "qelim (Imp p q) = (\<lambda>qe. imp (qelim p qe) (qelim q qe))"
@@ -3541,7 +3541,7 @@
"msubstpos (And p q) c t = And (msubstpos p c t) (msubstpos q c t)"
| "msubstpos (Or p q) c t = Or (msubstpos p c t) (msubstpos q c t)"
| "msubstpos (Eq (CNP 0 a r)) c t = msubsteq2 c t a r"
- | "msubstpos (NEq (CNP 0 a r)) c t = NOT (msubsteq2 c t a r)"
+ | "msubstpos (NEq (CNP 0 a r)) c t = Not (msubsteq2 c t a r)"
| "msubstpos (Lt (CNP 0 a r)) c t = msubstltpos c t a r"
| "msubstpos (Le (CNP 0 a r)) c t = msubstlepos c t a r"
| "msubstpos p c t = p"
@@ -3562,7 +3562,7 @@
"msubstneg (And p q) c t = And (msubstneg p c t) (msubstneg q c t)"
| "msubstneg (Or p q) c t = Or (msubstneg p c t) (msubstneg q c t)"
| "msubstneg (Eq (CNP 0 a r)) c t = msubsteq2 c t a r"
- | "msubstneg (NEq (CNP 0 a r)) c t = NOT (msubsteq2 c t a r)"
+ | "msubstneg (NEq (CNP 0 a r)) c t = Not (msubsteq2 c t a r)"
| "msubstneg (Lt (CNP 0 a r)) c t = msubstltneg c t a r"
| "msubstneg (Le (CNP 0 a r)) c t = msubstleneg c t a r"
| "msubstneg p c t = p"
@@ -3973,8 +3973,8 @@
fun fm_of_term fs ps \<^term>\<open>True\<close> = @{code T}
| fm_of_term fs ps \<^term>\<open>False\<close> = @{code F}
- | fm_of_term fs ps (Const (\<^const_name>\<open>Not\<close>, _) $ p) =
- @{code NOT} (fm_of_term fs ps p)
+ | fm_of_term fs ps (Const (\<^const_name>\<open>HOL.Not\<close>, _) $ p) =
+ @{code Not} (fm_of_term fs ps p)
| fm_of_term fs ps (Const (\<^const_name>\<open>HOL.conj\<close>, _) $ p $ q) =
@{code And} (fm_of_term fs ps p, fm_of_term fs ps q)
| fm_of_term fs ps (Const (\<^const_name>\<open>HOL.disj\<close>, _) $ p $ q) =
@@ -4039,7 +4039,7 @@
fun term_of_fm T fs ps @{code T} = \<^term>\<open>True\<close>
| term_of_fm T fs ps @{code F} = \<^term>\<open>False\<close>
- | term_of_fm T fs ps (@{code NOT} p) = \<^term>\<open>Not\<close> $ term_of_fm T fs ps p
+ | term_of_fm T fs ps (@{code Not} p) = \<^term>\<open>HOL.Not\<close> $ term_of_fm T fs ps p
| term_of_fm T fs ps (@{code And} (p, q)) =
\<^term>\<open>HOL.conj\<close> $ term_of_fm T fs ps p $ term_of_fm T fs ps q
| term_of_fm T fs ps (@{code Or} (p, q)) =
--- a/src/HOL/Divides.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/Divides.thy Mon Aug 02 10:01:06 2021 +0000
@@ -1106,14 +1106,6 @@
\<open>divmod (Num.BitM m) (Num.Bit0 Num.One) = (numeral m - 1, (1 :: int))\<close>
by (cases m) simp_all
-lemma bit_numeral_Bit0_Suc_iff [simp]:
- \<open>bit (numeral (Num.Bit0 m) :: int) (Suc n) \<longleftrightarrow> bit (numeral m :: int) n\<close>
- by (simp add: bit_Suc)
-
-lemma bit_numeral_Bit1_Suc_iff [simp]:
- \<open>bit (numeral (Num.Bit1 m) :: int) (Suc n) \<longleftrightarrow> bit (numeral m :: int) n\<close>
- by (simp add: bit_Suc)
-
lemma div_positive_int:
"k div l > 0" if "k \<ge> l" and "l > 0" for k l :: int
using that div_positive [of l k] by blast
@@ -1225,113 +1217,6 @@
code_module Divides \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith
-subsection \<open>More on bit operations\<close>
-
-lemma take_bit_incr_eq:
- \<open>take_bit n (k + 1) = 1 + take_bit n k\<close> if \<open>take_bit n k \<noteq> 2 ^ n - 1\<close>
- for k :: int
-proof -
- from that have \<open>2 ^ n \<noteq> k mod 2 ^ n + 1\<close>
- by (simp add: take_bit_eq_mod)
- moreover have \<open>k mod 2 ^ n < 2 ^ n\<close>
- by simp
- ultimately have *: \<open>k mod 2 ^ n + 1 < 2 ^ n\<close>
- by linarith
- have \<open>(k + 1) mod 2 ^ n = (k mod 2 ^ n + 1) mod 2 ^ n\<close>
- by (simp add: mod_simps)
- also have \<open>\<dots> = k mod 2 ^ n + 1\<close>
- using * by (simp add: zmod_trivial_iff)
- finally have \<open>(k + 1) mod 2 ^ n = k mod 2 ^ n + 1\<close> .
- then show ?thesis
- by (simp add: take_bit_eq_mod)
-qed
-
-lemma take_bit_decr_eq:
- \<open>take_bit n (k - 1) = take_bit n k - 1\<close> if \<open>take_bit n k \<noteq> 0\<close>
- for k :: int
-proof -
- from that have \<open>k mod 2 ^ n \<noteq> 0\<close>
- by (simp add: take_bit_eq_mod)
- moreover have \<open>k mod 2 ^ n \<ge> 0\<close> \<open>k mod 2 ^ n < 2 ^ n\<close>
- by simp_all
- ultimately have *: \<open>k mod 2 ^ n > 0\<close>
- by linarith
- have \<open>(k - 1) mod 2 ^ n = (k mod 2 ^ n - 1) mod 2 ^ n\<close>
- by (simp add: mod_simps)
- also have \<open>\<dots> = k mod 2 ^ n - 1\<close>
- by (simp add: zmod_trivial_iff)
- (use \<open>k mod 2 ^ n < 2 ^ n\<close> * in linarith)
- finally have \<open>(k - 1) mod 2 ^ n = k mod 2 ^ n - 1\<close> .
- then show ?thesis
- by (simp add: take_bit_eq_mod)
-qed
-
-lemma take_bit_int_greater_eq:
- \<open>k + 2 ^ n \<le> take_bit n k\<close> if \<open>k < 0\<close> for k :: int
-proof -
- have \<open>k + 2 ^ n \<le> take_bit n (k + 2 ^ n)\<close>
- proof (cases \<open>k > - (2 ^ n)\<close>)
- case False
- then have \<open>k + 2 ^ n \<le> 0\<close>
- by simp
- also note take_bit_nonnegative
- finally show ?thesis .
- next
- case True
- with that have \<open>0 \<le> k + 2 ^ n\<close> and \<open>k + 2 ^ n < 2 ^ n\<close>
- 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:
- \<open>take_bit n k \<le> k - 2 ^ n\<close> if \<open>2 ^ n \<le> k\<close> and \<open>n > 0\<close> for k :: int
- using that zmod_le_nonneg_dividend [of \<open>k - 2 ^ n\<close> \<open>2 ^ n\<close>]
- by (simp add: take_bit_eq_mod)
-
-lemma take_bit_int_less_eq_self_iff:
- \<open>take_bit n k \<le> k \<longleftrightarrow> 0 \<le> k\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
- for k :: int
-proof
- assume ?P
- show ?Q
- proof (rule ccontr)
- assume \<open>\<not> 0 \<le> k\<close>
- then have \<open>k < 0\<close>
- by simp
- with \<open>?P\<close>
- have \<open>take_bit n k < 0\<close>
- 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:
- \<open>take_bit n k < k \<longleftrightarrow> 2 ^ n \<le> k\<close>
- 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 \<open>2 ^ n\<close> k])
-
-lemma take_bit_int_greater_self_iff:
- \<open>k < take_bit n k \<longleftrightarrow> k < 0\<close>
- for k :: int
- using take_bit_int_less_eq_self_iff [of n k] by auto
-
-lemma take_bit_int_greater_eq_self_iff:
- \<open>k \<le> take_bit n k \<longleftrightarrow> k < 2 ^ n\<close>
- 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 \<open>2 ^ n\<close>])
-
-
subsection \<open>Lemmas of doubtful value\<close>
lemma div_geq: "m div n = Suc ((m - n) div n)" if "0 < n" and " \<not> m < n" for m n :: nat
--- a/src/HOL/Fun.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/Fun.thy Mon Aug 02 10:01:06 2021 +0000
@@ -93,7 +93,7 @@
lemma (in group_add) minus_comp_minus [simp]: "uminus \<circ> uminus = id"
by (simp add: fun_eq_iff)
-lemma (in boolean_algebra) minus_comp_minus [simp]: "uminus \<circ> uminus = id"
+lemma (in Lattices.boolean_algebra) minus_comp_minus [simp]: "uminus \<circ> uminus = id"
by (simp add: fun_eq_iff)
code_printing
--- a/src/HOL/GCD.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/GCD.thy Mon Aug 02 10:01:06 2021 +0000
@@ -31,7 +31,7 @@
section \<open>Greatest common divisor and least common multiple\<close>
theory GCD
- imports Groups_List
+ imports Groups_List Code_Numeral
begin
subsection \<open>Abstract bounded quasi semilattices as common foundation\<close>
--- a/src/HOL/Groups_List.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/Groups_List.thy Mon Aug 02 10:01:06 2021 +0000
@@ -403,80 +403,6 @@
end
-context semiring_bit_shifts
-begin
-
-lemma horner_sum_bit_eq_take_bit:
- \<open>horner_sum of_bool 2 (map (bit a) [0..<n]) = take_bit n a\<close>
-proof (induction a arbitrary: n rule: bits_induct)
- case (stable a)
- moreover have \<open>bit a = (\<lambda>_. odd a)\<close>
- using stable by (simp add: stable_imp_bit_iff_odd fun_eq_iff)
- moreover have \<open>{q. q < n} = {0..<n}\<close>
- 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 \<open>map (bit (of_bool b + 2 * a)) [0..<Suc m] = b # map (bit (of_bool b + 2 * a)) [Suc 0..<Suc m]\<close>
- by (simp only: upt_conv_Cons) simp
- also have \<open>\<dots> = b # map (bit a) [0..<m]\<close>
- 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]:
- \<open>bit (horner_sum of_bool 2 bs) n \<longleftrightarrow> n < length bs \<and> bs ! n\<close>
-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:
- \<open>take_bit n (horner_sum of_bool 2 bs) = horner_sum of_bool 2 (take n bs)\<close>
- by (auto simp add: bit_eq_iff bit_take_bit_iff bit_horner_sum_bit_iff)
-
-end
-
-lemma horner_sum_of_bool_2_less:
- \<open>(horner_sum of_bool 2 bs :: int) < 2 ^ length bs\<close>
-proof -
- have \<open>(\<Sum>n = 0..<length bs. of_bool (bs ! n) * (2::int) ^ n) \<le> (\<Sum>n = 0..<length bs. 2 ^ n)\<close>
- by (rule sum_mono) simp
- also have \<open>\<dots> = 2 ^ length bs - 1\<close>
- by (induction bs) simp_all
- finally show ?thesis
- by (simp add: horner_sum_eq_sum)
-qed
-
subsection \<open>Further facts about \<^const>\<open>List.n_lists\<close>\<close>
--- a/src/HOL/HOLCF/IOA/Pred.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/HOLCF/IOA/Pred.thy Mon Aug 02 10:01:06 2021 +0000
@@ -18,8 +18,8 @@
definition valid :: "'a predicate \<Rightarrow> bool" ("\<TTurnstile> _" [9] 8)
where "(\<TTurnstile> P) \<longleftrightarrow> (\<forall>s. (s \<Turnstile> P))"
-definition NOT :: "'a predicate \<Rightarrow> 'a predicate" ("\<^bold>\<not> _" [40] 40)
- where "NOT P s \<longleftrightarrow> \<not> P s"
+definition Not :: "'a predicate \<Rightarrow> 'a predicate" ("\<^bold>\<not> _" [40] 40)
+ where NOT_def: "Not P s \<longleftrightarrow> \<not> P s"
definition AND :: "'a predicate \<Rightarrow> 'a predicate \<Rightarrow> 'a predicate" (infixr "\<^bold>\<and>" 35)
where "(P \<^bold>\<and> Q) s \<longleftrightarrow> P s \<and> Q s"
--- a/src/HOL/IMP/OO.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/IMP/OO.thy Mon Aug 02 10:01:06 2021 +0000
@@ -55,7 +55,7 @@
ve = (\<lambda>x. null)(''this'' := or, ''param'' := pr);
me \<turnstile> (me m,ve,sn\<^sub>3) \<Rightarrow> (r,ve',sn\<^sub>4) \<rbrakk>
\<Longrightarrow>
- me \<turnstile> (oe\<bullet>m<pe>,c\<^sub>1) \<Rightarrow> (r,ve\<^sub>3,sn\<^sub>4)" |
+ me \<turnstile> (oe\<bullet>m<pe>,c\<^sub>1) \<Rightarrow> (r,ve\<^sub>3,sn\<^sub>4)" for or |
Seq:
"\<lbrakk> me \<turnstile> (e\<^sub>1,c\<^sub>1) \<Rightarrow> (r,c\<^sub>2); me \<turnstile> (e\<^sub>2,c\<^sub>2) \<Rightarrow> c\<^sub>3 \<rbrakk> \<Longrightarrow>
me \<turnstile> (e\<^sub>1; e\<^sub>2,c\<^sub>1) \<Rightarrow> c\<^sub>3" |
--- a/src/HOL/Library/Bit_Operations.thy Sun Aug 01 23:18:13 2021 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,2105 +0,0 @@
-(* Author: Florian Haftmann, TUM
-*)
-
-section \<open>Bit operations in suitable algebraic structures\<close>
-
-theory Bit_Operations
- imports
- Main
- "HOL-Library.Boolean_Algebra"
-begin
-
-subsection \<open>Bit operations\<close>
-
-class semiring_bit_operations = semiring_bit_shifts +
- fixes "and" :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close> (infixr \<open>AND\<close> 64)
- and or :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close> (infixr \<open>OR\<close> 59)
- and xor :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close> (infixr \<open>XOR\<close> 59)
- and mask :: \<open>nat \<Rightarrow> 'a\<close>
- and set_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
- and unset_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
- and flip_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
- assumes bit_and_iff [bit_simps]: \<open>bit (a AND b) n \<longleftrightarrow> bit a n \<and> bit b n\<close>
- and bit_or_iff [bit_simps]: \<open>bit (a OR b) n \<longleftrightarrow> bit a n \<or> bit b n\<close>
- and bit_xor_iff [bit_simps]: \<open>bit (a XOR b) n \<longleftrightarrow> bit a n \<noteq> bit b n\<close>
- and mask_eq_exp_minus_1: \<open>mask n = 2 ^ n - 1\<close>
- and set_bit_eq_or: \<open>set_bit n a = a OR push_bit n 1\<close>
- and bit_unset_bit_iff [bit_simps]: \<open>bit (unset_bit m a) n \<longleftrightarrow> bit a n \<and> m \<noteq> n\<close>
- and flip_bit_eq_xor: \<open>flip_bit n a = a XOR push_bit n 1\<close>
-begin
-
-text \<open>
- We want the bitwise operations to bind slightly weaker
- than \<open>+\<close> and \<open>-\<close>.
- For the sake of code generation
- the operations \<^const>\<open>and\<close>, \<^const>\<open>or\<close> and \<^const>\<open>xor\<close>
- are specified as definitional class operations.
-\<close>
-
-sublocale "and": semilattice \<open>(AND)\<close>
- by standard (auto simp add: bit_eq_iff bit_and_iff)
-
-sublocale or: semilattice_neutr \<open>(OR)\<close> 0
- by standard (auto simp add: bit_eq_iff bit_or_iff)
-
-sublocale xor: comm_monoid \<open>(XOR)\<close> 0
- by standard (auto simp add: bit_eq_iff bit_xor_iff)
-
-lemma even_and_iff:
- \<open>even (a AND b) \<longleftrightarrow> even a \<or> even b\<close>
- using bit_and_iff [of a b 0] by auto
-
-lemma even_or_iff:
- \<open>even (a OR b) \<longleftrightarrow> even a \<and> even b\<close>
- using bit_or_iff [of a b 0] by auto
-
-lemma even_xor_iff:
- \<open>even (a XOR b) \<longleftrightarrow> (even a \<longleftrightarrow> even b)\<close>
- using bit_xor_iff [of a b 0] by auto
-
-lemma zero_and_eq [simp]:
- \<open>0 AND a = 0\<close>
- by (simp add: bit_eq_iff bit_and_iff)
-
-lemma and_zero_eq [simp]:
- \<open>a AND 0 = 0\<close>
- by (simp add: bit_eq_iff bit_and_iff)
-
-lemma one_and_eq:
- \<open>1 AND a = a mod 2\<close>
- by (simp add: bit_eq_iff bit_and_iff) (auto simp add: bit_1_iff)
-
-lemma and_one_eq:
- \<open>a AND 1 = a mod 2\<close>
- using one_and_eq [of a] by (simp add: ac_simps)
-
-lemma one_or_eq:
- \<open>1 OR a = a + of_bool (even a)\<close>
- 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:
- \<open>a OR 1 = a + of_bool (even a)\<close>
- using one_or_eq [of a] by (simp add: ac_simps)
-
-lemma one_xor_eq:
- \<open>1 XOR a = a + of_bool (even a) - of_bool (odd a)\<close>
- 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:
- \<open>a XOR 1 = a + of_bool (even a) - of_bool (odd a)\<close>
- using one_xor_eq [of a] by (simp add: ac_simps)
-
-lemma take_bit_and [simp]:
- \<open>take_bit n (a AND b) = take_bit n a AND take_bit n b\<close>
- by (auto simp add: bit_eq_iff bit_take_bit_iff bit_and_iff)
-
-lemma take_bit_or [simp]:
- \<open>take_bit n (a OR b) = take_bit n a OR take_bit n b\<close>
- by (auto simp add: bit_eq_iff bit_take_bit_iff bit_or_iff)
-
-lemma take_bit_xor [simp]:
- \<open>take_bit n (a XOR b) = take_bit n a XOR take_bit n b\<close>
- by (auto simp add: bit_eq_iff bit_take_bit_iff bit_xor_iff)
-
-lemma push_bit_and [simp]:
- \<open>push_bit n (a AND b) = push_bit n a AND push_bit n b\<close>
- by (rule bit_eqI) (auto simp add: bit_push_bit_iff bit_and_iff)
-
-lemma push_bit_or [simp]:
- \<open>push_bit n (a OR b) = push_bit n a OR push_bit n b\<close>
- by (rule bit_eqI) (auto simp add: bit_push_bit_iff bit_or_iff)
-
-lemma push_bit_xor [simp]:
- \<open>push_bit n (a XOR b) = push_bit n a XOR push_bit n b\<close>
- by (rule bit_eqI) (auto simp add: bit_push_bit_iff bit_xor_iff)
-
-lemma drop_bit_and [simp]:
- \<open>drop_bit n (a AND b) = drop_bit n a AND drop_bit n b\<close>
- by (rule bit_eqI) (auto simp add: bit_drop_bit_eq bit_and_iff)
-
-lemma drop_bit_or [simp]:
- \<open>drop_bit n (a OR b) = drop_bit n a OR drop_bit n b\<close>
- by (rule bit_eqI) (auto simp add: bit_drop_bit_eq bit_or_iff)
-
-lemma drop_bit_xor [simp]:
- \<open>drop_bit n (a XOR b) = drop_bit n a XOR drop_bit n b\<close>
- by (rule bit_eqI) (auto simp add: bit_drop_bit_eq bit_xor_iff)
-
-lemma bit_mask_iff [bit_simps]:
- \<open>bit (mask m) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n < m\<close>
- by (simp add: mask_eq_exp_minus_1 bit_mask_iff)
-
-lemma even_mask_iff:
- \<open>even (mask n) \<longleftrightarrow> n = 0\<close>
- using bit_mask_iff [of n 0] by auto
-
-lemma mask_0 [simp]:
- \<open>mask 0 = 0\<close>
- by (simp add: mask_eq_exp_minus_1)
-
-lemma mask_Suc_0 [simp]:
- \<open>mask (Suc 0) = 1\<close>
- by (simp add: mask_eq_exp_minus_1 add_implies_diff sym)
-
-lemma mask_Suc_exp:
- \<open>mask (Suc n) = 2 ^ n OR mask n\<close>
- 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:
- \<open>mask (Suc n) = 1 OR 2 * mask n\<close>
-proof (rule bit_eqI)
- fix q
- assume \<open>2 ^ q \<noteq> 0\<close>
- show \<open>bit (mask (Suc n)) q \<longleftrightarrow> bit (1 OR 2 * mask n) q\<close>
- 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:
- \<open>mask (numeral n) = 1 + 2 * mask (pred_numeral n)\<close>
- by (simp add: numeral_eq_Suc mask_Suc_double one_or_eq ac_simps)
-
-lemma take_bit_mask [simp]:
- \<open>take_bit m (mask n) = mask (min m n)\<close>
- by (rule bit_eqI) (simp add: bit_simps)
-
-lemma take_bit_eq_mask:
- \<open>take_bit n a = a AND mask n\<close>
- by (rule bit_eqI)
- (auto simp add: bit_take_bit_iff bit_and_iff bit_mask_iff)
-
-lemma or_eq_0_iff:
- \<open>a OR b = 0 \<longleftrightarrow> a = 0 \<and> b = 0\<close>
- by (auto simp add: bit_eq_iff bit_or_iff)
-
-lemma disjunctive_add:
- \<open>a + b = a OR b\<close> if \<open>\<And>n. \<not> bit a n \<or> \<not> bit b n\<close>
- by (rule bit_eqI) (use that in \<open>simp add: bit_disjunctive_add_iff bit_or_iff\<close>)
-
-lemma bit_iff_and_drop_bit_eq_1:
- \<open>bit a n \<longleftrightarrow> drop_bit n a AND 1 = 1\<close>
- 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:
- \<open>bit a n \<longleftrightarrow> a AND push_bit n 1 \<noteq> 0\<close>
- apply (cases \<open>2 ^ n = 0\<close>)
- 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]:
- \<open>bit (set_bit m a) n \<longleftrightarrow> bit a n \<or> (m = n \<and> 2 ^ n \<noteq> 0)\<close>
- by (auto simp add: set_bit_def push_bit_of_1 bit_or_iff bit_exp_iff)
-
-lemma even_set_bit_iff:
- \<open>even (set_bit m a) \<longleftrightarrow> even a \<and> m \<noteq> 0\<close>
- using bit_set_bit_iff [of m a 0] by auto
-
-lemma even_unset_bit_iff:
- \<open>even (unset_bit m a) \<longleftrightarrow> even a \<or> m = 0\<close>
- using bit_unset_bit_iff [of m a 0] by auto
-
-lemma and_exp_eq_0_iff_not_bit:
- \<open>a AND 2 ^ n = 0 \<longleftrightarrow> \<not> bit a n\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
-proof
- assume ?Q
- then show ?P
- by (auto intro: bit_eqI simp add: bit_simps)
-next
- assume ?P
- show ?Q
- proof (rule notI)
- assume \<open>bit a n\<close>
- then have \<open>a AND 2 ^ n = 2 ^ n\<close>
- by (auto intro: bit_eqI simp add: bit_simps)
- with \<open>?P\<close> show False
- using \<open>bit a n\<close> 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]:
- \<open>bit (flip_bit m a) n \<longleftrightarrow> (m = n \<longleftrightarrow> \<not> bit a n) \<and> 2 ^ n \<noteq> 0\<close>
- 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:
- \<open>even (flip_bit m a) \<longleftrightarrow> \<not> (even a \<longleftrightarrow> m = 0)\<close>
- using bit_flip_bit_iff [of m a 0] by auto
-
-lemma set_bit_0 [simp]:
- \<open>set_bit 0 a = 1 + 2 * (a div 2)\<close>
-proof (rule bit_eqI)
- fix m
- assume *: \<open>2 ^ m \<noteq> 0\<close>
- then show \<open>bit (set_bit 0 a) m = bit (1 + 2 * (a div 2)) m\<close>
- 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:
- \<open>set_bit (Suc n) a = a mod 2 + 2 * set_bit n (a div 2)\<close>
-proof (rule bit_eqI)
- fix m
- assume *: \<open>2 ^ m \<noteq> 0\<close>
- show \<open>bit (set_bit (Suc n) a) m \<longleftrightarrow> bit (a mod 2 + 2 * set_bit n (a div 2)) m\<close>
- proof (cases m)
- case 0
- then show ?thesis
- by (simp add: even_set_bit_iff)
- next
- case (Suc m)
- with * have \<open>2 ^ m \<noteq> 0\<close>
- 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 \<open>2 ^ m \<noteq> 0\<close> bit_Suc)
- qed
-qed
-
-lemma unset_bit_0 [simp]:
- \<open>unset_bit 0 a = 2 * (a div 2)\<close>
-proof (rule bit_eqI)
- fix m
- assume *: \<open>2 ^ m \<noteq> 0\<close>
- then show \<open>bit (unset_bit 0 a) m = bit (2 * (a div 2)) m\<close>
- by (simp add: bit_unset_bit_iff bit_double_iff)
- (cases m, simp_all add: bit_Suc)
-qed
-
-lemma unset_bit_Suc:
- \<open>unset_bit (Suc n) a = a mod 2 + 2 * unset_bit n (a div 2)\<close>
-proof (rule bit_eqI)
- fix m
- assume *: \<open>2 ^ m \<noteq> 0\<close>
- then show \<open>bit (unset_bit (Suc n) a) m \<longleftrightarrow> bit (a mod 2 + 2 * unset_bit n (a div 2)) m\<close>
- 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]:
- \<open>flip_bit 0 a = of_bool (even a) + 2 * (a div 2)\<close>
-proof (rule bit_eqI)
- fix m
- assume *: \<open>2 ^ m \<noteq> 0\<close>
- then show \<open>bit (flip_bit 0 a) m = bit (of_bool (even a) + 2 * (a div 2)) m\<close>
- 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:
- \<open>flip_bit (Suc n) a = a mod 2 + 2 * flip_bit n (a div 2)\<close>
-proof (rule bit_eqI)
- fix m
- assume *: \<open>2 ^ m \<noteq> 0\<close>
- show \<open>bit (flip_bit (Suc n) a) m \<longleftrightarrow> bit (a mod 2 + 2 * flip_bit n (a div 2)) m\<close>
- proof (cases m)
- case 0
- then show ?thesis
- by (simp add: even_flip_bit_iff)
- next
- case (Suc m)
- with * have \<open>2 ^ m \<noteq> 0\<close>
- 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 \<open>2 ^ m \<noteq> 0\<close> bit_Suc)
- qed
-qed
-
-lemma flip_bit_eq_if:
- \<open>flip_bit n a = (if bit a n then unset_bit else set_bit) n a\<close>
- 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:
- \<open>take_bit n (set_bit m a) = (if n \<le> m then take_bit n a else set_bit m (take_bit n a))\<close>
- by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_set_bit_iff)
-
-lemma take_bit_unset_bit_eq:
- \<open>take_bit n (unset_bit m a) = (if n \<le> m then take_bit n a else unset_bit m (take_bit n a))\<close>
- by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_unset_bit_iff)
-
-lemma take_bit_flip_bit_eq:
- \<open>take_bit n (flip_bit m a) = (if n \<le> m then take_bit n a else flip_bit m (take_bit n a))\<close>
- 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 :: \<open>'a \<Rightarrow> 'a\<close> (\<open>NOT\<close>)
- assumes bit_not_iff [bit_simps]: \<open>\<And>n. bit (NOT a) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> \<not> bit a n\<close>
- assumes minus_eq_not_minus_1: \<open>- a = NOT (a - 1)\<close>
-begin
-
-text \<open>
- For the sake of code generation \<^const>\<open>not\<close> is specified as
- definitional class operation. Note that \<^const>\<open>not\<close> has no
- sensible definition for unlimited but only positive bit strings
- (type \<^typ>\<open>nat\<close>).
-\<close>
-
-lemma bits_minus_1_mod_2_eq [simp]:
- \<open>(- 1) mod 2 = 1\<close>
- by (simp add: mod_2_eq_odd)
-
-lemma not_eq_complement:
- \<open>NOT a = - a - 1\<close>
- using minus_eq_not_minus_1 [of \<open>a + 1\<close>] by simp
-
-lemma minus_eq_not_plus_1:
- \<open>- a = NOT a + 1\<close>
- using not_eq_complement [of a] by simp
-
-lemma bit_minus_iff [bit_simps]:
- \<open>bit (- a) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> \<not> bit (a - 1) n\<close>
- by (simp add: minus_eq_not_minus_1 bit_not_iff)
-
-lemma even_not_iff [simp]:
- \<open>even (NOT a) \<longleftrightarrow> odd a\<close>
- using bit_not_iff [of a 0] by auto
-
-lemma bit_not_exp_iff [bit_simps]:
- \<open>bit (NOT (2 ^ m)) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n \<noteq> m\<close>
- by (auto simp add: bit_not_iff bit_exp_iff)
-
-lemma bit_minus_1_iff [simp]:
- \<open>bit (- 1) n \<longleftrightarrow> 2 ^ n \<noteq> 0\<close>
- by (simp add: bit_minus_iff)
-
-lemma bit_minus_exp_iff [bit_simps]:
- \<open>bit (- (2 ^ m)) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n \<ge> m\<close>
- by (auto simp add: bit_simps simp flip: mask_eq_exp_minus_1)
-
-lemma bit_minus_2_iff [simp]:
- \<open>bit (- 2) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n > 0\<close>
- by (simp add: bit_minus_iff bit_1_iff)
-
-lemma not_one [simp]:
- \<open>NOT 1 = - 2\<close>
- by (simp add: bit_eq_iff bit_not_iff) (simp add: bit_1_iff)
-
-sublocale "and": semilattice_neutr \<open>(AND)\<close> \<open>- 1\<close>
- by standard (rule bit_eqI, simp add: bit_and_iff)
-
-sublocale bit: boolean_algebra \<open>(AND)\<close> \<open>(OR)\<close> NOT 0 \<open>- 1\<close>
- rewrites \<open>bit.xor = (XOR)\<close>
-proof -
- interpret bit: boolean_algebra \<open>(AND)\<close> \<open>(OR)\<close> NOT 0 \<open>- 1\<close>
- by standard (auto simp add: bit_and_iff bit_or_iff bit_not_iff intro: bit_eqI)
- show \<open>boolean_algebra (AND) (OR) NOT 0 (- 1)\<close>
- by standard
- show \<open>boolean_algebra.xor (AND) (OR) NOT = (XOR)\<close>
- 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:
- \<open>a AND b = NOT (NOT a OR NOT b)\<close>
- by simp
-
-lemma or_eq_not_not_and:
- \<open>a OR b = NOT (NOT a AND NOT b)\<close>
- by simp
-
-lemma not_add_distrib:
- \<open>NOT (a + b) = NOT a - b\<close>
- by (simp add: not_eq_complement algebra_simps)
-
-lemma not_diff_distrib:
- \<open>NOT (a - b) = NOT a + b\<close>
- using not_add_distrib [of a \<open>- b\<close>] by simp
-
-lemma (in ring_bit_operations) and_eq_minus_1_iff:
- \<open>a AND b = - 1 \<longleftrightarrow> a = - 1 \<and> b = - 1\<close>
-proof
- assume \<open>a = - 1 \<and> b = - 1\<close>
- then show \<open>a AND b = - 1\<close>
- by simp
-next
- assume \<open>a AND b = - 1\<close>
- have *: \<open>bit a n\<close> \<open>bit b n\<close> if \<open>2 ^ n \<noteq> 0\<close> for n
- proof -
- from \<open>a AND b = - 1\<close>
- have \<open>bit (a AND b) n = bit (- 1) n\<close>
- by (simp add: bit_eq_iff)
- then show \<open>bit a n\<close> \<open>bit b n\<close>
- using that by (simp_all add: bit_and_iff)
- qed
- have \<open>a = - 1\<close>
- by (rule bit_eqI) (simp add: *)
- moreover have \<open>b = - 1\<close>
- by (rule bit_eqI) (simp add: *)
- ultimately show \<open>a = - 1 \<and> b = - 1\<close>
- by simp
-qed
-
-lemma disjunctive_diff:
- \<open>a - b = a AND NOT b\<close> if \<open>\<And>n. bit b n \<Longrightarrow> bit a n\<close>
-proof -
- have \<open>NOT a + b = NOT a OR b\<close>
- by (rule disjunctive_add) (auto simp add: bit_not_iff dest: that)
- then have \<open>NOT (NOT a + b) = NOT (NOT a OR b)\<close>
- by simp
- then show ?thesis
- by (simp add: not_add_distrib)
-qed
-
-lemma push_bit_minus:
- \<open>push_bit n (- a) = - push_bit n a\<close>
- by (simp add: push_bit_eq_mult)
-
-lemma take_bit_not_take_bit:
- \<open>take_bit n (NOT (take_bit n a)) = take_bit n (NOT a)\<close>
- by (auto simp add: bit_eq_iff bit_take_bit_iff bit_not_iff)
-
-lemma take_bit_not_iff:
- \<open>take_bit n (NOT a) = take_bit n (NOT b) \<longleftrightarrow> take_bit n a = take_bit n b\<close>
- 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:
- \<open>take_bit n (NOT a) = mask n - take_bit n a\<close>
-proof -
- have \<open>take_bit n (NOT a) = take_bit n (NOT (take_bit n a))\<close>
- by (simp add: take_bit_not_take_bit)
- also have \<open>\<dots> = mask n AND NOT (take_bit n a)\<close>
- by (simp add: take_bit_eq_mask ac_simps)
- also have \<open>\<dots> = mask n - take_bit n a\<close>
- 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:
- \<open>mask n = take_bit n (- 1)\<close>
- by (simp add: bit_eq_iff bit_mask_iff bit_take_bit_iff conj_commute)
-
-lemma take_bit_minus_one_eq_mask:
- \<open>take_bit n (- 1) = mask n\<close>
- by (simp add: mask_eq_take_bit_minus_one)
-
-lemma minus_exp_eq_not_mask:
- \<open>- (2 ^ n) = NOT (mask n)\<close>
- 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:
- \<open>push_bit n (- 1) = NOT (mask n)\<close>
- by (simp add: push_bit_eq_mult minus_exp_eq_not_mask)
-
-lemma take_bit_not_mask_eq_0:
- \<open>take_bit m (NOT (mask n)) = 0\<close> if \<open>n \<ge> m\<close>
- by (rule bit_eqI) (use that in \<open>simp add: bit_take_bit_iff bit_not_iff bit_mask_iff\<close>)
-
-lemma unset_bit_eq_and_not:
- \<open>unset_bit n a = a AND NOT (push_bit n 1)\<close>
- by (rule bit_eqI) (auto simp add: bit_simps)
-
-lemmas unset_bit_def = unset_bit_eq_and_not
-
-end
-
-
-subsection \<open>Instance \<^typ>\<open>int\<close>\<close>
-
-lemma int_bit_bound:
- fixes k :: int
- obtains n where \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k n\<close>
- and \<open>n > 0 \<Longrightarrow> bit k (n - 1) \<noteq> bit k n\<close>
-proof -
- obtain q where *: \<open>\<And>m. q \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k q\<close>
- proof (cases \<open>k \<ge> 0\<close>)
- case True
- moreover from power_gt_expt [of 2 \<open>nat k\<close>]
- have \<open>nat k < 2 ^ nat k\<close>
- by simp
- then have \<open>int (nat k) < int (2 ^ nat k)\<close>
- by (simp only: of_nat_less_iff)
- ultimately have *: \<open>k div 2 ^ nat k = 0\<close>
- by simp
- show thesis
- proof (rule that [of \<open>nat k\<close>])
- fix m
- assume \<open>nat k \<le> m\<close>
- then show \<open>bit k m \<longleftrightarrow> bit k (nat k)\<close>
- 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 \<open>nat (- k)\<close>]
- have \<open>nat (- k) < 2 ^ nat (- k)\<close>
- by simp
- then have \<open>int (nat (- k)) < int (2 ^ nat (- k))\<close>
- by (simp only: of_nat_less_iff)
- ultimately have \<open>- k div - (2 ^ nat (- k)) = - 1\<close>
- by (subst div_pos_neg_trivial) simp_all
- then have *: \<open>k div 2 ^ nat (- k) = - 1\<close>
- by simp
- show thesis
- proof (rule that [of \<open>nat (- k)\<close>])
- fix m
- assume \<open>nat (- k) \<le> m\<close>
- then show \<open>bit k m \<longleftrightarrow> bit k (nat (- k))\<close>
- 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 \<open>\<forall>m. bit k m \<longleftrightarrow> bit k q\<close>)
- case True
- then have \<open>bit k 0 \<longleftrightarrow> bit k q\<close>
- by blast
- with True that [of 0] show thesis
- by simp
- next
- case False
- then obtain r where **: \<open>bit k r \<noteq> bit k q\<close>
- by blast
- have \<open>r < q\<close>
- by (rule ccontr) (use * [of r] ** in simp)
- define N where \<open>N = {n. n < q \<and> bit k n \<noteq> bit k q}\<close>
- moreover have \<open>finite N\<close> \<open>r \<in> N\<close>
- using ** N_def \<open>r < q\<close> by auto
- moreover define n where \<open>n = Suc (Max N)\<close>
- ultimately have \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k n\<close>
- apply auto
- apply (metis (full_types, lifting) "*" Max_ge_iff Suc_n_not_le_n \<open>finite N\<close> all_not_in_conv mem_Collect_eq not_le)
- apply (metis "*" Max_ge Suc_n_not_le_n \<open>finite N\<close> linorder_not_less mem_Collect_eq)
- apply (metis "*" Max_ge Suc_n_not_le_n \<open>finite N\<close> linorder_not_less mem_Collect_eq)
- apply (metis (full_types, lifting) "*" Max_ge_iff Suc_n_not_le_n \<open>finite N\<close> all_not_in_conv mem_Collect_eq not_le)
- done
- have \<open>bit k (Max N) \<noteq> bit k n\<close>
- by (metis (mono_tags, lifting) "*" Max_in N_def \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m = bit k n\<close> \<open>finite N\<close> \<open>r \<in> N\<close> empty_iff le_cases mem_Collect_eq)
- show thesis apply (rule that [of n])
- using \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m = bit k n\<close> apply blast
- using \<open>bit k (Max N) \<noteq> bit k n\<close> n_def by auto
- qed
-qed
-
-instantiation int :: ring_bit_operations
-begin
-
-definition not_int :: \<open>int \<Rightarrow> int\<close>
- where \<open>not_int k = - k - 1\<close>
-
-lemma not_int_rec:
- \<open>NOT k = of_bool (even k) + 2 * NOT (k div 2)\<close> for k :: int
- by (auto simp add: not_int_def elim: oddE)
-
-lemma even_not_iff_int:
- \<open>even (NOT k) \<longleftrightarrow> odd k\<close> for k :: int
- by (simp add: not_int_def)
-
-lemma not_int_div_2:
- \<open>NOT k div 2 = NOT (k div 2)\<close> for k :: int
- by (simp add: not_int_def)
-
-lemma bit_not_int_iff [bit_simps]:
- \<open>bit (NOT k) n \<longleftrightarrow> \<not> bit k n\<close>
- for k :: int
- by (simp add: bit_not_int_iff' not_int_def)
-
-function and_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>
- where \<open>(k::int) AND l = (if k \<in> {0, - 1} \<and> l \<in> {0, - 1}
- then - of_bool (odd k \<and> odd l)
- else of_bool (odd k \<and> odd l) + 2 * ((k div 2) AND (l div 2)))\<close>
- by auto
-
-termination
- by (relation \<open>measure (\<lambda>(k, l). nat (\<bar>k\<bar> + \<bar>l\<bar>))\<close>) auto
-
-declare and_int.simps [simp del]
-
-lemma and_int_rec:
- \<open>k AND l = of_bool (odd k \<and> odd l) + 2 * ((k div 2) AND (l div 2))\<close>
- for k l :: int
-proof (cases \<open>k \<in> {0, - 1} \<and> l \<in> {0, - 1}\<close>)
- 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:
- \<open>bit (k AND l) n \<longleftrightarrow> bit k n \<and> bit l n\<close> 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:
- \<open>even (k AND l) \<longleftrightarrow> even k \<or> even l\<close> for k l :: int
- using bit_and_int_iff [of k l 0] by auto
-
-definition or_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>
- where \<open>k OR l = NOT (NOT k AND NOT l)\<close> for k l :: int
-
-lemma or_int_rec:
- \<open>k OR l = of_bool (odd k \<or> odd l) + 2 * ((k div 2) OR (l div 2))\<close>
- for k l :: int
- using and_int_rec [of \<open>NOT k\<close> \<open>NOT l\<close>]
- 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:
- \<open>bit (k OR l) n \<longleftrightarrow> bit k n \<or> bit l n\<close> for k l :: int
- by (simp add: or_int_def bit_not_int_iff bit_and_int_iff)
-
-definition xor_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close>
- where \<open>k XOR l = k AND NOT l OR NOT k AND l\<close> for k l :: int
-
-lemma xor_int_rec:
- \<open>k XOR l = of_bool (odd k \<noteq> odd l) + 2 * ((k div 2) XOR (l div 2))\<close>
- for k l :: int
- by (simp add: xor_int_def or_int_rec [of \<open>k AND NOT l\<close> \<open>NOT k AND l\<close>] even_and_iff_int even_not_iff_int)
- (simp add: and_int_rec [of \<open>NOT k\<close> \<open>l\<close>] and_int_rec [of \<open>k\<close> \<open>NOT l\<close>] not_int_div_2)
-
-lemma bit_xor_int_iff:
- \<open>bit (k XOR l) n \<longleftrightarrow> bit k n \<noteq> bit l n\<close> 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 :: \<open>nat \<Rightarrow> int\<close>
- where \<open>mask n = (2 :: int) ^ n - 1\<close>
-
-definition set_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
- where \<open>set_bit n k = k OR push_bit n 1\<close> for k :: int
-
-definition unset_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
- where \<open>unset_bit n k = k AND NOT (push_bit n 1)\<close> for k :: int
-
-definition flip_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
- where \<open>flip_bit n k = k XOR push_bit n 1\<close> for k :: int
-
-instance proof
- fix k l :: int and m n :: nat
- show \<open>- k = NOT (k - 1)\<close>
- by (simp add: not_int_def)
- show \<open>bit (k AND l) n \<longleftrightarrow> bit k n \<and> bit l n\<close>
- by (fact bit_and_int_iff)
- show \<open>bit (k OR l) n \<longleftrightarrow> bit k n \<or> bit l n\<close>
- by (fact bit_or_int_iff)
- show \<open>bit (k XOR l) n \<longleftrightarrow> bit k n \<noteq> bit l n\<close>
- by (fact bit_xor_int_iff)
- show \<open>bit (unset_bit m k) n \<longleftrightarrow> bit k n \<and> m \<noteq> n\<close>
- proof -
- have \<open>unset_bit m k = k AND NOT (push_bit m 1)\<close>
- by (simp add: unset_bit_int_def)
- also have \<open>NOT (push_bit m 1 :: int) = - (push_bit m 1 + 1)\<close>
- 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:
- \<open>mask n div 2 = (mask (n - 1) :: int)\<close>
- by (cases n) (simp_all add: mask_eq_exp_minus_1 algebra_simps)
-
-lemma mask_nonnegative_int [simp]:
- \<open>mask n \<ge> (0::int)\<close>
- by (simp add: mask_eq_exp_minus_1)
-
-lemma not_mask_negative_int [simp]:
- \<open>\<not> mask n < (0::int)\<close>
- by (simp add: not_less)
-
-lemma not_nonnegative_int_iff [simp]:
- \<open>NOT k \<ge> 0 \<longleftrightarrow> k < 0\<close> for k :: int
- by (simp add: not_int_def)
-
-lemma not_negative_int_iff [simp]:
- \<open>NOT k < 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
- by (subst Not_eq_iff [symmetric]) (simp add: not_less not_le)
-
-lemma and_nonnegative_int_iff [simp]:
- \<open>k AND l \<ge> 0 \<longleftrightarrow> k \<ge> 0 \<or> l \<ge> 0\<close> 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 \<open>k * 2\<close> l] by (simp add: pos_imp_zdiv_nonneg_iff)
-next
- case (odd k)
- from odd have \<open>0 \<le> k AND l div 2 \<longleftrightarrow> 0 \<le> k \<or> 0 \<le> l div 2\<close>
- by simp
- then have \<open>0 \<le> (1 + k * 2) div 2 AND l div 2 \<longleftrightarrow> 0 \<le> (1 + k * 2) div 2\<or> 0 \<le> l div 2\<close>
- by simp
- with and_int_rec [of \<open>1 + k * 2\<close> l]
- show ?case
- by auto
-qed
-
-lemma and_negative_int_iff [simp]:
- \<open>k AND l < 0 \<longleftrightarrow> k < 0 \<and> l < 0\<close> for k l :: int
- by (subst Not_eq_iff [symmetric]) (simp add: not_less)
-
-lemma and_less_eq:
- \<open>k AND l \<le> k\<close> if \<open>l < 0\<close> 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 \<open>l div 2\<close>] even.hyps even.prems
- show ?case
- by (simp add: and_int_rec [of _ l])
-next
- case (odd k)
- from odd.IH [of \<open>l div 2\<close>] odd.hyps odd.prems
- show ?case
- by (simp add: and_int_rec [of _ l])
-qed
-
-lemma or_nonnegative_int_iff [simp]:
- \<open>k OR l \<ge> 0 \<longleftrightarrow> k \<ge> 0 \<and> l \<ge> 0\<close> for k l :: int
- by (simp only: or_eq_not_not_and not_nonnegative_int_iff) simp
-
-lemma or_negative_int_iff [simp]:
- \<open>k OR l < 0 \<longleftrightarrow> k < 0 \<or> l < 0\<close> for k l :: int
- by (subst Not_eq_iff [symmetric]) (simp add: not_less)
-
-lemma or_greater_eq:
- \<open>k OR l \<ge> k\<close> if \<open>l \<ge> 0\<close> 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 \<open>l div 2\<close>] even.hyps even.prems
- show ?case
- by (simp add: or_int_rec [of _ l])
-next
- case (odd k)
- from odd.IH [of \<open>l div 2\<close>] odd.hyps odd.prems
- show ?case
- by (simp add: or_int_rec [of _ l])
-qed
-
-lemma xor_nonnegative_int_iff [simp]:
- \<open>k XOR l \<ge> 0 \<longleftrightarrow> (k \<ge> 0 \<longleftrightarrow> l \<ge> 0)\<close> for k l :: int
- by (simp only: bit.xor_def or_nonnegative_int_iff) auto
-
-lemma xor_negative_int_iff [simp]:
- \<open>k XOR l < 0 \<longleftrightarrow> (k < 0) \<noteq> (l < 0)\<close> for k l :: int
- by (subst Not_eq_iff [symmetric]) (auto simp add: not_less)
-
-lemma OR_upper: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
- fixes x y :: int
- assumes \<open>0 \<le> x\<close> \<open>x < 2 ^ n\<close> \<open>y < 2 ^ n\<close>
- shows \<open>x OR y < 2 ^ n\<close>
-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 \<open>n - 1\<close> \<open>y div 2\<close>] even.prems even.hyps
- show ?case
- by (cases n) (auto simp add: or_int_rec [of \<open>_ * 2\<close>] elim: oddE)
-next
- case (odd x)
- from odd.IH [of \<open>n - 1\<close> \<open>y div 2\<close>] odd.prems odd.hyps
- show ?case
- by (cases n) (auto simp add: or_int_rec [of \<open>1 + _ * 2\<close>], linarith)
-qed
-
-lemma XOR_upper: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
- fixes x y :: int
- assumes \<open>0 \<le> x\<close> \<open>x < 2 ^ n\<close> \<open>y < 2 ^ n\<close>
- shows \<open>x XOR y < 2 ^ n\<close>
-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 \<open>n - 1\<close> \<open>y div 2\<close>] even.prems even.hyps
- show ?case
- by (cases n) (auto simp add: xor_int_rec [of \<open>_ * 2\<close>] elim: oddE)
-next
- case (odd x)
- from odd.IH [of \<open>n - 1\<close> \<open>y div 2\<close>] odd.prems odd.hyps
- show ?case
- by (cases n) (auto simp add: xor_int_rec [of \<open>1 + _ * 2\<close>])
-qed
-
-lemma AND_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
- fixes x y :: int
- assumes \<open>0 \<le> x\<close>
- shows \<open>0 \<le> x AND y\<close>
- using assms by simp
-
-lemma OR_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
- fixes x y :: int
- assumes \<open>0 \<le> x\<close> \<open>0 \<le> y\<close>
- shows \<open>0 \<le> x OR y\<close>
- using assms by simp
-
-lemma XOR_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
- fixes x y :: int
- assumes \<open>0 \<le> x\<close> \<open>0 \<le> y\<close>
- shows \<open>0 \<le> x XOR y\<close>
- using assms by simp
-
-lemma AND_upper1 [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
- fixes x y :: int
- assumes \<open>0 \<le> x\<close>
- shows \<open>x AND y \<le> x\<close>
-using assms proof (induction x arbitrary: y rule: int_bit_induct)
- case (odd k)
- then have \<open>k AND y div 2 \<le> k\<close>
- by simp
- then show ?case
- by (simp add: and_int_rec [of \<open>1 + _ * 2\<close>])
-qed (simp_all add: and_int_rec [of \<open>_ * 2\<close>])
-
-lemmas AND_upper1' [simp] = order_trans [OF AND_upper1] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
-lemmas AND_upper1'' [simp] = order_le_less_trans [OF AND_upper1] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
-
-lemma AND_upper2 [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
- fixes x y :: int
- assumes \<open>0 \<le> y\<close>
- shows \<open>x AND y \<le> y\<close>
- using assms AND_upper1 [of y x] by (simp add: ac_simps)
-
-lemmas AND_upper2' [simp] = order_trans [OF AND_upper2] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
-lemmas AND_upper2'' [simp] = order_le_less_trans [OF AND_upper2] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close>
-
-lemma plus_and_or: \<open>(x AND y) + (x OR y) = x + y\<close> 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 \<open>y div 2\<close>]
- 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 \<open>y div 2\<close>]
- 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]:
- \<open>set_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
- by (simp add: set_bit_def)
-
-lemma set_bit_negative_int_iff [simp]:
- \<open>set_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
- by (simp add: set_bit_def)
-
-lemma unset_bit_nonnegative_int_iff [simp]:
- \<open>unset_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
- by (simp add: unset_bit_def)
-
-lemma unset_bit_negative_int_iff [simp]:
- \<open>unset_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
- by (simp add: unset_bit_def)
-
-lemma flip_bit_nonnegative_int_iff [simp]:
- \<open>flip_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
- by (simp add: flip_bit_def)
-
-lemma flip_bit_negative_int_iff [simp]:
- \<open>flip_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
- by (simp add: flip_bit_def)
-
-lemma set_bit_greater_eq:
- \<open>set_bit n k \<ge> k\<close> for k :: int
- by (simp add: set_bit_def or_greater_eq)
-
-lemma unset_bit_less_eq:
- \<open>unset_bit n k \<le> k\<close> for k :: int
- by (simp add: unset_bit_def and_less_eq)
-
-lemma set_bit_eq:
- \<open>set_bit n k = k + of_bool (\<not> bit k n) * 2 ^ n\<close> for k :: int
-proof (rule bit_eqI)
- fix m
- show \<open>bit (set_bit n k) m \<longleftrightarrow> bit (k + of_bool (\<not> bit k n) * 2 ^ n) m\<close>
- proof (cases \<open>m = n\<close>)
- 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:
- \<open>unset_bit n k = k - of_bool (bit k n) * 2 ^ n\<close> for k :: int
-proof (rule bit_eqI)
- fix m
- show \<open>bit (unset_bit n k) m \<longleftrightarrow> bit (k - of_bool (bit k n) * 2 ^ n) m\<close>
- proof (cases \<open>m = n\<close>)
- 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 \<open>2 ^ n\<close> \<open>- (2 ^ n)\<close> 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:
- \<open>take_bit n k = mask n \<longleftrightarrow> take_bit n (k + 1) = 0\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
- for k :: int
-proof
- assume ?P
- then have \<open>take_bit n (take_bit n k + take_bit n 1) = 0\<close>
- by (simp add: mask_eq_exp_minus_1)
- then show ?Q
- by (simp only: take_bit_add)
-next
- assume ?Q
- then have \<open>take_bit n (k + 1) - 1 = - 1\<close>
- by simp
- then have \<open>take_bit n (take_bit n (k + 1) - 1) = take_bit n (- 1)\<close>
- by simp
- moreover have \<open>take_bit n (take_bit n (k + 1) - 1) = take_bit n k\<close>
- 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:
- \<open>take_bit n k = mask n \<longleftrightarrow> 2 ^ n dvd k + 1\<close>
- 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:
- \<open>even (of_int k) \<longleftrightarrow> even k\<close>
- by (induction k rule: int_bit_induct) simp_all
-
-lemma bit_of_int_iff [bit_simps]:
- \<open>bit (of_int k) n \<longleftrightarrow> (2::'a) ^ n \<noteq> 0 \<and> bit k n\<close>
-proof (cases \<open>(2::'a) ^ n = 0\<close>)
- case True
- then show ?thesis
- by (simp add: exp_eq_0_imp_not_bit)
-next
- case False
- then have \<open>bit (of_int k) n \<longleftrightarrow> bit k n\<close>
- 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 \<open>of_int k\<close> n] Parity.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 \<open>of_int k\<close> n]
- by (cases n) (auto simp add: ac_simps bit_double_iff even_bit_succ_iff Parity.bit_Suc dest: mult_not_zero)
- qed
- with False show ?thesis
- by simp
-qed
-
-lemma push_bit_of_int:
- \<open>push_bit n (of_int k) = of_int (push_bit n k)\<close>
- by (simp add: push_bit_eq_mult semiring_bit_shifts_class.push_bit_eq_mult)
-
-lemma of_int_push_bit:
- \<open>of_int (push_bit n k) = push_bit n (of_int k)\<close>
- by (simp add: push_bit_eq_mult semiring_bit_shifts_class.push_bit_eq_mult)
-
-lemma take_bit_of_int:
- \<open>take_bit n (of_int k) = of_int (take_bit n k)\<close>
- by (rule bit_eqI) (simp add: bit_take_bit_iff Parity.bit_take_bit_iff bit_of_int_iff)
-
-lemma of_int_take_bit:
- \<open>of_int (take_bit n k) = take_bit n (of_int k)\<close>
- by (rule bit_eqI) (simp add: bit_take_bit_iff Parity.bit_take_bit_iff bit_of_int_iff)
-
-lemma of_int_not_eq:
- \<open>of_int (NOT k) = NOT (of_int k)\<close>
- by (rule bit_eqI) (simp add: bit_not_iff Bit_Operations.bit_not_iff bit_of_int_iff)
-
-lemma of_int_and_eq:
- \<open>of_int (k AND l) = of_int k AND of_int l\<close>
- by (rule bit_eqI) (simp add: bit_of_int_iff bit_and_iff Bit_Operations.bit_and_iff)
-
-lemma of_int_or_eq:
- \<open>of_int (k OR l) = of_int k OR of_int l\<close>
- by (rule bit_eqI) (simp add: bit_of_int_iff bit_or_iff Bit_Operations.bit_or_iff)
-
-lemma of_int_xor_eq:
- \<open>of_int (k XOR l) = of_int k XOR of_int l\<close>
- by (rule bit_eqI) (simp add: bit_of_int_iff bit_xor_iff Bit_Operations.bit_xor_iff)
-
-lemma of_int_mask_eq:
- \<open>of_int (mask n) = mask n\<close>
- by (induction n) (simp_all add: mask_Suc_double Bit_Operations.mask_Suc_double of_int_or_eq)
-
-end
-
-lemma minus_numeral_inc_eq:
- \<open>- numeral (Num.inc n) = NOT (numeral n :: int)\<close>
- by (simp add: not_int_def sub_inc_One_eq add_One)
-
-lemma sub_one_eq_not_neg:
- \<open>Num.sub n num.One = NOT (- numeral n :: int)\<close>
- by (simp add: not_int_def)
-
-lemma int_not_numerals [simp]:
- \<open>NOT (numeral (Num.Bit0 n) :: int) = - numeral (Num.Bit1 n)\<close>
- \<open>NOT (numeral (Num.Bit1 n) :: int) = - numeral (Num.inc (num.Bit1 n))\<close>
- \<open>NOT (numeral (Num.BitM n) :: int) = - numeral (num.Bit0 n)\<close>
- \<open>NOT (- numeral (Num.Bit0 n) :: int) = numeral (Num.BitM n)\<close>
- \<open>NOT (- numeral (Num.Bit1 n) :: int) = numeral (Num.Bit0 n)\<close>
- by (simp_all add: not_int_def add_One inc_BitM_eq)
-
-text \<open>FIXME: The rule sets below are very large (24 rules for each
- operator). Is there a simpler way to do this?\<close>
-
-context
-begin
-
-private lemma eqI:
- \<open>k = l\<close>
- if num: \<open>\<And>n. bit k (numeral n) \<longleftrightarrow> bit l (numeral n)\<close>
- and even: \<open>even k \<longleftrightarrow> even l\<close>
- for k l :: int
-proof (rule bit_eqI)
- fix n
- show \<open>bit k n \<longleftrightarrow> bit l n\<close>
- proof (cases n)
- case 0
- with even show ?thesis
- by simp
- next
- case (Suc n)
- with num [of \<open>num_of_nat (Suc n)\<close>] show ?thesis
- by (simp only: numeral_num_of_nat)
- qed
-qed
-
-lemma int_and_numerals [simp]:
- \<open>numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND numeral y)\<close>
- \<open>numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (2 :: int) * (numeral x AND numeral y)\<close>
- \<open>numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND numeral y)\<close>
- \<open>numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x AND numeral y)\<close>
- \<open>numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND - numeral y)\<close>
- \<open>numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (2 :: int) * (numeral x AND - numeral (y + Num.One))\<close>
- \<open>numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (numeral x AND - numeral y)\<close>
- \<open>numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x AND - numeral (y + Num.One))\<close>
- \<open>- numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (2 :: int) * (- numeral x AND numeral y)\<close>
- \<open>- numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (2 :: int) * (- numeral x AND numeral y)\<close>
- \<open>- numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (2 :: int) * (- numeral (x + Num.One) AND numeral y)\<close>
- \<open>- numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) AND numeral y)\<close>
- \<open>- numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x AND - numeral y)\<close>
- \<open>- numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (2 :: int) * (- numeral x AND - numeral (y + Num.One))\<close>
- \<open>- numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (2 :: int) * (- numeral (x + Num.One) AND - numeral y)\<close>
- \<open>- numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) AND - numeral (y + Num.One))\<close>
- \<open>(1::int) AND numeral (Num.Bit0 y) = 0\<close>
- \<open>(1::int) AND numeral (Num.Bit1 y) = 1\<close>
- \<open>(1::int) AND - numeral (Num.Bit0 y) = 0\<close>
- \<open>(1::int) AND - numeral (Num.Bit1 y) = 1\<close>
- \<open>numeral (Num.Bit0 x) AND (1::int) = 0\<close>
- \<open>numeral (Num.Bit1 x) AND (1::int) = 1\<close>
- \<open>- numeral (Num.Bit0 x) AND (1::int) = 0\<close>
- \<open>- numeral (Num.Bit1 x) AND (1::int) = 1\<close>
- 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]:
- \<open>numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (2 :: int) * (numeral x OR numeral y)\<close>
- \<open>numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR numeral y)\<close>
- \<open>numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x OR numeral y)\<close>
- \<open>numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR numeral y)\<close>
- \<open>numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (2 :: int) * (numeral x OR - numeral y)\<close>
- \<open>numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR - numeral (y + Num.One))\<close>
- \<open>numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x OR - numeral y)\<close>
- \<open>numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x OR - numeral (y + Num.One))\<close>
- \<open>- numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (2 :: int) * (- numeral x OR numeral y)\<close>
- \<open>- numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x OR numeral y)\<close>
- \<open>- numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR numeral y)\<close>
- \<open>- numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR numeral y)\<close>
- \<open>- numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x OR - numeral y)\<close>
- \<open>- numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x OR - numeral (y + Num.One))\<close>
- \<open>- numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR - numeral y)\<close>
- \<open>- numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral (x + Num.One) OR - numeral (y + Num.One))\<close>
- \<open>(1::int) OR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)\<close>
- \<open>(1::int) OR numeral (Num.Bit1 y) = numeral (Num.Bit1 y)\<close>
- \<open>(1::int) OR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)\<close>
- \<open>(1::int) OR - numeral (Num.Bit1 y) = - numeral (Num.Bit1 y)\<close>
- \<open>numeral (Num.Bit0 x) OR (1::int) = numeral (Num.Bit1 x)\<close>
- \<open>numeral (Num.Bit1 x) OR (1::int) = numeral (Num.Bit1 x)\<close>
- \<open>- numeral (Num.Bit0 x) OR (1::int) = - numeral (Num.BitM x)\<close>
- \<open>- numeral (Num.Bit1 x) OR (1::int) = - numeral (Num.Bit1 x)\<close>
- 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]:
- \<open>numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (2 :: int) * (numeral x XOR numeral y)\<close>
- \<open>numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x XOR numeral y)\<close>
- \<open>numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x XOR numeral y)\<close>
- \<open>numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (2 :: int) * (numeral x XOR numeral y)\<close>
- \<open>numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (2 :: int) * (numeral x XOR - numeral y)\<close>
- \<open>numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (numeral x XOR - numeral (y + Num.One))\<close>
- \<open>numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (numeral x XOR - numeral y)\<close>
- \<open>numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (2 :: int) * (numeral x XOR - numeral (y + Num.One))\<close>
- \<open>- numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (2 :: int) * (- numeral x XOR numeral y)\<close>
- \<open>- numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x XOR numeral y)\<close>
- \<open>- numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) XOR numeral y)\<close>
- \<open>- numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (2 :: int) * (- numeral (x + Num.One) XOR numeral y)\<close>
- \<open>- numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (2 :: int) * (- numeral x XOR - numeral y)\<close>
- \<open>- numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = 1 + (2 :: int) * (- numeral x XOR - numeral (y + Num.One))\<close>
- \<open>- numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = 1 + (2 :: int) * (- numeral (x + Num.One) XOR - numeral y)\<close>
- \<open>- numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (2 :: int) * (- numeral (x + Num.One) XOR - numeral (y + Num.One))\<close>
- \<open>(1::int) XOR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)\<close>
- \<open>(1::int) XOR numeral (Num.Bit1 y) = numeral (Num.Bit0 y)\<close>
- \<open>(1::int) XOR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)\<close>
- \<open>(1::int) XOR - numeral (Num.Bit1 y) = - numeral (Num.Bit0 (y + Num.One))\<close>
- \<open>numeral (Num.Bit0 x) XOR (1::int) = numeral (Num.Bit1 x)\<close>
- \<open>numeral (Num.Bit1 x) XOR (1::int) = numeral (Num.Bit0 x)\<close>
- \<open>- numeral (Num.Bit0 x) XOR (1::int) = - numeral (Num.BitM x)\<close>
- \<open>- numeral (Num.Bit1 x) XOR (1::int) = - numeral (Num.Bit0 (x + Num.One))\<close>
- 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 \<open>Bit concatenation\<close>
-
-definition concat_bit :: \<open>nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> int\<close>
- where \<open>concat_bit n k l = take_bit n k OR push_bit n l\<close>
-
-lemma bit_concat_bit_iff [bit_simps]:
- \<open>bit (concat_bit m k l) n \<longleftrightarrow> n < m \<and> bit k n \<or> m \<le> n \<and> bit l (n - m)\<close>
- 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:
- \<open>concat_bit n k l = take_bit n k + push_bit n l\<close>
- 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]:
- \<open>concat_bit 0 k l = l\<close>
- by (simp add: concat_bit_def)
-
-lemma concat_bit_Suc:
- \<open>concat_bit (Suc n) k l = k mod 2 + 2 * concat_bit n (k div 2) l\<close>
- by (simp add: concat_bit_eq take_bit_Suc push_bit_double)
-
-lemma concat_bit_of_zero_1 [simp]:
- \<open>concat_bit n 0 l = push_bit n l\<close>
- by (simp add: concat_bit_def)
-
-lemma concat_bit_of_zero_2 [simp]:
- \<open>concat_bit n k 0 = take_bit n k\<close>
- by (simp add: concat_bit_def take_bit_eq_mask)
-
-lemma concat_bit_nonnegative_iff [simp]:
- \<open>concat_bit n k l \<ge> 0 \<longleftrightarrow> l \<ge> 0\<close>
- by (simp add: concat_bit_def)
-
-lemma concat_bit_negative_iff [simp]:
- \<open>concat_bit n k l < 0 \<longleftrightarrow> l < 0\<close>
- by (simp add: concat_bit_def)
-
-lemma concat_bit_assoc:
- \<open>concat_bit n k (concat_bit m l r) = concat_bit (m + n) (concat_bit n k l) r\<close>
- by (rule bit_eqI) (auto simp add: bit_concat_bit_iff ac_simps)
-
-lemma concat_bit_assoc_sym:
- \<open>concat_bit m (concat_bit n k l) r = concat_bit (min m n) k (concat_bit (m - n) l r)\<close>
- by (rule bit_eqI) (auto simp add: bit_concat_bit_iff ac_simps min_def)
-
-lemma concat_bit_eq_iff:
- \<open>concat_bit n k l = concat_bit n r s
- \<longleftrightarrow> take_bit n k = take_bit n r \<and> l = s\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
-proof
- assume ?Q
- then show ?P
- by (simp add: concat_bit_def)
-next
- assume ?P
- then have *: \<open>bit (concat_bit n k l) m = bit (concat_bit n r s) m\<close> for m
- by (simp add: bit_eq_iff)
- have \<open>take_bit n k = take_bit n r\<close>
- proof (rule bit_eqI)
- fix m
- from * [of m]
- show \<open>bit (take_bit n k) m \<longleftrightarrow> bit (take_bit n r) m\<close>
- by (auto simp add: bit_take_bit_iff bit_concat_bit_iff)
- qed
- moreover have \<open>push_bit n l = push_bit n s\<close>
- proof (rule bit_eqI)
- fix m
- from * [of m]
- show \<open>bit (push_bit n l) m \<longleftrightarrow> bit (push_bit n s) m\<close>
- by (auto simp add: bit_push_bit_iff bit_concat_bit_iff)
- qed
- then have \<open>l = s\<close>
- by (simp add: push_bit_eq_mult)
- ultimately show ?Q
- by (simp add: concat_bit_def)
-qed
-
-lemma take_bit_concat_bit_eq:
- \<open>take_bit m (concat_bit n k l) = concat_bit (min m n) k (take_bit (m - n) l)\<close>
- by (rule bit_eqI)
- (auto simp add: bit_take_bit_iff bit_concat_bit_iff min_def)
-
-lemma concat_bit_take_bit_eq:
- \<open>concat_bit n (take_bit n b) = concat_bit n b\<close>
- by (simp add: concat_bit_def [abs_def])
-
-
-subsection \<open>Taking bits with sign propagation\<close>
-
-context ring_bit_operations
-begin
-
-definition signed_take_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
- where \<open>signed_take_bit n a = take_bit n a OR (of_bool (bit a n) * NOT (mask n))\<close>
-
-lemma signed_take_bit_eq_if_positive:
- \<open>signed_take_bit n a = take_bit n a\<close> if \<open>\<not> bit a n\<close>
- using that by (simp add: signed_take_bit_def)
-
-lemma signed_take_bit_eq_if_negative:
- \<open>signed_take_bit n a = take_bit n a OR NOT (mask n)\<close> if \<open>bit a n\<close>
- using that by (simp add: signed_take_bit_def)
-
-lemma even_signed_take_bit_iff:
- \<open>even (signed_take_bit m a) \<longleftrightarrow> even a\<close>
- 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]:
- \<open>bit (signed_take_bit m a) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> bit a (min m n)\<close>
- 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]:
- \<open>signed_take_bit 0 a = - (a mod 2)\<close>
- by (simp add: signed_take_bit_def odd_iff_mod_2_eq_one)
-
-lemma signed_take_bit_Suc:
- \<open>signed_take_bit (Suc n) a = a mod 2 + 2 * signed_take_bit n (a div 2)\<close>
-proof (rule bit_eqI)
- fix m
- assume *: \<open>2 ^ m \<noteq> 0\<close>
- show \<open>bit (signed_take_bit (Suc n) a) m \<longleftrightarrow>
- bit (a mod 2 + 2 * signed_take_bit n (a div 2)) m\<close>
- proof (cases m)
- case 0
- then show ?thesis
- by (simp add: even_signed_take_bit_iff)
- next
- case (Suc m)
- with * have \<open>2 ^ m \<noteq> 0\<close>
- 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]:
- \<open>signed_take_bit n 0 = 0\<close>
- by (simp add: signed_take_bit_def)
-
-lemma signed_take_bit_of_minus_1 [simp]:
- \<open>signed_take_bit n (- 1) = - 1\<close>
- 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]:
- \<open>signed_take_bit (Suc n) 1 = 1\<close>
- by (simp add: signed_take_bit_Suc)
-
-lemma signed_take_bit_rec:
- \<open>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))\<close>
- by (cases n) (simp_all add: signed_take_bit_Suc)
-
-lemma signed_take_bit_eq_iff_take_bit_eq:
- \<open>signed_take_bit n a = signed_take_bit n b \<longleftrightarrow> take_bit (Suc n) a = take_bit (Suc n) b\<close>
-proof -
- have \<open>bit (signed_take_bit n a) = bit (signed_take_bit n b) \<longleftrightarrow> bit (take_bit (Suc n) a) = bit (take_bit (Suc n) b)\<close>
- 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]:
- \<open>signed_take_bit m (signed_take_bit n a) = signed_take_bit (min m n) a\<close>
-proof (rule bit_eqI)
- fix q
- show \<open>bit (signed_take_bit m (signed_take_bit n a)) q \<longleftrightarrow>
- bit (signed_take_bit (min m n) a) q\<close>
- 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:
- \<open>signed_take_bit m (take_bit n a) = (if n \<le> m then take_bit n else signed_take_bit m) a\<close>
- by (rule bit_eqI) (auto simp add: bit_signed_take_bit_iff min_def bit_take_bit_iff)
-
-lemma take_bit_signed_take_bit:
- \<open>take_bit m (signed_take_bit n a) = take_bit m a\<close> if \<open>m \<le> Suc n\<close>
- 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 \<open>Modulus centered around 0\<close>
-
-lemma signed_take_bit_eq_concat_bit:
- \<open>signed_take_bit n k = concat_bit n k (- of_bool (bit k n))\<close>
- by (simp add: concat_bit_def signed_take_bit_def push_bit_minus_one_eq_not_mask)
-
-lemma signed_take_bit_add:
- \<open>signed_take_bit n (signed_take_bit n k + signed_take_bit n l) = signed_take_bit n (k + l)\<close>
- for k l :: int
-proof -
- have \<open>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)\<close>
- 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:
- \<open>signed_take_bit n (signed_take_bit n k - signed_take_bit n l) = signed_take_bit n (k - l)\<close>
- for k l :: int
-proof -
- have \<open>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)\<close>
- 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:
- \<open>signed_take_bit n (- signed_take_bit n k) = signed_take_bit n (- k)\<close>
- for k :: int
-proof -
- have \<open>take_bit (Suc n)
- (- take_bit (Suc n) (signed_take_bit n k)) =
- take_bit (Suc n) (- k)\<close>
- 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:
- \<open>signed_take_bit n (signed_take_bit n k * signed_take_bit n l) = signed_take_bit n (k * l)\<close>
- for k l :: int
-proof -
- have \<open>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)\<close>
- 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:
- \<open>signed_take_bit n k = take_bit (Suc n) k - 2 ^ Suc n * of_bool (bit k n)\<close>
- for k :: int
-proof (cases \<open>bit k n\<close>)
- case True
- have \<open>signed_take_bit n k = take_bit (Suc n) k OR NOT (mask (Suc n))\<close>
- 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 \<open>signed_take_bit n k = take_bit (Suc n) k + NOT (mask (Suc n))\<close>
- 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:
- \<open>signed_take_bit n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n\<close>
- for k :: int
-proof -
- have *: \<open>take_bit n k OR 2 ^ n = take_bit n k + 2 ^ n\<close>
- by (simp add: disjunctive_add bit_exp_iff bit_take_bit_iff)
- have \<open>take_bit n k - 2 ^ n = take_bit n k + NOT (mask n)\<close>
- by (simp add: minus_exp_eq_not_mask)
- also have \<open>\<dots> = take_bit n k OR NOT (mask n)\<close>
- by (rule disjunctive_add)
- (simp add: bit_exp_iff bit_take_bit_iff bit_not_iff bit_mask_iff)
- finally have **: \<open>take_bit n k - 2 ^ n = take_bit n k OR NOT (mask n)\<close> .
- have \<open>take_bit (Suc n) (k + 2 ^ n) = take_bit (Suc n) (take_bit (Suc n) k + take_bit (Suc n) (2 ^ n))\<close>
- by (simp only: take_bit_add)
- also have \<open>take_bit (Suc n) k = 2 ^ n * of_bool (bit k n) + take_bit n k\<close>
- by (simp add: take_bit_Suc_from_most)
- finally have \<open>take_bit (Suc n) (k + 2 ^ n) = take_bit (Suc n) (2 ^ (n + of_bool (bit k n)) + take_bit n k)\<close>
- by (simp add: ac_simps)
- also have \<open>2 ^ (n + of_bool (bit k n)) + take_bit n k = 2 ^ (n + of_bool (bit k n)) OR take_bit n k\<close>
- 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]:
- \<open>0 \<le> signed_take_bit n k \<longleftrightarrow> \<not> bit k n\<close>
- for k :: int
- by (simp add: signed_take_bit_def not_less concat_bit_def)
-
-lemma signed_take_bit_negative_iff [simp]:
- \<open>signed_take_bit n k < 0 \<longleftrightarrow> bit k n\<close>
- 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]:
- \<open>- (2 ^ n) \<le> signed_take_bit n k\<close>
- for k :: int
- by (simp add: signed_take_bit_eq_take_bit_shift)
-
-lemma signed_take_bit_int_less_exp [simp]:
- \<open>signed_take_bit n k < 2 ^ n\<close>
- for k :: int
- using take_bit_int_less_exp [of \<open>Suc n\<close>]
- by (simp add: signed_take_bit_eq_take_bit_shift)
-
-lemma signed_take_bit_int_eq_self_iff:
- \<open>signed_take_bit n k = k \<longleftrightarrow> - (2 ^ n) \<le> k \<and> k < 2 ^ n\<close>
- 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:
- \<open>signed_take_bit n k = k\<close> if \<open>- (2 ^ n) \<le> k\<close> \<open>k < 2 ^ n\<close>
- for k :: int
- using that by (simp add: signed_take_bit_int_eq_self_iff)
-
-lemma signed_take_bit_int_less_eq_self_iff:
- \<open>signed_take_bit n k \<le> k \<longleftrightarrow> - (2 ^ n) \<le> k\<close>
- 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:
- \<open>signed_take_bit n k < k \<longleftrightarrow> 2 ^ n \<le> k\<close>
- 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:
- \<open>k < signed_take_bit n k \<longleftrightarrow> k < - (2 ^ n)\<close>
- 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:
- \<open>k \<le> signed_take_bit n k \<longleftrightarrow> k < 2 ^ n\<close>
- 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:
- \<open>k + 2 ^ Suc n \<le> signed_take_bit n k\<close> if \<open>k < - (2 ^ n)\<close>
- for k :: int
- using that take_bit_int_greater_eq [of \<open>k + 2 ^ n\<close> \<open>Suc n\<close>]
- by (simp add: signed_take_bit_eq_take_bit_shift)
-
-lemma signed_take_bit_int_less_eq:
- \<open>signed_take_bit n k \<le> k - 2 ^ Suc n\<close> if \<open>k \<ge> 2 ^ n\<close>
- for k :: int
- using that take_bit_int_less_eq [of \<open>Suc n\<close> \<open>k + 2 ^ n\<close>]
- by (simp add: signed_take_bit_eq_take_bit_shift)
-
-lemma signed_take_bit_Suc_bit0 [simp]:
- \<open>signed_take_bit (Suc n) (numeral (Num.Bit0 k)) = signed_take_bit n (numeral k) * (2 :: int)\<close>
- by (simp add: signed_take_bit_Suc)
-
-lemma signed_take_bit_Suc_bit1 [simp]:
- \<open>signed_take_bit (Suc n) (numeral (Num.Bit1 k)) = signed_take_bit n (numeral k) * 2 + (1 :: int)\<close>
- by (simp add: signed_take_bit_Suc)
-
-lemma signed_take_bit_Suc_minus_bit0 [simp]:
- \<open>signed_take_bit (Suc n) (- numeral (Num.Bit0 k)) = signed_take_bit n (- numeral k) * (2 :: int)\<close>
- by (simp add: signed_take_bit_Suc)
-
-lemma signed_take_bit_Suc_minus_bit1 [simp]:
- \<open>signed_take_bit (Suc n) (- numeral (Num.Bit1 k)) = signed_take_bit n (- numeral k - 1) * 2 + (1 :: int)\<close>
- by (simp add: signed_take_bit_Suc)
-
-lemma signed_take_bit_numeral_bit0 [simp]:
- \<open>signed_take_bit (numeral l) (numeral (Num.Bit0 k)) = signed_take_bit (pred_numeral l) (numeral k) * (2 :: int)\<close>
- by (simp add: signed_take_bit_rec)
-
-lemma signed_take_bit_numeral_bit1 [simp]:
- \<open>signed_take_bit (numeral l) (numeral (Num.Bit1 k)) = signed_take_bit (pred_numeral l) (numeral k) * 2 + (1 :: int)\<close>
- by (simp add: signed_take_bit_rec)
-
-lemma signed_take_bit_numeral_minus_bit0 [simp]:
- \<open>signed_take_bit (numeral l) (- numeral (Num.Bit0 k)) = signed_take_bit (pred_numeral l) (- numeral k) * (2 :: int)\<close>
- by (simp add: signed_take_bit_rec)
-
-lemma signed_take_bit_numeral_minus_bit1 [simp]:
- \<open>signed_take_bit (numeral l) (- numeral (Num.Bit1 k)) = signed_take_bit (pred_numeral l) (- numeral k - 1) * 2 + (1 :: int)\<close>
- by (simp add: signed_take_bit_rec)
-
-lemma signed_take_bit_code [code]:
- \<open>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)\<close>
-proof -
- have *: \<open>take_bit (Suc n) a + push_bit n (- 2) =
- take_bit (Suc n) a OR NOT (mask (Suc n))\<close>
- 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 \<open>Instance \<^typ>\<open>nat\<close>\<close>
-
-instantiation nat :: semiring_bit_operations
-begin
-
-definition and_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
- where \<open>m AND n = nat (int m AND int n)\<close> for m n :: nat
-
-definition or_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
- where \<open>m OR n = nat (int m OR int n)\<close> for m n :: nat
-
-definition xor_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
- where \<open>m XOR n = nat (int m XOR int n)\<close> for m n :: nat
-
-definition mask_nat :: \<open>nat \<Rightarrow> nat\<close>
- where \<open>mask n = (2 :: nat) ^ n - 1\<close>
-
-definition set_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
- where \<open>set_bit m n = n OR push_bit m 1\<close> for m n :: nat
-
-definition unset_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
- where \<open>unset_bit m n = (if bit n m then n - push_bit m 1 else n)\<close> for m n :: nat
-
-definition flip_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
- where \<open>flip_bit m n = n XOR push_bit m 1\<close> for m n :: nat
-
-instance proof
- fix m n q :: nat
- show \<open>bit (m AND n) q \<longleftrightarrow> bit m q \<and> bit n q\<close>
- by (simp add: and_nat_def bit_simps)
- show \<open>bit (m OR n) q \<longleftrightarrow> bit m q \<or> bit n q\<close>
- by (simp add: or_nat_def bit_simps)
- show \<open>bit (m XOR n) q \<longleftrightarrow> bit m q \<noteq> bit n q\<close>
- by (simp add: xor_nat_def bit_simps)
- show \<open>bit (unset_bit m n) q \<longleftrightarrow> bit n q \<and> m \<noteq> q\<close>
- proof (cases \<open>bit n m\<close>)
- case False
- then show ?thesis by (auto simp add: unset_bit_nat_def)
- next
- case True
- have \<open>push_bit m (drop_bit m n) + take_bit m n = n\<close>
- by (fact bits_ident)
- also from \<open>bit n m\<close> have \<open>drop_bit m n = 2 * drop_bit (Suc m) n + 1\<close>
- by (simp add: drop_bit_Suc drop_bit_half even_drop_bit_iff_not_bit ac_simps)
- finally have \<open>push_bit m (2 * drop_bit (Suc m) n) + take_bit m n + push_bit m 1 = n\<close>
- by (simp only: push_bit_add ac_simps)
- then have \<open>n - push_bit m 1 = push_bit m (2 * drop_bit (Suc m) n) + take_bit m n\<close>
- by simp
- then have \<open>n - push_bit m 1 = push_bit m (2 * drop_bit (Suc m) n) OR take_bit m n\<close>
- by (simp add: or_nat_def bit_simps flip: disjunctive_add)
- with \<open>bit n m\<close> 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:
- \<open>m AND n = of_bool (odd m \<and> odd n) + 2 * ((m div 2) AND (n div 2))\<close> for m n :: nat
- by (simp add: and_nat_def and_int_rec [of \<open>int m\<close> \<open>int n\<close>] zdiv_int nat_add_distrib nat_mult_distrib)
-
-lemma or_nat_rec:
- \<open>m OR n = of_bool (odd m \<or> odd n) + 2 * ((m div 2) OR (n div 2))\<close> for m n :: nat
- by (simp add: or_nat_def or_int_rec [of \<open>int m\<close> \<open>int n\<close>] zdiv_int nat_add_distrib nat_mult_distrib)
-
-lemma xor_nat_rec:
- \<open>m XOR n = of_bool (odd m \<noteq> odd n) + 2 * ((m div 2) XOR (n div 2))\<close> for m n :: nat
- by (simp add: xor_nat_def xor_int_rec [of \<open>int m\<close> \<open>int n\<close>] zdiv_int nat_add_distrib nat_mult_distrib)
-
-lemma Suc_0_and_eq [simp]:
- \<open>Suc 0 AND n = n mod 2\<close>
- using one_and_eq [of n] by simp
-
-lemma and_Suc_0_eq [simp]:
- \<open>n AND Suc 0 = n mod 2\<close>
- using and_one_eq [of n] by simp
-
-lemma Suc_0_or_eq:
- \<open>Suc 0 OR n = n + of_bool (even n)\<close>
- using one_or_eq [of n] by simp
-
-lemma or_Suc_0_eq:
- \<open>n OR Suc 0 = n + of_bool (even n)\<close>
- using or_one_eq [of n] by simp
-
-lemma Suc_0_xor_eq:
- \<open>Suc 0 XOR n = n + of_bool (even n) - of_bool (odd n)\<close>
- using one_xor_eq [of n] by simp
-
-lemma xor_Suc_0_eq:
- \<open>n XOR Suc 0 = n + of_bool (even n) - of_bool (odd n)\<close>
- using xor_one_eq [of n] by simp
-
-context semiring_bit_operations
-begin
-
-lemma of_nat_and_eq:
- \<open>of_nat (m AND n) = of_nat m AND of_nat n\<close>
- by (rule bit_eqI) (simp add: bit_of_nat_iff bit_and_iff Bit_Operations.bit_and_iff)
-
-lemma of_nat_or_eq:
- \<open>of_nat (m OR n) = of_nat m OR of_nat n\<close>
- by (rule bit_eqI) (simp add: bit_of_nat_iff bit_or_iff Bit_Operations.bit_or_iff)
-
-lemma of_nat_xor_eq:
- \<open>of_nat (m XOR n) = of_nat m XOR of_nat n\<close>
- 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:
- \<open>of_nat (mask n) = mask n\<close>
- by (induction n) (simp_all add: mask_Suc_double Bit_Operations.mask_Suc_double of_nat_or_eq)
-
-end
-
-lemma Suc_mask_eq_exp:
- \<open>Suc (mask n) = 2 ^ n\<close>
- by (simp add: mask_eq_exp_minus_1)
-
-lemma less_eq_mask:
- \<open>n \<le> mask n\<close>
- 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:
- \<open>n < mask n\<close> if \<open>Suc 0 < n\<close>
-proof -
- define m where \<open>m = n - 2\<close>
- with that have *: \<open>n = m + 2\<close>
- by simp
- have \<open>Suc (Suc (Suc m)) < 4 * 2 ^ m\<close>
- by (induction m) simp_all
- then have \<open>Suc (m + 2) < Suc (mask (m + 2))\<close>
- by (simp add: Suc_mask_eq_exp)
- then have \<open>m + 2 < mask (m + 2)\<close>
- by (simp add: less_le)
- with * show ?thesis
- by simp
-qed
-
-
-subsection \<open>Symbolic computations on numeral expressions\<close> \<^marker>\<open>contributor \<open>Andreas Lochbihler\<close>\<close>
-
-fun and_num :: \<open>num \<Rightarrow> num \<Rightarrow> num option\<close>
-where
- \<open>and_num num.One num.One = Some num.One\<close>
-| \<open>and_num num.One (num.Bit0 n) = None\<close>
-| \<open>and_num num.One (num.Bit1 n) = Some num.One\<close>
-| \<open>and_num (num.Bit0 m) num.One = None\<close>
-| \<open>and_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (and_num m n)\<close>
-| \<open>and_num (num.Bit0 m) (num.Bit1 n) = map_option num.Bit0 (and_num m n)\<close>
-| \<open>and_num (num.Bit1 m) num.One = Some num.One\<close>
-| \<open>and_num (num.Bit1 m) (num.Bit0 n) = map_option num.Bit0 (and_num m n)\<close>
-| \<open>and_num (num.Bit1 m) (num.Bit1 n) = (case and_num m n of None \<Rightarrow> Some num.One | Some n' \<Rightarrow> Some (num.Bit1 n'))\<close>
-
-fun and_not_num :: \<open>num \<Rightarrow> num \<Rightarrow> num option\<close>
-where
- \<open>and_not_num num.One num.One = None\<close>
-| \<open>and_not_num num.One (num.Bit0 n) = Some num.One\<close>
-| \<open>and_not_num num.One (num.Bit1 n) = None\<close>
-| \<open>and_not_num (num.Bit0 m) num.One = Some (num.Bit0 m)\<close>
-| \<open>and_not_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (and_not_num m n)\<close>
-| \<open>and_not_num (num.Bit0 m) (num.Bit1 n) = map_option num.Bit0 (and_not_num m n)\<close>
-| \<open>and_not_num (num.Bit1 m) num.One = Some (num.Bit0 m)\<close>
-| \<open>and_not_num (num.Bit1 m) (num.Bit0 n) = (case and_not_num m n of None \<Rightarrow> Some num.One | Some n' \<Rightarrow> Some (num.Bit1 n'))\<close>
-| \<open>and_not_num (num.Bit1 m) (num.Bit1 n) = map_option num.Bit0 (and_not_num m n)\<close>
-
-fun or_num :: \<open>num \<Rightarrow> num \<Rightarrow> num\<close>
-where
- \<open>or_num num.One num.One = num.One\<close>
-| \<open>or_num num.One (num.Bit0 n) = num.Bit1 n\<close>
-| \<open>or_num num.One (num.Bit1 n) = num.Bit1 n\<close>
-| \<open>or_num (num.Bit0 m) num.One = num.Bit1 m\<close>
-| \<open>or_num (num.Bit0 m) (num.Bit0 n) = num.Bit0 (or_num m n)\<close>
-| \<open>or_num (num.Bit0 m) (num.Bit1 n) = num.Bit1 (or_num m n)\<close>
-| \<open>or_num (num.Bit1 m) num.One = num.Bit1 m\<close>
-| \<open>or_num (num.Bit1 m) (num.Bit0 n) = num.Bit1 (or_num m n)\<close>
-| \<open>or_num (num.Bit1 m) (num.Bit1 n) = num.Bit1 (or_num m n)\<close>
-
-fun or_not_num_neg :: \<open>num \<Rightarrow> num \<Rightarrow> num\<close>
-where
- \<open>or_not_num_neg num.One num.One = num.One\<close>
-| \<open>or_not_num_neg num.One (num.Bit0 m) = num.Bit1 m\<close>
-| \<open>or_not_num_neg num.One (num.Bit1 m) = num.Bit1 m\<close>
-| \<open>or_not_num_neg (num.Bit0 n) num.One = num.Bit0 num.One\<close>
-| \<open>or_not_num_neg (num.Bit0 n) (num.Bit0 m) = Num.BitM (or_not_num_neg n m)\<close>
-| \<open>or_not_num_neg (num.Bit0 n) (num.Bit1 m) = num.Bit0 (or_not_num_neg n m)\<close>
-| \<open>or_not_num_neg (num.Bit1 n) num.One = num.One\<close>
-| \<open>or_not_num_neg (num.Bit1 n) (num.Bit0 m) = Num.BitM (or_not_num_neg n m)\<close>
-| \<open>or_not_num_neg (num.Bit1 n) (num.Bit1 m) = Num.BitM (or_not_num_neg n m)\<close>
-
-fun xor_num :: \<open>num \<Rightarrow> num \<Rightarrow> num option\<close>
-where
- \<open>xor_num num.One num.One = None\<close>
-| \<open>xor_num num.One (num.Bit0 n) = Some (num.Bit1 n)\<close>
-| \<open>xor_num num.One (num.Bit1 n) = Some (num.Bit0 n)\<close>
-| \<open>xor_num (num.Bit0 m) num.One = Some (num.Bit1 m)\<close>
-| \<open>xor_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (xor_num m n)\<close>
-| \<open>xor_num (num.Bit0 m) (num.Bit1 n) = Some (case xor_num m n of None \<Rightarrow> num.One | Some n' \<Rightarrow> num.Bit1 n')\<close>
-| \<open>xor_num (num.Bit1 m) num.One = Some (num.Bit0 m)\<close>
-| \<open>xor_num (num.Bit1 m) (num.Bit0 n) = Some (case xor_num m n of None \<Rightarrow> num.One | Some n' \<Rightarrow> num.Bit1 n')\<close>
-| \<open>xor_num (num.Bit1 m) (num.Bit1 n) = map_option num.Bit0 (xor_num m n)\<close>
-
-lemma int_numeral_and_num:
- \<open>numeral m AND numeral n = (case and_num m n of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close>
- by (induction m n rule: and_num.induct) (simp_all split: option.split)
-
-lemma and_num_eq_None_iff:
- \<open>and_num m n = None \<longleftrightarrow> numeral m AND numeral n = (0::int)\<close>
- by (simp add: int_numeral_and_num split: option.split)
-
-lemma and_num_eq_Some_iff:
- \<open>and_num m n = Some q \<longleftrightarrow> numeral m AND numeral n = (numeral q :: int)\<close>
- by (simp add: int_numeral_and_num split: option.split)
-
-lemma int_numeral_and_not_num:
- \<open>numeral m AND NOT (numeral n) = (case and_not_num m n of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close>
- 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:
- \<open>NOT (numeral m) AND numeral n = (case and_not_num n m of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close>
- using int_numeral_and_not_num [of n m] by (simp add: ac_simps)
-
-lemma and_not_num_eq_None_iff:
- \<open>and_not_num m n = None \<longleftrightarrow> numeral m AND NOT (numeral n) = (0::int)\<close>
- by (simp add: int_numeral_and_not_num split: option.split)
-
-lemma and_not_num_eq_Some_iff:
- \<open>and_not_num m n = Some q \<longleftrightarrow> numeral m AND NOT (numeral n) = (numeral q :: int)\<close>
- by (simp add: int_numeral_and_not_num split: option.split)
-
-lemma int_numeral_or_num:
- \<open>numeral m OR numeral n = (numeral (or_num m n) :: int)\<close>
- by (induction m n rule: or_num.induct) simp_all
-
-lemma numeral_or_num_eq:
- \<open>numeral (or_num m n) = (numeral m OR numeral n :: int)\<close>
- by (simp add: int_numeral_or_num)
-
-lemma int_numeral_or_not_num_neg:
- \<open>numeral m OR NOT (numeral n :: int) = - numeral (or_not_num_neg m n)\<close>
- 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:
- \<open>NOT (numeral m) OR (numeral n :: int) = - numeral (or_not_num_neg n m)\<close>
- using int_numeral_or_not_num_neg [of n m] by (simp add: ac_simps)
-
-lemma numeral_or_not_num_eq:
- \<open>numeral (or_not_num_neg m n) = - (numeral m OR NOT (numeral n :: int))\<close>
- using int_numeral_or_not_num_neg [of m n] by simp
-
-lemma int_numeral_xor_num:
- \<open>numeral m XOR numeral n = (case xor_num m n of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close>
- by (induction m n rule: xor_num.induct) (simp_all split: option.split)
-
-lemma xor_num_eq_None_iff:
- \<open>xor_num m n = None \<longleftrightarrow> numeral m XOR numeral n = (0::int)\<close>
- by (simp add: int_numeral_xor_num split: option.split)
-
-lemma xor_num_eq_Some_iff:
- \<open>xor_num m n = Some q \<longleftrightarrow> numeral m XOR numeral n = (numeral q :: int)\<close>
- by (simp add: int_numeral_xor_num split: option.split)
-
-
-subsection \<open>Instances for \<^typ>\<open>integer\<close> and \<^typ>\<open>natural\<close>\<close>
-
-unbundle integer.lifting natural.lifting
-
-instantiation integer :: ring_bit_operations
-begin
-
-lift_definition not_integer :: \<open>integer \<Rightarrow> integer\<close>
- is not .
-
-lift_definition and_integer :: \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close>
- is \<open>and\<close> .
-
-lift_definition or_integer :: \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close>
- is or .
-
-lift_definition xor_integer :: \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close>
- is xor .
-
-lift_definition mask_integer :: \<open>nat \<Rightarrow> integer\<close>
- is mask .
-
-lift_definition set_bit_integer :: \<open>nat \<Rightarrow> integer \<Rightarrow> integer\<close>
- is set_bit .
-
-lift_definition unset_bit_integer :: \<open>nat \<Rightarrow> integer \<Rightarrow> integer\<close>
- is unset_bit .
-
-lift_definition flip_bit_integer :: \<open>nat \<Rightarrow> integer \<Rightarrow> integer\<close>
- is flip_bit .
-
-instance by (standard; transfer)
- (simp_all add: minus_eq_not_minus_1 mask_eq_exp_minus_1
- bit_not_iff bit_and_iff bit_or_iff bit_xor_iff
- set_bit_def bit_unset_bit_iff flip_bit_def)
-
-end
-
-lemma [code]:
- \<open>mask n = 2 ^ n - (1::integer)\<close>
- by (simp add: mask_eq_exp_minus_1)
-
-instantiation natural :: semiring_bit_operations
-begin
-
-lift_definition and_natural :: \<open>natural \<Rightarrow> natural \<Rightarrow> natural\<close>
- is \<open>and\<close> .
-
-lift_definition or_natural :: \<open>natural \<Rightarrow> natural \<Rightarrow> natural\<close>
- is or .
-
-lift_definition xor_natural :: \<open>natural \<Rightarrow> natural \<Rightarrow> natural\<close>
- is xor .
-
-lift_definition mask_natural :: \<open>nat \<Rightarrow> natural\<close>
- is mask .
-
-lift_definition set_bit_natural :: \<open>nat \<Rightarrow> natural \<Rightarrow> natural\<close>
- is set_bit .
-
-lift_definition unset_bit_natural :: \<open>nat \<Rightarrow> natural \<Rightarrow> natural\<close>
- is unset_bit .
-
-lift_definition flip_bit_natural :: \<open>nat \<Rightarrow> natural \<Rightarrow> natural\<close>
- is flip_bit .
-
-instance by (standard; transfer)
- (simp_all add: mask_eq_exp_minus_1
- bit_and_iff bit_or_iff bit_xor_iff
- set_bit_def bit_unset_bit_iff flip_bit_def)
-
-end
-
-lemma [code]:
- \<open>integer_of_natural (mask n) = mask n\<close>
- by transfer (simp add: mask_eq_exp_minus_1 of_nat_diff)
-
-lifting_update integer.lifting
-lifting_forget integer.lifting
-
-lifting_update natural.lifting
-lifting_forget natural.lifting
-
-
-subsection \<open>Key ideas of bit operations\<close>
-
-text \<open>
- 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>\<open>int\<close>,
- bounded bit values by quotient types over \<^typ>\<open>int\<close>.
-
- \<^item> (A special case are idealized unbounded bit values ending
- in @{term [source] 0} which can be represented by type \<^typ>\<open>nat\<close> but
- only support a restricted set of operations).
-
- \<^item> From this idea follows that
-
- \<^item> multiplication by \<^term>\<open>2 :: int\<close> is a bit shift to the left and
-
- \<^item> division by \<^term>\<open>2 :: int\<close> 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>\<open>(2 :: int) ^ n \<noteq> 0\<close>
- represents that \<^term>\<open>n\<close> is \<^emph>\<open>not\<close> 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>\<open>n\<close>th bit: \<^term>\<open>(2 :: int) ^ n\<close>
-
- \<^item> Bit mask upto bit \<^term>\<open>n\<close>: @{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>\<open>0::int\<close>: @{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]}
-\<close>
-
-code_identifier
- type_class semiring_bits \<rightharpoonup>
- (SML) Bit_Operations.semiring_bits and (OCaml) Bit_Operations.semiring_bits and (Haskell) Bit_Operations.semiring_bits and (Scala) Bit_Operations.semiring_bits
-| class_relation semiring_bits < semiring_parity \<rightharpoonup>
- (SML) Bit_Operations.semiring_parity_semiring_bits and (OCaml) Bit_Operations.semiring_parity_semiring_bits and (Haskell) Bit_Operations.semiring_parity_semiring_bits and (Scala) Bit_Operations.semiring_parity_semiring_bits
-| constant bit \<rightharpoonup>
- (SML) Bit_Operations.bit and (OCaml) Bit_Operations.bit and (Haskell) Bit_Operations.bit and (Scala) Bit_Operations.bit
-| class_instance nat :: semiring_bits \<rightharpoonup>
- (SML) Bit_Operations.semiring_bits_nat and (OCaml) Bit_Operations.semiring_bits_nat and (Haskell) Bit_Operations.semiring_bits_nat and (Scala) Bit_Operations.semiring_bits_nat
-| class_instance int :: semiring_bits \<rightharpoonup>
- (SML) Bit_Operations.semiring_bits_int and (OCaml) Bit_Operations.semiring_bits_int and (Haskell) Bit_Operations.semiring_bits_int and (Scala) Bit_Operations.semiring_bits_int
-| type_class semiring_bit_shifts \<rightharpoonup>
- (SML) Bit_Operations.semiring_bit_shifts and (OCaml) Bit_Operations.semiring_bit_shifts and (Haskell) Bit_Operations.semiring_bits and (Scala) Bit_Operations.semiring_bit_shifts
-| class_relation semiring_bit_shifts < semiring_bits \<rightharpoonup>
- (SML) Bit_Operations.semiring_bits_semiring_bit_shifts and (OCaml) Bit_Operations.semiring_bits_semiring_bit_shifts and (Haskell) Bit_Operations.semiring_bits_semiring_bit_shifts and (Scala) Bit_Operations.semiring_bits_semiring_bit_shifts
-| constant push_bit \<rightharpoonup>
- (SML) Bit_Operations.push_bit and (OCaml) Bit_Operations.push_bit and (Haskell) Bit_Operations.push_bit and (Scala) Bit_Operations.push_bit
-| constant drop_bit \<rightharpoonup>
- (SML) Bit_Operations.drop_bit and (OCaml) Bit_Operations.drop_bit and (Haskell) Bit_Operations.drop_bit and (Scala) Bit_Operations.drop_bit
-| constant take_bit \<rightharpoonup>
- (SML) Bit_Operations.take_bit and (OCaml) Bit_Operations.take_bit and (Haskell) Bit_Operations.take_bit and (Scala) Bit_Operations.take_bit
-| class_instance nat :: semiring_bit_shifts \<rightharpoonup>
- (SML) Bit_Operations.semiring_bit_shifts and (OCaml) Bit_Operations.semiring_bit_shifts and (Haskell) Bit_Operations.semiring_bit_shifts and (Scala) Bit_Operations.semiring_bit_shifts
-| class_instance int :: semiring_bit_shifts \<rightharpoonup>
- (SML) Bit_Operations.semiring_bit_shifts and (OCaml) Bit_Operations.semiring_bit_shifts and (Haskell) Bit_Operations.semiring_bit_shifts and (Scala) Bit_Operations.semiring_bit_shifts
-
-no_notation
- "and" (infixr \<open>AND\<close> 64)
- and or (infixr \<open>OR\<close> 59)
- and xor (infixr \<open>XOR\<close> 59)
-
-bundle bit_operations_syntax
-begin
-
-notation
- "and" (infixr \<open>AND\<close> 64)
- and or (infixr \<open>OR\<close> 59)
- and xor (infixr \<open>XOR\<close> 59)
-
-end
-
-end
--- a/src/HOL/Library/Boolean_Algebra.thy Sun Aug 01 23:18:13 2021 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,286 +0,0 @@
-(* Title: HOL/Library/Boolean_Algebra.thy
- Author: Brian Huffman
-*)
-
-section \<open>Boolean Algebras\<close>
-
-theory Boolean_Algebra
- imports Main
-begin
-
-locale boolean_algebra = conj: abel_semigroup "(\<^bold>\<sqinter>)" + disj: abel_semigroup "(\<^bold>\<squnion>)"
- for conj :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<^bold>\<sqinter>" 70)
- and disj :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<^bold>\<squnion>" 65) +
- fixes compl :: "'a \<Rightarrow> 'a" ("\<sim> _" [81] 80)
- and zero :: "'a" ("\<zero>")
- and one :: "'a" ("\<one>")
- assumes conj_disj_distrib: "x \<^bold>\<sqinter> (y \<^bold>\<squnion> z) = (x \<^bold>\<sqinter> y) \<^bold>\<squnion> (x \<^bold>\<sqinter> z)"
- and disj_conj_distrib: "x \<^bold>\<squnion> (y \<^bold>\<sqinter> z) = (x \<^bold>\<squnion> y) \<^bold>\<sqinter> (x \<^bold>\<squnion> z)"
- and conj_one_right: "x \<^bold>\<sqinter> \<one> = x"
- and disj_zero_right: "x \<^bold>\<squnion> \<zero> = x"
- and conj_cancel_right [simp]: "x \<^bold>\<sqinter> \<sim> x = \<zero>"
- and disj_cancel_right [simp]: "x \<^bold>\<squnion> \<sim> x = \<one>"
-begin
-
-sublocale conj: semilattice_neutr "(\<^bold>\<sqinter>)" "\<one>"
-proof
- show "x \<^bold>\<sqinter> \<one> = x" for x
- by (fact conj_one_right)
- show "x \<^bold>\<sqinter> x = x" for x
- proof -
- have "x \<^bold>\<sqinter> x = (x \<^bold>\<sqinter> x) \<^bold>\<squnion> \<zero>"
- by (simp add: disj_zero_right)
- also have "\<dots> = (x \<^bold>\<sqinter> x) \<^bold>\<squnion> (x \<^bold>\<sqinter> \<sim> x)"
- by simp
- also have "\<dots> = x \<^bold>\<sqinter> (x \<^bold>\<squnion> \<sim> x)"
- by (simp only: conj_disj_distrib)
- also have "\<dots> = x \<^bold>\<sqinter> \<one>"
- by simp
- also have "\<dots> = x"
- by (simp add: conj_one_right)
- finally show ?thesis .
- qed
-qed
-
-sublocale disj: semilattice_neutr "(\<^bold>\<squnion>)" "\<zero>"
-proof
- show "x \<^bold>\<squnion> \<zero> = x" for x
- by (fact disj_zero_right)
- show "x \<^bold>\<squnion> x = x" for x
- proof -
- have "x \<^bold>\<squnion> x = (x \<^bold>\<squnion> x) \<^bold>\<sqinter> \<one>"
- by simp
- also have "\<dots> = (x \<^bold>\<squnion> x) \<^bold>\<sqinter> (x \<^bold>\<squnion> \<sim> x)"
- by simp
- also have "\<dots> = x \<^bold>\<squnion> (x \<^bold>\<sqinter> \<sim> x)"
- by (simp only: disj_conj_distrib)
- also have "\<dots> = x \<^bold>\<squnion> \<zero>"
- by simp
- also have "\<dots> = x"
- by (simp add: disj_zero_right)
- finally show ?thesis .
- qed
-qed
-
-
-subsection \<open>Complement\<close>
-
-lemma complement_unique:
- assumes 1: "a \<^bold>\<sqinter> x = \<zero>"
- assumes 2: "a \<^bold>\<squnion> x = \<one>"
- assumes 3: "a \<^bold>\<sqinter> y = \<zero>"
- assumes 4: "a \<^bold>\<squnion> y = \<one>"
- shows "x = y"
-proof -
- from 1 3 have "(a \<^bold>\<sqinter> x) \<^bold>\<squnion> (x \<^bold>\<sqinter> y) = (a \<^bold>\<sqinter> y) \<^bold>\<squnion> (x \<^bold>\<sqinter> y)"
- by simp
- then have "(x \<^bold>\<sqinter> a) \<^bold>\<squnion> (x \<^bold>\<sqinter> y) = (y \<^bold>\<sqinter> a) \<^bold>\<squnion> (y \<^bold>\<sqinter> x)"
- by (simp add: ac_simps)
- then have "x \<^bold>\<sqinter> (a \<^bold>\<squnion> y) = y \<^bold>\<sqinter> (a \<^bold>\<squnion> x)"
- by (simp add: conj_disj_distrib)
- with 2 4 have "x \<^bold>\<sqinter> \<one> = y \<^bold>\<sqinter> \<one>"
- by simp
- then show "x = y"
- by simp
-qed
-
-lemma compl_unique: "x \<^bold>\<sqinter> y = \<zero> \<Longrightarrow> x \<^bold>\<squnion> y = \<one> \<Longrightarrow> \<sim> x = y"
- by (rule complement_unique [OF conj_cancel_right disj_cancel_right])
-
-lemma double_compl [simp]: "\<sim> (\<sim> x) = x"
-proof (rule compl_unique)
- show "\<sim> x \<^bold>\<sqinter> x = \<zero>"
- by (simp only: conj_cancel_right conj.commute)
- show "\<sim> x \<^bold>\<squnion> x = \<one>"
- by (simp only: disj_cancel_right disj.commute)
-qed
-
-lemma compl_eq_compl_iff [simp]: "\<sim> x = \<sim> y \<longleftrightarrow> x = y"
- by (rule inj_eq [OF inj_on_inverseI]) (rule double_compl)
-
-
-subsection \<open>Conjunction\<close>
-
-lemma conj_zero_right [simp]: "x \<^bold>\<sqinter> \<zero> = \<zero>"
- using conj.left_idem conj_cancel_right by fastforce
-
-lemma compl_one [simp]: "\<sim> \<one> = \<zero>"
- by (rule compl_unique [OF conj_zero_right disj_zero_right])
-
-lemma conj_zero_left [simp]: "\<zero> \<^bold>\<sqinter> x = \<zero>"
- by (subst conj.commute) (rule conj_zero_right)
-
-lemma conj_cancel_left [simp]: "\<sim> x \<^bold>\<sqinter> x = \<zero>"
- by (subst conj.commute) (rule conj_cancel_right)
-
-lemma conj_disj_distrib2: "(y \<^bold>\<squnion> z) \<^bold>\<sqinter> x = (y \<^bold>\<sqinter> x) \<^bold>\<squnion> (z \<^bold>\<sqinter> x)"
- by (simp only: conj.commute conj_disj_distrib)
-
-lemmas conj_disj_distribs = conj_disj_distrib conj_disj_distrib2
-
-lemma conj_assoc: "(x \<^bold>\<sqinter> y) \<^bold>\<sqinter> z = x \<^bold>\<sqinter> (y \<^bold>\<sqinter> z)"
- by (fact ac_simps)
-
-lemma conj_commute: "x \<^bold>\<sqinter> y = y \<^bold>\<sqinter> x"
- by (fact ac_simps)
-
-lemmas conj_left_commute = conj.left_commute
-lemmas conj_ac = conj.assoc conj.commute conj.left_commute
-
-lemma conj_one_left: "\<one> \<^bold>\<sqinter> x = x"
- by (fact conj.left_neutral)
-
-lemma conj_left_absorb: "x \<^bold>\<sqinter> (x \<^bold>\<sqinter> y) = x \<^bold>\<sqinter> y"
- by (fact conj.left_idem)
-
-lemma conj_absorb: "x \<^bold>\<sqinter> x = x"
- by (fact conj.idem)
-
-
-subsection \<open>Disjunction\<close>
-
-interpretation dual: boolean_algebra "(\<^bold>\<squnion>)" "(\<^bold>\<sqinter>)" compl \<one> \<zero>
- apply standard
- apply (rule disj_conj_distrib)
- apply (rule conj_disj_distrib)
- apply simp_all
- done
-
-lemma compl_zero [simp]: "\<sim> \<zero> = \<one>"
- by (fact dual.compl_one)
-
-lemma disj_one_right [simp]: "x \<^bold>\<squnion> \<one> = \<one>"
- by (fact dual.conj_zero_right)
-
-lemma disj_one_left [simp]: "\<one> \<^bold>\<squnion> x = \<one>"
- by (fact dual.conj_zero_left)
-
-lemma disj_cancel_left [simp]: "\<sim> x \<^bold>\<squnion> x = \<one>"
- by (rule dual.conj_cancel_left)
-
-lemma disj_conj_distrib2: "(y \<^bold>\<sqinter> z) \<^bold>\<squnion> x = (y \<^bold>\<squnion> x) \<^bold>\<sqinter> (z \<^bold>\<squnion> x)"
- by (rule dual.conj_disj_distrib2)
-
-lemmas disj_conj_distribs = disj_conj_distrib disj_conj_distrib2
-
-lemma disj_assoc: "(x \<^bold>\<squnion> y) \<^bold>\<squnion> z = x \<^bold>\<squnion> (y \<^bold>\<squnion> z)"
- by (fact ac_simps)
-
-lemma disj_commute: "x \<^bold>\<squnion> y = y \<^bold>\<squnion> x"
- by (fact ac_simps)
-
-lemmas disj_left_commute = disj.left_commute
-
-lemmas disj_ac = disj.assoc disj.commute disj.left_commute
-
-lemma disj_zero_left: "\<zero> \<^bold>\<squnion> x = x"
- by (fact disj.left_neutral)
-
-lemma disj_left_absorb: "x \<^bold>\<squnion> (x \<^bold>\<squnion> y) = x \<^bold>\<squnion> y"
- by (fact disj.left_idem)
-
-lemma disj_absorb: "x \<^bold>\<squnion> x = x"
- by (fact disj.idem)
-
-
-subsection \<open>De Morgan's Laws\<close>
-
-lemma de_Morgan_conj [simp]: "\<sim> (x \<^bold>\<sqinter> y) = \<sim> x \<^bold>\<squnion> \<sim> y"
-proof (rule compl_unique)
- have "(x \<^bold>\<sqinter> y) \<^bold>\<sqinter> (\<sim> x \<^bold>\<squnion> \<sim> y) = ((x \<^bold>\<sqinter> y) \<^bold>\<sqinter> \<sim> x) \<^bold>\<squnion> ((x \<^bold>\<sqinter> y) \<^bold>\<sqinter> \<sim> y)"
- by (rule conj_disj_distrib)
- also have "\<dots> = (y \<^bold>\<sqinter> (x \<^bold>\<sqinter> \<sim> x)) \<^bold>\<squnion> (x \<^bold>\<sqinter> (y \<^bold>\<sqinter> \<sim> y))"
- by (simp only: conj_ac)
- finally show "(x \<^bold>\<sqinter> y) \<^bold>\<sqinter> (\<sim> x \<^bold>\<squnion> \<sim> y) = \<zero>"
- by (simp only: conj_cancel_right conj_zero_right disj_zero_right)
-next
- have "(x \<^bold>\<sqinter> y) \<^bold>\<squnion> (\<sim> x \<^bold>\<squnion> \<sim> y) = (x \<^bold>\<squnion> (\<sim> x \<^bold>\<squnion> \<sim> y)) \<^bold>\<sqinter> (y \<^bold>\<squnion> (\<sim> x \<^bold>\<squnion> \<sim> y))"
- by (rule disj_conj_distrib2)
- also have "\<dots> = (\<sim> y \<^bold>\<squnion> (x \<^bold>\<squnion> \<sim> x)) \<^bold>\<sqinter> (\<sim> x \<^bold>\<squnion> (y \<^bold>\<squnion> \<sim> y))"
- by (simp only: disj_ac)
- finally show "(x \<^bold>\<sqinter> y) \<^bold>\<squnion> (\<sim> x \<^bold>\<squnion> \<sim> y) = \<one>"
- by (simp only: disj_cancel_right disj_one_right conj_one_right)
-qed
-
-lemma de_Morgan_disj [simp]: "\<sim> (x \<^bold>\<squnion> y) = \<sim> x \<^bold>\<sqinter> \<sim> y"
- using dual.boolean_algebra_axioms by (rule boolean_algebra.de_Morgan_conj)
-
-
-subsection \<open>Symmetric Difference\<close>
-
-definition xor :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<oplus>" 65)
- where "x \<oplus> y = (x \<^bold>\<sqinter> \<sim> y) \<^bold>\<squnion> (\<sim> x \<^bold>\<sqinter> y)"
-
-sublocale xor: comm_monoid xor \<zero>
-proof
- fix x y z :: 'a
- let ?t = "(x \<^bold>\<sqinter> y \<^bold>\<sqinter> z) \<^bold>\<squnion> (x \<^bold>\<sqinter> \<sim> y \<^bold>\<sqinter> \<sim> z) \<^bold>\<squnion> (\<sim> x \<^bold>\<sqinter> y \<^bold>\<sqinter> \<sim> z) \<^bold>\<squnion> (\<sim> x \<^bold>\<sqinter> \<sim> y \<^bold>\<sqinter> z)"
- have "?t \<^bold>\<squnion> (z \<^bold>\<sqinter> x \<^bold>\<sqinter> \<sim> x) \<^bold>\<squnion> (z \<^bold>\<sqinter> y \<^bold>\<sqinter> \<sim> y) = ?t \<^bold>\<squnion> (x \<^bold>\<sqinter> y \<^bold>\<sqinter> \<sim> y) \<^bold>\<squnion> (x \<^bold>\<sqinter> z \<^bold>\<sqinter> \<sim> z)"
- by (simp only: conj_cancel_right conj_zero_right)
- then show "(x \<oplus> y) \<oplus> z = x \<oplus> (y \<oplus> z)"
- by (simp only: xor_def de_Morgan_disj de_Morgan_conj double_compl)
- (simp only: conj_disj_distribs conj_ac disj_ac)
- show "x \<oplus> y = y \<oplus> x"
- by (simp only: xor_def conj_commute disj_commute)
- show "x \<oplus> \<zero> = x"
- by (simp add: xor_def)
-qed
-
-lemmas xor_assoc = xor.assoc
-lemmas xor_commute = xor.commute
-lemmas xor_left_commute = xor.left_commute
-
-lemmas xor_ac = xor.assoc xor.commute xor.left_commute
-
-lemma xor_def2: "x \<oplus> y = (x \<^bold>\<squnion> y) \<^bold>\<sqinter> (\<sim> x \<^bold>\<squnion> \<sim> y)"
- using conj.commute conj_disj_distrib2 disj.commute xor_def by auto
-
-lemma xor_zero_right [simp]: "x \<oplus> \<zero> = x"
- by (simp only: xor_def compl_zero conj_one_right conj_zero_right disj_zero_right)
-
-lemma xor_zero_left [simp]: "\<zero> \<oplus> x = x"
- by (subst xor_commute) (rule xor_zero_right)
-
-lemma xor_one_right [simp]: "x \<oplus> \<one> = \<sim> x"
- by (simp only: xor_def compl_one conj_zero_right conj_one_right disj_zero_left)
-
-lemma xor_one_left [simp]: "\<one> \<oplus> x = \<sim> x"
- by (subst xor_commute) (rule xor_one_right)
-
-lemma xor_self [simp]: "x \<oplus> x = \<zero>"
- by (simp only: xor_def conj_cancel_right conj_cancel_left disj_zero_right)
-
-lemma xor_left_self [simp]: "x \<oplus> (x \<oplus> y) = y"
- by (simp only: xor_assoc [symmetric] xor_self xor_zero_left)
-
-lemma xor_compl_left [simp]: "\<sim> x \<oplus> y = \<sim> (x \<oplus> y)"
- by (metis xor_assoc xor_one_left)
-
-lemma xor_compl_right [simp]: "x \<oplus> \<sim> y = \<sim> (x \<oplus> y)"
- using xor_commute xor_compl_left by auto
-
-lemma xor_cancel_right: "x \<oplus> \<sim> x = \<one>"
- by (simp only: xor_compl_right xor_self compl_zero)
-
-lemma xor_cancel_left: "\<sim> x \<oplus> x = \<one>"
- by (simp only: xor_compl_left xor_self compl_zero)
-
-lemma conj_xor_distrib: "x \<^bold>\<sqinter> (y \<oplus> z) = (x \<^bold>\<sqinter> y) \<oplus> (x \<^bold>\<sqinter> z)"
-proof -
- have *: "(x \<^bold>\<sqinter> y \<^bold>\<sqinter> \<sim> z) \<^bold>\<squnion> (x \<^bold>\<sqinter> \<sim> y \<^bold>\<sqinter> z) =
- (y \<^bold>\<sqinter> x \<^bold>\<sqinter> \<sim> x) \<^bold>\<squnion> (z \<^bold>\<sqinter> x \<^bold>\<sqinter> \<sim> x) \<^bold>\<squnion> (x \<^bold>\<sqinter> y \<^bold>\<sqinter> \<sim> z) \<^bold>\<squnion> (x \<^bold>\<sqinter> \<sim> y \<^bold>\<sqinter> z)"
- by (simp only: conj_cancel_right conj_zero_right disj_zero_left)
- then show "x \<^bold>\<sqinter> (y \<oplus> z) = (x \<^bold>\<sqinter> y) \<oplus> (x \<^bold>\<sqinter> z)"
- by (simp (no_asm_use) only:
- xor_def de_Morgan_disj de_Morgan_conj double_compl
- conj_disj_distribs conj_ac disj_ac)
-qed
-
-lemma conj_xor_distrib2: "(y \<oplus> z) \<^bold>\<sqinter> x = (y \<^bold>\<sqinter> x) \<oplus> (z \<^bold>\<sqinter> x)"
- by (simp add: conj.commute conj_xor_distrib)
-
-lemmas conj_xor_distribs = conj_xor_distrib conj_xor_distrib2
-
-end
-
-end
--- a/src/HOL/Library/Library.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/Library/Library.thy Mon Aug 02 10:01:06 2021 +0000
@@ -4,10 +4,8 @@
AList
Adhoc_Overloading
BigO
- Bit_Operations
BNF_Axiomatization
BNF_Corec
- Boolean_Algebra
Bourbaki_Witt_Fixpoint
Char_ord
Code_Cardinality
--- a/src/HOL/Library/Linear_Temporal_Logic_on_Streams.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/Library/Linear_Temporal_Logic_on_Streams.thy Mon Aug 02 10:01:06 2021 +0000
@@ -39,7 +39,7 @@
abbreviation (input) AND (infix "aand" 60)
where "\<phi> aand \<psi> \<equiv> \<lambda> xs. \<phi> xs \<and> \<psi> xs"
-abbreviation (input) "not \<phi> \<equiv> \<lambda> xs. \<not> \<phi> xs"
+abbreviation (input) not where "not \<phi> \<equiv> \<lambda> xs. \<not> \<phi> xs"
abbreviation (input) "true \<equiv> \<lambda> xs. True"
--- a/src/HOL/Library/Word.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/Library/Word.thy Mon Aug 02 10:01:06 2021 +0000
@@ -7,8 +7,6 @@
theory Word
imports
"HOL-Library.Type_Length"
- "HOL-Library.Boolean_Algebra"
- "HOL-Library.Bit_Operations"
begin
subsection \<open>Preliminaries\<close>
@@ -1131,18 +1129,18 @@
with \<open>2 ^ m \<noteq> 0\<close> have \<open>2 ^ (m - n) \<noteq> 0\<close>
by (metis (full_types) diff_add exp_add_not_zero_imp)
with True show ?thesis
- by (simp add: bit_unsigned_iff bit_push_bit_iff Parity.bit_push_bit_iff bit_take_bit_iff not_le exp_eq_zero_iff ac_simps)
+ by (simp add: bit_unsigned_iff bit_push_bit_iff Bit_Operations.bit_push_bit_iff bit_take_bit_iff not_le ac_simps)
next
case False
then show ?thesis
- by (simp add: not_le bit_unsigned_iff bit_push_bit_iff Parity.bit_push_bit_iff bit_take_bit_iff)
+ by (simp add: not_le bit_unsigned_iff bit_push_bit_iff Bit_Operations.bit_push_bit_iff bit_take_bit_iff)
qed
qed
lemma unsigned_take_bit_eq:
\<open>unsigned (take_bit n w) = take_bit n (unsigned w)\<close>
for w :: \<open>'b::len word\<close>
- by (rule bit_eqI) (simp add: bit_unsigned_iff bit_take_bit_iff Parity.bit_take_bit_iff)
+ by (rule bit_eqI) (simp add: bit_unsigned_iff bit_take_bit_iff Bit_Operations.bit_take_bit_iff)
end
@@ -1152,7 +1150,7 @@
lemma unsigned_drop_bit_eq:
\<open>unsigned (drop_bit n w) = drop_bit n (take_bit LENGTH('b) (unsigned w))\<close>
for w :: \<open>'b::len word\<close>
- by (rule bit_eqI) (auto simp add: bit_unsigned_iff bit_take_bit_iff bit_drop_bit_eq Parity.bit_drop_bit_eq dest: bit_imp_le_length)
+ by (rule bit_eqI) (auto simp add: bit_unsigned_iff bit_take_bit_iff bit_drop_bit_eq Bit_Operations.bit_drop_bit_eq dest: bit_imp_le_length)
end
@@ -1198,7 +1196,7 @@
\<open>unsigned (NOT w) = take_bit LENGTH('b) (NOT (unsigned w))\<close>
for w :: \<open>'b::len word\<close>
by (rule bit_eqI)
- (simp add: bit_unsigned_iff bit_take_bit_iff bit_not_iff Bit_Operations.bit_not_iff exp_eq_zero_iff not_le)
+ (simp add: bit_unsigned_iff bit_take_bit_iff bit_not_iff Bit_Operations.bit_not_iff not_le)
end
@@ -1264,15 +1262,14 @@
moreover from \<open>2 ^ q \<noteq> 0\<close> have \<open>2 ^ (q - n) \<noteq> 0\<close>
by (rule exp_not_zero_imp_exp_diff_not_zero)
ultimately show ?thesis
- by (auto simp add: bit_signed_iff bit_signed_take_bit_iff bit_push_bit_iff Parity.bit_push_bit_iff
- min_def * exp_eq_zero_iff le_diff_conv2)
+ by (auto simp add: bit_signed_iff bit_signed_take_bit_iff bit_push_bit_iff Bit_Operations.bit_push_bit_iff
+ min_def * le_diff_conv2)
next
case False
then show ?thesis
using exp_not_zero_imp_exp_diff_not_zero [of m n]
- by (auto simp add: bit_signed_iff bit_signed_take_bit_iff bit_push_bit_iff Parity.bit_push_bit_iff
- min_def not_le not_less * le_diff_conv2 less_diff_conv2 Parity.exp_eq_0_imp_not_bit exp_eq_0_imp_not_bit
- exp_eq_zero_iff)
+ by (auto simp add: bit_signed_iff bit_signed_take_bit_iff bit_push_bit_iff Bit_Operations.bit_push_bit_iff
+ min_def not_le not_less * le_diff_conv2 less_diff_conv2 Bit_Operations.exp_eq_0_imp_not_bit exp_eq_0_imp_not_bit)
qed
qed
@@ -1302,13 +1299,11 @@
have \<open>2 ^ Suc q \<noteq> 0\<close>
using exp_add_not_zero_imp_right by blast
ultimately show ?thesis
- by (simp add: * bit_signed_iff bit_not_iff bit_signed_take_bit_iff Bit_Operations.bit_not_iff min_def
- exp_eq_zero_iff)
+ by (simp add: * bit_signed_iff bit_not_iff bit_signed_take_bit_iff Bit_Operations.bit_not_iff min_def)
next
case False
then show ?thesis
- by (auto simp add: * bit_signed_iff bit_not_iff bit_signed_take_bit_iff Bit_Operations.bit_not_iff min_def
- exp_eq_zero_iff)
+ by (auto simp add: * bit_signed_iff bit_not_iff bit_signed_take_bit_iff Bit_Operations.bit_not_iff min_def)
qed
qed
@@ -1411,7 +1406,7 @@
lemma unsigned_ucast_eq:
\<open>unsigned (ucast w :: 'c::len word) = take_bit LENGTH('c) (unsigned w)\<close>
for w :: \<open>'b::len word\<close>
- by (rule bit_eqI) (simp add: bit_unsigned_iff Word.bit_unsigned_iff bit_take_bit_iff exp_eq_zero_iff not_le)
+ by (rule bit_eqI) (simp add: bit_unsigned_iff Word.bit_unsigned_iff bit_take_bit_iff not_le)
end
@@ -1428,7 +1423,7 @@
by (simp add: min_def)
(metis (mono_tags) diff_diff_cancel exp_not_zero_imp_exp_diff_not_zero)
then show \<open>bit (signed (ucast w :: 'c::len word)) n \<longleftrightarrow> bit (signed_take_bit (LENGTH('c) - Suc 0) (unsigned w)) n\<close>
- by (simp add: bit_signed_iff bit_unsigned_iff Word.bit_unsigned_iff bit_signed_take_bit_iff exp_eq_zero_iff not_le)
+ by (simp add: bit_signed_iff bit_unsigned_iff Word.bit_unsigned_iff bit_signed_take_bit_iff not_le)
qed
lemma signed_scast_eq:
@@ -1441,7 +1436,7 @@
by (simp add: min_def)
(metis (mono_tags) diff_diff_cancel exp_not_zero_imp_exp_diff_not_zero)
then show \<open>bit (signed (scast w :: 'c::len word)) n \<longleftrightarrow> bit (signed_take_bit (LENGTH('c) - Suc 0) (signed w)) n\<close>
- by (simp add: bit_signed_iff bit_unsigned_iff Word.bit_signed_iff bit_signed_take_bit_iff exp_eq_zero_iff not_le)
+ by (simp add: bit_signed_iff bit_unsigned_iff Word.bit_signed_iff bit_signed_take_bit_iff not_le)
qed
end
@@ -2104,8 +2099,7 @@
(drop_bit (n mod LENGTH('a)) (uint w))
(uint (take_bit (n mod LENGTH('a)) w))\<close>
for w :: \<open>'a::len word\<close>
- apply transfer
- by (simp add: min.absorb2 take_bit_concat_bit_eq)
+ by transfer (simp add: take_bit_concat_bit_eq)
lemma [code]:
\<open>Word.the_int (word_rotr n w) = concat_bit (LENGTH('a) - n mod LENGTH('a))
@@ -2347,7 +2341,7 @@
\<close>
lemma bit_ucast_iff:
- \<open>bit (ucast a :: 'a::len word) n \<longleftrightarrow> n < LENGTH('a::len) \<and> Parity.bit a n\<close>
+ \<open>bit (ucast a :: 'a::len word) n \<longleftrightarrow> n < LENGTH('a::len) \<and> bit a n\<close>
by transfer (simp add: bit_take_bit_iff)
lemma ucast_id [simp]: "ucast w = w"
@@ -2358,7 +2352,7 @@
lemma ucast_mask_eq:
\<open>ucast (mask n :: 'b word) = mask (min LENGTH('b::len) n)\<close>
- by (simp add: bit_eq_iff) (auto simp add: bit_mask_iff bit_ucast_iff exp_eq_zero_iff)
+ by (simp add: bit_eq_iff) (auto simp add: bit_mask_iff bit_ucast_iff)
\<comment> \<open>literal u(s)cast\<close>
lemma ucast_bintr [simp]:
@@ -2491,7 +2485,7 @@
then show ?thesis
apply transfer
apply (simp add: take_bit_drop_bit)
- by (simp add: bit_iff_odd_drop_bit drop_bit_take_bit min.absorb2 odd_iff_mod_2_eq_one)
+ by (simp add: bit_iff_odd_drop_bit drop_bit_take_bit odd_iff_mod_2_eq_one)
qed auto
@@ -3623,7 +3617,7 @@
lemma minus_1_eq_mask:
\<open>- 1 = (mask LENGTH('a) :: 'a::len word)\<close>
- by (rule bit_eqI) (simp add: bit_exp_iff bit_mask_iff exp_eq_zero_iff)
+ by (rule bit_eqI) (simp add: bit_exp_iff bit_mask_iff)
lemma mask_eq_decr_exp:
\<open>mask n = 2 ^ n - (1 :: 'a::len word)\<close>
@@ -3638,7 +3632,7 @@
qualified lemma bit_mask_iff [bit_simps]:
\<open>bit (mask m :: 'a::len word) n \<longleftrightarrow> n < min LENGTH('a) m\<close>
- by (simp add: bit_mask_iff exp_eq_zero_iff not_le)
+ by (simp add: bit_mask_iff not_le)
end
@@ -4218,10 +4212,10 @@
by (metis add.commute diff_add_cancel word_rec_Suc)
lemma word_rec_in: "f (word_rec z (\<lambda>_. f) n) = word_rec (f z) (\<lambda>_. f) n"
- by (induct n) (simp_all add: word_rec_Suc)
+ by (induct n) simp_all
lemma word_rec_in2: "f n (word_rec z f n) = word_rec (f 0 z) (f \<circ> (+) 1) n"
- by (induct n) (simp_all add: word_rec_Suc)
+ by (induct n) simp_all
lemma word_rec_twice:
"m \<le> n \<Longrightarrow> word_rec z f n = word_rec (word_rec z f (n - m)) (f \<circ> (+) (n - m)) m"
--- a/src/HOL/Library/Z2.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/Library/Z2.thy Mon Aug 02 10:01:06 2021 +0000
@@ -5,11 +5,11 @@
section \<open>The Field of Integers mod 2\<close>
theory Z2
-imports Main "HOL-Library.Bit_Operations"
+imports Main
begin
text \<open>
- Note that in most cases \<^typ>\<open>bool\<close> is appropriate hen a binary type is needed; the
+ Note that in most cases \<^typ>\<open>bool\<close> is appropriate when a binary type is needed; the
type provided here, for historical reasons named \<^text>\<open>bit\<close>, is only needed if proper
field operations are required.
\<close>
--- a/src/HOL/List.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/List.thy Mon Aug 02 10:01:06 2021 +0000
@@ -5,7 +5,7 @@
section \<open>The datatype of finite lists\<close>
theory List
-imports Sledgehammer Code_Numeral Lifting_Set
+imports Sledgehammer Lifting_Set
begin
datatype (set: 'a) list =
--- a/src/HOL/Main.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/Main.thy Mon Aug 02 10:01:06 2021 +0000
@@ -71,4 +71,6 @@
"_SUP1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_./ _)" [0, 10] 10)
"_SUP" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
+thy_deps
+
end
--- a/src/HOL/Nominal/Examples/Class1.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/Nominal/Examples/Class1.thy Mon Aug 02 10:01:06 2021 +0000
@@ -8,6 +8,8 @@
text \<open>types\<close>
+no_notation not ("NOT")
+
nominal_datatype ty =
PR "string"
| NOT "ty"
--- a/src/HOL/Numeral_Simprocs.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/Numeral_Simprocs.thy Mon Aug 02 10:01:06 2021 +0000
@@ -299,14 +299,4 @@
Numeral_Simprocs.field_divide_cancel_numeral_factor])
\<close>
-lemma bit_numeral_int_simps [simp]:
- \<open>bit (1 :: int) (numeral n) \<longleftrightarrow> bit (0 :: int) (pred_numeral n)\<close>
- \<open>bit (numeral (num.Bit0 w) :: int) (numeral n) \<longleftrightarrow> bit (numeral w :: int) (pred_numeral n)\<close>
- \<open>bit (numeral (num.Bit1 w) :: int) (numeral n) \<longleftrightarrow> bit (numeral w :: int) (pred_numeral n)\<close>
- \<open>bit (numeral (Num.BitM w) :: int) (numeral n) \<longleftrightarrow> \<not> bit (- numeral w :: int) (pred_numeral n)\<close>
- \<open>bit (- numeral (num.Bit0 w) :: int) (numeral n) \<longleftrightarrow> bit (- numeral w :: int) (pred_numeral n)\<close>
- \<open>bit (- numeral (num.Bit1 w) :: int) (numeral n) \<longleftrightarrow> \<not> bit (numeral w :: int) (pred_numeral n)\<close>
- \<open>bit (- numeral (Num.BitM w) :: int) (numeral n) \<longleftrightarrow> bit (- (numeral w) :: int) (pred_numeral n)\<close>
- by (simp_all add: bit_1_iff numeral_eq_Suc bit_Suc add_One sub_inc_One_eq bit_minus_int_iff)
-
end
--- a/src/HOL/Parity.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/Parity.thy Mon Aug 02 10:01:06 2021 +0000
@@ -699,1333 +699,6 @@
end
-
-subsection \<open>Abstract bit structures\<close>
-
-class semiring_bits = semiring_parity +
- assumes bits_induct [case_names stable rec]:
- \<open>(\<And>a. a div 2 = a \<Longrightarrow> P a)
- \<Longrightarrow> (\<And>a b. P a \<Longrightarrow> (of_bool b + 2 * a) div 2 = a \<Longrightarrow> P (of_bool b + 2 * a))
- \<Longrightarrow> P a\<close>
- assumes bits_div_0 [simp]: \<open>0 div a = 0\<close>
- and bits_div_by_1 [simp]: \<open>a div 1 = a\<close>
- and bits_mod_div_trivial [simp]: \<open>a mod b div b = 0\<close>
- and even_succ_div_2 [simp]: \<open>even a \<Longrightarrow> (1 + a) div 2 = a div 2\<close>
- and even_mask_div_iff: \<open>even ((2 ^ m - 1) div 2 ^ n) \<longleftrightarrow> 2 ^ n = 0 \<or> m \<le> n\<close>
- and exp_div_exp_eq: \<open>2 ^ m div 2 ^ n = of_bool (2 ^ m \<noteq> 0 \<and> m \<ge> n) * 2 ^ (m - n)\<close>
- and div_exp_eq: \<open>a div 2 ^ m div 2 ^ n = a div 2 ^ (m + n)\<close>
- and mod_exp_eq: \<open>a mod 2 ^ m mod 2 ^ n = a mod 2 ^ min m n\<close>
- and mult_exp_mod_exp_eq: \<open>m \<le> n \<Longrightarrow> (a * 2 ^ m) mod (2 ^ n) = (a mod 2 ^ (n - m)) * 2 ^ m\<close>
- and div_exp_mod_exp_eq: \<open>a div 2 ^ n mod 2 ^ m = a mod (2 ^ (n + m)) div 2 ^ n\<close>
- and even_mult_exp_div_exp_iff: \<open>even (a * 2 ^ m div 2 ^ n) \<longleftrightarrow> m > n \<or> 2 ^ n = 0 \<or> (m \<le> n \<and> even (a div 2 ^ (n - m)))\<close>
- fixes bit :: \<open>'a \<Rightarrow> nat \<Rightarrow> bool\<close>
- assumes bit_iff_odd: \<open>bit a n \<longleftrightarrow> odd (a div 2 ^ n)\<close>
-begin
-
-text \<open>
- Having \<^const>\<open>bit\<close> as definitional class operation
- takes into account that specific instances can be implemented
- differently wrt. code generation.
-\<close>
-
-lemma bits_div_by_0 [simp]:
- \<open>a div 0 = 0\<close>
- by (metis add_cancel_right_right bits_mod_div_trivial mod_mult_div_eq mult_not_zero)
-
-lemma bits_1_div_2 [simp]:
- \<open>1 div 2 = 0\<close>
- using even_succ_div_2 [of 0] by simp
-
-lemma bits_1_div_exp [simp]:
- \<open>1 div 2 ^ n = of_bool (n = 0)\<close>
- using div_exp_eq [of 1 1] by (cases n) simp_all
-
-lemma even_succ_div_exp [simp]:
- \<open>(1 + a) div 2 ^ n = a div 2 ^ n\<close> if \<open>even a\<close> and \<open>n > 0\<close>
-proof (cases n)
- case 0
- with that show ?thesis
- by simp
-next
- case (Suc n)
- with \<open>even a\<close> have \<open>(1 + a) div 2 ^ Suc n = a div 2 ^ Suc n\<close>
- proof (induction n)
- case 0
- then show ?case
- by simp
- next
- case (Suc n)
- then show ?case
- using div_exp_eq [of _ 1 \<open>Suc n\<close>, symmetric]
- by simp
- qed
- with Suc show ?thesis
- by simp
-qed
-
-lemma even_succ_mod_exp [simp]:
- \<open>(1 + a) mod 2 ^ n = 1 + (a mod 2 ^ n)\<close> if \<open>even a\<close> and \<open>n > 0\<close>
- using div_mult_mod_eq [of \<open>1 + a\<close> \<open>2 ^ n\<close>] that
- apply simp
- by (metis local.add.left_commute local.add_left_cancel local.div_mult_mod_eq)
-
-lemma bits_mod_by_1 [simp]:
- \<open>a mod 1 = 0\<close>
- using div_mult_mod_eq [of a 1] by simp
-
-lemma bits_mod_0 [simp]:
- \<open>0 mod a = 0\<close>
- using div_mult_mod_eq [of 0 a] by simp
-
-lemma bits_one_mod_two_eq_one [simp]:
- \<open>1 mod 2 = 1\<close>
- by (simp add: mod2_eq_if)
-
-lemma bit_0 [simp]:
- \<open>bit a 0 \<longleftrightarrow> odd a\<close>
- by (simp add: bit_iff_odd)
-
-lemma bit_Suc:
- \<open>bit a (Suc n) \<longleftrightarrow> bit (a div 2) n\<close>
- using div_exp_eq [of a 1 n] by (simp add: bit_iff_odd)
-
-lemma bit_rec:
- \<open>bit a n \<longleftrightarrow> (if n = 0 then odd a else bit (a div 2) (n - 1))\<close>
- by (cases n) (simp_all add: bit_Suc)
-
-lemma bit_0_eq [simp]:
- \<open>bit 0 = bot\<close>
- by (simp add: fun_eq_iff bit_iff_odd)
-
-context
- fixes a
- assumes stable: \<open>a div 2 = a\<close>
-begin
-
-lemma bits_stable_imp_add_self:
- \<open>a + a mod 2 = 0\<close>
-proof -
- have \<open>a div 2 * 2 + a mod 2 = a\<close>
- by (fact div_mult_mod_eq)
- then have \<open>a * 2 + a mod 2 = a\<close>
- by (simp add: stable)
- then show ?thesis
- by (simp add: mult_2_right ac_simps)
-qed
-
-lemma stable_imp_bit_iff_odd:
- \<open>bit a n \<longleftrightarrow> odd a\<close>
- by (induction n) (simp_all add: stable bit_Suc)
-
-end
-
-lemma bit_iff_idd_imp_stable:
- \<open>a div 2 = a\<close> if \<open>\<And>n. bit a n \<longleftrightarrow> odd a\<close>
-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]: \<open>b = odd a\<close>
- by (simp add: rec.hyps bit_Suc)
- from rec.hyps have hyp: \<open>(of_bool (odd a) + 2 * a) div 2 = a\<close>
- by simp
- have \<open>bit a n \<longleftrightarrow> odd a\<close> for n
- using rec.prems [of \<open>Suc n\<close>] by (simp add: hyp bit_Suc)
- then have \<open>a div 2 = a\<close>
- by (rule rec.IH)
- then have \<open>of_bool (odd a) + 2 * a = 2 * (a div 2) + of_bool (odd a)\<close>
- by (simp add: ac_simps)
- also have \<open>\<dots> = a\<close>
- using mult_div_mod_eq [of 2 a]
- by (simp add: of_bool_odd_eq_mod_2)
- finally show ?case
- using \<open>a div 2 = a\<close> by (simp add: hyp)
-qed
-
-lemma exp_eq_0_imp_not_bit:
- \<open>\<not> bit a n\<close> if \<open>2 ^ n = 0\<close>
- using that by (simp add: bit_iff_odd)
-
-lemma bit_eqI:
- \<open>a = b\<close> if \<open>\<And>n. 2 ^ n \<noteq> 0 \<Longrightarrow> bit a n \<longleftrightarrow> bit b n\<close>
-proof -
- have \<open>bit a n \<longleftrightarrow> bit b n\<close> for n
- proof (cases \<open>2 ^ n = 0\<close>)
- 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 **: \<open>even b \<longleftrightarrow> even a\<close>
- by simp
- have \<open>b div 2 = b\<close>
- proof (rule bit_iff_idd_imp_stable)
- fix n
- from stable have *: \<open>bit b n \<longleftrightarrow> bit a n\<close>
- by simp
- also have \<open>bit a n \<longleftrightarrow> odd a\<close>
- using stable by (simp add: stable_imp_bit_iff_odd)
- finally show \<open>bit b n \<longleftrightarrow> odd b\<close>
- by (simp add: **)
- qed
- from ** have \<open>a mod 2 = b mod 2\<close>
- by (simp add: mod2_eq_if)
- then have \<open>a mod 2 + (a + b) = b mod 2 + (a + b)\<close>
- by simp
- then have \<open>a + a mod 2 + b = b + b mod 2 + a\<close>
- by (simp add: ac_simps)
- with \<open>a div 2 = a\<close> \<open>b div 2 = b\<close> show ?case
- by (simp add: bits_stable_imp_add_self)
- next
- case (rec a p)
- from rec.prems [of 0] have [simp]: \<open>p = odd b\<close>
- by simp
- from rec.hyps have \<open>bit a n \<longleftrightarrow> bit (b div 2) n\<close> for n
- using rec.prems [of \<open>Suc n\<close>] by (simp add: bit_Suc)
- then have \<open>a = b div 2\<close>
- by (rule rec.IH)
- then have \<open>2 * a = 2 * (b div 2)\<close>
- by simp
- then have \<open>b mod 2 + 2 * a = b mod 2 + 2 * (b div 2)\<close>
- by simp
- also have \<open>\<dots> = b\<close>
- by (fact mod_mult_div_eq)
- finally show ?case
- by (auto simp add: mod2_eq_if)
- qed
-qed
-
-lemma bit_eq_iff:
- \<open>a = b \<longleftrightarrow> (\<forall>n. bit a n \<longleftrightarrow> bit b n)\<close>
- by (auto intro: bit_eqI)
-
-named_theorems bit_simps \<open>Simplification rules for \<^const>\<open>bit\<close>\<close>
-
-lemma bit_exp_iff [bit_simps]:
- \<open>bit (2 ^ m) n \<longleftrightarrow> 2 ^ m \<noteq> 0 \<and> m = n\<close>
- by (auto simp add: bit_iff_odd exp_div_exp_eq)
-
-lemma bit_1_iff [bit_simps]:
- \<open>bit 1 n \<longleftrightarrow> 1 \<noteq> 0 \<and> n = 0\<close>
- using bit_exp_iff [of 0 n] by simp
-
-lemma bit_2_iff [bit_simps]:
- \<open>bit 2 n \<longleftrightarrow> 2 \<noteq> 0 \<and> n = 1\<close>
- using bit_exp_iff [of 1 n] by auto
-
-lemma even_bit_succ_iff:
- \<open>bit (1 + a) n \<longleftrightarrow> bit a n \<or> n = 0\<close> if \<open>even a\<close>
- using that by (cases \<open>n = 0\<close>) (simp_all add: bit_iff_odd)
-
-lemma odd_bit_iff_bit_pred:
- \<open>bit a n \<longleftrightarrow> bit (a - 1) n \<or> n = 0\<close> if \<open>odd a\<close>
-proof -
- from \<open>odd a\<close> obtain b where \<open>a = 2 * b + 1\<close> ..
- moreover have \<open>bit (2 * b) n \<or> n = 0 \<longleftrightarrow> bit (1 + 2 * b) n\<close>
- using even_bit_succ_iff by simp
- ultimately show ?thesis by (simp add: ac_simps)
-qed
-
-lemma bit_double_iff [bit_simps]:
- \<open>bit (2 * a) n \<longleftrightarrow> bit a (n - 1) \<and> n \<noteq> 0 \<and> 2 ^ n \<noteq> 0\<close>
- 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:
- \<open>a = b \<longleftrightarrow> (even a \<longleftrightarrow> even b) \<and> a div 2 = b div 2\<close> (is \<open>?P = ?Q\<close>)
-proof
- assume ?P
- then show ?Q
- by simp
-next
- assume ?Q
- then have \<open>even a \<longleftrightarrow> even b\<close> and \<open>a div 2 = b div 2\<close>
- by simp_all
- show ?P
- proof (rule bit_eqI)
- fix n
- show \<open>bit a n \<longleftrightarrow> bit b n\<close>
- proof (cases n)
- case 0
- with \<open>even a \<longleftrightarrow> even b\<close> show ?thesis
- by simp
- next
- case (Suc n)
- moreover from \<open>a div 2 = b div 2\<close> have \<open>bit (a div 2) n = bit (b div 2) n\<close>
- by simp
- ultimately show ?thesis
- by (simp add: bit_Suc)
- qed
- qed
-qed
-
-lemma bit_mod_2_iff [simp]:
- \<open>bit (a mod 2) n \<longleftrightarrow> n = 0 \<and> odd a\<close>
- by (cases a rule: parity_cases) (simp_all add: bit_iff_odd)
-
-lemma bit_mask_iff:
- \<open>bit (2 ^ m - 1) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n < m\<close>
- by (simp add: bit_iff_odd even_mask_div_iff not_le)
-
-lemma bit_Numeral1_iff [simp]:
- \<open>bit (numeral Num.One) n \<longleftrightarrow> n = 0\<close>
- by (simp add: bit_rec)
-
-lemma exp_add_not_zero_imp:
- \<open>2 ^ m \<noteq> 0\<close> and \<open>2 ^ n \<noteq> 0\<close> if \<open>2 ^ (m + n) \<noteq> 0\<close>
-proof -
- have \<open>\<not> (2 ^ m = 0 \<or> 2 ^ n = 0)\<close>
- proof (rule notI)
- assume \<open>2 ^ m = 0 \<or> 2 ^ n = 0\<close>
- then have \<open>2 ^ (m + n) = 0\<close>
- by (rule disjE) (simp_all add: power_add)
- with that show False ..
- qed
- then show \<open>2 ^ m \<noteq> 0\<close> and \<open>2 ^ n \<noteq> 0\<close>
- by simp_all
-qed
-
-lemma bit_disjunctive_add_iff:
- \<open>bit (a + b) n \<longleftrightarrow> bit a n \<or> bit b n\<close>
- if \<open>\<And>n. \<not> bit a n \<or> \<not> bit b n\<close>
-proof (cases \<open>2 ^ n = 0\<close>)
- 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: \<open>even a \<or> even b\<close>
- by auto
- have bit: \<open>\<not> bit (a div 2) n \<or> \<not> bit (b div 2) n\<close> for n
- using Suc.prems(1) [of \<open>Suc n\<close>] by (simp add: bit_Suc)
- from Suc.prems(2) have \<open>2 * 2 ^ n \<noteq> 0\<close> \<open>2 ^ n \<noteq> 0\<close>
- by (auto simp add: mult_2)
- have \<open>a + b = (a div 2 * 2 + a mod 2) + (b div 2 * 2 + b mod 2)\<close>
- using div_mult_mod_eq [of a 2] div_mult_mod_eq [of b 2] by simp
- also have \<open>\<dots> = of_bool (odd a \<or> odd b) + 2 * (a div 2 + b div 2)\<close>
- using even by (auto simp add: algebra_simps mod2_eq_if)
- finally have \<open>bit ((a + b) div 2) n \<longleftrightarrow> bit (a div 2 + b div 2) n\<close>
- using \<open>2 * 2 ^ n \<noteq> 0\<close> by simp (simp_all flip: bit_Suc add: bit_double_iff)
- also have \<open>\<dots> \<longleftrightarrow> bit (a div 2) n \<or> bit (b div 2) n\<close>
- using bit \<open>2 ^ n \<noteq> 0\<close> by (rule Suc.IH)
- finally show ?case
- by (simp add: bit_Suc)
- qed
-qed
-
-lemma
- exp_add_not_zero_imp_left: \<open>2 ^ m \<noteq> 0\<close>
- and exp_add_not_zero_imp_right: \<open>2 ^ n \<noteq> 0\<close>
- if \<open>2 ^ (m + n) \<noteq> 0\<close>
-proof -
- have \<open>\<not> (2 ^ m = 0 \<or> 2 ^ n = 0)\<close>
- proof (rule notI)
- assume \<open>2 ^ m = 0 \<or> 2 ^ n = 0\<close>
- then have \<open>2 ^ (m + n) = 0\<close>
- by (rule disjE) (simp_all add: power_add)
- with that show False ..
- qed
- then show \<open>2 ^ m \<noteq> 0\<close> and \<open>2 ^ n \<noteq> 0\<close>
- by simp_all
-qed
-
-lemma exp_not_zero_imp_exp_diff_not_zero:
- \<open>2 ^ (n - m) \<noteq> 0\<close> if \<open>2 ^ n \<noteq> 0\<close>
-proof (cases \<open>m \<le> n\<close>)
- case True
- moreover define q where \<open>q = n - m\<close>
- ultimately have \<open>n = m + q\<close>
- 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: "\<And>n. P n \<Longrightarrow> n > 0 \<Longrightarrow> P (2 * n)"
- and odd: "\<And>n. P n \<Longrightarrow> 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 \<noteq> 1"
- by auto
- with \<open>n \<noteq> 0\<close> have "n div 2 > 0"
- by simp
- with \<open>even n\<close> 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 :: \<open>nat \<Rightarrow> nat \<Rightarrow> bool\<close>
- where \<open>bit_nat m n \<longleftrightarrow> odd (m div 2 ^ n)\<close>
-
-instance
-proof
- show \<open>P n\<close> if stable: \<open>\<And>n. n div 2 = n \<Longrightarrow> P n\<close>
- and rec: \<open>\<And>n b. P n \<Longrightarrow> (of_bool b + 2 * n) div 2 = n \<Longrightarrow> P (of_bool b + 2 * n)\<close>
- 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 \<open>q mod 2 ^ m mod 2 ^ n = q mod 2 ^ min m n\<close>
- 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 \<open>(q * 2 ^ m) mod (2 ^ n) = (q mod 2 ^ (n - m)) * 2 ^ m\<close> if \<open>m \<le> n\<close>
- 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)
- apply (simp add: mult.commute)
- done
- show \<open>even ((2 ^ m - (1::nat)) div 2 ^ n) \<longleftrightarrow> 2 ^ n = (0::nat) \<or> m \<le> n\<close>
- for m n :: nat
- using even_mask_div_iff' [where ?'a = nat, of m n] by simp
- show \<open>even (q * 2 ^ m div 2 ^ n) \<longleftrightarrow> n < m \<or> (2::nat) ^ n = 0 \<or> m \<le> n \<and> even (q div 2 ^ (n - m))\<close>
- 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: "\<And>k. P k \<Longrightarrow> k \<noteq> 0 \<Longrightarrow> P (k * 2)"
- and odd_int: "\<And>k. P k \<Longrightarrow> k \<noteq> - 1 \<Longrightarrow> P (1 + (k * 2))" for k :: int
-proof (cases "k \<ge> 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 \<open>simp_all add: algebra_simps\<close>)
- also have "\<dots> = - 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 \<open>simp_all add: algebra_simps\<close>)
- also have "\<dots> = - 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]:
- \<open>bit (of_bool b) n \<longleftrightarrow> b \<and> n = 0\<close>
- by (simp add: bit_1_iff)
-
-lemma even_of_nat_iff:
- \<open>even (of_nat n) \<longleftrightarrow> even n\<close>
- by (induction n rule: nat_bit_induct) simp_all
-
-lemma bit_of_nat_iff [bit_simps]:
- \<open>bit (of_nat m) n \<longleftrightarrow> (2::'a) ^ n \<noteq> 0 \<and> bit m n\<close>
-proof (cases \<open>(2::'a) ^ n = 0\<close>)
- case True
- then show ?thesis
- by (simp add: exp_eq_0_imp_not_bit)
-next
- case False
- then have \<open>bit (of_nat m) n \<longleftrightarrow> bit m n\<close>
- 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 Parity.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 Parity.bit_Suc dest: mult_not_zero)
- qed
- with False show ?thesis
- by simp
-qed
-
-end
-
-instantiation int :: semiring_bits
-begin
-
-definition bit_int :: \<open>int \<Rightarrow> nat \<Rightarrow> bool\<close>
- where \<open>bit_int k n \<longleftrightarrow> odd (k div 2 ^ n)\<close>
-
-instance
-proof
- show \<open>P k\<close> if stable: \<open>\<And>k. k div 2 = k \<Longrightarrow> P k\<close>
- and rec: \<open>\<And>k b. P k \<Longrightarrow> (of_bool b + 2 * k) div 2 = k \<Longrightarrow> P (of_bool b + 2 * k)\<close>
- 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 \<open>- 1\<close>] 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 \<open>(2::int) ^ m div 2 ^ n = of_bool ((2::int) ^ m \<noteq> 0 \<and> n \<le> m) * 2 ^ (m - n)\<close>
- for m n :: nat
- proof (cases \<open>m < n\<close>)
- case True
- then have \<open>n = m + (n - m)\<close>
- by simp
- then have \<open>(2::int) ^ m div 2 ^ n = (2::int) ^ m div 2 ^ (m + (n - m))\<close>
- by simp
- also have \<open>\<dots> = (2::int) ^ m div (2 ^ m * 2 ^ (n - m))\<close>
- by (simp add: power_add)
- also have \<open>\<dots> = (2::int) ^ m div 2 ^ m div 2 ^ (n - m)\<close>
- by (simp add: zdiv_zmult2_eq)
- finally show ?thesis using \<open>m < n\<close> by simp
- next
- case False
- then show ?thesis
- by (simp add: power_diff)
- qed
- show \<open>k mod 2 ^ m mod 2 ^ n = k mod 2 ^ min m n\<close>
- for m n :: nat and k :: int
- using mod_exp_eq [of \<open>nat k\<close> 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 \<open>2 ^ m\<close>])
- apply (subst zmod_zmult2_eq) apply simp_all
- done
- show \<open>(k * 2 ^ m) mod (2 ^ n) = (k mod 2 ^ (n - m)) * 2 ^ m\<close>
- if \<open>m \<le> n\<close> 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)
- apply (simp add: ac_simps)
- done
- show \<open>even ((2 ^ m - (1::int)) div 2 ^ n) \<longleftrightarrow> 2 ^ n = (0::int) \<or> m \<le> n\<close>
- for m n :: nat
- using even_mask_div_iff' [where ?'a = int, of m n] by simp
- show \<open>even (k * 2 ^ m div 2 ^ n) \<longleftrightarrow> n < m \<or> (2::int) ^ n = 0 \<or> m \<le> n \<and> even (k div 2 ^ (n - m))\<close>
- 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 :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
- assumes push_bit_eq_mult: \<open>push_bit n a = a * 2 ^ n\<close>
- fixes drop_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
- assumes drop_bit_eq_div: \<open>drop_bit n a = a div 2 ^ n\<close>
- fixes take_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close>
- assumes take_bit_eq_mod: \<open>take_bit n a = a mod 2 ^ n\<close>
-begin
-
-text \<open>
- Logically, \<^const>\<open>push_bit\<close>,
- \<^const>\<open>drop_bit\<close> and \<^const>\<open>take_bit\<close> are just aliases; having them
- as separate operations makes proofs easier, otherwise proof automation
- would fiddle with concrete expressions \<^term>\<open>2 ^ n\<close> 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.
-\<close>
-
-lemma bit_iff_odd_drop_bit:
- \<open>bit a n \<longleftrightarrow> odd (drop_bit n a)\<close>
- by (simp add: bit_iff_odd drop_bit_eq_div)
-
-lemma even_drop_bit_iff_not_bit:
- \<open>even (drop_bit n a) \<longleftrightarrow> \<not> bit a n\<close>
- by (simp add: bit_iff_odd_drop_bit)
-
-lemma div_push_bit_of_1_eq_drop_bit:
- \<open>a div push_bit n 1 = drop_bit n a\<close>
- 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]:
- \<open>push_bit (numeral l) (numeral k) = push_bit (pred_numeral l) (numeral (Num.Bit0 k))\<close>
- 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:
- \<open>take_bit (Suc n) a = take_bit n (a div 2) * 2 + a mod 2\<close>
-proof -
- have \<open>take_bit (Suc n) (a div 2 * 2 + of_bool (odd a)) = take_bit n (a div 2) * 2 + of_bool (odd a)\<close>
- using even_succ_mod_exp [of \<open>2 * (a div 2)\<close> \<open>Suc n\<close>]
- mult_exp_mod_exp_eq [of 1 \<open>Suc n\<close> \<open>a div 2\<close>]
- 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:
- \<open>take_bit n a = (if n = 0 then 0 else take_bit (n - 1) (a div 2) * 2 + a mod 2)\<close>
- by (cases n) (simp_all add: take_bit_Suc)
-
-lemma take_bit_Suc_0 [simp]:
- \<open>take_bit (Suc 0) a = a mod 2\<close>
- 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 \<and> b)"
- by (cases n) simp_all
-
-lemma even_take_bit_eq [simp]:
- \<open>even (take_bit n a) \<longleftrightarrow> n = 0 \<or> even a\<close>
- 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 \<open>m + n\<close> 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 \<le> 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 \<open>a * 2 ^ n\<close> 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 \<le> 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 \<open>m = n + q\<close>
- by (auto simp add: not_le dest: less_imp_Suc_add)
- then have \<open>drop_bit m (take_bit n a) = 0\<close>
- using div_exp_eq [of \<open>a mod 2 ^ n\<close> 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]:
- \<open>even (push_bit n a) \<longleftrightarrow> n \<noteq> 0 \<or> even a\<close>
- by (simp add: push_bit_eq_mult) auto
-
-lemma bit_push_bit_iff [bit_simps]:
- \<open>bit (push_bit m a) n \<longleftrightarrow> m \<le> n \<and> 2 ^ n \<noteq> 0 \<and> bit a (n - m)\<close>
- by (auto simp add: bit_iff_odd push_bit_eq_mult even_mult_exp_div_exp_iff)
-
-lemma bit_drop_bit_eq [bit_simps]:
- \<open>bit (drop_bit n a) = bit a \<circ> (+) n\<close>
- by (simp add: bit_iff_odd fun_eq_iff ac_simps flip: drop_bit_eq_div)
-
-lemma bit_take_bit_iff [bit_simps]:
- \<open>bit (take_bit m a) n \<longleftrightarrow> n < m \<and> bit a n\<close>
- by (simp add: bit_iff_odd drop_bit_take_bit not_le flip: drop_bit_eq_div)
-
-lemma stable_imp_drop_bit_eq:
- \<open>drop_bit n a = a\<close>
- if \<open>a div 2 = a\<close>
- by (induction n) (simp_all add: that drop_bit_Suc)
-
-lemma stable_imp_take_bit_eq:
- \<open>take_bit n a = (if even a then 0 else 2 ^ n - 1)\<close>
- if \<open>a div 2 = a\<close>
-proof (rule bit_eqI)
- fix m
- assume \<open>2 ^ m \<noteq> 0\<close>
- with that show \<open>bit (take_bit n a) m \<longleftrightarrow> bit (if even a then 0 else 2 ^ n - 1) m\<close>
- by (simp add: bit_take_bit_iff bit_mask_iff stable_imp_bit_iff_odd)
-qed
-
-lemma exp_dvdE:
- assumes \<open>2 ^ n dvd a\<close>
- obtains b where \<open>a = push_bit n b\<close>
-proof -
- from assms obtain b where \<open>a = 2 ^ n * b\<close> ..
- then have \<open>a = push_bit n b\<close>
- by (simp add: push_bit_eq_mult ac_simps)
- with that show thesis .
-qed
-
-lemma take_bit_eq_0_iff:
- \<open>take_bit n a = 0 \<longleftrightarrow> 2 ^ n dvd a\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
-proof
- assume ?P
- then show ?Q
- by (simp add: take_bit_eq_mod mod_0_imp_dvd)
-next
- assume ?Q
- then obtain b where \<open>a = push_bit n b\<close>
- by (rule exp_dvdE)
- then show ?P
- by (simp add: take_bit_push_bit)
-qed
-
-lemma take_bit_tightened:
- \<open>take_bit m a = take_bit m b\<close> if \<open>take_bit n a = take_bit n b\<close> and \<open>m \<le> n\<close>
-proof -
- from that have \<open>take_bit m (take_bit n a) = take_bit m (take_bit n b)\<close>
- by simp
- then have \<open>take_bit (min m n) a = take_bit (min m n) b\<close>
- by simp
- with that show ?thesis
- by (simp add: min_def)
-qed
-
-lemma take_bit_eq_self_iff_drop_bit_eq_0:
- \<open>take_bit n a = a \<longleftrightarrow> drop_bit n a = 0\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
-proof
- assume ?P
- show ?Q
- proof (rule bit_eqI)
- fix m
- from \<open>?P\<close> have \<open>a = take_bit n a\<close> ..
- also have \<open>\<not> bit (take_bit n a) (n + m)\<close>
- unfolding bit_simps
- by (simp add: bit_simps)
- finally show \<open>bit (drop_bit n a) m \<longleftrightarrow> bit 0 m\<close>
- by (simp add: bit_simps)
- qed
-next
- assume ?Q
- show ?P
- proof (rule bit_eqI)
- fix m
- from \<open>?Q\<close> have \<open>\<not> bit (drop_bit n a) (m - n)\<close>
- by simp
- then have \<open> \<not> bit a (n + (m - n))\<close>
- by (simp add: bit_simps)
- then show \<open>bit (take_bit n a) m \<longleftrightarrow> bit a m\<close>
- by (cases \<open>m < n\<close>) (auto simp add: bit_simps)
- qed
-qed
-
-lemma drop_bit_exp_eq:
- \<open>drop_bit m (2 ^ n) = of_bool (m \<le> n \<and> 2 ^ n \<noteq> 0) * 2 ^ (n - m)\<close>
- by (rule bit_eqI) (auto simp add: bit_simps)
-
-end
-
-instantiation nat :: semiring_bit_shifts
-begin
-
-definition push_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
- where \<open>push_bit_nat n m = m * 2 ^ n\<close>
-
-definition drop_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
- where \<open>drop_bit_nat n m = m div 2 ^ n\<close>
-
-definition take_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
- where \<open>take_bit_nat n m = m mod 2 ^ n\<close>
-
-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:
- \<open>push_bit n (of_nat m) = of_nat (push_bit n m)\<close>
- by (simp add: push_bit_eq_mult semiring_bit_shifts_class.push_bit_eq_mult)
-
-lemma of_nat_push_bit:
- \<open>of_nat (push_bit m n) = push_bit m (of_nat n)\<close>
- by (simp add: push_bit_eq_mult semiring_bit_shifts_class.push_bit_eq_mult)
-
-lemma take_bit_of_nat:
- \<open>take_bit n (of_nat m) = of_nat (take_bit n m)\<close>
- by (rule bit_eqI) (simp add: bit_take_bit_iff Parity.bit_take_bit_iff bit_of_nat_iff)
-
-lemma of_nat_take_bit:
- \<open>of_nat (take_bit n m) = take_bit n (of_nat m)\<close>
- by (rule bit_eqI) (simp add: bit_take_bit_iff Parity.bit_take_bit_iff bit_of_nat_iff)
-
-end
-
-instantiation int :: semiring_bit_shifts
-begin
-
-definition push_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
- where \<open>push_bit_int n k = k * 2 ^ n\<close>
-
-definition drop_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
- where \<open>drop_bit_int n k = k div 2 ^ n\<close>
-
-definition take_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close>
- where \<open>take_bit_int n k = k mod 2 ^ n\<close>
-
-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:
- \<open>bit (push_bit m q) n \<longleftrightarrow> m \<le> n \<and> bit q (n - m)\<close> for q :: nat
- by (auto simp add: bit_push_bit_iff)
-
-lemma bit_push_bit_iff_int:
- \<open>bit (push_bit m k) n \<longleftrightarrow> m \<le> n \<and> bit k (n - m)\<close> for k :: int
- by (auto simp add: bit_push_bit_iff)
-
-lemma take_bit_nat_less_exp [simp]:
- \<open>take_bit n m < 2 ^ n\<close> for n m ::nat
- by (simp add: take_bit_eq_mod)
-
-lemma take_bit_nonnegative [simp]:
- \<open>take_bit n k \<ge> 0\<close> for k :: int
- by (simp add: take_bit_eq_mod)
-
-lemma not_take_bit_negative [simp]:
- \<open>\<not> take_bit n k < 0\<close> for k :: int
- by (simp add: not_less)
-
-lemma take_bit_int_less_exp [simp]:
- \<open>take_bit n k < 2 ^ n\<close> for k :: int
- by (simp add: take_bit_eq_mod)
-
-lemma take_bit_nat_eq_self_iff:
- \<open>take_bit n m = m \<longleftrightarrow> m < 2 ^ n\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
- 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:
- \<open>take_bit n m = m\<close> if \<open>m < 2 ^ n\<close> for m n :: nat
- using that by (simp add: take_bit_nat_eq_self_iff)
-
-lemma take_bit_int_eq_self_iff:
- \<open>take_bit n k = k \<longleftrightarrow> 0 \<le> k \<and> k < 2 ^ n\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
- 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:
- \<open>take_bit n k = k\<close> if \<open>0 \<le> k\<close> \<open>k < 2 ^ n\<close> for k :: int
- using that by (simp add: take_bit_int_eq_self_iff)
-
-lemma take_bit_nat_less_eq_self [simp]:
- \<open>take_bit n m \<le> m\<close> for n m :: nat
- by (simp add: take_bit_eq_mod)
-
-lemma take_bit_nat_less_self_iff:
- \<open>take_bit n m < m \<longleftrightarrow> 2 ^ n \<le> m\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
- for m n :: nat
-proof
- assume ?P
- then have \<open>take_bit n m \<noteq> m\<close>
- by simp
- then show \<open>?Q\<close>
- by (simp add: take_bit_nat_eq_self_iff)
-next
- have \<open>take_bit n m < 2 ^ n\<close>
- 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]:
- \<open>take_bit m (2 ^ n) = of_bool (n < m) * 2 ^ n\<close>
- by (simp add: take_bit_eq_mod exp_mod_exp)
-
-lemma take_bit_of_2 [simp]:
- \<open>take_bit n 2 = of_bool (2 \<le> n) * 2\<close>
- using take_bit_of_exp [of n 1] by simp
-
-lemma take_bit_of_mask:
- \<open>take_bit m (2 ^ n - 1) = 2 ^ min m n - 1\<close>
- by (simp add: take_bit_eq_mod mask_mod_exp)
-
-lemma push_bit_eq_0_iff [simp]:
- "push_bit n a = 0 \<longleftrightarrow> 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 \<longleftrightarrow> n = 0"
- by (simp add: take_bit_eq_mod)
-
-lemma take_bit_Suc_1 [simp]:
- \<open>take_bit (Suc n) 1 = 1\<close>
- by (simp add: take_bit_Suc)
-
-lemma take_bit_Suc_bit0 [simp]:
- \<open>take_bit (Suc n) (numeral (Num.Bit0 k)) = take_bit n (numeral k) * 2\<close>
- by (simp add: take_bit_Suc numeral_Bit0_div_2)
-
-lemma take_bit_Suc_bit1 [simp]:
- \<open>take_bit (Suc n) (numeral (Num.Bit1 k)) = take_bit n (numeral k) * 2 + 1\<close>
- by (simp add: take_bit_Suc numeral_Bit1_div_2 mod_2_eq_odd)
-
-lemma take_bit_numeral_1 [simp]:
- \<open>take_bit (numeral l) 1 = 1\<close>
- by (simp add: take_bit_rec [of \<open>numeral l\<close> 1])
-
-lemma take_bit_numeral_bit0 [simp]:
- \<open>take_bit (numeral l) (numeral (Num.Bit0 k)) = take_bit (pred_numeral l) (numeral k) * 2\<close>
- by (simp add: take_bit_rec numeral_Bit0_div_2)
-
-lemma take_bit_numeral_bit1 [simp]:
- \<open>take_bit (numeral l) (numeral (Num.Bit1 k)) = take_bit (pred_numeral l) (numeral k) * 2 + 1\<close>
- by (simp add: take_bit_rec numeral_Bit1_div_2 mod_2_eq_odd)
-
-lemma drop_bit_Suc_bit0 [simp]:
- \<open>drop_bit (Suc n) (numeral (Num.Bit0 k)) = drop_bit n (numeral k)\<close>
- by (simp add: drop_bit_Suc numeral_Bit0_div_2)
-
-lemma drop_bit_Suc_bit1 [simp]:
- \<open>drop_bit (Suc n) (numeral (Num.Bit1 k)) = drop_bit n (numeral k)\<close>
- by (simp add: drop_bit_Suc numeral_Bit1_div_2)
-
-lemma drop_bit_numeral_bit0 [simp]:
- \<open>drop_bit (numeral l) (numeral (Num.Bit0 k)) = drop_bit (pred_numeral l) (numeral k)\<close>
- by (simp add: drop_bit_rec numeral_Bit0_div_2)
-
-lemma drop_bit_numeral_bit1 [simp]:
- \<open>drop_bit (numeral l) (numeral (Num.Bit1 k)) = drop_bit (pred_numeral l) (numeral k)\<close>
- 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 Parity.drop_bit_eq_div of_nat_div [of m "2 ^ n"])
-
-lemma bit_of_nat_iff_bit [bit_simps]:
- \<open>bit (of_nat m) n \<longleftrightarrow> bit m n\<close>
-proof -
- have \<open>even (m div 2 ^ n) \<longleftrightarrow> even (of_nat (m div 2 ^ n))\<close>
- by simp
- also have \<open>of_nat (m div 2 ^ n) = of_nat m div of_nat (2 ^ n)\<close>
- 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:
- \<open>of_nat (drop_bit m n) = drop_bit m (of_nat n)\<close>
- 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]:
- \<open>bit (push_bit m (of_nat r)) n \<longleftrightarrow> m \<le> n \<and> bit (of_nat r) (n - m)\<close>
- by (auto simp add: bit_push_bit_iff)
-
-end
-
-instance nat :: unique_euclidean_semiring_with_bit_shifts ..
-
-instance int :: unique_euclidean_semiring_with_bit_shifts ..
-
-lemma bit_numeral_int_iff [bit_simps]:
- \<open>bit (numeral m :: int) n \<longleftrightarrow> bit (numeral m :: nat) n\<close>
- using bit_of_nat_iff_bit [of \<open>numeral m\<close> n] by simp
-
-lemma bit_not_int_iff':
- \<open>bit (- k - 1) n \<longleftrightarrow> \<not> bit k n\<close>
- for k :: int
-proof (induction n arbitrary: k)
- case 0
- show ?case
- by simp
-next
- case (Suc n)
- have \<open>- k - 1 = - (k + 2) + 1\<close>
- by simp
- also have \<open>(- (k + 2) + 1) div 2 = - (k div 2) - 1\<close>
- proof (cases \<open>even k\<close>)
- case True
- then have \<open>- k div 2 = - (k div 2)\<close>
- by rule (simp flip: mult_minus_right)
- with True show ?thesis
- by simp
- next
- case False
- have \<open>4 = 2 * (2::int)\<close>
- by simp
- also have \<open>2 * 2 div 2 = (2::int)\<close>
- by (simp only: nonzero_mult_div_cancel_left)
- finally have *: \<open>4 div 2 = (2::int)\<close> .
- from False obtain l where k: \<open>k = 2 * l + 1\<close> ..
- then have \<open>- k - 2 = 2 * - (l + 2) + 1\<close>
- by simp
- then have \<open>(- k - 2) div 2 + 1 = - (k div 2) - 1\<close>
- by (simp flip: mult_minus_right add: *) (simp add: k)
- with False show ?thesis
- by simp
- qed
- finally have \<open>(- k - 1) div 2 = - (k div 2) - 1\<close> .
- with Suc show ?case
- by (simp add: bit_Suc)
-qed
-
-lemma bit_minus_int_iff [bit_simps]:
- \<open>bit (- k) n \<longleftrightarrow> \<not> bit (k - 1) n\<close>
- for k :: int
- using bit_not_int_iff' [of \<open>k - 1\<close>] by simp
-
-lemma bit_nat_iff [bit_simps]:
- \<open>bit (nat k) n \<longleftrightarrow> k \<ge> 0 \<and> bit k n\<close>
-proof (cases \<open>k \<ge> 0\<close>)
- case True
- moreover define m where \<open>m = nat k\<close>
- ultimately have \<open>k = int m\<close>
- by simp
- then show ?thesis
- by (simp add: bit_simps)
-next
- case False
- then show ?thesis
- by simp
-qed
-
-lemma push_bit_nat_eq:
- \<open>push_bit n (nat k) = nat (push_bit n k)\<close>
- by (cases \<open>k \<ge> 0\<close>) (simp_all add: push_bit_eq_mult nat_mult_distrib not_le mult_nonneg_nonpos2)
-
-lemma drop_bit_nat_eq:
- \<open>drop_bit n (nat k) = nat (drop_bit n k)\<close>
- apply (cases \<open>k \<ge> 0\<close>)
- 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:
- \<open>take_bit n (nat k) = nat (take_bit n k)\<close> if \<open>k \<ge> 0\<close>
- using that by (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq)
-
-lemma nat_take_bit_eq:
- \<open>nat (take_bit n k) = take_bit n (nat k)\<close>
- if \<open>k \<ge> 0\<close>
- using that by (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq)
-
-lemma not_exp_less_eq_0_int [simp]:
- \<open>\<not> 2 ^ n \<le> (0::int)\<close>
- by (simp add: power_le_zero_eq)
-
-lemma half_nonnegative_int_iff [simp]:
- \<open>k div 2 \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
-proof (cases \<open>k \<ge> 0\<close>)
- case True
- then show ?thesis
- by (auto simp add: divide_int_def sgn_1_pos)
-next
- case False
- then show ?thesis
- apply (auto simp add: divide_int_def not_le elim!: evenE)
- apply (simp only: minus_mult_right)
- apply (subst (asm) nat_mult_distrib)
- apply simp_all
- done
-qed
-
-lemma half_negative_int_iff [simp]:
- \<open>k div 2 < 0 \<longleftrightarrow> k < 0\<close> 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:
- \<open>- 1 div (2 :: int) ^ n = - 1\<close>
- by (induction n) (use div_exp_eq [symmetric, of \<open>- 1 :: int\<close> 1] in \<open>simp_all add: ac_simps\<close>)
-
-lemma drop_bit_minus_one [simp]:
- \<open>drop_bit n (- 1 :: int) = - 1\<close>
- by (simp add: drop_bit_eq_div minus_1_div_exp_eq_int)
-
-lemma take_bit_Suc_from_most:
- \<open>take_bit (Suc n) k = 2 ^ n * of_bool (bit k n) + take_bit n k\<close> 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:
- \<open>take_bit n (- take_bit n k) = take_bit n (- k)\<close>
- for k :: int
- by (simp add: take_bit_eq_mod mod_minus_eq)
-
-lemma take_bit_diff:
- \<open>take_bit n (take_bit n k - take_bit n l) = take_bit n (k - l)\<close>
- for k l :: int
- by (simp add: take_bit_eq_mod mod_diff_eq)
-
-lemma bit_imp_take_bit_positive:
- \<open>0 < take_bit m k\<close> if \<open>n < m\<close> and \<open>bit k n\<close> for k :: int
-proof (rule ccontr)
- assume \<open>\<not> 0 < take_bit m k\<close>
- then have \<open>take_bit m k = 0\<close>
- by (auto simp add: not_less intro: order_antisym)
- then have \<open>bit (take_bit m k) n = bit 0 n\<close>
- by simp
- with that show False
- by (simp add: bit_take_bit_iff)
-qed
-
-lemma take_bit_mult:
- \<open>take_bit n (take_bit n k * take_bit n l) = take_bit n (k * l)\<close>
- 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]:
- \<open>of_nat (nat (take_bit n k)) = of_int (take_bit n k)\<close>
- by simp
-
-lemma take_bit_minus_small_eq:
- \<open>take_bit n (- k) = 2 ^ n - k\<close> if \<open>0 < k\<close> \<open>k \<le> 2 ^ n\<close> for k :: int
-proof -
- define m where \<open>m = nat k\<close>
- with that have \<open>k = int m\<close> and \<open>0 < m\<close> and \<open>m \<le> 2 ^ n\<close>
- by simp_all
- have \<open>(2 ^ n - m) mod 2 ^ n = 2 ^ n - m\<close>
- using \<open>0 < m\<close> by simp
- then have \<open>int ((2 ^ n - m) mod 2 ^ n) = int (2 ^ n - m)\<close>
- by simp
- then have \<open>(2 ^ n - int m) mod 2 ^ n = 2 ^ n - int m\<close>
- using \<open>m \<le> 2 ^ n\<close> by (simp only: of_nat_mod of_nat_diff) simp
- with \<open>k = int m\<close> have \<open>(2 ^ n - k) mod 2 ^ n = 2 ^ n - k\<close>
- by simp
- then show ?thesis
- by (simp add: take_bit_eq_mod)
-qed
-
-lemma drop_bit_push_bit_int:
- \<open>drop_bit m (push_bit n k) = drop_bit (m - n) (push_bit (n - m) k)\<close> for k :: int
- by (cases \<open>m \<le> n\<close>) (auto simp add: mult.left_commute [of _ \<open>2 ^ n\<close>] mult.commute [of _ \<open>2 ^ n\<close>] 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]:
- \<open>push_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
- by (simp add: push_bit_eq_mult zero_le_mult_iff)
-
-lemma push_bit_negative_int_iff [simp]:
- \<open>push_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
- by (subst Not_eq_iff [symmetric]) (simp add: not_less)
-
-lemma drop_bit_nonnegative_int_iff [simp]:
- \<open>drop_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int
- by (induction n) (simp_all add: drop_bit_Suc drop_bit_half)
-
-lemma drop_bit_negative_int_iff [simp]:
- \<open>drop_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int
- by (subst Not_eq_iff [symmetric]) (simp add: not_less)
-
code_identifier
code_module Parity \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith
--- a/src/HOL/Random.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/Random.thy Mon Aug 02 10:01:06 2021 +0000
@@ -3,7 +3,7 @@
section \<open>A HOL random engine\<close>
theory Random
-imports List Groups_List
+imports List Groups_List Code_Numeral
begin
subsection \<open>Auxiliary functions\<close>
--- a/src/HOL/Set.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/Set.thy Mon Aug 02 10:01:06 2021 +0000
@@ -7,7 +7,7 @@
section \<open>Set theory for higher-order logic\<close>
theory Set
- imports Lattices
+ imports Lattices Boolean_Algebra
begin
subsection \<open>Sets as predicates\<close>
--- a/src/HOL/Set_Interval.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/Set_Interval.thy Mon Aug 02 10:01:06 2021 +0000
@@ -2140,31 +2140,6 @@
by (subst sum_subtractf_nat) auto
-context unique_euclidean_semiring_with_bit_shifts
-begin
-
-lemma take_bit_sum:
- "take_bit n a = (\<Sum>k = 0..<n. push_bit k (of_bool (bit a k)))"
- for n :: nat
-proof (induction n arbitrary: a)
- case 0
- then show ?case
- by simp
-next
- case (Suc n)
- have "(\<Sum>k = 0..<Suc n. push_bit k (of_bool (bit a k))) =
- of_bool (odd a) + (\<Sum>k = Suc 0..<Suc n. push_bit k (of_bool (bit a k)))"
- by (simp add: sum.atLeast_Suc_lessThan ac_simps)
- also have "(\<Sum>k = Suc 0..<Suc n. push_bit k (of_bool (bit a k)))
- = (\<Sum>k = 0..<n. push_bit k (of_bool (bit (a div 2) k))) * 2"
- by (simp only: sum.atLeast_Suc_lessThan_Suc_shift) (simp add: sum_distrib_right push_bit_double drop_bit_Suc bit_Suc)
- finally show ?case
- using Suc [of "a div 2"] by (simp add: ac_simps take_bit_Suc mod_2_eq_odd)
-qed
-
-end
-
-
subsubsection \<open>Shifting bounds\<close>
context comm_monoid_add
--- a/src/HOL/String.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/String.thy Mon Aug 02 10:01:06 2021 +0000
@@ -3,7 +3,7 @@
section \<open>Character and string types\<close>
theory String
-imports Enum
+imports Enum Bit_Operations Code_Numeral
begin
subsection \<open>Strings as list of bytes\<close>
--- a/src/HOL/ex/Meson_Test.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/ex/Meson_Test.thy Mon Aug 02 10:01:06 2021 +0000
@@ -10,7 +10,7 @@
below and constants declared in HOL!
\<close>
-hide_const (open) implies union inter subset quotient sum
+hide_const (open) implies union inter subset quotient sum or
text \<open>
Test data for the MESON proof procedure
@@ -1286,7 +1286,7 @@
(\<forall>X Y. value(X::'a,truth) & value(Y::'a,falsity) --> value(xor(X::'a,Y),truth)) &
(\<forall>X Y. value(X::'a,falsity) & value(Y::'a,truth) --> value(xor(X::'a,Y),truth)) &
(\<forall>X Y. value(X::'a,falsity) & value(Y::'a,falsity) --> value(xor(X::'a,Y),falsity)) &
- (\<forall>Value. ~value(xor(xor(xor(xor(truth::'a,falsity),falsity),truth),falsity),Value)) --> False"
+ (\<forall>Value. ~value(xor(xor(xor(xor(truth::'a,falsity),falsity),truth),falsity),Value)) --> False" for xor
by meson
(*19116 inferences so far. Searching to depth 16. 15.9 secs*)
--- a/src/HOL/ex/Reflection_Examples.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/ex/Reflection_Examples.thy Mon Aug 02 10:01:06 2021 +0000
@@ -66,7 +66,7 @@
apply reify
oops
-datatype fm = And fm fm | Or fm fm | Imp fm fm | Iff fm fm | NOT fm | At nat
+datatype fm = And fm fm | Or fm fm | Imp fm fm | Iff fm fm | Not fm | At nat
primrec Ifm :: "fm \<Rightarrow> bool list \<Rightarrow> bool"
where
@@ -75,7 +75,7 @@
| "Ifm (Or p q) vs \<longleftrightarrow> Ifm p vs \<or> Ifm q vs"
| "Ifm (Imp p q) vs \<longleftrightarrow> Ifm p vs \<longrightarrow> Ifm q vs"
| "Ifm (Iff p q) vs \<longleftrightarrow> Ifm p vs = Ifm q vs"
-| "Ifm (NOT p) vs \<longleftrightarrow> \<not> Ifm p vs"
+| "Ifm (Not p) vs \<longleftrightarrow> \<not> Ifm p vs"
lemma "Q \<longrightarrow> (D \<and> F \<and> ((\<not> D) \<and> (\<not> F)))"
apply (reify Ifm.simps)
@@ -93,7 +93,7 @@
primrec fmsize :: "fm \<Rightarrow> nat"
where
"fmsize (At n) = 1"
-| "fmsize (NOT p) = 1 + fmsize p"
+| "fmsize (Not p) = 1 + fmsize p"
| "fmsize (And p q) = 1 + fmsize p + fmsize q"
| "fmsize (Or p q) = 1 + fmsize p + fmsize q"
| "fmsize (Imp p q) = 2 + fmsize p + fmsize q"
@@ -106,14 +106,14 @@
"nnf (At n) = At n"
| "nnf (And p q) = And (nnf p) (nnf q)"
| "nnf (Or p q) = Or (nnf p) (nnf q)"
-| "nnf (Imp p q) = Or (nnf (NOT p)) (nnf q)"
-| "nnf (Iff p q) = Or (And (nnf p) (nnf q)) (And (nnf (NOT p)) (nnf (NOT q)))"
-| "nnf (NOT (And p q)) = Or (nnf (NOT p)) (nnf (NOT q))"
-| "nnf (NOT (Or p q)) = And (nnf (NOT p)) (nnf (NOT q))"
-| "nnf (NOT (Imp p q)) = And (nnf p) (nnf (NOT q))"
-| "nnf (NOT (Iff p q)) = Or (And (nnf p) (nnf (NOT q))) (And (nnf (NOT p)) (nnf q))"
-| "nnf (NOT (NOT p)) = nnf p"
-| "nnf (NOT p) = NOT p"
+| "nnf (Imp p q) = Or (nnf (Not p)) (nnf q)"
+| "nnf (Iff p q) = Or (And (nnf p) (nnf q)) (And (nnf (Not p)) (nnf (Not q)))"
+| "nnf (Not (And p q)) = Or (nnf (Not p)) (nnf (Not q))"
+| "nnf (Not (Or p q)) = And (nnf (Not p)) (nnf (Not q))"
+| "nnf (Not (Imp p q)) = And (nnf p) (nnf (Not q))"
+| "nnf (Not (Iff p q)) = Or (And (nnf p) (nnf (Not q))) (And (nnf (Not p)) (nnf q))"
+| "nnf (Not (Not p)) = nnf p"
+| "nnf (Not p) = Not p"
text \<open>The correctness theorem of \<^const>\<open>nnf\<close>: it preserves the semantics of \<^typ>\<open>fm\<close>\<close>
lemma nnf [reflection]:
--- a/src/HOL/ex/Tree23.thy Sun Aug 01 23:18:13 2021 +0200
+++ b/src/HOL/ex/Tree23.thy Mon Aug 02 10:01:06 2021 +0000
@@ -8,6 +8,8 @@
imports Main
begin
+hide_const (open) or
+
text\<open>This is a very direct translation of some of the functions in table.ML
in the Isabelle source code. That source is due to Makarius Wenzel and Stefan
Berghofer.