| author | nipkow | 
| Thu, 09 Dec 2021 08:32:29 +0100 | |
| changeset 74888 | 1c50ddcf6a01 | 
| parent 74618 | 43142ac556e6 | 
| child 75085 | ccc3a72210e6 | 
| permissions | -rw-r--r-- | 
| 71042 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 1 | (* Author: Florian Haftmann, TUM | 
| 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 2 | *) | 
| 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 3 | |
| 71956 | 4 | section \<open>Bit operations in suitable algebraic structures\<close> | 
| 71042 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 5 | |
| 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 6 | theory Bit_Operations | 
| 74101 | 7 | imports Presburger Groups_List | 
| 8 | begin | |
| 9 | ||
| 10 | subsection \<open>Abstract bit structures\<close> | |
| 11 | ||
| 12 | class semiring_bits = semiring_parity + | |
| 13 | assumes bits_induct [case_names stable rec]: | |
| 14 | \<open>(\<And>a. a div 2 = a \<Longrightarrow> P a) | |
| 15 | \<Longrightarrow> (\<And>a b. P a \<Longrightarrow> (of_bool b + 2 * a) div 2 = a \<Longrightarrow> P (of_bool b + 2 * a)) | |
| 16 | \<Longrightarrow> P a\<close> | |
| 17 | assumes bits_div_0 [simp]: \<open>0 div a = 0\<close> | |
| 18 | and bits_div_by_1 [simp]: \<open>a div 1 = a\<close> | |
| 19 | and bits_mod_div_trivial [simp]: \<open>a mod b div b = 0\<close> | |
| 20 | and even_succ_div_2 [simp]: \<open>even a \<Longrightarrow> (1 + a) div 2 = a div 2\<close> | |
| 21 | and even_mask_div_iff: \<open>even ((2 ^ m - 1) div 2 ^ n) \<longleftrightarrow> 2 ^ n = 0 \<or> m \<le> n\<close> | |
| 22 | 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> | |
| 23 | and div_exp_eq: \<open>a div 2 ^ m div 2 ^ n = a div 2 ^ (m + n)\<close> | |
| 24 | and mod_exp_eq: \<open>a mod 2 ^ m mod 2 ^ n = a mod 2 ^ min m n\<close> | |
| 25 | 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> | |
| 26 | and div_exp_mod_exp_eq: \<open>a div 2 ^ n mod 2 ^ m = a mod (2 ^ (n + m)) div 2 ^ n\<close> | |
| 27 | 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> | |
| 28 | fixes bit :: \<open>'a \<Rightarrow> nat \<Rightarrow> bool\<close> | |
| 29 | assumes bit_iff_odd: \<open>bit a n \<longleftrightarrow> odd (a div 2 ^ n)\<close> | |
| 30 | begin | |
| 31 | ||
| 32 | text \<open> | |
| 33 | Having \<^const>\<open>bit\<close> as definitional class operation | |
| 34 | takes into account that specific instances can be implemented | |
| 35 | differently wrt. code generation. | |
| 36 | \<close> | |
| 37 | ||
| 38 | lemma bits_div_by_0 [simp]: | |
| 39 | \<open>a div 0 = 0\<close> | |
| 40 | by (metis add_cancel_right_right bits_mod_div_trivial mod_mult_div_eq mult_not_zero) | |
| 41 | ||
| 42 | lemma bits_1_div_2 [simp]: | |
| 43 | \<open>1 div 2 = 0\<close> | |
| 44 | using even_succ_div_2 [of 0] by simp | |
| 45 | ||
| 46 | lemma bits_1_div_exp [simp]: | |
| 47 | \<open>1 div 2 ^ n = of_bool (n = 0)\<close> | |
| 48 | using div_exp_eq [of 1 1] by (cases n) simp_all | |
| 49 | ||
| 50 | lemma even_succ_div_exp [simp]: | |
| 51 | \<open>(1 + a) div 2 ^ n = a div 2 ^ n\<close> if \<open>even a\<close> and \<open>n > 0\<close> | |
| 52 | proof (cases n) | |
| 53 | case 0 | |
| 54 | with that show ?thesis | |
| 55 | by simp | |
| 56 | next | |
| 57 | case (Suc n) | |
| 58 | with \<open>even a\<close> have \<open>(1 + a) div 2 ^ Suc n = a div 2 ^ Suc n\<close> | |
| 59 | proof (induction n) | |
| 60 | case 0 | |
| 61 | then show ?case | |
| 62 | by simp | |
| 63 | next | |
| 64 | case (Suc n) | |
| 65 | then show ?case | |
| 66 | using div_exp_eq [of _ 1 \<open>Suc n\<close>, symmetric] | |
| 67 | by simp | |
| 68 | qed | |
| 69 | with Suc show ?thesis | |
| 70 | by simp | |
| 71 | qed | |
| 72 | ||
| 73 | lemma even_succ_mod_exp [simp]: | |
| 74 | \<open>(1 + a) mod 2 ^ n = 1 + (a mod 2 ^ n)\<close> if \<open>even a\<close> and \<open>n > 0\<close> | |
| 75 | using div_mult_mod_eq [of \<open>1 + a\<close> \<open>2 ^ n\<close>] that | |
| 76 | apply simp | |
| 77 | by (metis local.add.left_commute local.add_left_cancel local.div_mult_mod_eq) | |
| 78 | ||
| 79 | lemma bits_mod_by_1 [simp]: | |
| 80 | \<open>a mod 1 = 0\<close> | |
| 81 | using div_mult_mod_eq [of a 1] by simp | |
| 82 | ||
| 83 | lemma bits_mod_0 [simp]: | |
| 84 | \<open>0 mod a = 0\<close> | |
| 85 | using div_mult_mod_eq [of 0 a] by simp | |
| 86 | ||
| 87 | lemma bits_one_mod_two_eq_one [simp]: | |
| 88 | \<open>1 mod 2 = 1\<close> | |
| 89 | by (simp add: mod2_eq_if) | |
| 90 | ||
| 91 | lemma bit_0 [simp]: | |
| 92 | \<open>bit a 0 \<longleftrightarrow> odd a\<close> | |
| 93 | by (simp add: bit_iff_odd) | |
| 94 | ||
| 95 | lemma bit_Suc: | |
| 96 | \<open>bit a (Suc n) \<longleftrightarrow> bit (a div 2) n\<close> | |
| 97 | using div_exp_eq [of a 1 n] by (simp add: bit_iff_odd) | |
| 98 | ||
| 99 | lemma bit_rec: | |
| 100 | \<open>bit a n \<longleftrightarrow> (if n = 0 then odd a else bit (a div 2) (n - 1))\<close> | |
| 101 | by (cases n) (simp_all add: bit_Suc) | |
| 102 | ||
| 103 | lemma bit_0_eq [simp]: | |
| 104 | \<open>bit 0 = bot\<close> | |
| 105 | by (simp add: fun_eq_iff bit_iff_odd) | |
| 106 | ||
| 107 | context | |
| 108 | fixes a | |
| 109 | assumes stable: \<open>a div 2 = a\<close> | |
| 110 | begin | |
| 111 | ||
| 112 | lemma bits_stable_imp_add_self: | |
| 113 | \<open>a + a mod 2 = 0\<close> | |
| 114 | proof - | |
| 115 | have \<open>a div 2 * 2 + a mod 2 = a\<close> | |
| 116 | by (fact div_mult_mod_eq) | |
| 117 | then have \<open>a * 2 + a mod 2 = a\<close> | |
| 118 | by (simp add: stable) | |
| 119 | then show ?thesis | |
| 120 | by (simp add: mult_2_right ac_simps) | |
| 121 | qed | |
| 122 | ||
| 123 | lemma stable_imp_bit_iff_odd: | |
| 124 | \<open>bit a n \<longleftrightarrow> odd a\<close> | |
| 125 | by (induction n) (simp_all add: stable bit_Suc) | |
| 126 | ||
| 127 | end | |
| 128 | ||
| 129 | lemma bit_iff_idd_imp_stable: | |
| 130 | \<open>a div 2 = a\<close> if \<open>\<And>n. bit a n \<longleftrightarrow> odd a\<close> | |
| 131 | using that proof (induction a rule: bits_induct) | |
| 132 | case (stable a) | |
| 133 | then show ?case | |
| 134 | by simp | |
| 135 | next | |
| 136 | case (rec a b) | |
| 137 | from rec.prems [of 1] have [simp]: \<open>b = odd a\<close> | |
| 138 | by (simp add: rec.hyps bit_Suc) | |
| 139 | from rec.hyps have hyp: \<open>(of_bool (odd a) + 2 * a) div 2 = a\<close> | |
| 140 | by simp | |
| 141 | have \<open>bit a n \<longleftrightarrow> odd a\<close> for n | |
| 142 | using rec.prems [of \<open>Suc n\<close>] by (simp add: hyp bit_Suc) | |
| 143 | then have \<open>a div 2 = a\<close> | |
| 144 | by (rule rec.IH) | |
| 145 | then have \<open>of_bool (odd a) + 2 * a = 2 * (a div 2) + of_bool (odd a)\<close> | |
| 146 | by (simp add: ac_simps) | |
| 147 | also have \<open>\<dots> = a\<close> | |
| 148 | using mult_div_mod_eq [of 2 a] | |
| 149 | by (simp add: of_bool_odd_eq_mod_2) | |
| 150 | finally show ?case | |
| 151 | using \<open>a div 2 = a\<close> by (simp add: hyp) | |
| 152 | qed | |
| 153 | ||
| 154 | lemma exp_eq_0_imp_not_bit: | |
| 155 | \<open>\<not> bit a n\<close> if \<open>2 ^ n = 0\<close> | |
| 156 | using that by (simp add: bit_iff_odd) | |
| 157 | ||
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 158 | definition | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 159 | possible_bit :: "'a itself \<Rightarrow> nat \<Rightarrow> bool" | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 160 | where | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 161 | "possible_bit tyrep n = (2 ^ n \<noteq> (0 :: 'a))" | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 162 | |
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 163 | lemma possible_bit_0[simp]: | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 164 | "possible_bit ty 0" | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 165 | by (simp add: possible_bit_def) | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 166 | |
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 167 | lemma fold_possible_bit: | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 168 |   "2 ^ n = (0 :: 'a) \<longleftrightarrow> \<not> possible_bit TYPE('a) n"
 | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 169 | by (simp add: possible_bit_def) | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 170 | |
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 171 | lemmas impossible_bit = exp_eq_0_imp_not_bit[simplified fold_possible_bit] | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 172 | |
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 173 | lemma bit_imp_possible_bit: | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 174 |   "bit a n \<Longrightarrow> possible_bit TYPE('a) n"
 | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 175 | by (rule ccontr) (simp add: impossible_bit) | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 176 | |
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 177 | lemma possible_bit_less_imp: | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 178 | "possible_bit tyrep i \<Longrightarrow> j \<le> i \<Longrightarrow> possible_bit tyrep j" | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 179 | using power_add[of "2 :: 'a" j "i - j"] | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 180 | by (clarsimp simp: possible_bit_def eq_commute[where a=0]) | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 181 | |
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 182 | lemma possible_bit_min[simp]: | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 183 | "possible_bit tyrep (min i j) \<longleftrightarrow> possible_bit tyrep i \<or> possible_bit tyrep j" | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 184 | by (auto simp: min_def elim: possible_bit_less_imp) | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 185 | |
| 74101 | 186 | lemma bit_eqI: | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 187 |   \<open>a = b\<close> if \<open>\<And>n. possible_bit TYPE('a) n \<Longrightarrow> bit a n \<longleftrightarrow> bit b n\<close>
 | 
| 74101 | 188 | proof - | 
| 189 | have \<open>bit a n \<longleftrightarrow> bit b n\<close> for n | |
| 190 | proof (cases \<open>2 ^ n = 0\<close>) | |
| 191 | case True | |
| 192 | then show ?thesis | |
| 193 | by (simp add: exp_eq_0_imp_not_bit) | |
| 194 | next | |
| 195 | case False | |
| 196 | then show ?thesis | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 197 | by (rule that[unfolded possible_bit_def]) | 
| 74101 | 198 | qed | 
| 199 | then show ?thesis proof (induction a arbitrary: b rule: bits_induct) | |
| 200 | case (stable a) | |
| 201 | from stable(2) [of 0] have **: \<open>even b \<longleftrightarrow> even a\<close> | |
| 202 | by simp | |
| 203 | have \<open>b div 2 = b\<close> | |
| 204 | proof (rule bit_iff_idd_imp_stable) | |
| 205 | fix n | |
| 206 | from stable have *: \<open>bit b n \<longleftrightarrow> bit a n\<close> | |
| 207 | by simp | |
| 208 | also have \<open>bit a n \<longleftrightarrow> odd a\<close> | |
| 209 | using stable by (simp add: stable_imp_bit_iff_odd) | |
| 210 | finally show \<open>bit b n \<longleftrightarrow> odd b\<close> | |
| 211 | by (simp add: **) | |
| 212 | qed | |
| 213 | from ** have \<open>a mod 2 = b mod 2\<close> | |
| 214 | by (simp add: mod2_eq_if) | |
| 215 | then have \<open>a mod 2 + (a + b) = b mod 2 + (a + b)\<close> | |
| 216 | by simp | |
| 217 | then have \<open>a + a mod 2 + b = b + b mod 2 + a\<close> | |
| 218 | by (simp add: ac_simps) | |
| 219 | with \<open>a div 2 = a\<close> \<open>b div 2 = b\<close> show ?case | |
| 220 | by (simp add: bits_stable_imp_add_self) | |
| 221 | next | |
| 222 | case (rec a p) | |
| 223 | from rec.prems [of 0] have [simp]: \<open>p = odd b\<close> | |
| 224 | by simp | |
| 225 | from rec.hyps have \<open>bit a n \<longleftrightarrow> bit (b div 2) n\<close> for n | |
| 226 | using rec.prems [of \<open>Suc n\<close>] by (simp add: bit_Suc) | |
| 227 | then have \<open>a = b div 2\<close> | |
| 228 | by (rule rec.IH) | |
| 229 | then have \<open>2 * a = 2 * (b div 2)\<close> | |
| 230 | by simp | |
| 231 | then have \<open>b mod 2 + 2 * a = b mod 2 + 2 * (b div 2)\<close> | |
| 232 | by simp | |
| 233 | also have \<open>\<dots> = b\<close> | |
| 234 | by (fact mod_mult_div_eq) | |
| 235 | finally show ?case | |
| 236 | by (auto simp add: mod2_eq_if) | |
| 237 | qed | |
| 238 | qed | |
| 239 | ||
| 240 | lemma bit_eq_iff: | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 241 |   \<open>a = b \<longleftrightarrow> (\<forall>n. possible_bit TYPE('a) n \<longrightarrow> bit a n \<longleftrightarrow> bit b n)\<close>
 | 
| 74101 | 242 | by (auto intro: bit_eqI) | 
| 243 | ||
| 244 | named_theorems bit_simps \<open>Simplification rules for \<^const>\<open>bit\<close>\<close> | |
| 245 | ||
| 246 | lemma bit_exp_iff [bit_simps]: | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 247 |   \<open>bit (2 ^ m) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> m = n\<close>
 | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 248 | by (auto simp add: bit_iff_odd exp_div_exp_eq possible_bit_def) | 
| 74101 | 249 | |
| 250 | lemma bit_1_iff [bit_simps]: | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 251 | \<open>bit 1 n \<longleftrightarrow> n = 0\<close> | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 252 | using bit_exp_iff [of 0 n] | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 253 | by auto | 
| 74101 | 254 | |
| 255 | lemma bit_2_iff [bit_simps]: | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 256 |   \<open>bit 2 n \<longleftrightarrow> possible_bit TYPE('a) 1 \<and> n = 1\<close>
 | 
| 74101 | 257 | using bit_exp_iff [of 1 n] by auto | 
| 258 | ||
| 259 | lemma even_bit_succ_iff: | |
| 260 | \<open>bit (1 + a) n \<longleftrightarrow> bit a n \<or> n = 0\<close> if \<open>even a\<close> | |
| 261 | using that by (cases \<open>n = 0\<close>) (simp_all add: bit_iff_odd) | |
| 262 | ||
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 263 | lemma bit_double_iff [bit_simps]: | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 264 |   \<open>bit (2 * a) n \<longleftrightarrow> bit a (n - 1) \<and> n \<noteq> 0 \<and> possible_bit TYPE('a) n\<close>
 | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 265 | using even_mult_exp_div_exp_iff [of a 1 n] | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 266 | by (cases n, auto simp add: bit_iff_odd ac_simps possible_bit_def) | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 267 | |
| 74101 | 268 | lemma odd_bit_iff_bit_pred: | 
| 269 | \<open>bit a n \<longleftrightarrow> bit (a - 1) n \<or> n = 0\<close> if \<open>odd a\<close> | |
| 270 | proof - | |
| 271 | from \<open>odd a\<close> obtain b where \<open>a = 2 * b + 1\<close> .. | |
| 272 | moreover have \<open>bit (2 * b) n \<or> n = 0 \<longleftrightarrow> bit (1 + 2 * b) n\<close> | |
| 273 | using even_bit_succ_iff by simp | |
| 274 | ultimately show ?thesis by (simp add: ac_simps) | |
| 275 | qed | |
| 276 | ||
| 277 | lemma bit_eq_rec: | |
| 278 | \<open>a = b \<longleftrightarrow> (even a \<longleftrightarrow> even b) \<and> a div 2 = b div 2\<close> (is \<open>?P = ?Q\<close>) | |
| 279 | proof | |
| 280 | assume ?P | |
| 281 | then show ?Q | |
| 282 | by simp | |
| 283 | next | |
| 284 | assume ?Q | |
| 285 | then have \<open>even a \<longleftrightarrow> even b\<close> and \<open>a div 2 = b div 2\<close> | |
| 286 | by simp_all | |
| 287 | show ?P | |
| 288 | proof (rule bit_eqI) | |
| 289 | fix n | |
| 290 | show \<open>bit a n \<longleftrightarrow> bit b n\<close> | |
| 291 | proof (cases n) | |
| 292 | case 0 | |
| 293 | with \<open>even a \<longleftrightarrow> even b\<close> show ?thesis | |
| 294 | by simp | |
| 295 | next | |
| 296 | case (Suc n) | |
| 297 | moreover from \<open>a div 2 = b div 2\<close> have \<open>bit (a div 2) n = bit (b div 2) n\<close> | |
| 298 | by simp | |
| 299 | ultimately show ?thesis | |
| 300 | by (simp add: bit_Suc) | |
| 301 | qed | |
| 302 | qed | |
| 303 | qed | |
| 304 | ||
| 305 | lemma bit_mod_2_iff [simp]: | |
| 306 | \<open>bit (a mod 2) n \<longleftrightarrow> n = 0 \<and> odd a\<close> | |
| 307 | by (cases a rule: parity_cases) (simp_all add: bit_iff_odd) | |
| 308 | ||
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 309 | lemma bit_mask_sub_iff: | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 310 |   \<open>bit (2 ^ m - 1) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> n < m\<close>
 | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 311 | by (simp add: bit_iff_odd even_mask_div_iff not_le possible_bit_def) | 
| 74101 | 312 | |
| 313 | lemma exp_add_not_zero_imp: | |
| 314 | \<open>2 ^ m \<noteq> 0\<close> and \<open>2 ^ n \<noteq> 0\<close> if \<open>2 ^ (m + n) \<noteq> 0\<close> | |
| 315 | proof - | |
| 316 | have \<open>\<not> (2 ^ m = 0 \<or> 2 ^ n = 0)\<close> | |
| 317 | proof (rule notI) | |
| 318 | assume \<open>2 ^ m = 0 \<or> 2 ^ n = 0\<close> | |
| 319 | then have \<open>2 ^ (m + n) = 0\<close> | |
| 320 | by (rule disjE) (simp_all add: power_add) | |
| 321 | with that show False .. | |
| 322 | qed | |
| 323 | then show \<open>2 ^ m \<noteq> 0\<close> and \<open>2 ^ n \<noteq> 0\<close> | |
| 324 | by simp_all | |
| 325 | qed | |
| 326 | ||
| 327 | lemma bit_disjunctive_add_iff: | |
| 328 | \<open>bit (a + b) n \<longleftrightarrow> bit a n \<or> bit b n\<close> | |
| 329 | if \<open>\<And>n. \<not> bit a n \<or> \<not> bit b n\<close> | |
| 330 | proof (cases \<open>2 ^ n = 0\<close>) | |
| 331 | case True | |
| 332 | then show ?thesis | |
| 333 | by (simp add: exp_eq_0_imp_not_bit) | |
| 334 | next | |
| 335 | case False | |
| 336 | with that show ?thesis proof (induction n arbitrary: a b) | |
| 337 | case 0 | |
| 338 | from "0.prems"(1) [of 0] show ?case | |
| 339 | by auto | |
| 340 | next | |
| 341 | case (Suc n) | |
| 342 | from Suc.prems(1) [of 0] have even: \<open>even a \<or> even b\<close> | |
| 343 | by auto | |
| 344 | have bit: \<open>\<not> bit (a div 2) n \<or> \<not> bit (b div 2) n\<close> for n | |
| 345 | using Suc.prems(1) [of \<open>Suc n\<close>] by (simp add: bit_Suc) | |
| 346 | from Suc.prems(2) have \<open>2 * 2 ^ n \<noteq> 0\<close> \<open>2 ^ n \<noteq> 0\<close> | |
| 347 | by (auto simp add: mult_2) | |
| 348 | have \<open>a + b = (a div 2 * 2 + a mod 2) + (b div 2 * 2 + b mod 2)\<close> | |
| 349 | using div_mult_mod_eq [of a 2] div_mult_mod_eq [of b 2] by simp | |
| 350 | also have \<open>\<dots> = of_bool (odd a \<or> odd b) + 2 * (a div 2 + b div 2)\<close> | |
| 351 | using even by (auto simp add: algebra_simps mod2_eq_if) | |
| 352 | finally have \<open>bit ((a + b) div 2) n \<longleftrightarrow> bit (a div 2 + b div 2) n\<close> | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 353 | using \<open>2 * 2 ^ n \<noteq> 0\<close> by simp (simp_all flip: bit_Suc add: bit_double_iff possible_bit_def) | 
| 74101 | 354 | also have \<open>\<dots> \<longleftrightarrow> bit (a div 2) n \<or> bit (b div 2) n\<close> | 
| 355 | using bit \<open>2 ^ n \<noteq> 0\<close> by (rule Suc.IH) | |
| 356 | finally show ?case | |
| 357 | by (simp add: bit_Suc) | |
| 358 | qed | |
| 359 | qed | |
| 360 | ||
| 361 | lemma | |
| 362 | exp_add_not_zero_imp_left: \<open>2 ^ m \<noteq> 0\<close> | |
| 363 | and exp_add_not_zero_imp_right: \<open>2 ^ n \<noteq> 0\<close> | |
| 364 | if \<open>2 ^ (m + n) \<noteq> 0\<close> | |
| 365 | proof - | |
| 366 | have \<open>\<not> (2 ^ m = 0 \<or> 2 ^ n = 0)\<close> | |
| 367 | proof (rule notI) | |
| 368 | assume \<open>2 ^ m = 0 \<or> 2 ^ n = 0\<close> | |
| 369 | then have \<open>2 ^ (m + n) = 0\<close> | |
| 370 | by (rule disjE) (simp_all add: power_add) | |
| 371 | with that show False .. | |
| 372 | qed | |
| 373 | then show \<open>2 ^ m \<noteq> 0\<close> and \<open>2 ^ n \<noteq> 0\<close> | |
| 374 | by simp_all | |
| 375 | qed | |
| 376 | ||
| 377 | lemma exp_not_zero_imp_exp_diff_not_zero: | |
| 378 | \<open>2 ^ (n - m) \<noteq> 0\<close> if \<open>2 ^ n \<noteq> 0\<close> | |
| 379 | proof (cases \<open>m \<le> n\<close>) | |
| 380 | case True | |
| 381 | moreover define q where \<open>q = n - m\<close> | |
| 382 | ultimately have \<open>n = m + q\<close> | |
| 383 | by simp | |
| 384 | with that show ?thesis | |
| 385 | by (simp add: exp_add_not_zero_imp_right) | |
| 386 | next | |
| 387 | case False | |
| 388 | with that show ?thesis | |
| 389 | by simp | |
| 390 | qed | |
| 391 | ||
| 392 | end | |
| 393 | ||
| 394 | lemma nat_bit_induct [case_names zero even odd]: | |
| 395 | "P n" if zero: "P 0" | |
| 396 | and even: "\<And>n. P n \<Longrightarrow> n > 0 \<Longrightarrow> P (2 * n)" | |
| 397 | and odd: "\<And>n. P n \<Longrightarrow> P (Suc (2 * n))" | |
| 398 | proof (induction n rule: less_induct) | |
| 399 | case (less n) | |
| 400 | show "P n" | |
| 401 | proof (cases "n = 0") | |
| 402 | case True with zero show ?thesis by simp | |
| 403 | next | |
| 404 | case False | |
| 405 | with less have hyp: "P (n div 2)" by simp | |
| 406 | show ?thesis | |
| 407 | proof (cases "even n") | |
| 408 | case True | |
| 409 | then have "n \<noteq> 1" | |
| 410 | by auto | |
| 411 | with \<open>n \<noteq> 0\<close> have "n div 2 > 0" | |
| 412 | by simp | |
| 413 | with \<open>even n\<close> hyp even [of "n div 2"] show ?thesis | |
| 414 | by simp | |
| 415 | next | |
| 416 | case False | |
| 417 | with hyp odd [of "n div 2"] show ?thesis | |
| 418 | by simp | |
| 419 | qed | |
| 420 | qed | |
| 421 | qed | |
| 422 | ||
| 423 | instantiation nat :: semiring_bits | |
| 424 | begin | |
| 425 | ||
| 426 | definition bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> bool\<close> | |
| 427 | where \<open>bit_nat m n \<longleftrightarrow> odd (m div 2 ^ n)\<close> | |
| 428 | ||
| 429 | instance | |
| 430 | proof | |
| 431 | show \<open>P n\<close> if stable: \<open>\<And>n. n div 2 = n \<Longrightarrow> P n\<close> | |
| 432 | 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> | |
| 433 | for P and n :: nat | |
| 434 | proof (induction n rule: nat_bit_induct) | |
| 435 | case zero | |
| 436 | from stable [of 0] show ?case | |
| 437 | by simp | |
| 438 | next | |
| 439 | case (even n) | |
| 440 | with rec [of n False] show ?case | |
| 441 | by simp | |
| 442 | next | |
| 443 | case (odd n) | |
| 444 | with rec [of n True] show ?case | |
| 445 | by simp | |
| 446 | qed | |
| 447 | show \<open>q mod 2 ^ m mod 2 ^ n = q mod 2 ^ min m n\<close> | |
| 448 | for q m n :: nat | |
| 449 | apply (auto simp add: less_iff_Suc_add power_add mod_mod_cancel split: split_min_lin) | |
| 450 | apply (metis div_mult2_eq mod_div_trivial mod_eq_self_iff_div_eq_0 mod_mult_self2_is_0 power_commutes) | |
| 451 | done | |
| 452 | show \<open>(q * 2 ^ m) mod (2 ^ n) = (q mod 2 ^ (n - m)) * 2 ^ m\<close> if \<open>m \<le> n\<close> | |
| 453 | for q m n :: nat | |
| 454 | using that | |
| 455 | apply (auto simp add: mod_mod_cancel div_mult2_eq power_add mod_mult2_eq le_iff_add split: split_min_lin) | |
| 456 | done | |
| 457 | show \<open>even ((2 ^ m - (1::nat)) div 2 ^ n) \<longleftrightarrow> 2 ^ n = (0::nat) \<or> m \<le> n\<close> | |
| 458 | for m n :: nat | |
| 459 | using even_mask_div_iff' [where ?'a = nat, of m n] by simp | |
| 460 | 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> | |
| 461 | for m n q r :: nat | |
| 462 | apply (auto simp add: not_less power_add ac_simps dest!: le_Suc_ex) | |
| 463 | apply (metis (full_types) dvd_mult dvd_mult_imp_div dvd_power_iff_le not_less not_less_eq order_refl power_Suc) | |
| 464 | done | |
| 465 | qed (auto simp add: div_mult2_eq mod_mult2_eq power_add power_diff bit_nat_def) | |
| 466 | ||
| 467 | end | |
| 468 | ||
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 469 | lemma possible_bit_nat[simp]: | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 470 | "possible_bit TYPE(nat) n" | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 471 | by (simp add: possible_bit_def) | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 472 | |
| 74497 | 473 | lemma not_bit_Suc_0_Suc [simp]: | 
| 474 | \<open>\<not> bit (Suc 0) (Suc n)\<close> | |
| 475 | by (simp add: bit_Suc) | |
| 476 | ||
| 477 | lemma not_bit_Suc_0_numeral [simp]: | |
| 478 | \<open>\<not> bit (Suc 0) (numeral n)\<close> | |
| 479 | by (simp add: numeral_eq_Suc) | |
| 480 | ||
| 74101 | 481 | lemma int_bit_induct [case_names zero minus even odd]: | 
| 482 | "P k" if zero_int: "P 0" | |
| 483 | and minus_int: "P (- 1)" | |
| 484 | and even_int: "\<And>k. P k \<Longrightarrow> k \<noteq> 0 \<Longrightarrow> P (k * 2)" | |
| 485 | and odd_int: "\<And>k. P k \<Longrightarrow> k \<noteq> - 1 \<Longrightarrow> P (1 + (k * 2))" for k :: int | |
| 486 | proof (cases "k \<ge> 0") | |
| 487 | case True | |
| 488 | define n where "n = nat k" | |
| 489 | with True have "k = int n" | |
| 490 | by simp | |
| 491 | then show "P k" | |
| 492 | proof (induction n arbitrary: k rule: nat_bit_induct) | |
| 493 | case zero | |
| 494 | then show ?case | |
| 495 | by (simp add: zero_int) | |
| 496 | next | |
| 497 | case (even n) | |
| 498 | have "P (int n * 2)" | |
| 499 | by (rule even_int) (use even in simp_all) | |
| 500 | with even show ?case | |
| 501 | by (simp add: ac_simps) | |
| 502 | next | |
| 503 | case (odd n) | |
| 504 | have "P (1 + (int n * 2))" | |
| 505 | by (rule odd_int) (use odd in simp_all) | |
| 506 | with odd show ?case | |
| 507 | by (simp add: ac_simps) | |
| 508 | qed | |
| 509 | next | |
| 510 | case False | |
| 511 | define n where "n = nat (- k - 1)" | |
| 512 | with False have "k = - int n - 1" | |
| 513 | by simp | |
| 514 | then show "P k" | |
| 515 | proof (induction n arbitrary: k rule: nat_bit_induct) | |
| 516 | case zero | |
| 517 | then show ?case | |
| 518 | by (simp add: minus_int) | |
| 519 | next | |
| 520 | case (even n) | |
| 521 | have "P (1 + (- int (Suc n) * 2))" | |
| 522 | by (rule odd_int) (use even in \<open>simp_all add: algebra_simps\<close>) | |
| 523 | also have "\<dots> = - int (2 * n) - 1" | |
| 524 | by (simp add: algebra_simps) | |
| 525 | finally show ?case | |
| 526 | using even.prems by simp | |
| 527 | next | |
| 528 | case (odd n) | |
| 529 | have "P (- int (Suc n) * 2)" | |
| 530 | by (rule even_int) (use odd in \<open>simp_all add: algebra_simps\<close>) | |
| 531 | also have "\<dots> = - int (Suc (2 * n)) - 1" | |
| 532 | by (simp add: algebra_simps) | |
| 533 | finally show ?case | |
| 534 | using odd.prems by simp | |
| 535 | qed | |
| 536 | qed | |
| 537 | ||
| 538 | context semiring_bits | |
| 539 | begin | |
| 540 | ||
| 541 | lemma bit_of_bool_iff [bit_simps]: | |
| 542 | \<open>bit (of_bool b) n \<longleftrightarrow> b \<and> n = 0\<close> | |
| 543 | by (simp add: bit_1_iff) | |
| 544 | ||
| 545 | lemma bit_of_nat_iff [bit_simps]: | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 546 |   \<open>bit (of_nat m) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> bit m n\<close>
 | 
| 74101 | 547 | proof (cases \<open>(2::'a) ^ n = 0\<close>) | 
| 548 | case True | |
| 549 | then show ?thesis | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 550 | by (simp add: exp_eq_0_imp_not_bit possible_bit_def) | 
| 74101 | 551 | next | 
| 552 | case False | |
| 553 | then have \<open>bit (of_nat m) n \<longleftrightarrow> bit m n\<close> | |
| 554 | proof (induction m arbitrary: n rule: nat_bit_induct) | |
| 555 | case zero | |
| 556 | then show ?case | |
| 557 | by simp | |
| 558 | next | |
| 559 | case (even m) | |
| 560 | then show ?case | |
| 561 | by (cases n) | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 562 | (auto simp add: bit_double_iff Bit_Operations.bit_double_iff possible_bit_def dest: mult_not_zero) | 
| 74101 | 563 | next | 
| 564 | case (odd m) | |
| 565 | then show ?case | |
| 566 | by (cases n) | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 567 | (auto simp add: bit_double_iff even_bit_succ_iff possible_bit_def Bit_Operations.bit_Suc dest: mult_not_zero) | 
| 74101 | 568 | qed | 
| 569 | with False show ?thesis | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 570 | by (simp add: possible_bit_def) | 
| 74101 | 571 | qed | 
| 572 | ||
| 573 | end | |
| 574 | ||
| 575 | instantiation int :: semiring_bits | |
| 576 | begin | |
| 577 | ||
| 578 | definition bit_int :: \<open>int \<Rightarrow> nat \<Rightarrow> bool\<close> | |
| 579 | where \<open>bit_int k n \<longleftrightarrow> odd (k div 2 ^ n)\<close> | |
| 580 | ||
| 581 | instance | |
| 582 | proof | |
| 583 | show \<open>P k\<close> if stable: \<open>\<And>k. k div 2 = k \<Longrightarrow> P k\<close> | |
| 584 | 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> | |
| 585 | for P and k :: int | |
| 586 | proof (induction k rule: int_bit_induct) | |
| 587 | case zero | |
| 588 | from stable [of 0] show ?case | |
| 589 | by simp | |
| 590 | next | |
| 591 | case minus | |
| 592 | from stable [of \<open>- 1\<close>] show ?case | |
| 593 | by simp | |
| 594 | next | |
| 595 | case (even k) | |
| 596 | with rec [of k False] show ?case | |
| 597 | by (simp add: ac_simps) | |
| 598 | next | |
| 599 | case (odd k) | |
| 600 | with rec [of k True] show ?case | |
| 601 | by (simp add: ac_simps) | |
| 602 | qed | |
| 603 | show \<open>(2::int) ^ m div 2 ^ n = of_bool ((2::int) ^ m \<noteq> 0 \<and> n \<le> m) * 2 ^ (m - n)\<close> | |
| 604 | for m n :: nat | |
| 605 | proof (cases \<open>m < n\<close>) | |
| 606 | case True | |
| 607 | then have \<open>n = m + (n - m)\<close> | |
| 608 | by simp | |
| 609 | then have \<open>(2::int) ^ m div 2 ^ n = (2::int) ^ m div 2 ^ (m + (n - m))\<close> | |
| 610 | by simp | |
| 611 | also have \<open>\<dots> = (2::int) ^ m div (2 ^ m * 2 ^ (n - m))\<close> | |
| 612 | by (simp add: power_add) | |
| 613 | also have \<open>\<dots> = (2::int) ^ m div 2 ^ m div 2 ^ (n - m)\<close> | |
| 614 | by (simp add: zdiv_zmult2_eq) | |
| 615 | finally show ?thesis using \<open>m < n\<close> by simp | |
| 616 | next | |
| 617 | case False | |
| 618 | then show ?thesis | |
| 619 | by (simp add: power_diff) | |
| 620 | qed | |
| 621 | show \<open>k mod 2 ^ m mod 2 ^ n = k mod 2 ^ min m n\<close> | |
| 622 | for m n :: nat and k :: int | |
| 623 | using mod_exp_eq [of \<open>nat k\<close> m n] | |
| 624 | apply (auto simp add: mod_mod_cancel zdiv_zmult2_eq power_add zmod_zmult2_eq le_iff_add split: split_min_lin) | |
| 625 | apply (auto simp add: less_iff_Suc_add mod_mod_cancel power_add) | |
| 626 | apply (simp only: flip: mult.left_commute [of \<open>2 ^ m\<close>]) | |
| 627 | apply (subst zmod_zmult2_eq) apply simp_all | |
| 628 | done | |
| 629 | show \<open>(k * 2 ^ m) mod (2 ^ n) = (k mod 2 ^ (n - m)) * 2 ^ m\<close> | |
| 630 | if \<open>m \<le> n\<close> for m n :: nat and k :: int | |
| 631 | using that | |
| 632 | apply (auto simp add: power_add zmod_zmult2_eq le_iff_add split: split_min_lin) | |
| 633 | done | |
| 634 | show \<open>even ((2 ^ m - (1::int)) div 2 ^ n) \<longleftrightarrow> 2 ^ n = (0::int) \<or> m \<le> n\<close> | |
| 635 | for m n :: nat | |
| 636 | using even_mask_div_iff' [where ?'a = int, of m n] by simp | |
| 637 | 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> | |
| 638 | for m n :: nat and k l :: int | |
| 639 | apply (auto simp add: not_less power_add ac_simps dest!: le_Suc_ex) | |
| 640 | apply (metis Suc_leI dvd_mult dvd_mult_imp_div dvd_power_le dvd_refl power.simps(2)) | |
| 641 | done | |
| 642 | qed (auto simp add: zdiv_zmult2_eq zmod_zmult2_eq power_add power_diff not_le bit_int_def) | |
| 643 | ||
| 644 | end | |
| 645 | ||
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 646 | lemma possible_bit_int[simp]: | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 647 | "possible_bit TYPE(int) n" | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 648 | by (simp add: possible_bit_def) | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 649 | |
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 650 | lemma bit_not_int_iff': | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 651 | \<open>bit (- k - 1) n \<longleftrightarrow> \<not> bit k n\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 652 | for k :: int | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 653 | proof (induction n arbitrary: k) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 654 | case 0 | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 655 | show ?case | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 656 | by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 657 | next | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 658 | case (Suc n) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 659 | have \<open>- k - 1 = - (k + 2) + 1\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 660 | by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 661 | also have \<open>(- (k + 2) + 1) div 2 = - (k div 2) - 1\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 662 | proof (cases \<open>even k\<close>) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 663 | case True | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 664 | then have \<open>- k div 2 = - (k div 2)\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 665 | by rule (simp flip: mult_minus_right) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 666 | with True show ?thesis | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 667 | by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 668 | next | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 669 | case False | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 670 | have \<open>4 = 2 * (2::int)\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 671 | by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 672 | also have \<open>2 * 2 div 2 = (2::int)\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 673 | by (simp only: nonzero_mult_div_cancel_left) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 674 | finally have *: \<open>4 div 2 = (2::int)\<close> . | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 675 | from False obtain l where k: \<open>k = 2 * l + 1\<close> .. | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 676 | then have \<open>- k - 2 = 2 * - (l + 2) + 1\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 677 | by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 678 | then have \<open>(- k - 2) div 2 + 1 = - (k div 2) - 1\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 679 | by (simp flip: mult_minus_right add: *) (simp add: k) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 680 | with False show ?thesis | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 681 | by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 682 | qed | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 683 | finally have \<open>(- k - 1) div 2 = - (k div 2) - 1\<close> . | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 684 | with Suc show ?case | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 685 | by (simp add: bit_Suc) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 686 | qed | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 687 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 688 | lemma bit_nat_iff [bit_simps]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 689 | \<open>bit (nat k) n \<longleftrightarrow> k \<ge> 0 \<and> bit k n\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 690 | proof (cases \<open>k \<ge> 0\<close>) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 691 | case True | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 692 | moreover define m where \<open>m = nat k\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 693 | ultimately have \<open>k = int m\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 694 | by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 695 | then show ?thesis | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 696 | by (simp add: bit_simps) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 697 | next | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 698 | case False | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 699 | then show ?thesis | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 700 | by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 701 | qed | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 702 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 703 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 704 | subsection \<open>Bit operations\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 705 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 706 | class semiring_bit_operations = semiring_bits + | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 707 | fixes "and" :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close> (infixr \<open>AND\<close> 64) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 708 | and or :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close> (infixr \<open>OR\<close> 59) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 709 | and xor :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close> (infixr \<open>XOR\<close> 59) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 710 | and mask :: \<open>nat \<Rightarrow> 'a\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 711 | and set_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 712 | and unset_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 713 | and flip_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 714 | and push_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 715 | and drop_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 716 | and take_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 717 | assumes bit_and_iff [bit_simps]: \<open>bit (a AND b) n \<longleftrightarrow> bit a n \<and> bit b n\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 718 | and bit_or_iff [bit_simps]: \<open>bit (a OR b) n \<longleftrightarrow> bit a n \<or> bit b n\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 719 | and bit_xor_iff [bit_simps]: \<open>bit (a XOR b) n \<longleftrightarrow> bit a n \<noteq> bit b n\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 720 | and mask_eq_exp_minus_1: \<open>mask n = 2 ^ n - 1\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 721 | and set_bit_eq_or: \<open>set_bit n a = a OR push_bit n 1\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 722 | and bit_unset_bit_iff [bit_simps]: \<open>bit (unset_bit m a) n \<longleftrightarrow> bit a n \<and> m \<noteq> n\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 723 | and flip_bit_eq_xor: \<open>flip_bit n a = a XOR push_bit n 1\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 724 | and push_bit_eq_mult: \<open>push_bit n a = a * 2 ^ n\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 725 | and drop_bit_eq_div: \<open>drop_bit n a = a div 2 ^ n\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 726 | and take_bit_eq_mod: \<open>take_bit n a = a mod 2 ^ n\<close> | 
| 71042 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 727 | begin | 
| 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 728 | |
| 74101 | 729 | text \<open> | 
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 730 | We want the bitwise operations to bind slightly weaker | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 731 | than \<open>+\<close> and \<open>-\<close>. | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 732 | |
| 74101 | 733 | Logically, \<^const>\<open>push_bit\<close>, | 
| 734 | \<^const>\<open>drop_bit\<close> and \<^const>\<open>take_bit\<close> are just aliases; having them | |
| 735 | as separate operations makes proofs easier, otherwise proof automation | |
| 736 | would fiddle with concrete expressions \<^term>\<open>2 ^ n\<close> in a way obfuscating the basic | |
| 737 | algebraic relationships between those operations. | |
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 738 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 739 | For the sake of code generation operations | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 740 | are specified as definitional class operations, | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 741 | taking into account that specific instances of these can be implemented | 
| 74101 | 742 | differently wrt. code generation. | 
| 743 | \<close> | |
| 744 | ||
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 745 | sublocale "and": semilattice \<open>(AND)\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 746 | by standard (auto simp add: bit_eq_iff bit_and_iff) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 747 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 748 | sublocale or: semilattice_neutr \<open>(OR)\<close> 0 | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 749 | by standard (auto simp add: bit_eq_iff bit_or_iff) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 750 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 751 | sublocale xor: comm_monoid \<open>(XOR)\<close> 0 | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 752 | by standard (auto simp add: bit_eq_iff bit_xor_iff) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 753 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 754 | lemma even_and_iff: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 755 | \<open>even (a AND b) \<longleftrightarrow> even a \<or> even b\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 756 | using bit_and_iff [of a b 0] by auto | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 757 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 758 | lemma even_or_iff: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 759 | \<open>even (a OR b) \<longleftrightarrow> even a \<and> even b\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 760 | using bit_or_iff [of a b 0] by auto | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 761 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 762 | lemma even_xor_iff: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 763 | \<open>even (a XOR b) \<longleftrightarrow> (even a \<longleftrightarrow> even b)\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 764 | using bit_xor_iff [of a b 0] by auto | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 765 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 766 | lemma zero_and_eq [simp]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 767 | \<open>0 AND a = 0\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 768 | by (simp add: bit_eq_iff bit_and_iff) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 769 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 770 | lemma and_zero_eq [simp]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 771 | \<open>a AND 0 = 0\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 772 | by (simp add: bit_eq_iff bit_and_iff) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 773 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 774 | lemma one_and_eq: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 775 | \<open>1 AND a = a mod 2\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 776 | by (simp add: bit_eq_iff bit_and_iff) (auto simp add: bit_1_iff) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 777 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 778 | lemma and_one_eq: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 779 | \<open>a AND 1 = a mod 2\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 780 | using one_and_eq [of a] by (simp add: ac_simps) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 781 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 782 | lemma one_or_eq: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 783 | \<open>1 OR a = a + of_bool (even a)\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 784 | by (simp add: bit_eq_iff bit_or_iff add.commute [of _ 1] even_bit_succ_iff) (auto simp add: bit_1_iff) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 785 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 786 | lemma or_one_eq: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 787 | \<open>a OR 1 = a + of_bool (even a)\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 788 | using one_or_eq [of a] by (simp add: ac_simps) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 789 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 790 | lemma one_xor_eq: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 791 | \<open>1 XOR a = a + of_bool (even a) - of_bool (odd a)\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 792 | 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) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 793 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 794 | lemma xor_one_eq: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 795 | \<open>a XOR 1 = a + of_bool (even a) - of_bool (odd a)\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 796 | using one_xor_eq [of a] by (simp add: ac_simps) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 797 | |
| 74163 | 798 | lemma xor_self_eq [simp]: | 
| 799 | \<open>a XOR a = 0\<close> | |
| 800 | by (rule bit_eqI) (simp add: bit_simps) | |
| 801 | ||
| 74101 | 802 | lemma bit_iff_odd_drop_bit: | 
| 803 | \<open>bit a n \<longleftrightarrow> odd (drop_bit n a)\<close> | |
| 804 | by (simp add: bit_iff_odd drop_bit_eq_div) | |
| 805 | ||
| 806 | lemma even_drop_bit_iff_not_bit: | |
| 807 | \<open>even (drop_bit n a) \<longleftrightarrow> \<not> bit a n\<close> | |
| 808 | by (simp add: bit_iff_odd_drop_bit) | |
| 809 | ||
| 810 | lemma div_push_bit_of_1_eq_drop_bit: | |
| 811 | \<open>a div push_bit n 1 = drop_bit n a\<close> | |
| 812 | by (simp add: push_bit_eq_mult drop_bit_eq_div) | |
| 813 | ||
| 814 | lemma bits_ident: | |
| 815 | "push_bit n (drop_bit n a) + take_bit n a = a" | |
| 816 | using div_mult_mod_eq by (simp add: push_bit_eq_mult take_bit_eq_mod drop_bit_eq_div) | |
| 817 | ||
| 818 | lemma push_bit_push_bit [simp]: | |
| 819 | "push_bit m (push_bit n a) = push_bit (m + n) a" | |
| 820 | by (simp add: push_bit_eq_mult power_add ac_simps) | |
| 821 | ||
| 822 | lemma push_bit_0_id [simp]: | |
| 823 | "push_bit 0 = id" | |
| 824 | by (simp add: fun_eq_iff push_bit_eq_mult) | |
| 825 | ||
| 826 | lemma push_bit_of_0 [simp]: | |
| 827 | "push_bit n 0 = 0" | |
| 828 | by (simp add: push_bit_eq_mult) | |
| 829 | ||
| 74592 | 830 | lemma push_bit_of_1 [simp]: | 
| 74101 | 831 | "push_bit n 1 = 2 ^ n" | 
| 832 | by (simp add: push_bit_eq_mult) | |
| 833 | ||
| 834 | lemma push_bit_Suc [simp]: | |
| 835 | "push_bit (Suc n) a = push_bit n (a * 2)" | |
| 836 | by (simp add: push_bit_eq_mult ac_simps) | |
| 837 | ||
| 838 | lemma push_bit_double: | |
| 839 | "push_bit n (a * 2) = push_bit n a * 2" | |
| 840 | by (simp add: push_bit_eq_mult ac_simps) | |
| 841 | ||
| 842 | lemma push_bit_add: | |
| 843 | "push_bit n (a + b) = push_bit n a + push_bit n b" | |
| 844 | by (simp add: push_bit_eq_mult algebra_simps) | |
| 845 | ||
| 846 | lemma push_bit_numeral [simp]: | |
| 847 | \<open>push_bit (numeral l) (numeral k) = push_bit (pred_numeral l) (numeral (Num.Bit0 k))\<close> | |
| 848 | by (simp add: numeral_eq_Suc mult_2_right) (simp add: numeral_Bit0) | |
| 849 | ||
| 850 | lemma take_bit_0 [simp]: | |
| 851 | "take_bit 0 a = 0" | |
| 852 | by (simp add: take_bit_eq_mod) | |
| 853 | ||
| 854 | lemma take_bit_Suc: | |
| 855 | \<open>take_bit (Suc n) a = take_bit n (a div 2) * 2 + a mod 2\<close> | |
| 856 | proof - | |
| 857 | 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> | |
| 858 | using even_succ_mod_exp [of \<open>2 * (a div 2)\<close> \<open>Suc n\<close>] | |
| 859 | mult_exp_mod_exp_eq [of 1 \<open>Suc n\<close> \<open>a div 2\<close>] | |
| 860 | by (auto simp add: take_bit_eq_mod ac_simps) | |
| 861 | then show ?thesis | |
| 862 | using div_mult_mod_eq [of a 2] by (simp add: mod_2_eq_odd) | |
| 863 | qed | |
| 864 | ||
| 865 | lemma take_bit_rec: | |
| 866 | \<open>take_bit n a = (if n = 0 then 0 else take_bit (n - 1) (a div 2) * 2 + a mod 2)\<close> | |
| 867 | by (cases n) (simp_all add: take_bit_Suc) | |
| 868 | ||
| 869 | lemma take_bit_Suc_0 [simp]: | |
| 870 | \<open>take_bit (Suc 0) a = a mod 2\<close> | |
| 871 | by (simp add: take_bit_eq_mod) | |
| 872 | ||
| 873 | lemma take_bit_of_0 [simp]: | |
| 874 | "take_bit n 0 = 0" | |
| 875 | by (simp add: take_bit_eq_mod) | |
| 876 | ||
| 877 | lemma take_bit_of_1 [simp]: | |
| 878 | "take_bit n 1 = of_bool (n > 0)" | |
| 879 | by (cases n) (simp_all add: take_bit_Suc) | |
| 880 | ||
| 881 | lemma drop_bit_of_0 [simp]: | |
| 882 | "drop_bit n 0 = 0" | |
| 883 | by (simp add: drop_bit_eq_div) | |
| 884 | ||
| 885 | lemma drop_bit_of_1 [simp]: | |
| 886 | "drop_bit n 1 = of_bool (n = 0)" | |
| 887 | by (simp add: drop_bit_eq_div) | |
| 888 | ||
| 889 | lemma drop_bit_0 [simp]: | |
| 890 | "drop_bit 0 = id" | |
| 891 | by (simp add: fun_eq_iff drop_bit_eq_div) | |
| 892 | ||
| 893 | lemma drop_bit_Suc: | |
| 894 | "drop_bit (Suc n) a = drop_bit n (a div 2)" | |
| 895 | using div_exp_eq [of a 1] by (simp add: drop_bit_eq_div) | |
| 896 | ||
| 897 | lemma drop_bit_rec: | |
| 898 | "drop_bit n a = (if n = 0 then a else drop_bit (n - 1) (a div 2))" | |
| 899 | by (cases n) (simp_all add: drop_bit_Suc) | |
| 900 | ||
| 901 | lemma drop_bit_half: | |
| 902 | "drop_bit n (a div 2) = drop_bit n a div 2" | |
| 903 | by (induction n arbitrary: a) (simp_all add: drop_bit_Suc) | |
| 904 | ||
| 905 | lemma drop_bit_of_bool [simp]: | |
| 906 | "drop_bit n (of_bool b) = of_bool (n = 0 \<and> b)" | |
| 907 | by (cases n) simp_all | |
| 908 | ||
| 909 | lemma even_take_bit_eq [simp]: | |
| 910 | \<open>even (take_bit n a) \<longleftrightarrow> n = 0 \<or> even a\<close> | |
| 911 | by (simp add: take_bit_rec [of n a]) | |
| 912 | ||
| 913 | lemma take_bit_take_bit [simp]: | |
| 914 | "take_bit m (take_bit n a) = take_bit (min m n) a" | |
| 915 | by (simp add: take_bit_eq_mod mod_exp_eq ac_simps) | |
| 916 | ||
| 917 | lemma drop_bit_drop_bit [simp]: | |
| 918 | "drop_bit m (drop_bit n a) = drop_bit (m + n) a" | |
| 919 | by (simp add: drop_bit_eq_div power_add div_exp_eq ac_simps) | |
| 920 | ||
| 921 | lemma push_bit_take_bit: | |
| 922 | "push_bit m (take_bit n a) = take_bit (m + n) (push_bit m a)" | |
| 923 | apply (simp add: push_bit_eq_mult take_bit_eq_mod power_add ac_simps) | |
| 924 | using mult_exp_mod_exp_eq [of m \<open>m + n\<close> a] apply (simp add: ac_simps power_add) | |
| 925 | done | |
| 926 | ||
| 927 | lemma take_bit_push_bit: | |
| 928 | "take_bit m (push_bit n a) = push_bit n (take_bit (m - n) a)" | |
| 929 | proof (cases "m \<le> n") | |
| 930 | case True | |
| 931 | then show ?thesis | |
| 932 | apply (simp add:) | |
| 933 | apply (simp_all add: push_bit_eq_mult take_bit_eq_mod) | |
| 934 | apply (auto dest!: le_Suc_ex simp add: power_add ac_simps) | |
| 935 | using mult_exp_mod_exp_eq [of m m \<open>a * 2 ^ n\<close> for n] | |
| 936 | apply (simp add: ac_simps) | |
| 937 | done | |
| 938 | next | |
| 939 | case False | |
| 940 | then show ?thesis | |
| 941 | using push_bit_take_bit [of n "m - n" a] | |
| 942 | by simp | |
| 943 | qed | |
| 944 | ||
| 945 | lemma take_bit_drop_bit: | |
| 946 | "take_bit m (drop_bit n a) = drop_bit n (take_bit (m + n) a)" | |
| 947 | by (simp add: drop_bit_eq_div take_bit_eq_mod ac_simps div_exp_mod_exp_eq) | |
| 948 | ||
| 949 | lemma drop_bit_take_bit: | |
| 950 | "drop_bit m (take_bit n a) = take_bit (n - m) (drop_bit m a)" | |
| 951 | proof (cases "m \<le> n") | |
| 952 | case True | |
| 953 | then show ?thesis | |
| 954 | using take_bit_drop_bit [of "n - m" m a] by simp | |
| 955 | next | |
| 956 | case False | |
| 957 | then obtain q where \<open>m = n + q\<close> | |
| 958 | by (auto simp add: not_le dest: less_imp_Suc_add) | |
| 959 | then have \<open>drop_bit m (take_bit n a) = 0\<close> | |
| 960 | using div_exp_eq [of \<open>a mod 2 ^ n\<close> n q] | |
| 961 | by (simp add: take_bit_eq_mod drop_bit_eq_div) | |
| 962 | with False show ?thesis | |
| 963 | by simp | |
| 964 | qed | |
| 965 | ||
| 966 | lemma even_push_bit_iff [simp]: | |
| 967 | \<open>even (push_bit n a) \<longleftrightarrow> n \<noteq> 0 \<or> even a\<close> | |
| 968 | by (simp add: push_bit_eq_mult) auto | |
| 969 | ||
| 970 | lemma bit_push_bit_iff [bit_simps]: | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 971 |   \<open>bit (push_bit m a) n \<longleftrightarrow> m \<le> n \<and> possible_bit TYPE('a) n \<and> bit a (n - m)\<close>
 | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 972 | by (auto simp add: bit_iff_odd push_bit_eq_mult even_mult_exp_div_exp_iff possible_bit_def) | 
| 74101 | 973 | |
| 974 | lemma bit_drop_bit_eq [bit_simps]: | |
| 975 | \<open>bit (drop_bit n a) = bit a \<circ> (+) n\<close> | |
| 976 | by (simp add: bit_iff_odd fun_eq_iff ac_simps flip: drop_bit_eq_div) | |
| 977 | ||
| 978 | lemma bit_take_bit_iff [bit_simps]: | |
| 979 | \<open>bit (take_bit m a) n \<longleftrightarrow> n < m \<and> bit a n\<close> | |
| 980 | by (simp add: bit_iff_odd drop_bit_take_bit not_le flip: drop_bit_eq_div) | |
| 981 | ||
| 982 | lemma stable_imp_drop_bit_eq: | |
| 983 | \<open>drop_bit n a = a\<close> | |
| 984 | if \<open>a div 2 = a\<close> | |
| 985 | by (induction n) (simp_all add: that drop_bit_Suc) | |
| 986 | ||
| 987 | lemma stable_imp_take_bit_eq: | |
| 988 | \<open>take_bit n a = (if even a then 0 else 2 ^ n - 1)\<close> | |
| 989 | if \<open>a div 2 = a\<close> | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 990 | proof (rule bit_eqI[unfolded possible_bit_def]) | 
| 74101 | 991 | fix m | 
| 992 | assume \<open>2 ^ m \<noteq> 0\<close> | |
| 993 | with that show \<open>bit (take_bit n a) m \<longleftrightarrow> bit (if even a then 0 else 2 ^ n - 1) m\<close> | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 994 | by (simp add: bit_take_bit_iff bit_mask_sub_iff possible_bit_def stable_imp_bit_iff_odd) | 
| 74101 | 995 | qed | 
| 996 | ||
| 997 | lemma exp_dvdE: | |
| 998 | assumes \<open>2 ^ n dvd a\<close> | |
| 999 | obtains b where \<open>a = push_bit n b\<close> | |
| 1000 | proof - | |
| 1001 | from assms obtain b where \<open>a = 2 ^ n * b\<close> .. | |
| 1002 | then have \<open>a = push_bit n b\<close> | |
| 1003 | by (simp add: push_bit_eq_mult ac_simps) | |
| 1004 | with that show thesis . | |
| 1005 | qed | |
| 1006 | ||
| 1007 | lemma take_bit_eq_0_iff: | |
| 1008 | \<open>take_bit n a = 0 \<longleftrightarrow> 2 ^ n dvd a\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) | |
| 1009 | proof | |
| 1010 | assume ?P | |
| 1011 | then show ?Q | |
| 1012 | by (simp add: take_bit_eq_mod mod_0_imp_dvd) | |
| 1013 | next | |
| 1014 | assume ?Q | |
| 1015 | then obtain b where \<open>a = push_bit n b\<close> | |
| 1016 | by (rule exp_dvdE) | |
| 1017 | then show ?P | |
| 1018 | by (simp add: take_bit_push_bit) | |
| 1019 | qed | |
| 1020 | ||
| 1021 | lemma take_bit_tightened: | |
| 1022 | \<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> | |
| 1023 | proof - | |
| 1024 | from that have \<open>take_bit m (take_bit n a) = take_bit m (take_bit n b)\<close> | |
| 1025 | by simp | |
| 1026 | then have \<open>take_bit (min m n) a = take_bit (min m n) b\<close> | |
| 1027 | by simp | |
| 1028 | with that show ?thesis | |
| 1029 | by (simp add: min_def) | |
| 1030 | qed | |
| 1031 | ||
| 1032 | lemma take_bit_eq_self_iff_drop_bit_eq_0: | |
| 1033 | \<open>take_bit n a = a \<longleftrightarrow> drop_bit n a = 0\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) | |
| 1034 | proof | |
| 1035 | assume ?P | |
| 1036 | show ?Q | |
| 1037 | proof (rule bit_eqI) | |
| 1038 | fix m | |
| 1039 | from \<open>?P\<close> have \<open>a = take_bit n a\<close> .. | |
| 1040 | also have \<open>\<not> bit (take_bit n a) (n + m)\<close> | |
| 1041 | unfolding bit_simps | |
| 1042 | by (simp add: bit_simps) | |
| 1043 | finally show \<open>bit (drop_bit n a) m \<longleftrightarrow> bit 0 m\<close> | |
| 1044 | by (simp add: bit_simps) | |
| 1045 | qed | |
| 1046 | next | |
| 1047 | assume ?Q | |
| 1048 | show ?P | |
| 1049 | proof (rule bit_eqI) | |
| 1050 | fix m | |
| 1051 | from \<open>?Q\<close> have \<open>\<not> bit (drop_bit n a) (m - n)\<close> | |
| 1052 | by simp | |
| 1053 | then have \<open> \<not> bit a (n + (m - n))\<close> | |
| 1054 | by (simp add: bit_simps) | |
| 1055 | then show \<open>bit (take_bit n a) m \<longleftrightarrow> bit a m\<close> | |
| 1056 | by (cases \<open>m < n\<close>) (auto simp add: bit_simps) | |
| 1057 | qed | |
| 1058 | qed | |
| 1059 | ||
| 1060 | lemma drop_bit_exp_eq: | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1061 |   \<open>drop_bit m (2 ^ n) = of_bool (m \<le> n \<and> possible_bit TYPE('a) n) * 2 ^ (n - m)\<close>
 | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1062 | by (auto simp add: bit_eq_iff bit_simps) | 
| 74101 | 1063 | |
| 71409 | 1064 | lemma take_bit_and [simp]: | 
| 1065 | \<open>take_bit n (a AND b) = take_bit n a AND take_bit n b\<close> | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1066 | by (auto simp add: bit_eq_iff bit_simps) | 
| 71409 | 1067 | |
| 1068 | lemma take_bit_or [simp]: | |
| 1069 | \<open>take_bit n (a OR b) = take_bit n a OR take_bit n b\<close> | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1070 | by (auto simp add: bit_eq_iff bit_simps) | 
| 71409 | 1071 | |
| 1072 | lemma take_bit_xor [simp]: | |
| 1073 | \<open>take_bit n (a XOR b) = take_bit n a XOR take_bit n b\<close> | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1074 | by (auto simp add: bit_eq_iff bit_simps) | 
| 71409 | 1075 | |
| 72239 | 1076 | lemma push_bit_and [simp]: | 
| 1077 | \<open>push_bit n (a AND b) = push_bit n a AND push_bit n b\<close> | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1078 | by (auto simp add: bit_eq_iff bit_simps) | 
| 72239 | 1079 | |
| 1080 | lemma push_bit_or [simp]: | |
| 1081 | \<open>push_bit n (a OR b) = push_bit n a OR push_bit n b\<close> | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1082 | by (auto simp add: bit_eq_iff bit_simps) | 
| 72239 | 1083 | |
| 1084 | lemma push_bit_xor [simp]: | |
| 1085 | \<open>push_bit n (a XOR b) = push_bit n a XOR push_bit n b\<close> | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1086 | by (auto simp add: bit_eq_iff bit_simps) | 
| 72239 | 1087 | |
| 1088 | lemma drop_bit_and [simp]: | |
| 1089 | \<open>drop_bit n (a AND b) = drop_bit n a AND drop_bit n b\<close> | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1090 | by (auto simp add: bit_eq_iff bit_simps) | 
| 72239 | 1091 | |
| 1092 | lemma drop_bit_or [simp]: | |
| 1093 | \<open>drop_bit n (a OR b) = drop_bit n a OR drop_bit n b\<close> | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1094 | by (auto simp add: bit_eq_iff bit_simps) | 
| 72239 | 1095 | |
| 1096 | lemma drop_bit_xor [simp]: | |
| 1097 | \<open>drop_bit n (a XOR b) = drop_bit n a XOR drop_bit n b\<close> | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1098 | by (auto simp add: bit_eq_iff bit_simps) | 
| 72239 | 1099 | |
| 72611 
c7bc3e70a8c7
official collection for bit projection simplifications
 haftmann parents: 
72512diff
changeset | 1100 | lemma bit_mask_iff [bit_simps]: | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1101 |   \<open>bit (mask m) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> n < m\<close>
 | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1102 | by (simp add: mask_eq_exp_minus_1 bit_mask_sub_iff) | 
| 71823 | 1103 | |
| 1104 | lemma even_mask_iff: | |
| 1105 | \<open>even (mask n) \<longleftrightarrow> n = 0\<close> | |
| 1106 | using bit_mask_iff [of n 0] by auto | |
| 1107 | ||
| 72082 | 1108 | lemma mask_0 [simp]: | 
| 71823 | 1109 | \<open>mask 0 = 0\<close> | 
| 1110 | by (simp add: mask_eq_exp_minus_1) | |
| 1111 | ||
| 72082 | 1112 | lemma mask_Suc_0 [simp]: | 
| 1113 | \<open>mask (Suc 0) = 1\<close> | |
| 1114 | by (simp add: mask_eq_exp_minus_1 add_implies_diff sym) | |
| 1115 | ||
| 1116 | lemma mask_Suc_exp: | |
| 71823 | 1117 | \<open>mask (Suc n) = 2 ^ n OR mask n\<close> | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1118 | by (auto simp add: bit_eq_iff bit_simps) | 
| 71823 | 1119 | |
| 1120 | lemma mask_Suc_double: | |
| 72082 | 1121 | \<open>mask (Suc n) = 1 OR 2 * mask n\<close> | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1122 | by (auto simp add: bit_eq_iff bit_simps elim: possible_bit_less_imp) | 
| 71823 | 1123 | |
| 72082 | 1124 | lemma mask_numeral: | 
| 1125 | \<open>mask (numeral n) = 1 + 2 * mask (pred_numeral n)\<close> | |
| 1126 | by (simp add: numeral_eq_Suc mask_Suc_double one_or_eq ac_simps) | |
| 1127 | ||
| 74592 | 1128 | lemma take_bit_of_mask [simp]: | 
| 72830 | 1129 | \<open>take_bit m (mask n) = mask (min m n)\<close> | 
| 1130 | by (rule bit_eqI) (simp add: bit_simps) | |
| 1131 | ||
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71956diff
changeset | 1132 | lemma take_bit_eq_mask: | 
| 71823 | 1133 | \<open>take_bit n a = a AND mask n\<close> | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1134 | by (auto simp add: bit_eq_iff bit_simps) | 
| 71823 | 1135 | |
| 72281 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 haftmann parents: 
72262diff
changeset | 1136 | lemma or_eq_0_iff: | 
| 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 haftmann parents: 
72262diff
changeset | 1137 | \<open>a OR b = 0 \<longleftrightarrow> a = 0 \<and> b = 0\<close> | 
| 72792 | 1138 | by (auto simp add: bit_eq_iff bit_or_iff) | 
| 72281 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 haftmann parents: 
72262diff
changeset | 1139 | |
| 72239 | 1140 | lemma disjunctive_add: | 
| 1141 | \<open>a + b = a OR b\<close> if \<open>\<And>n. \<not> bit a n \<or> \<not> bit b n\<close> | |
| 1142 | by (rule bit_eqI) (use that in \<open>simp add: bit_disjunctive_add_iff bit_or_iff\<close>) | |
| 1143 | ||
| 72508 | 1144 | lemma bit_iff_and_drop_bit_eq_1: | 
| 1145 | \<open>bit a n \<longleftrightarrow> drop_bit n a AND 1 = 1\<close> | |
| 1146 | by (simp add: bit_iff_odd_drop_bit and_one_eq odd_iff_mod_2_eq_one) | |
| 1147 | ||
| 1148 | lemma bit_iff_and_push_bit_not_eq_0: | |
| 1149 | \<open>bit a n \<longleftrightarrow> a AND push_bit n 1 \<noteq> 0\<close> | |
| 1150 | apply (cases \<open>2 ^ n = 0\<close>) | |
| 74592 | 1151 | apply (simp_all add: bit_eq_iff bit_and_iff bit_push_bit_iff exp_eq_0_imp_not_bit) | 
| 72508 | 1152 | apply (simp_all add: bit_exp_iff) | 
| 1153 | done | |
| 1154 | ||
| 73682 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1155 | lemmas set_bit_def = set_bit_eq_or | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1156 | |
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1157 | lemma bit_set_bit_iff [bit_simps]: | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1158 |   \<open>bit (set_bit m a) n \<longleftrightarrow> bit a n \<or> (m = n \<and> possible_bit TYPE('a) n)\<close>
 | 
| 74592 | 1159 | by (auto simp add: set_bit_def bit_or_iff bit_exp_iff) | 
| 73682 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1160 | |
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1161 | lemma even_set_bit_iff: | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1162 | \<open>even (set_bit m a) \<longleftrightarrow> even a \<and> m \<noteq> 0\<close> | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1163 | using bit_set_bit_iff [of m a 0] by auto | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1164 | |
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1165 | lemma even_unset_bit_iff: | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1166 | \<open>even (unset_bit m a) \<longleftrightarrow> even a \<or> m = 0\<close> | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1167 | using bit_unset_bit_iff [of m a 0] by auto | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1168 | |
| 73789 | 1169 | lemma and_exp_eq_0_iff_not_bit: | 
| 1170 | \<open>a AND 2 ^ n = 0 \<longleftrightarrow> \<not> bit a n\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1171 | using bit_imp_possible_bit[of a n] | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1172 | by (auto simp add: bit_eq_iff bit_simps) | 
| 73789 | 1173 | |
| 73682 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1174 | lemmas flip_bit_def = flip_bit_eq_xor | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1175 | |
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1176 | lemma bit_flip_bit_iff [bit_simps]: | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1177 |   \<open>bit (flip_bit m a) n \<longleftrightarrow> (m = n \<longleftrightarrow> \<not> bit a n) \<and> possible_bit TYPE('a) n\<close>
 | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1178 | by (auto simp add: bit_eq_iff bit_simps flip_bit_eq_xor bit_imp_possible_bit) | 
| 73682 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1179 | |
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1180 | lemma even_flip_bit_iff: | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1181 | \<open>even (flip_bit m a) \<longleftrightarrow> \<not> (even a \<longleftrightarrow> m = 0)\<close> | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1182 | using bit_flip_bit_iff [of m a 0] by (auto simp: possible_bit_def) | 
| 73682 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1183 | |
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1184 | lemma set_bit_0 [simp]: | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1185 | \<open>set_bit 0 a = 1 + 2 * (a div 2)\<close> | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1186 | by (auto simp add: bit_eq_iff bit_simps even_bit_succ_iff simp flip: bit_Suc) | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1187 | |
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1188 | lemma bit_sum_mult_2_cases: | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1189 | assumes a: "\<forall>j. \<not> bit a (Suc j)" | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1190 | shows "bit (a + 2 * b) n = (if n = 0 then odd a else bit (2 * b) n)" | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1191 | proof - | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1192 | have a_eq: "bit a i \<longleftrightarrow> i = 0 \<and> odd a" for i | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1193 | by (cases i, simp_all add: a) | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1194 | show ?thesis | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1195 | by (simp add: disjunctive_add[simplified disj_imp] a_eq bit_simps) | 
| 73682 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1196 | qed | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1197 | |
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1198 | lemma set_bit_Suc: | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1199 | \<open>set_bit (Suc n) a = a mod 2 + 2 * set_bit n (a div 2)\<close> | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1200 | by (auto simp add: bit_eq_iff bit_sum_mult_2_cases bit_simps bit_Suc[symmetric] | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1201 | elim: possible_bit_less_imp) | 
| 73682 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1202 | |
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1203 | lemma unset_bit_0 [simp]: | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1204 | \<open>unset_bit 0 a = 2 * (a div 2)\<close> | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1205 | by (auto simp add: bit_eq_iff bit_simps even_bit_succ_iff simp flip: bit_Suc) | 
| 73682 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1206 | |
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1207 | lemma unset_bit_Suc: | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1208 | \<open>unset_bit (Suc n) a = a mod 2 + 2 * unset_bit n (a div 2)\<close> | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1209 | by (auto simp add: bit_eq_iff bit_sum_mult_2_cases bit_simps bit_Suc[symmetric] | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1210 | elim: possible_bit_less_imp) | 
| 73682 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1211 | |
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1212 | lemma flip_bit_0 [simp]: | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1213 | \<open>flip_bit 0 a = of_bool (even a) + 2 * (a div 2)\<close> | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1214 | by (auto simp add: bit_eq_iff bit_simps even_bit_succ_iff simp flip: bit_Suc) | 
| 73682 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1215 | |
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1216 | lemma flip_bit_Suc: | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1217 | \<open>flip_bit (Suc n) a = a mod 2 + 2 * flip_bit n (a div 2)\<close> | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1218 | by (auto simp add: bit_eq_iff bit_sum_mult_2_cases bit_simps bit_Suc[symmetric] | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1219 | elim: possible_bit_less_imp) | 
| 73682 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1220 | |
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1221 | lemma flip_bit_eq_if: | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1222 | \<open>flip_bit n a = (if bit a n then unset_bit else set_bit) n a\<close> | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1223 | by (rule bit_eqI) (auto simp add: bit_set_bit_iff bit_unset_bit_iff bit_flip_bit_iff) | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1224 | |
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1225 | lemma take_bit_set_bit_eq: | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1226 | \<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> | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1227 | by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_set_bit_iff) | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1228 | |
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1229 | lemma take_bit_unset_bit_eq: | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1230 | \<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> | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1231 | by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_unset_bit_iff) | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1232 | |
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1233 | lemma take_bit_flip_bit_eq: | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1234 | \<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> | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1235 | by (rule bit_eqI) (auto simp add: bit_take_bit_iff bit_flip_bit_iff) | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1236 | |
| 74497 | 1237 | lemma not_bit_1_Suc [simp]: | 
| 1238 | \<open>\<not> bit 1 (Suc n)\<close> | |
| 1239 | by (simp add: bit_Suc) | |
| 1240 | ||
| 1241 | lemma push_bit_Suc_numeral [simp]: | |
| 1242 | \<open>push_bit (Suc n) (numeral k) = push_bit n (numeral (Num.Bit0 k))\<close> | |
| 1243 | by (simp add: numeral_eq_Suc mult_2_right) (simp add: numeral_Bit0) | |
| 1244 | ||
| 74592 | 1245 | lemma mask_eq_0_iff [simp]: | 
| 1246 | \<open>mask n = 0 \<longleftrightarrow> n = 0\<close> | |
| 1247 | by (cases n) (simp_all add: mask_Suc_double or_eq_0_iff) | |
| 1248 | ||
| 71042 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 1249 | end | 
| 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 1250 | |
| 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 1251 | class ring_bit_operations = semiring_bit_operations + ring_parity + | 
| 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 1252 | fixes not :: \<open>'a \<Rightarrow> 'a\<close> (\<open>NOT\<close>) | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1253 | assumes bit_not_iff_eq: \<open>\<And>n. bit (NOT a) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> \<not> bit a n\<close> | 
| 71409 | 1254 | assumes minus_eq_not_minus_1: \<open>- a = NOT (a - 1)\<close> | 
| 71042 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 1255 | begin | 
| 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 1256 | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1257 | lemmas bit_not_iff[bit_simps] = bit_not_iff_eq[unfolded fold_possible_bit] | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1258 | |
| 71409 | 1259 | text \<open> | 
| 1260 | For the sake of code generation \<^const>\<open>not\<close> is specified as | |
| 1261 | definitional class operation. Note that \<^const>\<open>not\<close> has no | |
| 1262 | sensible definition for unlimited but only positive bit strings | |
| 1263 | (type \<^typ>\<open>nat\<close>). | |
| 1264 | \<close> | |
| 1265 | ||
| 71186 | 1266 | lemma bits_minus_1_mod_2_eq [simp]: | 
| 1267 | \<open>(- 1) mod 2 = 1\<close> | |
| 1268 | by (simp add: mod_2_eq_odd) | |
| 1269 | ||
| 71409 | 1270 | lemma not_eq_complement: | 
| 1271 | \<open>NOT a = - a - 1\<close> | |
| 1272 | using minus_eq_not_minus_1 [of \<open>a + 1\<close>] by simp | |
| 1273 | ||
| 1274 | lemma minus_eq_not_plus_1: | |
| 1275 | \<open>- a = NOT a + 1\<close> | |
| 1276 | using not_eq_complement [of a] by simp | |
| 1277 | ||
| 72611 
c7bc3e70a8c7
official collection for bit projection simplifications
 haftmann parents: 
72512diff
changeset | 1278 | lemma bit_minus_iff [bit_simps]: | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1279 |   \<open>bit (- a) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> \<not> bit (a - 1) n\<close>
 | 
| 71409 | 1280 | by (simp add: minus_eq_not_minus_1 bit_not_iff) | 
| 1281 | ||
| 71418 | 1282 | lemma even_not_iff [simp]: | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1283 | \<open>even (NOT a) \<longleftrightarrow> odd a\<close> | 
| 71418 | 1284 | using bit_not_iff [of a 0] by auto | 
| 1285 | ||
| 72611 
c7bc3e70a8c7
official collection for bit projection simplifications
 haftmann parents: 
72512diff
changeset | 1286 | lemma bit_not_exp_iff [bit_simps]: | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1287 |   \<open>bit (NOT (2 ^ m)) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> n \<noteq> m\<close>
 | 
| 71409 | 1288 | by (auto simp add: bit_not_iff bit_exp_iff) | 
| 1289 | ||
| 71186 | 1290 | lemma bit_minus_1_iff [simp]: | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1291 |   \<open>bit (- 1) n \<longleftrightarrow> possible_bit TYPE('a) n\<close>
 | 
| 71409 | 1292 | by (simp add: bit_minus_iff) | 
| 1293 | ||
| 72611 
c7bc3e70a8c7
official collection for bit projection simplifications
 haftmann parents: 
72512diff
changeset | 1294 | lemma bit_minus_exp_iff [bit_simps]: | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1295 |   \<open>bit (- (2 ^ m)) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> n \<ge> m\<close>
 | 
| 72611 
c7bc3e70a8c7
official collection for bit projection simplifications
 haftmann parents: 
72512diff
changeset | 1296 | by (auto simp add: bit_simps simp flip: mask_eq_exp_minus_1) | 
| 71409 | 1297 | |
| 1298 | lemma bit_minus_2_iff [simp]: | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1299 |   \<open>bit (- 2) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> n > 0\<close>
 | 
| 71409 | 1300 | by (simp add: bit_minus_iff bit_1_iff) | 
| 71186 | 1301 | |
| 74495 | 1302 | lemma not_one_eq [simp]: | 
| 73969 
ca2a35c0fe6e
operations for symbolic computation of bit operations
 haftmann parents: 
73871diff
changeset | 1303 | \<open>NOT 1 = - 2\<close> | 
| 71418 | 1304 | by (simp add: bit_eq_iff bit_not_iff) (simp add: bit_1_iff) | 
| 1305 | ||
| 1306 | sublocale "and": semilattice_neutr \<open>(AND)\<close> \<open>- 1\<close> | |
| 72239 | 1307 | by standard (rule bit_eqI, simp add: bit_and_iff) | 
| 71418 | 1308 | |
| 74123 
7c5842b06114
clarified abstract and concrete boolean algebras
 haftmann parents: 
74108diff
changeset | 1309 | sublocale bit: abstract_boolean_algebra \<open>(AND)\<close> \<open>(OR)\<close> NOT 0 \<open>- 1\<close> | 
| 
7c5842b06114
clarified abstract and concrete boolean algebras
 haftmann parents: 
74108diff
changeset | 1310 | by standard (auto simp add: bit_and_iff bit_or_iff bit_not_iff intro: bit_eqI) | 
| 
7c5842b06114
clarified abstract and concrete boolean algebras
 haftmann parents: 
74108diff
changeset | 1311 | |
| 
7c5842b06114
clarified abstract and concrete boolean algebras
 haftmann parents: 
74108diff
changeset | 1312 | sublocale bit: abstract_boolean_algebra_sym_diff \<open>(AND)\<close> \<open>(OR)\<close> NOT 0 \<open>- 1\<close> \<open>(XOR)\<close> | 
| 
7c5842b06114
clarified abstract and concrete boolean algebras
 haftmann parents: 
74108diff
changeset | 1313 | apply standard | 
| 
7c5842b06114
clarified abstract and concrete boolean algebras
 haftmann parents: 
74108diff
changeset | 1314 | apply (rule bit_eqI) | 
| 
7c5842b06114
clarified abstract and concrete boolean algebras
 haftmann parents: 
74108diff
changeset | 1315 | apply (auto simp add: bit_simps) | 
| 
7c5842b06114
clarified abstract and concrete boolean algebras
 haftmann parents: 
74108diff
changeset | 1316 | done | 
| 71042 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 1317 | |
| 71802 | 1318 | lemma and_eq_not_not_or: | 
| 1319 | \<open>a AND b = NOT (NOT a OR NOT b)\<close> | |
| 1320 | by simp | |
| 1321 | ||
| 1322 | lemma or_eq_not_not_and: | |
| 1323 | \<open>a OR b = NOT (NOT a AND NOT b)\<close> | |
| 1324 | by simp | |
| 1325 | ||
| 72009 | 1326 | lemma not_add_distrib: | 
| 1327 | \<open>NOT (a + b) = NOT a - b\<close> | |
| 1328 | by (simp add: not_eq_complement algebra_simps) | |
| 1329 | ||
| 1330 | lemma not_diff_distrib: | |
| 1331 | \<open>NOT (a - b) = NOT a + b\<close> | |
| 1332 | using not_add_distrib [of a \<open>- b\<close>] by simp | |
| 1333 | ||
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1334 | lemma and_eq_minus_1_iff: | 
| 72281 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 haftmann parents: 
72262diff
changeset | 1335 | \<open>a AND b = - 1 \<longleftrightarrow> a = - 1 \<and> b = - 1\<close> | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1336 | by (auto simp: bit_eq_iff bit_simps) | 
| 72281 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 haftmann parents: 
72262diff
changeset | 1337 | |
| 72239 | 1338 | lemma disjunctive_diff: | 
| 1339 | \<open>a - b = a AND NOT b\<close> if \<open>\<And>n. bit b n \<Longrightarrow> bit a n\<close> | |
| 1340 | proof - | |
| 1341 | have \<open>NOT a + b = NOT a OR b\<close> | |
| 1342 | by (rule disjunctive_add) (auto simp add: bit_not_iff dest: that) | |
| 1343 | then have \<open>NOT (NOT a + b) = NOT (NOT a OR b)\<close> | |
| 1344 | by simp | |
| 1345 | then show ?thesis | |
| 1346 | by (simp add: not_add_distrib) | |
| 1347 | qed | |
| 1348 | ||
| 71412 | 1349 | lemma push_bit_minus: | 
| 1350 | \<open>push_bit n (- a) = - push_bit n a\<close> | |
| 1351 | by (simp add: push_bit_eq_mult) | |
| 1352 | ||
| 71409 | 1353 | lemma take_bit_not_take_bit: | 
| 1354 | \<open>take_bit n (NOT (take_bit n a)) = take_bit n (NOT a)\<close> | |
| 1355 | by (auto simp add: bit_eq_iff bit_take_bit_iff bit_not_iff) | |
| 71042 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 1356 | |
| 71418 | 1357 | lemma take_bit_not_iff: | 
| 73969 
ca2a35c0fe6e
operations for symbolic computation of bit operations
 haftmann parents: 
73871diff
changeset | 1358 | \<open>take_bit n (NOT a) = take_bit n (NOT b) \<longleftrightarrow> take_bit n a = take_bit n b\<close> | 
| 72239 | 1359 | apply (simp add: bit_eq_iff) | 
| 1360 | apply (simp add: bit_not_iff bit_take_bit_iff bit_exp_iff) | |
| 1361 | apply (use exp_eq_0_imp_not_bit in blast) | |
| 71418 | 1362 | done | 
| 1363 | ||
| 72262 | 1364 | lemma take_bit_not_eq_mask_diff: | 
| 1365 | \<open>take_bit n (NOT a) = mask n - take_bit n a\<close> | |
| 1366 | proof - | |
| 1367 | have \<open>take_bit n (NOT a) = take_bit n (NOT (take_bit n a))\<close> | |
| 1368 | by (simp add: take_bit_not_take_bit) | |
| 1369 | also have \<open>\<dots> = mask n AND NOT (take_bit n a)\<close> | |
| 1370 | by (simp add: take_bit_eq_mask ac_simps) | |
| 1371 | also have \<open>\<dots> = mask n - take_bit n a\<close> | |
| 1372 | by (subst disjunctive_diff) | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 1373 | (auto simp add: bit_take_bit_iff bit_mask_iff bit_imp_possible_bit) | 
| 72262 | 1374 | finally show ?thesis | 
| 1375 | by simp | |
| 1376 | qed | |
| 1377 | ||
| 72079 | 1378 | lemma mask_eq_take_bit_minus_one: | 
| 1379 | \<open>mask n = take_bit n (- 1)\<close> | |
| 1380 | by (simp add: bit_eq_iff bit_mask_iff bit_take_bit_iff conj_commute) | |
| 1381 | ||
| 74592 | 1382 | lemma take_bit_minus_one_eq_mask [simp]: | 
| 71922 | 1383 | \<open>take_bit n (- 1) = mask n\<close> | 
| 72079 | 1384 | by (simp add: mask_eq_take_bit_minus_one) | 
| 71922 | 1385 | |
| 72010 | 1386 | lemma minus_exp_eq_not_mask: | 
| 1387 | \<open>- (2 ^ n) = NOT (mask n)\<close> | |
| 1388 | by (rule bit_eqI) (simp add: bit_minus_iff bit_not_iff flip: mask_eq_exp_minus_1) | |
| 1389 | ||
| 74592 | 1390 | lemma push_bit_minus_one_eq_not_mask [simp]: | 
| 71922 | 1391 | \<open>push_bit n (- 1) = NOT (mask n)\<close> | 
| 72010 | 1392 | by (simp add: push_bit_eq_mult minus_exp_eq_not_mask) | 
| 1393 | ||
| 1394 | lemma take_bit_not_mask_eq_0: | |
| 1395 | \<open>take_bit m (NOT (mask n)) = 0\<close> if \<open>n \<ge> m\<close> | |
| 1396 | by (rule bit_eqI) (use that in \<open>simp add: bit_take_bit_iff bit_not_iff bit_mask_iff\<close>) | |
| 71922 | 1397 | |
| 73682 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1398 | lemma unset_bit_eq_and_not: | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1399 | \<open>unset_bit n a = a AND NOT (push_bit n 1)\<close> | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1400 | by (rule bit_eqI) (auto simp add: bit_simps) | 
| 71426 | 1401 | |
| 73682 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1402 | lemmas unset_bit_def = unset_bit_eq_and_not | 
| 71986 | 1403 | |
| 74497 | 1404 | lemma push_bit_Suc_minus_numeral [simp]: | 
| 1405 | \<open>push_bit (Suc n) (- numeral k) = push_bit n (- numeral (Num.Bit0 k))\<close> | |
| 1406 | apply (simp only: numeral_Bit0) | |
| 1407 | apply simp | |
| 1408 | apply (simp only: numeral_mult mult_2_right numeral_add) | |
| 1409 | done | |
| 1410 | ||
| 1411 | lemma push_bit_minus_numeral [simp]: | |
| 1412 | \<open>push_bit (numeral l) (- numeral k) = push_bit (pred_numeral l) (- numeral (Num.Bit0 k))\<close> | |
| 1413 | by (simp only: numeral_eq_Suc push_bit_Suc_minus_numeral) | |
| 1414 | ||
| 74592 | 1415 | lemma take_bit_Suc_minus_1_eq: | 
| 74498 | 1416 | \<open>take_bit (Suc n) (- 1) = 2 ^ Suc n - 1\<close> | 
| 74592 | 1417 | by (simp add: mask_eq_exp_minus_1) | 
| 1418 | ||
| 1419 | lemma take_bit_numeral_minus_1_eq: | |
| 74498 | 1420 | \<open>take_bit (numeral k) (- 1) = 2 ^ numeral k - 1\<close> | 
| 74592 | 1421 | by (simp add: mask_eq_exp_minus_1) | 
| 1422 | ||
| 1423 | lemma push_bit_mask_eq: | |
| 1424 | \<open>push_bit m (mask n) = mask (n + m) AND NOT (mask m)\<close> | |
| 1425 | apply (rule bit_eqI) | |
| 1426 | apply (auto simp add: bit_simps not_less possible_bit_def) | |
| 1427 | apply (drule sym [of 0]) | |
| 1428 | apply (simp only:) | |
| 1429 | using exp_not_zero_imp_exp_diff_not_zero apply (blast dest: exp_not_zero_imp_exp_diff_not_zero) | |
| 1430 | done | |
| 1431 | ||
| 1432 | lemma slice_eq_mask: | |
| 1433 | \<open>push_bit n (take_bit m (drop_bit n a)) = a AND mask (m + n) AND NOT (mask n)\<close> | |
| 1434 | by (rule bit_eqI) (auto simp add: bit_simps) | |
| 1435 | ||
| 1436 | lemma push_bit_numeral_minus_1 [simp]: | |
| 1437 | \<open>push_bit (numeral n) (- 1) = - (2 ^ numeral n)\<close> | |
| 1438 | by (simp add: push_bit_eq_mult) | |
| 74498 | 1439 | |
| 71042 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 1440 | end | 
| 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 1441 | |
| 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 1442 | |
| 71956 | 1443 | subsection \<open>Instance \<^typ>\<open>int\<close>\<close> | 
| 71042 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 1444 | |
| 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 1445 | instantiation int :: ring_bit_operations | 
| 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 1446 | begin | 
| 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 1447 | |
| 71420 | 1448 | definition not_int :: \<open>int \<Rightarrow> int\<close> | 
| 1449 | where \<open>not_int k = - k - 1\<close> | |
| 1450 | ||
| 1451 | lemma not_int_rec: | |
| 73969 
ca2a35c0fe6e
operations for symbolic computation of bit operations
 haftmann parents: 
73871diff
changeset | 1452 | \<open>NOT k = of_bool (even k) + 2 * NOT (k div 2)\<close> for k :: int | 
| 71420 | 1453 | by (auto simp add: not_int_def elim: oddE) | 
| 1454 | ||
| 1455 | lemma even_not_iff_int: | |
| 1456 | \<open>even (NOT k) \<longleftrightarrow> odd k\<close> for k :: int | |
| 1457 | by (simp add: not_int_def) | |
| 1458 | ||
| 1459 | lemma not_int_div_2: | |
| 1460 | \<open>NOT k div 2 = NOT (k div 2)\<close> for k :: int | |
| 74101 | 1461 | by (cases k) (simp_all add: not_int_def divide_int_def nat_add_distrib) | 
| 71042 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 1462 | |
| 74163 | 1463 | lemma bit_not_int_iff: | 
| 71186 | 1464 | \<open>bit (NOT k) n \<longleftrightarrow> \<not> bit k n\<close> | 
| 72488 | 1465 | for k :: int | 
| 1466 | by (simp add: bit_not_int_iff' not_int_def) | |
| 71186 | 1467 | |
| 71804 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1468 | function and_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close> | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1469 |   where \<open>(k::int) AND l = (if k \<in> {0, - 1} \<and> l \<in> {0, - 1}
 | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1470 | then - of_bool (odd k \<and> odd l) | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1471 | else of_bool (odd k \<and> odd l) + 2 * ((k div 2) AND (l div 2)))\<close> | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1472 | by auto | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1473 | |
| 74101 | 1474 | termination proof (relation \<open>measure (\<lambda>(k, l). nat (\<bar>k\<bar> + \<bar>l\<bar>))\<close>) | 
| 1475 | show \<open>wf (measure (\<lambda>(k, l). nat (\<bar>k\<bar> + \<bar>l\<bar>)))\<close> | |
| 1476 | by simp | |
| 1477 | show \<open>((k div 2, l div 2), k, l) \<in> measure (\<lambda>(k, l). nat (\<bar>k\<bar> + \<bar>l\<bar>))\<close> | |
| 1478 |     if \<open>\<not> (k \<in> {0, - 1} \<and> l \<in> {0, - 1})\<close> for k l
 | |
| 1479 | proof - | |
| 1480 | have less_eq: \<open>\<bar>k div 2\<bar> \<le> \<bar>k\<bar>\<close> for k :: int | |
| 1481 | by (cases k) (simp_all add: divide_int_def nat_add_distrib) | |
| 1482 |     have less: \<open>\<bar>k div 2\<bar> < \<bar>k\<bar>\<close> if \<open>k \<notin> {0, - 1}\<close> for k :: int
 | |
| 1483 | proof (cases k) | |
| 1484 | case (nonneg n) | |
| 1485 | with that show ?thesis | |
| 1486 | by (simp add: int_div_less_self) | |
| 1487 | next | |
| 1488 | case (neg n) | |
| 1489 | with that have \<open>n \<noteq> 0\<close> | |
| 1490 | by simp | |
| 1491 | then have \<open>n div 2 < n\<close> | |
| 1492 | by (simp add: div_less_iff_less_mult) | |
| 1493 | with neg that show ?thesis | |
| 1494 | by (simp add: divide_int_def nat_add_distrib) | |
| 1495 | qed | |
| 1496 |     from that have *: \<open>k \<notin> {0, - 1} \<or> l \<notin> {0, - 1}\<close>
 | |
| 1497 | by simp | |
| 1498 | then have \<open>0 < \<bar>k\<bar> + \<bar>l\<bar>\<close> | |
| 1499 | by auto | |
| 1500 | moreover from * have \<open>\<bar>k div 2\<bar> + \<bar>l div 2\<bar> < \<bar>k\<bar> + \<bar>l\<bar>\<close> | |
| 1501 | proof | |
| 1502 |       assume \<open>k \<notin> {0, - 1}\<close>
 | |
| 1503 | then have \<open>\<bar>k div 2\<bar> < \<bar>k\<bar>\<close> | |
| 1504 | by (rule less) | |
| 1505 | with less_eq [of l] show ?thesis | |
| 1506 | by auto | |
| 1507 | next | |
| 1508 |       assume \<open>l \<notin> {0, - 1}\<close>
 | |
| 1509 | then have \<open>\<bar>l div 2\<bar> < \<bar>l\<bar>\<close> | |
| 1510 | by (rule less) | |
| 1511 | with less_eq [of k] show ?thesis | |
| 1512 | by auto | |
| 1513 | qed | |
| 1514 | ultimately show ?thesis | |
| 1515 | by simp | |
| 1516 | qed | |
| 1517 | qed | |
| 71804 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1518 | |
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1519 | declare and_int.simps [simp del] | 
| 71802 | 1520 | |
| 71804 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1521 | lemma and_int_rec: | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1522 | \<open>k AND l = of_bool (odd k \<and> odd l) + 2 * ((k div 2) AND (l div 2))\<close> | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1523 | for k l :: int | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1524 | proof (cases \<open>k \<in> {0, - 1} \<and> l \<in> {0, - 1}\<close>)
 | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1525 | case True | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1526 | then show ?thesis | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1527 | by auto (simp_all add: and_int.simps) | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1528 | next | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1529 | case False | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1530 | then show ?thesis | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1531 | by (auto simp add: ac_simps and_int.simps [of k l]) | 
| 71802 | 1532 | qed | 
| 1533 | ||
| 71804 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1534 | lemma bit_and_int_iff: | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1535 | \<open>bit (k AND l) n \<longleftrightarrow> bit k n \<and> bit l n\<close> for k l :: int | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1536 | proof (induction n arbitrary: k l) | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1537 | case 0 | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1538 | then show ?case | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1539 | by (simp add: and_int_rec [of k l]) | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1540 | next | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1541 | case (Suc n) | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1542 | then show ?case | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1543 | by (simp add: and_int_rec [of k l] bit_Suc) | 
| 71802 | 1544 | qed | 
| 1545 | ||
| 71804 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1546 | lemma even_and_iff_int: | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1547 | \<open>even (k AND l) \<longleftrightarrow> even k \<or> even l\<close> for k l :: int | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1548 | using bit_and_int_iff [of k l 0] by auto | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1549 | |
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1550 | definition or_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close> | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1551 | where \<open>k OR l = NOT (NOT k AND NOT l)\<close> for k l :: int | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1552 | |
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1553 | lemma or_int_rec: | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1554 | \<open>k OR l = of_bool (odd k \<or> odd l) + 2 * ((k div 2) OR (l div 2))\<close> | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1555 | for k l :: int | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1556 | using and_int_rec [of \<open>NOT k\<close> \<open>NOT l\<close>] | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1557 | by (simp add: or_int_def even_not_iff_int not_int_div_2) | 
| 73535 | 1558 | (simp_all add: not_int_def) | 
| 71804 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1559 | |
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1560 | lemma bit_or_int_iff: | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1561 | \<open>bit (k OR l) n \<longleftrightarrow> bit k n \<or> bit l n\<close> for k l :: int | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1562 | by (simp add: or_int_def bit_not_int_iff bit_and_int_iff) | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1563 | |
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1564 | definition xor_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close> | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1565 | where \<open>k XOR l = k AND NOT l OR NOT k AND l\<close> for k l :: int | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1566 | |
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1567 | lemma xor_int_rec: | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1568 | \<open>k XOR l = of_bool (odd k \<noteq> odd l) + 2 * ((k div 2) XOR (l div 2))\<close> | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1569 | for k l :: int | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1570 | 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) | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1571 | (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) | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1572 | |
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1573 | lemma bit_xor_int_iff: | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1574 | \<open>bit (k XOR l) n \<longleftrightarrow> bit k n \<noteq> bit l n\<close> for k l :: int | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1575 | by (auto simp add: xor_int_def bit_or_int_iff bit_and_int_iff bit_not_int_iff) | 
| 71802 | 1576 | |
| 72082 | 1577 | definition mask_int :: \<open>nat \<Rightarrow> int\<close> | 
| 1578 | where \<open>mask n = (2 :: int) ^ n - 1\<close> | |
| 1579 | ||
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1580 | definition push_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1581 | where \<open>push_bit_int n k = k * 2 ^ n\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1582 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1583 | definition drop_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1584 | where \<open>drop_bit_int n k = k div 2 ^ n\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1585 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1586 | definition take_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1587 | where \<open>take_bit_int n k = k mod 2 ^ n\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1588 | |
| 73682 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1589 | definition set_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close> | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1590 | where \<open>set_bit n k = k OR push_bit n 1\<close> for k :: int | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1591 | |
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1592 | definition unset_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close> | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1593 | where \<open>unset_bit n k = k AND NOT (push_bit n 1)\<close> for k :: int | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1594 | |
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1595 | definition flip_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close> | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1596 | where \<open>flip_bit n k = k XOR push_bit n 1\<close> for k :: int | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1597 | |
| 71042 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 1598 | instance proof | 
| 73682 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1599 | fix k l :: int and m n :: nat | 
| 71409 | 1600 | show \<open>- k = NOT (k - 1)\<close> | 
| 1601 | by (simp add: not_int_def) | |
| 71186 | 1602 | show \<open>bit (k AND l) n \<longleftrightarrow> bit k n \<and> bit l n\<close> | 
| 71804 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1603 | by (fact bit_and_int_iff) | 
| 71186 | 1604 | show \<open>bit (k OR l) n \<longleftrightarrow> bit k n \<or> bit l n\<close> | 
| 71804 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1605 | by (fact bit_or_int_iff) | 
| 71186 | 1606 | show \<open>bit (k XOR l) n \<longleftrightarrow> bit k n \<noteq> bit l n\<close> | 
| 71804 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1607 | by (fact bit_xor_int_iff) | 
| 73682 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1608 | show \<open>bit (unset_bit m k) n \<longleftrightarrow> bit k n \<and> m \<noteq> n\<close> | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1609 | proof - | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1610 | have \<open>unset_bit m k = k AND NOT (push_bit m 1)\<close> | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1611 | by (simp add: unset_bit_int_def) | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1612 | also have \<open>NOT (push_bit m 1 :: int) = - (push_bit m 1 + 1)\<close> | 
| 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1613 | by (simp add: not_int_def) | 
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1614 | finally show ?thesis by (simp only: bit_simps bit_and_int_iff) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1615 | (auto simp add: bit_simps bit_not_int_iff' push_bit_int_def) | 
| 73682 
78044b2f001c
explicit type class operations for type-specific implementations
 haftmann parents: 
73535diff
changeset | 1616 | qed | 
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1617 | qed (simp_all add: bit_not_int_iff mask_int_def set_bit_int_def flip_bit_int_def | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1618 | push_bit_int_def drop_bit_int_def take_bit_int_def) | 
| 71042 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 1619 | |
| 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 1620 | end | 
| 
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
 haftmann parents: diff
changeset | 1621 | |
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1622 | lemma bit_push_bit_iff_int: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1623 | \<open>bit (push_bit m k) n \<longleftrightarrow> m \<le> n \<and> bit k (n - m)\<close> for k :: int | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1624 | by (auto simp add: bit_push_bit_iff) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1625 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1626 | lemma take_bit_nonnegative [simp]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1627 | \<open>take_bit n k \<ge> 0\<close> for k :: int | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1628 | by (simp add: take_bit_eq_mod) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1629 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1630 | lemma not_take_bit_negative [simp]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1631 | \<open>\<not> take_bit n k < 0\<close> for k :: int | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1632 | by (simp add: not_less) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1633 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1634 | lemma take_bit_int_less_exp [simp]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1635 | \<open>take_bit n k < 2 ^ n\<close> for k :: int | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1636 | by (simp add: take_bit_eq_mod) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1637 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1638 | lemma take_bit_int_eq_self_iff: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1639 | \<open>take_bit n k = k \<longleftrightarrow> 0 \<le> k \<and> k < 2 ^ n\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1640 | for k :: int | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1641 | proof | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1642 | assume ?P | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1643 | moreover note take_bit_int_less_exp [of n k] take_bit_nonnegative [of n k] | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1644 | ultimately show ?Q | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1645 | by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1646 | next | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1647 | assume ?Q | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1648 | then show ?P | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1649 | by (simp add: take_bit_eq_mod) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1650 | qed | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1651 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1652 | lemma take_bit_int_eq_self: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1653 | \<open>take_bit n k = k\<close> if \<open>0 \<le> k\<close> \<open>k < 2 ^ n\<close> for k :: int | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1654 | using that by (simp add: take_bit_int_eq_self_iff) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1655 | |
| 72241 | 1656 | lemma mask_half_int: | 
| 1657 | \<open>mask n div 2 = (mask (n - 1) :: int)\<close> | |
| 1658 | by (cases n) (simp_all add: mask_eq_exp_minus_1 algebra_simps) | |
| 1659 | ||
| 72028 | 1660 | lemma mask_nonnegative_int [simp]: | 
| 1661 | \<open>mask n \<ge> (0::int)\<close> | |
| 1662 | by (simp add: mask_eq_exp_minus_1) | |
| 1663 | ||
| 1664 | lemma not_mask_negative_int [simp]: | |
| 1665 | \<open>\<not> mask n < (0::int)\<close> | |
| 1666 | by (simp add: not_less) | |
| 1667 | ||
| 71802 | 1668 | lemma not_nonnegative_int_iff [simp]: | 
| 1669 | \<open>NOT k \<ge> 0 \<longleftrightarrow> k < 0\<close> for k :: int | |
| 1670 | by (simp add: not_int_def) | |
| 1671 | ||
| 1672 | lemma not_negative_int_iff [simp]: | |
| 1673 | \<open>NOT k < 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int | |
| 1674 | by (subst Not_eq_iff [symmetric]) (simp add: not_less not_le) | |
| 1675 | ||
| 1676 | lemma and_nonnegative_int_iff [simp]: | |
| 1677 | \<open>k AND l \<ge> 0 \<longleftrightarrow> k \<ge> 0 \<or> l \<ge> 0\<close> for k l :: int | |
| 1678 | proof (induction k arbitrary: l rule: int_bit_induct) | |
| 1679 | case zero | |
| 1680 | then show ?case | |
| 1681 | by simp | |
| 1682 | next | |
| 1683 | case minus | |
| 1684 | then show ?case | |
| 1685 | by simp | |
| 1686 | next | |
| 1687 | case (even k) | |
| 1688 | then show ?case | |
| 74101 | 1689 | using and_int_rec [of \<open>k * 2\<close> l] | 
| 1690 | by (simp add: pos_imp_zdiv_nonneg_iff zero_le_mult_iff) | |
| 71802 | 1691 | next | 
| 1692 | case (odd k) | |
| 1693 | from odd have \<open>0 \<le> k AND l div 2 \<longleftrightarrow> 0 \<le> k \<or> 0 \<le> l div 2\<close> | |
| 1694 | by simp | |
| 74101 | 1695 | 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> | 
| 71802 | 1696 | by simp | 
| 71804 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1697 | with and_int_rec [of \<open>1 + k * 2\<close> l] | 
| 71802 | 1698 | show ?case | 
| 74101 | 1699 | by (auto simp add: zero_le_mult_iff not_le) | 
| 71802 | 1700 | qed | 
| 1701 | ||
| 1702 | lemma and_negative_int_iff [simp]: | |
| 1703 | \<open>k AND l < 0 \<longleftrightarrow> k < 0 \<and> l < 0\<close> for k l :: int | |
| 1704 | by (subst Not_eq_iff [symmetric]) (simp add: not_less) | |
| 1705 | ||
| 72009 | 1706 | lemma and_less_eq: | 
| 1707 | \<open>k AND l \<le> k\<close> if \<open>l < 0\<close> for k l :: int | |
| 1708 | using that proof (induction k arbitrary: l rule: int_bit_induct) | |
| 1709 | case zero | |
| 1710 | then show ?case | |
| 1711 | by simp | |
| 1712 | next | |
| 1713 | case minus | |
| 1714 | then show ?case | |
| 1715 | by simp | |
| 1716 | next | |
| 1717 | case (even k) | |
| 1718 | from even.IH [of \<open>l div 2\<close>] even.hyps even.prems | |
| 1719 | show ?case | |
| 1720 | by (simp add: and_int_rec [of _ l]) | |
| 1721 | next | |
| 1722 | case (odd k) | |
| 1723 | from odd.IH [of \<open>l div 2\<close>] odd.hyps odd.prems | |
| 1724 | show ?case | |
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1725 | by (simp add: and_int_rec [of _ l]) linarith | 
| 72009 | 1726 | qed | 
| 1727 | ||
| 71802 | 1728 | lemma or_nonnegative_int_iff [simp]: | 
| 1729 | \<open>k OR l \<ge> 0 \<longleftrightarrow> k \<ge> 0 \<and> l \<ge> 0\<close> for k l :: int | |
| 1730 | by (simp only: or_eq_not_not_and not_nonnegative_int_iff) simp | |
| 1731 | ||
| 1732 | lemma or_negative_int_iff [simp]: | |
| 1733 | \<open>k OR l < 0 \<longleftrightarrow> k < 0 \<or> l < 0\<close> for k l :: int | |
| 1734 | by (subst Not_eq_iff [symmetric]) (simp add: not_less) | |
| 1735 | ||
| 72009 | 1736 | lemma or_greater_eq: | 
| 1737 | \<open>k OR l \<ge> k\<close> if \<open>l \<ge> 0\<close> for k l :: int | |
| 1738 | using that proof (induction k arbitrary: l rule: int_bit_induct) | |
| 1739 | case zero | |
| 1740 | then show ?case | |
| 1741 | by simp | |
| 1742 | next | |
| 1743 | case minus | |
| 1744 | then show ?case | |
| 1745 | by simp | |
| 1746 | next | |
| 1747 | case (even k) | |
| 1748 | from even.IH [of \<open>l div 2\<close>] even.hyps even.prems | |
| 1749 | show ?case | |
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1750 | by (simp add: or_int_rec [of _ l]) linarith | 
| 72009 | 1751 | next | 
| 1752 | case (odd k) | |
| 1753 | from odd.IH [of \<open>l div 2\<close>] odd.hyps odd.prems | |
| 1754 | show ?case | |
| 1755 | by (simp add: or_int_rec [of _ l]) | |
| 1756 | qed | |
| 1757 | ||
| 71802 | 1758 | lemma xor_nonnegative_int_iff [simp]: | 
| 1759 | \<open>k XOR l \<ge> 0 \<longleftrightarrow> (k \<ge> 0 \<longleftrightarrow> l \<ge> 0)\<close> for k l :: int | |
| 1760 | by (simp only: bit.xor_def or_nonnegative_int_iff) auto | |
| 1761 | ||
| 1762 | lemma xor_negative_int_iff [simp]: | |
| 1763 | \<open>k XOR l < 0 \<longleftrightarrow> (k < 0) \<noteq> (l < 0)\<close> for k l :: int | |
| 1764 | by (subst Not_eq_iff [symmetric]) (auto simp add: not_less) | |
| 1765 | ||
| 72488 | 1766 | lemma OR_upper: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close> | 
| 1767 | fixes x y :: int | |
| 73969 
ca2a35c0fe6e
operations for symbolic computation of bit operations
 haftmann parents: 
73871diff
changeset | 1768 | assumes \<open>0 \<le> x\<close> \<open>x < 2 ^ n\<close> \<open>y < 2 ^ n\<close> | 
| 
ca2a35c0fe6e
operations for symbolic computation of bit operations
 haftmann parents: 
73871diff
changeset | 1769 | shows \<open>x OR y < 2 ^ n\<close> | 
| 72488 | 1770 | using assms proof (induction x arbitrary: y n rule: int_bit_induct) | 
| 1771 | case zero | |
| 1772 | then show ?case | |
| 1773 | by simp | |
| 1774 | next | |
| 1775 | case minus | |
| 1776 | then show ?case | |
| 1777 | by simp | |
| 1778 | next | |
| 1779 | case (even x) | |
| 1780 | from even.IH [of \<open>n - 1\<close> \<open>y div 2\<close>] even.prems even.hyps | |
| 1781 | show ?case | |
| 1782 | by (cases n) (auto simp add: or_int_rec [of \<open>_ * 2\<close>] elim: oddE) | |
| 1783 | next | |
| 1784 | case (odd x) | |
| 1785 | from odd.IH [of \<open>n - 1\<close> \<open>y div 2\<close>] odd.prems odd.hyps | |
| 1786 | show ?case | |
| 1787 | by (cases n) (auto simp add: or_int_rec [of \<open>1 + _ * 2\<close>], linarith) | |
| 1788 | qed | |
| 1789 | ||
| 1790 | lemma XOR_upper: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close> | |
| 1791 | fixes x y :: int | |
| 73969 
ca2a35c0fe6e
operations for symbolic computation of bit operations
 haftmann parents: 
73871diff
changeset | 1792 | assumes \<open>0 \<le> x\<close> \<open>x < 2 ^ n\<close> \<open>y < 2 ^ n\<close> | 
| 
ca2a35c0fe6e
operations for symbolic computation of bit operations
 haftmann parents: 
73871diff
changeset | 1793 | shows \<open>x XOR y < 2 ^ n\<close> | 
| 72488 | 1794 | using assms proof (induction x arbitrary: y n rule: int_bit_induct) | 
| 1795 | case zero | |
| 1796 | then show ?case | |
| 1797 | by simp | |
| 1798 | next | |
| 1799 | case minus | |
| 1800 | then show ?case | |
| 1801 | by simp | |
| 1802 | next | |
| 1803 | case (even x) | |
| 1804 | from even.IH [of \<open>n - 1\<close> \<open>y div 2\<close>] even.prems even.hyps | |
| 1805 | show ?case | |
| 1806 | by (cases n) (auto simp add: xor_int_rec [of \<open>_ * 2\<close>] elim: oddE) | |
| 1807 | next | |
| 1808 | case (odd x) | |
| 1809 | from odd.IH [of \<open>n - 1\<close> \<open>y div 2\<close>] odd.prems odd.hyps | |
| 1810 | show ?case | |
| 1811 | by (cases n) (auto simp add: xor_int_rec [of \<open>1 + _ * 2\<close>]) | |
| 1812 | qed | |
| 1813 | ||
| 1814 | lemma AND_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close> | |
| 1815 | fixes x y :: int | |
| 73969 
ca2a35c0fe6e
operations for symbolic computation of bit operations
 haftmann parents: 
73871diff
changeset | 1816 | assumes \<open>0 \<le> x\<close> | 
| 
ca2a35c0fe6e
operations for symbolic computation of bit operations
 haftmann parents: 
73871diff
changeset | 1817 | shows \<open>0 \<le> x AND y\<close> | 
| 72488 | 1818 | using assms by simp | 
| 1819 | ||
| 1820 | lemma OR_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close> | |
| 1821 | fixes x y :: int | |
| 73969 
ca2a35c0fe6e
operations for symbolic computation of bit operations
 haftmann parents: 
73871diff
changeset | 1822 | assumes \<open>0 \<le> x\<close> \<open>0 \<le> y\<close> | 
| 
ca2a35c0fe6e
operations for symbolic computation of bit operations
 haftmann parents: 
73871diff
changeset | 1823 | shows \<open>0 \<le> x OR y\<close> | 
| 72488 | 1824 | using assms by simp | 
| 1825 | ||
| 1826 | lemma XOR_lower [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close> | |
| 1827 | fixes x y :: int | |
| 73969 
ca2a35c0fe6e
operations for symbolic computation of bit operations
 haftmann parents: 
73871diff
changeset | 1828 | assumes \<open>0 \<le> x\<close> \<open>0 \<le> y\<close> | 
| 
ca2a35c0fe6e
operations for symbolic computation of bit operations
 haftmann parents: 
73871diff
changeset | 1829 | shows \<open>0 \<le> x XOR y\<close> | 
| 72488 | 1830 | using assms by simp | 
| 1831 | ||
| 1832 | lemma AND_upper1 [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close> | |
| 1833 | fixes x y :: int | |
| 73969 
ca2a35c0fe6e
operations for symbolic computation of bit operations
 haftmann parents: 
73871diff
changeset | 1834 | assumes \<open>0 \<le> x\<close> | 
| 
ca2a35c0fe6e
operations for symbolic computation of bit operations
 haftmann parents: 
73871diff
changeset | 1835 | shows \<open>x AND y \<le> x\<close> | 
| 73535 | 1836 | using assms proof (induction x arbitrary: y rule: int_bit_induct) | 
| 1837 | case (odd k) | |
| 1838 | then have \<open>k AND y div 2 \<le> k\<close> | |
| 1839 | by simp | |
| 1840 | then show ?case | |
| 1841 | by (simp add: and_int_rec [of \<open>1 + _ * 2\<close>]) | |
| 1842 | qed (simp_all add: and_int_rec [of \<open>_ * 2\<close>]) | |
| 72488 | 1843 | |
| 1844 | lemmas AND_upper1' [simp] = order_trans [OF AND_upper1] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close> | |
| 1845 | lemmas AND_upper1'' [simp] = order_le_less_trans [OF AND_upper1] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close> | |
| 1846 | ||
| 1847 | lemma AND_upper2 [simp]: \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close> | |
| 1848 | fixes x y :: int | |
| 73969 
ca2a35c0fe6e
operations for symbolic computation of bit operations
 haftmann parents: 
73871diff
changeset | 1849 | assumes \<open>0 \<le> y\<close> | 
| 
ca2a35c0fe6e
operations for symbolic computation of bit operations
 haftmann parents: 
73871diff
changeset | 1850 | shows \<open>x AND y \<le> y\<close> | 
| 72488 | 1851 | using assms AND_upper1 [of y x] by (simp add: ac_simps) | 
| 1852 | ||
| 1853 | lemmas AND_upper2' [simp] = order_trans [OF AND_upper2] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close> | |
| 1854 | lemmas AND_upper2'' [simp] = order_le_less_trans [OF AND_upper2] \<^marker>\<open>contributor \<open>Stefan Berghofer\<close>\<close> | |
| 1855 | ||
| 1856 | lemma plus_and_or: \<open>(x AND y) + (x OR y) = x + y\<close> for x y :: int | |
| 1857 | proof (induction x arbitrary: y rule: int_bit_induct) | |
| 1858 | case zero | |
| 1859 | then show ?case | |
| 1860 | by simp | |
| 1861 | next | |
| 1862 | case minus | |
| 1863 | then show ?case | |
| 1864 | by simp | |
| 1865 | next | |
| 1866 | case (even x) | |
| 1867 | from even.IH [of \<open>y div 2\<close>] | |
| 1868 | show ?case | |
| 1869 | by (auto simp add: and_int_rec [of _ y] or_int_rec [of _ y] elim: oddE) | |
| 1870 | next | |
| 1871 | case (odd x) | |
| 1872 | from odd.IH [of \<open>y div 2\<close>] | |
| 1873 | show ?case | |
| 1874 | by (auto simp add: and_int_rec [of _ y] or_int_rec [of _ y] elim: oddE) | |
| 1875 | qed | |
| 1876 | ||
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1877 | lemma push_bit_minus_one: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1878 | "push_bit n (- 1 :: int) = - (2 ^ n)" | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1879 | by (simp add: push_bit_eq_mult) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1880 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1881 | lemma minus_1_div_exp_eq_int: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1882 | \<open>- 1 div (2 :: int) ^ n = - 1\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1883 | by (induction n) (use div_exp_eq [symmetric, of \<open>- 1 :: int\<close> 1] in \<open>simp_all add: ac_simps\<close>) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1884 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1885 | lemma drop_bit_minus_one [simp]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1886 | \<open>drop_bit n (- 1 :: int) = - 1\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1887 | by (simp add: drop_bit_eq_div minus_1_div_exp_eq_int) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1888 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1889 | lemma take_bit_Suc_from_most: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1890 | \<open>take_bit (Suc n) k = 2 ^ n * of_bool (bit k n) + take_bit n k\<close> for k :: int | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1891 | by (simp only: take_bit_eq_mod power_Suc2) (simp_all add: bit_iff_odd odd_iff_mod_2_eq_one zmod_zmult2_eq) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1892 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1893 | lemma take_bit_minus: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1894 | \<open>take_bit n (- take_bit n k) = take_bit n (- k)\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1895 | for k :: int | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1896 | by (simp add: take_bit_eq_mod mod_minus_eq) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1897 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1898 | lemma take_bit_diff: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1899 | \<open>take_bit n (take_bit n k - take_bit n l) = take_bit n (k - l)\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1900 | for k l :: int | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1901 | by (simp add: take_bit_eq_mod mod_diff_eq) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1902 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1903 | lemma bit_imp_take_bit_positive: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1904 | \<open>0 < take_bit m k\<close> if \<open>n < m\<close> and \<open>bit k n\<close> for k :: int | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1905 | proof (rule ccontr) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1906 | assume \<open>\<not> 0 < take_bit m k\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1907 | then have \<open>take_bit m k = 0\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1908 | by (auto simp add: not_less intro: order_antisym) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1909 | then have \<open>bit (take_bit m k) n = bit 0 n\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1910 | by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1911 | with that show False | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1912 | by (simp add: bit_take_bit_iff) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1913 | qed | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1914 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1915 | lemma take_bit_mult: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1916 | \<open>take_bit n (take_bit n k * take_bit n l) = take_bit n (k * l)\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1917 | for k l :: int | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1918 | by (simp add: take_bit_eq_mod mod_mult_eq) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1919 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1920 | lemma (in ring_1) of_nat_nat_take_bit_eq [simp]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1921 | \<open>of_nat (nat (take_bit n k)) = of_int (take_bit n k)\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1922 | by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1923 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1924 | lemma take_bit_minus_small_eq: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1925 | \<open>take_bit n (- k) = 2 ^ n - k\<close> if \<open>0 < k\<close> \<open>k \<le> 2 ^ n\<close> for k :: int | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1926 | proof - | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1927 | define m where \<open>m = nat k\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1928 | with that have \<open>k = int m\<close> and \<open>0 < m\<close> and \<open>m \<le> 2 ^ n\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1929 | by simp_all | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1930 | have \<open>(2 ^ n - m) mod 2 ^ n = 2 ^ n - m\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1931 | using \<open>0 < m\<close> by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1932 | then have \<open>int ((2 ^ n - m) mod 2 ^ n) = int (2 ^ n - m)\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1933 | by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1934 | then have \<open>(2 ^ n - int m) mod 2 ^ n = 2 ^ n - int m\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1935 | using \<open>m \<le> 2 ^ n\<close> by (simp only: of_nat_mod of_nat_diff) simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1936 | with \<open>k = int m\<close> have \<open>(2 ^ n - k) mod 2 ^ n = 2 ^ n - k\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1937 | by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1938 | then show ?thesis | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1939 | by (simp add: take_bit_eq_mod) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1940 | qed | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1941 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1942 | lemma drop_bit_push_bit_int: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1943 | \<open>drop_bit m (push_bit n k) = drop_bit (m - n) (push_bit (n - m) k)\<close> for k :: int | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1944 | 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 | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1945 | mult.commute [of k] drop_bit_eq_div push_bit_eq_mult not_le power_add dest!: le_Suc_ex less_imp_Suc_add) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1946 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1947 | lemma push_bit_nonnegative_int_iff [simp]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1948 | \<open>push_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1949 | by (simp add: push_bit_eq_mult zero_le_mult_iff power_le_zero_eq) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1950 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1951 | lemma push_bit_negative_int_iff [simp]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1952 | \<open>push_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1953 | by (subst Not_eq_iff [symmetric]) (simp add: not_less) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1954 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1955 | lemma drop_bit_nonnegative_int_iff [simp]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1956 | \<open>drop_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1957 | by (induction n) (auto simp add: drop_bit_Suc drop_bit_half) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1958 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1959 | lemma drop_bit_negative_int_iff [simp]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1960 | \<open>drop_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1961 | by (subst Not_eq_iff [symmetric]) (simp add: not_less) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 1962 | |
| 71802 | 1963 | lemma set_bit_nonnegative_int_iff [simp]: | 
| 1964 | \<open>set_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int | |
| 1965 | by (simp add: set_bit_def) | |
| 1966 | ||
| 1967 | lemma set_bit_negative_int_iff [simp]: | |
| 1968 | \<open>set_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int | |
| 1969 | by (simp add: set_bit_def) | |
| 1970 | ||
| 1971 | lemma unset_bit_nonnegative_int_iff [simp]: | |
| 1972 | \<open>unset_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int | |
| 1973 | by (simp add: unset_bit_def) | |
| 1974 | ||
| 1975 | lemma unset_bit_negative_int_iff [simp]: | |
| 1976 | \<open>unset_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int | |
| 1977 | by (simp add: unset_bit_def) | |
| 1978 | ||
| 1979 | lemma flip_bit_nonnegative_int_iff [simp]: | |
| 1980 | \<open>flip_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int | |
| 1981 | by (simp add: flip_bit_def) | |
| 1982 | ||
| 1983 | lemma flip_bit_negative_int_iff [simp]: | |
| 1984 | \<open>flip_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int | |
| 1985 | by (simp add: flip_bit_def) | |
| 1986 | ||
| 71986 | 1987 | lemma set_bit_greater_eq: | 
| 1988 | \<open>set_bit n k \<ge> k\<close> for k :: int | |
| 1989 | by (simp add: set_bit_def or_greater_eq) | |
| 1990 | ||
| 1991 | lemma unset_bit_less_eq: | |
| 1992 | \<open>unset_bit n k \<le> k\<close> for k :: int | |
| 1993 | by (simp add: unset_bit_def and_less_eq) | |
| 1994 | ||
| 72009 | 1995 | lemma set_bit_eq: | 
| 1996 | \<open>set_bit n k = k + of_bool (\<not> bit k n) * 2 ^ n\<close> for k :: int | |
| 1997 | proof (rule bit_eqI) | |
| 1998 | fix m | |
| 1999 | show \<open>bit (set_bit n k) m \<longleftrightarrow> bit (k + of_bool (\<not> bit k n) * 2 ^ n) m\<close> | |
| 2000 | proof (cases \<open>m = n\<close>) | |
| 2001 | case True | |
| 2002 | then show ?thesis | |
| 2003 | apply (simp add: bit_set_bit_iff) | |
| 2004 | apply (simp add: bit_iff_odd div_plus_div_distrib_dvd_right) | |
| 2005 | done | |
| 2006 | next | |
| 2007 | case False | |
| 2008 | then show ?thesis | |
| 2009 | apply (clarsimp simp add: bit_set_bit_iff) | |
| 2010 | apply (subst disjunctive_add) | |
| 2011 | apply (clarsimp simp add: bit_exp_iff) | |
| 2012 | apply (clarsimp simp add: bit_or_iff bit_exp_iff) | |
| 2013 | done | |
| 2014 | qed | |
| 2015 | qed | |
| 2016 | ||
| 2017 | lemma unset_bit_eq: | |
| 2018 | \<open>unset_bit n k = k - of_bool (bit k n) * 2 ^ n\<close> for k :: int | |
| 2019 | proof (rule bit_eqI) | |
| 2020 | fix m | |
| 2021 | show \<open>bit (unset_bit n k) m \<longleftrightarrow> bit (k - of_bool (bit k n) * 2 ^ n) m\<close> | |
| 2022 | proof (cases \<open>m = n\<close>) | |
| 2023 | case True | |
| 2024 | then show ?thesis | |
| 2025 | apply (simp add: bit_unset_bit_iff) | |
| 2026 | apply (simp add: bit_iff_odd) | |
| 2027 | using div_plus_div_distrib_dvd_right [of \<open>2 ^ n\<close> \<open>- (2 ^ n)\<close> k] | |
| 2028 | apply (simp add: dvd_neg_div) | |
| 2029 | done | |
| 2030 | next | |
| 2031 | case False | |
| 2032 | then show ?thesis | |
| 2033 | apply (clarsimp simp add: bit_unset_bit_iff) | |
| 2034 | apply (subst disjunctive_diff) | |
| 2035 | apply (clarsimp simp add: bit_exp_iff) | |
| 2036 | apply (clarsimp simp add: bit_and_iff bit_not_iff bit_exp_iff) | |
| 2037 | done | |
| 2038 | qed | |
| 2039 | qed | |
| 2040 | ||
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2041 | lemma and_int_unfold [code]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2042 | \<open>k AND l = (if k = 0 \<or> l = 0 then 0 else if k = - 1 then l else if l = - 1 then k | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2043 | else (k mod 2) * (l mod 2) + 2 * ((k div 2) AND (l div 2)))\<close> for k l :: int | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2044 | by (auto simp add: and_int_rec [of k l] zmult_eq_1_iff elim: oddE) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2045 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2046 | lemma or_int_unfold [code]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2047 | \<open>k OR l = (if k = - 1 \<or> l = - 1 then - 1 else if k = 0 then l else if l = 0 then k | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2048 | else max (k mod 2) (l mod 2) + 2 * ((k div 2) OR (l div 2)))\<close> for k l :: int | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2049 | by (auto simp add: or_int_rec [of k l] elim: oddE) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2050 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2051 | lemma xor_int_unfold [code]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2052 | \<open>k XOR l = (if k = - 1 then NOT l else if l = - 1 then NOT k else if k = 0 then l else if l = 0 then k | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2053 | else \<bar>k mod 2 - l mod 2\<bar> + 2 * ((k div 2) XOR (l div 2)))\<close> for k l :: int | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2054 | by (auto simp add: xor_int_rec [of k l] not_int_def elim!: oddE) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2055 | |
| 74163 | 2056 | lemma bit_minus_int_iff: | 
| 2057 | \<open>bit (- k) n \<longleftrightarrow> bit (NOT (k - 1)) n\<close> | |
| 2058 | for k :: int | |
| 2059 | by (simp add: bit_simps) | |
| 2060 | ||
| 74592 | 2061 | lemma take_bit_incr_eq: | 
| 2062 | \<open>take_bit n (k + 1) = 1 + take_bit n k\<close> if \<open>take_bit n k \<noteq> 2 ^ n - 1\<close> | |
| 2063 | for k :: int | |
| 2064 | proof - | |
| 2065 | from that have \<open>2 ^ n \<noteq> k mod 2 ^ n + 1\<close> | |
| 2066 | by (simp add: take_bit_eq_mod) | |
| 2067 | moreover have \<open>k mod 2 ^ n < 2 ^ n\<close> | |
| 2068 | by simp | |
| 2069 | ultimately have *: \<open>k mod 2 ^ n + 1 < 2 ^ n\<close> | |
| 2070 | by linarith | |
| 2071 | have \<open>(k + 1) mod 2 ^ n = (k mod 2 ^ n + 1) mod 2 ^ n\<close> | |
| 2072 | by (simp add: mod_simps) | |
| 2073 | also have \<open>\<dots> = k mod 2 ^ n + 1\<close> | |
| 2074 | using * by (simp add: zmod_trivial_iff) | |
| 2075 | finally have \<open>(k + 1) mod 2 ^ n = k mod 2 ^ n + 1\<close> . | |
| 2076 | then show ?thesis | |
| 2077 | by (simp add: take_bit_eq_mod) | |
| 2078 | qed | |
| 2079 | ||
| 2080 | lemma take_bit_decr_eq: | |
| 2081 | \<open>take_bit n (k - 1) = take_bit n k - 1\<close> if \<open>take_bit n k \<noteq> 0\<close> | |
| 2082 | for k :: int | |
| 2083 | proof - | |
| 2084 | from that have \<open>k mod 2 ^ n \<noteq> 0\<close> | |
| 2085 | by (simp add: take_bit_eq_mod) | |
| 2086 | moreover have \<open>k mod 2 ^ n \<ge> 0\<close> \<open>k mod 2 ^ n < 2 ^ n\<close> | |
| 2087 | by simp_all | |
| 2088 | ultimately have *: \<open>k mod 2 ^ n > 0\<close> | |
| 2089 | by linarith | |
| 2090 | have \<open>(k - 1) mod 2 ^ n = (k mod 2 ^ n - 1) mod 2 ^ n\<close> | |
| 2091 | by (simp add: mod_simps) | |
| 2092 | also have \<open>\<dots> = k mod 2 ^ n - 1\<close> | |
| 2093 | by (simp add: zmod_trivial_iff) | |
| 2094 | (use \<open>k mod 2 ^ n < 2 ^ n\<close> * in linarith) | |
| 2095 | finally have \<open>(k - 1) mod 2 ^ n = k mod 2 ^ n - 1\<close> . | |
| 2096 | then show ?thesis | |
| 2097 | by (simp add: take_bit_eq_mod) | |
| 2098 | qed | |
| 2099 | ||
| 2100 | lemma take_bit_int_greater_eq: | |
| 2101 | \<open>k + 2 ^ n \<le> take_bit n k\<close> if \<open>k < 0\<close> for k :: int | |
| 2102 | proof - | |
| 2103 | have \<open>k + 2 ^ n \<le> take_bit n (k + 2 ^ n)\<close> | |
| 2104 | proof (cases \<open>k > - (2 ^ n)\<close>) | |
| 2105 | case False | |
| 2106 | then have \<open>k + 2 ^ n \<le> 0\<close> | |
| 2107 | by simp | |
| 2108 | also note take_bit_nonnegative | |
| 2109 | finally show ?thesis . | |
| 2110 | next | |
| 2111 | case True | |
| 2112 | with that have \<open>0 \<le> k + 2 ^ n\<close> and \<open>k + 2 ^ n < 2 ^ n\<close> | |
| 2113 | by simp_all | |
| 2114 | then show ?thesis | |
| 2115 | by (simp only: take_bit_eq_mod mod_pos_pos_trivial) | |
| 2116 | qed | |
| 2117 | then show ?thesis | |
| 2118 | by (simp add: take_bit_eq_mod) | |
| 2119 | qed | |
| 2120 | ||
| 2121 | lemma take_bit_int_less_eq: | |
| 2122 | \<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 | |
| 2123 | using that zmod_le_nonneg_dividend [of \<open>k - 2 ^ n\<close> \<open>2 ^ n\<close>] | |
| 2124 | by (simp add: take_bit_eq_mod) | |
| 2125 | ||
| 2126 | lemma take_bit_int_less_eq_self_iff: | |
| 2127 | \<open>take_bit n k \<le> k \<longleftrightarrow> 0 \<le> k\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) | |
| 2128 | for k :: int | |
| 2129 | proof | |
| 2130 | assume ?P | |
| 2131 | show ?Q | |
| 2132 | proof (rule ccontr) | |
| 2133 | assume \<open>\<not> 0 \<le> k\<close> | |
| 2134 | then have \<open>k < 0\<close> | |
| 2135 | by simp | |
| 2136 | with \<open>?P\<close> | |
| 2137 | have \<open>take_bit n k < 0\<close> | |
| 2138 | by (rule le_less_trans) | |
| 2139 | then show False | |
| 2140 | by simp | |
| 2141 | qed | |
| 2142 | next | |
| 2143 | assume ?Q | |
| 2144 | then show ?P | |
| 2145 | by (simp add: take_bit_eq_mod zmod_le_nonneg_dividend) | |
| 2146 | qed | |
| 2147 | ||
| 2148 | lemma take_bit_int_less_self_iff: | |
| 2149 | \<open>take_bit n k < k \<longleftrightarrow> 2 ^ n \<le> k\<close> | |
| 2150 | for k :: int | |
| 2151 | by (auto simp add: less_le take_bit_int_less_eq_self_iff take_bit_int_eq_self_iff | |
| 2152 | intro: order_trans [of 0 \<open>2 ^ n\<close> k]) | |
| 2153 | ||
| 2154 | lemma take_bit_int_greater_self_iff: | |
| 2155 | \<open>k < take_bit n k \<longleftrightarrow> k < 0\<close> | |
| 2156 | for k :: int | |
| 2157 | using take_bit_int_less_eq_self_iff [of n k] by auto | |
| 2158 | ||
| 2159 | lemma take_bit_int_greater_eq_self_iff: | |
| 2160 | \<open>k \<le> take_bit n k \<longleftrightarrow> k < 2 ^ n\<close> | |
| 2161 | for k :: int | |
| 2162 | by (auto simp add: le_less take_bit_int_greater_self_iff take_bit_int_eq_self_iff | |
| 2163 | dest: sym not_sym intro: less_trans [of k 0 \<open>2 ^ n\<close>]) | |
| 2164 | ||
| 2165 | lemma not_exp_less_eq_0_int [simp]: | |
| 2166 | \<open>\<not> 2 ^ n \<le> (0::int)\<close> | |
| 2167 | by (simp add: power_le_zero_eq) | |
| 2168 | ||
| 2169 | lemma half_nonnegative_int_iff [simp]: | |
| 2170 | \<open>k div 2 \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int | |
| 2171 | proof (cases \<open>k \<ge> 0\<close>) | |
| 2172 | case True | |
| 2173 | then show ?thesis | |
| 2174 | by (auto simp add: divide_int_def sgn_1_pos) | |
| 2175 | next | |
| 2176 | case False | |
| 2177 | then show ?thesis | |
| 2178 | by (auto simp add: divide_int_def not_le elim!: evenE) | |
| 2179 | qed | |
| 2180 | ||
| 2181 | lemma half_negative_int_iff [simp]: | |
| 2182 | \<open>k div 2 < 0 \<longleftrightarrow> k < 0\<close> for k :: int | |
| 2183 | by (subst Not_eq_iff [symmetric]) (simp add: not_less) | |
| 2184 | ||
| 2185 | lemma int_bit_bound: | |
| 2186 | fixes k :: int | |
| 2187 | obtains n where \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k n\<close> | |
| 2188 | and \<open>n > 0 \<Longrightarrow> bit k (n - 1) \<noteq> bit k n\<close> | |
| 2189 | proof - | |
| 2190 | obtain q where *: \<open>\<And>m. q \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k q\<close> | |
| 2191 | proof (cases \<open>k \<ge> 0\<close>) | |
| 2192 | case True | |
| 2193 | moreover from power_gt_expt [of 2 \<open>nat k\<close>] | |
| 2194 | have \<open>nat k < 2 ^ nat k\<close> | |
| 2195 | by simp | |
| 2196 | then have \<open>int (nat k) < int (2 ^ nat k)\<close> | |
| 2197 | by (simp only: of_nat_less_iff) | |
| 2198 | ultimately have *: \<open>k div 2 ^ nat k = 0\<close> | |
| 2199 | by simp | |
| 2200 | show thesis | |
| 2201 | proof (rule that [of \<open>nat k\<close>]) | |
| 2202 | fix m | |
| 2203 | assume \<open>nat k \<le> m\<close> | |
| 2204 | then show \<open>bit k m \<longleftrightarrow> bit k (nat k)\<close> | |
| 2205 | by (auto simp add: * bit_iff_odd power_add zdiv_zmult2_eq dest!: le_Suc_ex) | |
| 2206 | qed | |
| 2207 | next | |
| 2208 | case False | |
| 2209 | moreover from power_gt_expt [of 2 \<open>nat (- k)\<close>] | |
| 2210 | have \<open>nat (- k) < 2 ^ nat (- k)\<close> | |
| 2211 | by simp | |
| 2212 | then have \<open>int (nat (- k)) < int (2 ^ nat (- k))\<close> | |
| 2213 | by (simp only: of_nat_less_iff) | |
| 2214 | ultimately have \<open>- k div - (2 ^ nat (- k)) = - 1\<close> | |
| 2215 | by (subst div_pos_neg_trivial) simp_all | |
| 2216 | then have *: \<open>k div 2 ^ nat (- k) = - 1\<close> | |
| 2217 | by simp | |
| 2218 | show thesis | |
| 2219 | proof (rule that [of \<open>nat (- k)\<close>]) | |
| 2220 | fix m | |
| 2221 | assume \<open>nat (- k) \<le> m\<close> | |
| 2222 | then show \<open>bit k m \<longleftrightarrow> bit k (nat (- k))\<close> | |
| 2223 | by (auto simp add: * bit_iff_odd power_add zdiv_zmult2_eq minus_1_div_exp_eq_int dest!: le_Suc_ex) | |
| 2224 | qed | |
| 2225 | qed | |
| 2226 | show thesis | |
| 2227 | proof (cases \<open>\<forall>m. bit k m \<longleftrightarrow> bit k q\<close>) | |
| 2228 | case True | |
| 2229 | then have \<open>bit k 0 \<longleftrightarrow> bit k q\<close> | |
| 2230 | by blast | |
| 2231 | with True that [of 0] show thesis | |
| 2232 | by simp | |
| 2233 | next | |
| 2234 | case False | |
| 2235 | then obtain r where **: \<open>bit k r \<noteq> bit k q\<close> | |
| 2236 | by blast | |
| 2237 | have \<open>r < q\<close> | |
| 2238 | by (rule ccontr) (use * [of r] ** in simp) | |
| 2239 |     define N where \<open>N = {n. n < q \<and> bit k n \<noteq> bit k q}\<close>
 | |
| 2240 | moreover have \<open>finite N\<close> \<open>r \<in> N\<close> | |
| 2241 | using ** N_def \<open>r < q\<close> by auto | |
| 2242 | moreover define n where \<open>n = Suc (Max N)\<close> | |
| 2243 | ultimately have \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m \<longleftrightarrow> bit k n\<close> | |
| 2244 | apply auto | |
| 2245 | 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) | |
| 2246 | apply (metis "*" Max_ge Suc_n_not_le_n \<open>finite N\<close> linorder_not_less mem_Collect_eq) | |
| 2247 | apply (metis "*" Max_ge Suc_n_not_le_n \<open>finite N\<close> linorder_not_less mem_Collect_eq) | |
| 2248 | 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) | |
| 2249 | done | |
| 2250 | have \<open>bit k (Max N) \<noteq> bit k n\<close> | |
| 2251 | 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) | |
| 2252 | show thesis apply (rule that [of n]) | |
| 2253 | using \<open>\<And>m. n \<le> m \<Longrightarrow> bit k m = bit k n\<close> apply blast | |
| 2254 | using \<open>bit k (Max N) \<noteq> bit k n\<close> n_def by auto | |
| 2255 | qed | |
| 2256 | qed | |
| 2257 | ||
| 2258 | lemma take_bit_tightened_less_eq_int: | |
| 2259 | \<open>take_bit m k \<le> take_bit n k\<close> if \<open>m \<le> n\<close> for k :: int | |
| 2260 | proof - | |
| 2261 | have \<open>take_bit m (take_bit n k) \<le> take_bit n k\<close> | |
| 2262 | by (simp only: take_bit_int_less_eq_self_iff take_bit_nonnegative) | |
| 2263 | with that show ?thesis | |
| 2264 | by simp | |
| 2265 | qed | |
| 2266 | ||
| 2267 | context ring_bit_operations | |
| 2268 | begin | |
| 2269 | ||
| 2270 | lemma even_of_int_iff: | |
| 2271 | \<open>even (of_int k) \<longleftrightarrow> even k\<close> | |
| 2272 | by (induction k rule: int_bit_induct) simp_all | |
| 2273 | ||
| 2274 | lemma bit_of_int_iff [bit_simps]: | |
| 2275 |   \<open>bit (of_int k) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> bit k n\<close>
 | |
| 2276 | proof (cases \<open>possible_bit TYPE('a) n\<close>)
 | |
| 2277 | case False | |
| 2278 | then show ?thesis | |
| 2279 | by (simp add: impossible_bit) | |
| 2280 | next | |
| 2281 | case True | |
| 2282 | then have \<open>bit (of_int k) n \<longleftrightarrow> bit k n\<close> | |
| 2283 | proof (induction k arbitrary: n rule: int_bit_induct) | |
| 2284 | case zero | |
| 2285 | then show ?case | |
| 2286 | by simp | |
| 2287 | next | |
| 2288 | case minus | |
| 2289 | then show ?case | |
| 2290 | by simp | |
| 2291 | next | |
| 2292 | case (even k) | |
| 2293 | then show ?case | |
| 2294 | using bit_double_iff [of \<open>of_int k\<close> n] Bit_Operations.bit_double_iff [of k n] | |
| 2295 | by (cases n) (auto simp add: ac_simps possible_bit_def dest: mult_not_zero) | |
| 2296 | next | |
| 2297 | case (odd k) | |
| 2298 | then show ?case | |
| 2299 | using bit_double_iff [of \<open>of_int k\<close> n] | |
| 2300 | by (cases n) (auto simp add: ac_simps bit_double_iff even_bit_succ_iff Bit_Operations.bit_Suc possible_bit_def dest: mult_not_zero) | |
| 2301 | qed | |
| 2302 | with True show ?thesis | |
| 2303 | by simp | |
| 2304 | qed | |
| 2305 | ||
| 2306 | lemma push_bit_of_int: | |
| 2307 | \<open>push_bit n (of_int k) = of_int (push_bit n k)\<close> | |
| 2308 | by (simp add: push_bit_eq_mult Bit_Operations.push_bit_eq_mult) | |
| 2309 | ||
| 2310 | lemma of_int_push_bit: | |
| 2311 | \<open>of_int (push_bit n k) = push_bit n (of_int k)\<close> | |
| 2312 | by (simp add: push_bit_eq_mult Bit_Operations.push_bit_eq_mult) | |
| 2313 | ||
| 2314 | lemma take_bit_of_int: | |
| 2315 | \<open>take_bit n (of_int k) = of_int (take_bit n k)\<close> | |
| 2316 | by (rule bit_eqI) (simp add: bit_take_bit_iff Bit_Operations.bit_take_bit_iff bit_of_int_iff) | |
| 2317 | ||
| 2318 | lemma of_int_take_bit: | |
| 2319 | \<open>of_int (take_bit n k) = take_bit n (of_int k)\<close> | |
| 2320 | by (rule bit_eqI) (simp add: bit_take_bit_iff Bit_Operations.bit_take_bit_iff bit_of_int_iff) | |
| 2321 | ||
| 2322 | lemma of_int_not_eq: | |
| 2323 | \<open>of_int (NOT k) = NOT (of_int k)\<close> | |
| 2324 | by (rule bit_eqI) (simp add: bit_not_iff Bit_Operations.bit_not_iff bit_of_int_iff) | |
| 2325 | ||
| 2326 | lemma of_int_not_numeral: | |
| 2327 | \<open>of_int (NOT (numeral k)) = NOT (numeral k)\<close> | |
| 2328 | by (simp add: local.of_int_not_eq) | |
| 2329 | ||
| 2330 | lemma of_int_and_eq: | |
| 2331 | \<open>of_int (k AND l) = of_int k AND of_int l\<close> | |
| 2332 | by (rule bit_eqI) (simp add: bit_of_int_iff bit_and_iff Bit_Operations.bit_and_iff) | |
| 2333 | ||
| 2334 | lemma of_int_or_eq: | |
| 2335 | \<open>of_int (k OR l) = of_int k OR of_int l\<close> | |
| 2336 | by (rule bit_eqI) (simp add: bit_of_int_iff bit_or_iff Bit_Operations.bit_or_iff) | |
| 2337 | ||
| 2338 | lemma of_int_xor_eq: | |
| 2339 | \<open>of_int (k XOR l) = of_int k XOR of_int l\<close> | |
| 2340 | by (rule bit_eqI) (simp add: bit_of_int_iff bit_xor_iff Bit_Operations.bit_xor_iff) | |
| 2341 | ||
| 2342 | lemma of_int_mask_eq: | |
| 2343 | \<open>of_int (mask n) = mask n\<close> | |
| 2344 | by (induction n) (simp_all add: mask_Suc_double Bit_Operations.mask_Suc_double of_int_or_eq) | |
| 2345 | ||
| 2346 | end | |
| 2347 | ||
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2348 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2349 | subsection \<open>Instance \<^typ>\<open>nat\<close>\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2350 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2351 | instantiation nat :: semiring_bit_operations | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2352 | begin | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2353 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2354 | definition and_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2355 | where \<open>m AND n = nat (int m AND int n)\<close> for m n :: nat | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2356 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2357 | definition or_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2358 | where \<open>m OR n = nat (int m OR int n)\<close> for m n :: nat | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2359 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2360 | definition xor_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2361 | where \<open>m XOR n = nat (int m XOR int n)\<close> for m n :: nat | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2362 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2363 | definition mask_nat :: \<open>nat \<Rightarrow> nat\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2364 | where \<open>mask n = (2 :: nat) ^ n - 1\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2365 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2366 | definition push_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2367 | where \<open>push_bit_nat n m = m * 2 ^ n\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2368 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2369 | definition drop_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2370 | where \<open>drop_bit_nat n m = m div 2 ^ n\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2371 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2372 | definition take_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2373 | where \<open>take_bit_nat n m = m mod 2 ^ n\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2374 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2375 | definition set_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2376 | where \<open>set_bit m n = n OR push_bit m 1\<close> for m n :: nat | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2377 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2378 | definition unset_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2379 | where \<open>unset_bit m n = nat (unset_bit m (int n))\<close> for m n :: nat | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2380 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2381 | definition flip_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2382 | where \<open>flip_bit m n = n XOR push_bit m 1\<close> for m n :: nat | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2383 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2384 | instance proof | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2385 | fix m n q :: nat | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2386 | show \<open>bit (m AND n) q \<longleftrightarrow> bit m q \<and> bit n q\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2387 | by (simp add: and_nat_def bit_simps) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2388 | show \<open>bit (m OR n) q \<longleftrightarrow> bit m q \<or> bit n q\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2389 | by (simp add: or_nat_def bit_simps) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2390 | show \<open>bit (m XOR n) q \<longleftrightarrow> bit m q \<noteq> bit n q\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2391 | by (simp add: xor_nat_def bit_simps) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2392 | show \<open>bit (unset_bit m n) q \<longleftrightarrow> bit n q \<and> m \<noteq> q\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2393 | by (simp add: unset_bit_nat_def bit_simps) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2394 | qed (simp_all add: mask_nat_def set_bit_nat_def flip_bit_nat_def push_bit_nat_def drop_bit_nat_def take_bit_nat_def) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2395 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2396 | end | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2397 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2398 | lemma take_bit_nat_less_exp [simp]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2399 | \<open>take_bit n m < 2 ^ n\<close> for n m ::nat | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2400 | by (simp add: take_bit_eq_mod) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2401 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2402 | lemma take_bit_nat_eq_self_iff: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2403 | \<open>take_bit n m = m \<longleftrightarrow> m < 2 ^ n\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2404 | for n m :: nat | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2405 | proof | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2406 | assume ?P | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2407 | moreover note take_bit_nat_less_exp [of n m] | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2408 | ultimately show ?Q | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2409 | by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2410 | next | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2411 | assume ?Q | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2412 | then show ?P | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2413 | by (simp add: take_bit_eq_mod) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2414 | qed | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2415 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2416 | lemma take_bit_nat_eq_self: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2417 | \<open>take_bit n m = m\<close> if \<open>m < 2 ^ n\<close> for m n :: nat | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2418 | using that by (simp add: take_bit_nat_eq_self_iff) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2419 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2420 | lemma take_bit_nat_less_eq_self [simp]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2421 | \<open>take_bit n m \<le> m\<close> for n m :: nat | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2422 | by (simp add: take_bit_eq_mod) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2423 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2424 | lemma take_bit_nat_less_self_iff: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2425 | \<open>take_bit n m < m \<longleftrightarrow> 2 ^ n \<le> m\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2426 | for m n :: nat | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2427 | proof | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2428 | assume ?P | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2429 | then have \<open>take_bit n m \<noteq> m\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2430 | by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2431 | then show \<open>?Q\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2432 | by (simp add: take_bit_nat_eq_self_iff) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2433 | next | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2434 | have \<open>take_bit n m < 2 ^ n\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2435 | by (fact take_bit_nat_less_exp) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2436 | also assume ?Q | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2437 | finally show ?P . | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2438 | qed | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2439 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2440 | lemma bit_push_bit_iff_nat: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2441 | \<open>bit (push_bit m q) n \<longleftrightarrow> m \<le> n \<and> bit q (n - m)\<close> for q :: nat | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2442 | by (auto simp add: bit_push_bit_iff) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2443 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2444 | lemma and_nat_rec: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2445 | \<open>m AND n = of_bool (odd m \<and> odd n) + 2 * ((m div 2) AND (n div 2))\<close> for m n :: nat | 
| 74592 | 2446 | 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) | 
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2447 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2448 | lemma or_nat_rec: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2449 | \<open>m OR n = of_bool (odd m \<or> odd n) + 2 * ((m div 2) OR (n div 2))\<close> for m n :: nat | 
| 74592 | 2450 | 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) | 
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2451 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2452 | lemma xor_nat_rec: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2453 | \<open>m XOR n = of_bool (odd m \<noteq> odd n) + 2 * ((m div 2) XOR (n div 2))\<close> for m n :: nat | 
| 74592 | 2454 | 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) | 
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2455 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2456 | lemma Suc_0_and_eq [simp]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2457 | \<open>Suc 0 AND n = n mod 2\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2458 | using one_and_eq [of n] by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2459 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2460 | lemma and_Suc_0_eq [simp]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2461 | \<open>n AND Suc 0 = n mod 2\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2462 | using and_one_eq [of n] by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2463 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2464 | lemma Suc_0_or_eq: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2465 | \<open>Suc 0 OR n = n + of_bool (even n)\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2466 | using one_or_eq [of n] by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2467 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2468 | lemma or_Suc_0_eq: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2469 | \<open>n OR Suc 0 = n + of_bool (even n)\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2470 | using or_one_eq [of n] by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2471 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2472 | lemma Suc_0_xor_eq: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2473 | \<open>Suc 0 XOR n = n + of_bool (even n) - of_bool (odd n)\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2474 | using one_xor_eq [of n] by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2475 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2476 | lemma xor_Suc_0_eq: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2477 | \<open>n XOR Suc 0 = n + of_bool (even n) - of_bool (odd n)\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2478 | using xor_one_eq [of n] by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2479 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2480 | lemma and_nat_unfold [code]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2481 | \<open>m AND n = (if m = 0 \<or> n = 0 then 0 else (m mod 2) * (n mod 2) + 2 * ((m div 2) AND (n div 2)))\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2482 | for m n :: nat | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2483 | by (auto simp add: and_nat_rec [of m n] elim: oddE) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2484 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2485 | lemma or_nat_unfold [code]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2486 | \<open>m OR n = (if m = 0 then n else if n = 0 then m | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2487 | else max (m mod 2) (n mod 2) + 2 * ((m div 2) OR (n div 2)))\<close> for m n :: nat | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2488 | by (auto simp add: or_nat_rec [of m n] elim: oddE) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2489 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2490 | lemma xor_nat_unfold [code]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2491 | \<open>m XOR n = (if m = 0 then n else if n = 0 then m | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2492 | else (m mod 2 + n mod 2) mod 2 + 2 * ((m div 2) XOR (n div 2)))\<close> for m n :: nat | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2493 | by (auto simp add: xor_nat_rec [of m n] elim!: oddE) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2494 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2495 | lemma [code]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2496 | \<open>unset_bit 0 m = 2 * (m div 2)\<close> | 
| 74163 | 2497 | \<open>unset_bit (Suc n) m = m mod 2 + 2 * unset_bit n (m div 2)\<close> for m n :: nat | 
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2498 | by (simp_all add: unset_bit_Suc) | 
| 74495 | 2499 | |
| 74592 | 2500 | lemma push_bit_of_Suc_0 [simp]: | 
| 2501 | \<open>push_bit n (Suc 0) = 2 ^ n\<close> | |
| 2502 | using push_bit_of_1 [where ?'a = nat] by simp | |
| 2503 | ||
| 2504 | lemma take_bit_of_Suc_0 [simp]: | |
| 2505 | \<open>take_bit n (Suc 0) = of_bool (0 < n)\<close> | |
| 2506 | using take_bit_of_1 [where ?'a = nat] by simp | |
| 2507 | ||
| 2508 | lemma drop_bit_of_Suc_0 [simp]: | |
| 2509 | \<open>drop_bit n (Suc 0) = of_bool (n = 0)\<close> | |
| 2510 | using drop_bit_of_1 [where ?'a = nat] by simp | |
| 2511 | ||
| 2512 | lemma Suc_mask_eq_exp: | |
| 2513 | \<open>Suc (mask n) = 2 ^ n\<close> | |
| 2514 | by (simp add: mask_eq_exp_minus_1) | |
| 2515 | ||
| 2516 | lemma less_eq_mask: | |
| 2517 | \<open>n \<le> mask n\<close> | |
| 2518 | by (simp add: mask_eq_exp_minus_1 le_diff_conv2) | |
| 2519 | (metis Suc_mask_eq_exp diff_Suc_1 diff_le_diff_pow diff_zero le_refl not_less_eq_eq power_0) | |
| 2520 | ||
| 2521 | lemma less_mask: | |
| 2522 | \<open>n < mask n\<close> if \<open>Suc 0 < n\<close> | |
| 2523 | proof - | |
| 2524 | define m where \<open>m = n - 2\<close> | |
| 2525 | with that have *: \<open>n = m + 2\<close> | |
| 2526 | by simp | |
| 2527 | have \<open>Suc (Suc (Suc m)) < 4 * 2 ^ m\<close> | |
| 2528 | by (induction m) simp_all | |
| 2529 | then have \<open>Suc (m + 2) < Suc (mask (m + 2))\<close> | |
| 2530 | by (simp add: Suc_mask_eq_exp) | |
| 2531 | then have \<open>m + 2 < mask (m + 2)\<close> | |
| 2532 | by (simp add: less_le) | |
| 2533 | with * show ?thesis | |
| 2534 | by simp | |
| 2535 | qed | |
| 2536 | ||
| 2537 | lemma mask_nat_less_exp [simp]: | |
| 2538 | \<open>(mask n :: nat) < 2 ^ n\<close> | |
| 2539 | by (simp add: mask_eq_exp_minus_1) | |
| 2540 | ||
| 2541 | lemma mask_nat_positive_iff [simp]: | |
| 2542 | \<open>(0::nat) < mask n \<longleftrightarrow> 0 < n\<close> | |
| 2543 | proof (cases \<open>n = 0\<close>) | |
| 2544 | case True | |
| 2545 | then show ?thesis | |
| 2546 | by simp | |
| 2547 | next | |
| 2548 | case False | |
| 2549 | then have \<open>0 < n\<close> | |
| 2550 | by simp | |
| 2551 | then have \<open>(0::nat) < mask n\<close> | |
| 2552 | using less_eq_mask [of n] by (rule order_less_le_trans) | |
| 2553 | with \<open>0 < n\<close> show ?thesis | |
| 2554 | by simp | |
| 2555 | qed | |
| 2556 | ||
| 2557 | lemma take_bit_tightened_less_eq_nat: | |
| 2558 | \<open>take_bit m q \<le> take_bit n q\<close> if \<open>m \<le> n\<close> for q :: nat | |
| 2559 | proof - | |
| 2560 | have \<open>take_bit m (take_bit n q) \<le> take_bit n q\<close> | |
| 2561 | by (rule take_bit_nat_less_eq_self) | |
| 2562 | with that show ?thesis | |
| 2563 | by simp | |
| 2564 | qed | |
| 2565 | ||
| 2566 | lemma push_bit_nat_eq: | |
| 2567 | \<open>push_bit n (nat k) = nat (push_bit n k)\<close> | |
| 2568 | by (cases \<open>k \<ge> 0\<close>) (simp_all add: push_bit_eq_mult nat_mult_distrib not_le mult_nonneg_nonpos2) | |
| 2569 | ||
| 2570 | lemma drop_bit_nat_eq: | |
| 2571 | \<open>drop_bit n (nat k) = nat (drop_bit n k)\<close> | |
| 2572 | apply (cases \<open>k \<ge> 0\<close>) | |
| 2573 | apply (simp_all add: drop_bit_eq_div nat_div_distrib nat_power_eq not_le) | |
| 2574 | apply (simp add: divide_int_def) | |
| 2575 | done | |
| 2576 | ||
| 2577 | lemma take_bit_nat_eq: | |
| 2578 | \<open>take_bit n (nat k) = nat (take_bit n k)\<close> if \<open>k \<ge> 0\<close> | |
| 2579 | using that by (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq) | |
| 2580 | ||
| 2581 | lemma nat_take_bit_eq: | |
| 2582 | \<open>nat (take_bit n k) = take_bit n (nat k)\<close> | |
| 2583 | if \<open>k \<ge> 0\<close> | |
| 2584 | using that by (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq) | |
| 2585 | ||
| 74495 | 2586 | context semiring_bit_operations | 
| 2587 | begin | |
| 2588 | ||
| 2589 | lemma push_bit_of_nat: | |
| 2590 | \<open>push_bit n (of_nat m) = of_nat (push_bit n m)\<close> | |
| 2591 | by (simp add: push_bit_eq_mult Bit_Operations.push_bit_eq_mult) | |
| 2592 | ||
| 2593 | lemma of_nat_push_bit: | |
| 2594 | \<open>of_nat (push_bit m n) = push_bit m (of_nat n)\<close> | |
| 2595 | by (simp add: push_bit_eq_mult Bit_Operations.push_bit_eq_mult) | |
| 2596 | ||
| 2597 | lemma take_bit_of_nat: | |
| 2598 | \<open>take_bit n (of_nat m) = of_nat (take_bit n m)\<close> | |
| 2599 | by (rule bit_eqI) (simp add: bit_take_bit_iff Bit_Operations.bit_take_bit_iff bit_of_nat_iff) | |
| 2600 | ||
| 2601 | lemma of_nat_take_bit: | |
| 2602 | \<open>of_nat (take_bit n m) = take_bit n (of_nat m)\<close> | |
| 2603 | by (rule bit_eqI) (simp add: bit_take_bit_iff Bit_Operations.bit_take_bit_iff bit_of_nat_iff) | |
| 2604 | ||
| 2605 | end | |
| 2606 | ||
| 74592 | 2607 | context semiring_bit_operations | 
| 2608 | begin | |
| 2609 | ||
| 2610 | lemma of_nat_and_eq: | |
| 2611 | \<open>of_nat (m AND n) = of_nat m AND of_nat n\<close> | |
| 2612 | by (rule bit_eqI) (simp add: bit_of_nat_iff bit_and_iff Bit_Operations.bit_and_iff) | |
| 2613 | ||
| 2614 | lemma of_nat_or_eq: | |
| 2615 | \<open>of_nat (m OR n) = of_nat m OR of_nat n\<close> | |
| 2616 | by (rule bit_eqI) (simp add: bit_of_nat_iff bit_or_iff Bit_Operations.bit_or_iff) | |
| 2617 | ||
| 2618 | lemma of_nat_xor_eq: | |
| 2619 | \<open>of_nat (m XOR n) = of_nat m XOR of_nat n\<close> | |
| 2620 | by (rule bit_eqI) (simp add: bit_of_nat_iff bit_xor_iff Bit_Operations.bit_xor_iff) | |
| 2621 | ||
| 2622 | lemma of_nat_mask_eq: | |
| 2623 | \<open>of_nat (mask n) = mask n\<close> | |
| 2624 | by (induction n) (simp_all add: mask_Suc_double Bit_Operations.mask_Suc_double of_nat_or_eq) | |
| 2625 | ||
| 2626 | end | |
| 2627 | ||
| 2628 | lemma nat_mask_eq: | |
| 2629 | \<open>nat (mask n) = mask n\<close> | |
| 2630 | by (simp add: nat_eq_iff of_nat_mask_eq) | |
| 2631 | ||
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2632 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2633 | subsection \<open>Common algebraic structure\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2634 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2635 | class unique_euclidean_semiring_with_bit_operations = | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2636 | unique_euclidean_semiring_with_nat + semiring_bit_operations | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2637 | begin | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2638 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2639 | lemma take_bit_of_exp [simp]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2640 | \<open>take_bit m (2 ^ n) = of_bool (n < m) * 2 ^ n\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2641 | by (simp add: take_bit_eq_mod exp_mod_exp) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2642 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2643 | lemma take_bit_of_2 [simp]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2644 | \<open>take_bit n 2 = of_bool (2 \<le> n) * 2\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2645 | using take_bit_of_exp [of n 1] by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2646 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2647 | lemma push_bit_eq_0_iff [simp]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2648 | "push_bit n a = 0 \<longleftrightarrow> a = 0" | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2649 | by (simp add: push_bit_eq_mult) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2650 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2651 | lemma take_bit_add: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2652 | "take_bit n (take_bit n a + take_bit n b) = take_bit n (a + b)" | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2653 | by (simp add: take_bit_eq_mod mod_simps) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2654 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2655 | lemma take_bit_of_1_eq_0_iff [simp]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2656 | "take_bit n 1 = 0 \<longleftrightarrow> n = 0" | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2657 | by (simp add: take_bit_eq_mod) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2658 | |
| 74497 | 2659 | lemma drop_bit_Suc_bit0 [simp]: | 
| 2660 | \<open>drop_bit (Suc n) (numeral (Num.Bit0 k)) = drop_bit n (numeral k)\<close> | |
| 2661 | by (simp add: drop_bit_Suc numeral_Bit0_div_2) | |
| 2662 | ||
| 2663 | lemma drop_bit_Suc_bit1 [simp]: | |
| 2664 | \<open>drop_bit (Suc n) (numeral (Num.Bit1 k)) = drop_bit n (numeral k)\<close> | |
| 2665 | by (simp add: drop_bit_Suc numeral_Bit1_div_2) | |
| 2666 | ||
| 2667 | lemma drop_bit_numeral_bit0 [simp]: | |
| 2668 | \<open>drop_bit (numeral l) (numeral (Num.Bit0 k)) = drop_bit (pred_numeral l) (numeral k)\<close> | |
| 2669 | by (simp add: drop_bit_rec numeral_Bit0_div_2) | |
| 2670 | ||
| 2671 | lemma drop_bit_numeral_bit1 [simp]: | |
| 2672 | \<open>drop_bit (numeral l) (numeral (Num.Bit1 k)) = drop_bit (pred_numeral l) (numeral k)\<close> | |
| 2673 | by (simp add: drop_bit_rec numeral_Bit1_div_2) | |
| 2674 | ||
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2675 | lemma take_bit_Suc_1 [simp]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2676 | \<open>take_bit (Suc n) 1 = 1\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2677 | by (simp add: take_bit_Suc) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2678 | |
| 74592 | 2679 | lemma take_bit_Suc_bit0: | 
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2680 | \<open>take_bit (Suc n) (numeral (Num.Bit0 k)) = take_bit n (numeral k) * 2\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2681 | by (simp add: take_bit_Suc numeral_Bit0_div_2) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2682 | |
| 74592 | 2683 | lemma take_bit_Suc_bit1: | 
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2684 | \<open>take_bit (Suc n) (numeral (Num.Bit1 k)) = take_bit n (numeral k) * 2 + 1\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2685 | by (simp add: take_bit_Suc numeral_Bit1_div_2 mod_2_eq_odd) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2686 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2687 | lemma take_bit_numeral_1 [simp]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2688 | \<open>take_bit (numeral l) 1 = 1\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2689 | by (simp add: take_bit_rec [of \<open>numeral l\<close> 1]) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2690 | |
| 74592 | 2691 | lemma take_bit_numeral_bit0: | 
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2692 | \<open>take_bit (numeral l) (numeral (Num.Bit0 k)) = take_bit (pred_numeral l) (numeral k) * 2\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2693 | by (simp add: take_bit_rec numeral_Bit0_div_2) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2694 | |
| 74592 | 2695 | lemma take_bit_numeral_bit1: | 
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2696 | \<open>take_bit (numeral l) (numeral (Num.Bit1 k)) = take_bit (pred_numeral l) (numeral k) * 2 + 1\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2697 | by (simp add: take_bit_rec numeral_Bit1_div_2 mod_2_eq_odd) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2698 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2699 | lemma bit_of_nat_iff_bit [bit_simps]: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2700 | \<open>bit (of_nat m) n \<longleftrightarrow> bit m n\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2701 | proof - | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2702 | have \<open>even (m div 2 ^ n) \<longleftrightarrow> even (of_nat (m div 2 ^ n))\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2703 | by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2704 | also have \<open>of_nat (m div 2 ^ n) = of_nat m div of_nat (2 ^ n)\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2705 | by (simp add: of_nat_div) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2706 | finally show ?thesis | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2707 | by (simp add: bit_iff_odd semiring_bits_class.bit_iff_odd) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2708 | qed | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2709 | |
| 74497 | 2710 | lemma drop_bit_of_nat: | 
| 2711 | "drop_bit n (of_nat m) = of_nat (drop_bit n m)" | |
| 2712 | by (simp add: drop_bit_eq_div Bit_Operations.drop_bit_eq_div of_nat_div [of m "2 ^ n"]) | |
| 2713 | ||
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2714 | lemma of_nat_drop_bit: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2715 | \<open>of_nat (drop_bit m n) = drop_bit m (of_nat n)\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2716 | by (simp add: drop_bit_eq_div Bit_Operations.drop_bit_eq_div of_nat_div) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2717 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2718 | lemma take_bit_sum: | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2719 | "take_bit n a = (\<Sum>k = 0..<n. push_bit k (of_bool (bit a k)))" | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2720 | for n :: nat | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2721 | proof (induction n arbitrary: a) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2722 | case 0 | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2723 | then show ?case | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2724 | by simp | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2725 | next | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2726 | case (Suc n) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2727 | have "(\<Sum>k = 0..<Suc n. push_bit k (of_bool (bit a k))) = | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2728 | of_bool (odd a) + (\<Sum>k = Suc 0..<Suc n. push_bit k (of_bool (bit a k)))" | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2729 | by (simp add: sum.atLeast_Suc_lessThan ac_simps) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2730 | also have "(\<Sum>k = Suc 0..<Suc n. push_bit k (of_bool (bit a k))) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2731 | = (\<Sum>k = 0..<n. push_bit k (of_bool (bit (a div 2) k))) * 2" | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2732 | by (simp only: sum.atLeast_Suc_lessThan_Suc_shift) (simp add: sum_distrib_right push_bit_double drop_bit_Suc bit_Suc) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2733 | finally show ?case | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2734 | using Suc [of "a div 2"] by (simp add: ac_simps take_bit_Suc mod_2_eq_odd) | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2735 | qed | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2736 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2737 | end | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2738 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2739 | instance nat :: unique_euclidean_semiring_with_bit_operations .. | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2740 | |
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2741 | instance int :: unique_euclidean_semiring_with_bit_operations .. | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2742 | |
| 74497 | 2743 | lemma drop_bit_Suc_minus_bit0 [simp]: | 
| 2744 | \<open>drop_bit (Suc n) (- numeral (Num.Bit0 k)) = drop_bit n (- numeral k :: int)\<close> | |
| 2745 | by (simp add: drop_bit_Suc numeral_Bit0_div_2) | |
| 2746 | ||
| 2747 | lemma drop_bit_Suc_minus_bit1 [simp]: | |
| 2748 | \<open>drop_bit (Suc n) (- numeral (Num.Bit1 k)) = drop_bit n (- numeral (Num.inc k) :: int)\<close> | |
| 2749 | by (simp add: drop_bit_Suc numeral_Bit1_div_2 add_One) | |
| 2750 | ||
| 2751 | lemma drop_bit_numeral_minus_bit0 [simp]: | |
| 2752 | \<open>drop_bit (numeral l) (- numeral (Num.Bit0 k)) = drop_bit (pred_numeral l) (- numeral k :: int)\<close> | |
| 2753 | by (simp add: numeral_eq_Suc numeral_Bit0_div_2) | |
| 2754 | ||
| 2755 | lemma drop_bit_numeral_minus_bit1 [simp]: | |
| 2756 | \<open>drop_bit (numeral l) (- numeral (Num.Bit1 k)) = drop_bit (pred_numeral l) (- numeral (Num.inc k) :: int)\<close> | |
| 2757 | by (simp add: numeral_eq_Suc numeral_Bit1_div_2) | |
| 2758 | ||
| 74592 | 2759 | lemma take_bit_Suc_minus_bit0: | 
| 74497 | 2760 | \<open>take_bit (Suc n) (- numeral (Num.Bit0 k)) = take_bit n (- numeral k) * (2 :: int)\<close> | 
| 2761 | by (simp add: take_bit_Suc numeral_Bit0_div_2) | |
| 2762 | ||
| 74592 | 2763 | lemma take_bit_Suc_minus_bit1: | 
| 74497 | 2764 | \<open>take_bit (Suc n) (- numeral (Num.Bit1 k)) = take_bit n (- numeral (Num.inc k)) * 2 + (1 :: int)\<close> | 
| 2765 | by (simp add: take_bit_Suc numeral_Bit1_div_2 add_One) | |
| 2766 | ||
| 74592 | 2767 | lemma take_bit_numeral_minus_bit0: | 
| 74497 | 2768 | \<open>take_bit (numeral l) (- numeral (Num.Bit0 k)) = take_bit (pred_numeral l) (- numeral k) * (2 :: int)\<close> | 
| 74592 | 2769 | by (simp add: numeral_eq_Suc numeral_Bit0_div_2 take_bit_Suc_minus_bit0) | 
| 2770 | ||
| 2771 | lemma take_bit_numeral_minus_bit1: | |
| 74497 | 2772 | \<open>take_bit (numeral l) (- numeral (Num.Bit1 k)) = take_bit (pred_numeral l) (- numeral (Num.inc k)) * 2 + (1 :: int)\<close> | 
| 74592 | 2773 | by (simp add: numeral_eq_Suc numeral_Bit1_div_2 take_bit_Suc_minus_bit1) | 
| 74497 | 2774 | |
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 2775 | |
| 74163 | 2776 | subsection \<open>Symbolic computations on numeral expressions\<close> | 
| 2777 | ||
| 2778 | context unique_euclidean_semiring_with_bit_operations | |
| 2779 | begin | |
| 2780 | ||
| 2781 | lemma bit_numeral_iff: | |
| 2782 | \<open>bit (numeral m) n \<longleftrightarrow> bit (numeral m :: nat) n\<close> | |
| 2783 | using bit_of_nat_iff_bit [of \<open>numeral m\<close> n] by simp | |
| 2784 | ||
| 2785 | lemma bit_numeral_Bit0_Suc_iff [simp]: | |
| 2786 | \<open>bit (numeral (Num.Bit0 m)) (Suc n) \<longleftrightarrow> bit (numeral m) n\<close> | |
| 2787 | by (simp add: bit_Suc numeral_Bit0_div_2) | |
| 2788 | ||
| 2789 | lemma bit_numeral_Bit1_Suc_iff [simp]: | |
| 2790 | \<open>bit (numeral (Num.Bit1 m)) (Suc n) \<longleftrightarrow> bit (numeral m) n\<close> | |
| 2791 | by (simp add: bit_Suc numeral_Bit1_div_2) | |
| 2792 | ||
| 2793 | lemma bit_numeral_rec: | |
| 2794 | \<open>bit (numeral (Num.Bit0 w)) n \<longleftrightarrow> (case n of 0 \<Rightarrow> False | Suc m \<Rightarrow> bit (numeral w) m)\<close> | |
| 2795 | \<open>bit (numeral (Num.Bit1 w)) n \<longleftrightarrow> (case n of 0 \<Rightarrow> True | Suc m \<Rightarrow> bit (numeral w) m)\<close> | |
| 2796 | by (cases n; simp)+ | |
| 2797 | ||
| 2798 | lemma bit_numeral_simps [simp]: | |
| 2799 | \<open>\<not> bit 1 (numeral n)\<close> | |
| 2800 | \<open>bit (numeral (Num.Bit0 w)) (numeral n) \<longleftrightarrow> bit (numeral w) (pred_numeral n)\<close> | |
| 2801 | \<open>bit (numeral (Num.Bit1 w)) (numeral n) \<longleftrightarrow> bit (numeral w) (pred_numeral n)\<close> | |
| 2802 | by (simp_all add: bit_1_iff numeral_eq_Suc) | |
| 2803 | ||
| 2804 | lemma and_numerals [simp]: | |
| 2805 | \<open>1 AND numeral (Num.Bit0 y) = 0\<close> | |
| 2806 | \<open>1 AND numeral (Num.Bit1 y) = 1\<close> | |
| 2807 | \<open>numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = 2 * (numeral x AND numeral y)\<close> | |
| 2808 | \<open>numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = 2 * (numeral x AND numeral y)\<close> | |
| 2809 | \<open>numeral (Num.Bit0 x) AND 1 = 0\<close> | |
| 2810 | \<open>numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = 2 * (numeral x AND numeral y)\<close> | |
| 2811 | \<open>numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = 1 + 2 * (numeral x AND numeral y)\<close> | |
| 2812 | \<open>numeral (Num.Bit1 x) AND 1 = 1\<close> | |
| 2813 | by (simp_all add: bit_eq_iff) (simp_all add: bit_simps bit_Suc bit_numeral_rec split: nat.splits) | |
| 2814 | ||
| 2815 | fun and_num :: \<open>num \<Rightarrow> num \<Rightarrow> num option\<close> \<^marker>\<open>contributor \<open>Andreas Lochbihler\<close>\<close> | |
| 2816 | where | |
| 2817 | \<open>and_num num.One num.One = Some num.One\<close> | |
| 2818 | | \<open>and_num num.One (num.Bit0 n) = None\<close> | |
| 2819 | | \<open>and_num num.One (num.Bit1 n) = Some num.One\<close> | |
| 2820 | | \<open>and_num (num.Bit0 m) num.One = None\<close> | |
| 2821 | | \<open>and_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (and_num m n)\<close> | |
| 2822 | | \<open>and_num (num.Bit0 m) (num.Bit1 n) = map_option num.Bit0 (and_num m n)\<close> | |
| 2823 | | \<open>and_num (num.Bit1 m) num.One = Some num.One\<close> | |
| 2824 | | \<open>and_num (num.Bit1 m) (num.Bit0 n) = map_option num.Bit0 (and_num m n)\<close> | |
| 2825 | | \<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> | |
| 2826 | ||
| 2827 | lemma numeral_and_num: | |
| 2828 | \<open>numeral m AND numeral n = (case and_num m n of None \<Rightarrow> 0 | Some n' \<Rightarrow> numeral n')\<close> | |
| 2829 | by (induction m n rule: and_num.induct) (simp_all add: split: option.split) | |
| 2830 | ||
| 2831 | lemma and_num_eq_None_iff: | |
| 2832 | \<open>and_num m n = None \<longleftrightarrow> numeral m AND numeral n = 0\<close> | |
| 2833 | by (simp add: numeral_and_num split: option.split) | |
| 2834 | ||
| 2835 | lemma and_num_eq_Some_iff: | |
| 2836 | \<open>and_num m n = Some q \<longleftrightarrow> numeral m AND numeral n = numeral q\<close> | |
| 2837 | by (simp add: numeral_and_num split: option.split) | |
| 2838 | ||
| 2839 | lemma or_numerals [simp]: | |
| 2840 | \<open>1 OR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)\<close> | |
| 2841 | \<open>1 OR numeral (Num.Bit1 y) = numeral (Num.Bit1 y)\<close> | |
| 2842 | \<open>numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = 2 * (numeral x OR numeral y)\<close> | |
| 2843 | \<open>numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = 1 + 2 * (numeral x OR numeral y)\<close> | |
| 2844 | \<open>numeral (Num.Bit0 x) OR 1 = numeral (Num.Bit1 x)\<close> | |
| 2845 | \<open>numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = 1 + 2 * (numeral x OR numeral y)\<close> | |
| 2846 | \<open>numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = 1 + 2 * (numeral x OR numeral y)\<close> | |
| 2847 | \<open>numeral (Num.Bit1 x) OR 1 = numeral (Num.Bit1 x)\<close> | |
| 2848 | by (simp_all add: bit_eq_iff) (simp_all add: bit_simps bit_Suc bit_numeral_rec split: nat.splits) | |
| 2849 | ||
| 2850 | fun or_num :: \<open>num \<Rightarrow> num \<Rightarrow> num\<close> \<^marker>\<open>contributor \<open>Andreas Lochbihler\<close>\<close> | |
| 2851 | where | |
| 2852 | \<open>or_num num.One num.One = num.One\<close> | |
| 2853 | | \<open>or_num num.One (num.Bit0 n) = num.Bit1 n\<close> | |
| 2854 | | \<open>or_num num.One (num.Bit1 n) = num.Bit1 n\<close> | |
| 2855 | | \<open>or_num (num.Bit0 m) num.One = num.Bit1 m\<close> | |
| 2856 | | \<open>or_num (num.Bit0 m) (num.Bit0 n) = num.Bit0 (or_num m n)\<close> | |
| 2857 | | \<open>or_num (num.Bit0 m) (num.Bit1 n) = num.Bit1 (or_num m n)\<close> | |
| 2858 | | \<open>or_num (num.Bit1 m) num.One = num.Bit1 m\<close> | |
| 2859 | | \<open>or_num (num.Bit1 m) (num.Bit0 n) = num.Bit1 (or_num m n)\<close> | |
| 2860 | | \<open>or_num (num.Bit1 m) (num.Bit1 n) = num.Bit1 (or_num m n)\<close> | |
| 2861 | ||
| 2862 | lemma numeral_or_num: | |
| 2863 | \<open>numeral m OR numeral n = numeral (or_num m n)\<close> | |
| 2864 | by (induction m n rule: or_num.induct) simp_all | |
| 2865 | ||
| 2866 | lemma numeral_or_num_eq: | |
| 2867 | \<open>numeral (or_num m n) = numeral m OR numeral n\<close> | |
| 2868 | by (simp add: numeral_or_num) | |
| 2869 | ||
| 2870 | lemma xor_numerals [simp]: | |
| 2871 | \<open>1 XOR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)\<close> | |
| 2872 | \<open>1 XOR numeral (Num.Bit1 y) = numeral (Num.Bit0 y)\<close> | |
| 2873 | \<open>numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = 2 * (numeral x XOR numeral y)\<close> | |
| 2874 | \<open>numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = 1 + 2 * (numeral x XOR numeral y)\<close> | |
| 2875 | \<open>numeral (Num.Bit0 x) XOR 1 = numeral (Num.Bit1 x)\<close> | |
| 2876 | \<open>numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = 1 + 2 * (numeral x XOR numeral y)\<close> | |
| 2877 | \<open>numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = 2 * (numeral x XOR numeral y)\<close> | |
| 2878 | \<open>numeral (Num.Bit1 x) XOR 1 = numeral (Num.Bit0 x)\<close> | |
| 2879 | by (simp_all add: bit_eq_iff) (simp_all add: bit_simps bit_Suc bit_numeral_rec split: nat.splits) | |
| 2880 | ||
| 2881 | fun xor_num :: \<open>num \<Rightarrow> num \<Rightarrow> num option\<close> \<^marker>\<open>contributor \<open>Andreas Lochbihler\<close>\<close> | |
| 2882 | where | |
| 2883 | \<open>xor_num num.One num.One = None\<close> | |
| 2884 | | \<open>xor_num num.One (num.Bit0 n) = Some (num.Bit1 n)\<close> | |
| 2885 | | \<open>xor_num num.One (num.Bit1 n) = Some (num.Bit0 n)\<close> | |
| 2886 | | \<open>xor_num (num.Bit0 m) num.One = Some (num.Bit1 m)\<close> | |
| 2887 | | \<open>xor_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (xor_num m n)\<close> | |
| 2888 | | \<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> | |
| 2889 | | \<open>xor_num (num.Bit1 m) num.One = Some (num.Bit0 m)\<close> | |
| 2890 | | \<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> | |
| 2891 | | \<open>xor_num (num.Bit1 m) (num.Bit1 n) = map_option num.Bit0 (xor_num m n)\<close> | |
| 2892 | ||
| 2893 | lemma numeral_xor_num: | |
| 2894 | \<open>numeral m XOR numeral n = (case xor_num m n of None \<Rightarrow> 0 | Some n' \<Rightarrow> numeral n')\<close> | |
| 2895 | by (induction m n rule: xor_num.induct) (simp_all split: option.split) | |
| 2896 | ||
| 2897 | lemma xor_num_eq_None_iff: | |
| 2898 | \<open>xor_num m n = None \<longleftrightarrow> numeral m XOR numeral n = 0\<close> | |
| 2899 | by (simp add: numeral_xor_num split: option.split) | |
| 2900 | ||
| 2901 | lemma xor_num_eq_Some_iff: | |
| 2902 | \<open>xor_num m n = Some q \<longleftrightarrow> numeral m XOR numeral n = numeral q\<close> | |
| 2903 | by (simp add: numeral_xor_num split: option.split) | |
| 2904 | ||
| 2905 | end | |
| 2906 | ||
| 74495 | 2907 | lemma bit_Suc_0_iff [bit_simps]: | 
| 2908 | \<open>bit (Suc 0) n \<longleftrightarrow> n = 0\<close> | |
| 2909 | using bit_1_iff [of n, where ?'a = nat] by simp | |
| 2910 | ||
| 2911 | lemma and_nat_numerals [simp]: | |
| 2912 | \<open>Suc 0 AND numeral (Num.Bit0 y) = 0\<close> | |
| 2913 | \<open>Suc 0 AND numeral (Num.Bit1 y) = 1\<close> | |
| 2914 | \<open>numeral (Num.Bit0 x) AND Suc 0 = 0\<close> | |
| 2915 | \<open>numeral (Num.Bit1 x) AND Suc 0 = 1\<close> | |
| 2916 | by (simp_all only: and_numerals flip: One_nat_def) | |
| 2917 | ||
| 2918 | lemma or_nat_numerals [simp]: | |
| 2919 | \<open>Suc 0 OR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)\<close> | |
| 2920 | \<open>Suc 0 OR numeral (Num.Bit1 y) = numeral (Num.Bit1 y)\<close> | |
| 2921 | \<open>numeral (Num.Bit0 x) OR Suc 0 = numeral (Num.Bit1 x)\<close> | |
| 2922 | \<open>numeral (Num.Bit1 x) OR Suc 0 = numeral (Num.Bit1 x)\<close> | |
| 2923 | by (simp_all only: or_numerals flip: One_nat_def) | |
| 2924 | ||
| 2925 | lemma xor_nat_numerals [simp]: | |
| 2926 | \<open>Suc 0 XOR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)\<close> | |
| 2927 | \<open>Suc 0 XOR numeral (Num.Bit1 y) = numeral (Num.Bit0 y)\<close> | |
| 2928 | \<open>numeral (Num.Bit0 x) XOR Suc 0 = numeral (Num.Bit1 x)\<close> | |
| 2929 | \<open>numeral (Num.Bit1 x) XOR Suc 0 = numeral (Num.Bit0 x)\<close> | |
| 2930 | by (simp_all only: xor_numerals flip: One_nat_def) | |
| 2931 | ||
| 74163 | 2932 | context ring_bit_operations | 
| 2933 | begin | |
| 2934 | ||
| 2935 | lemma minus_numeral_inc_eq: | |
| 2936 | \<open>- numeral (Num.inc n) = NOT (numeral n)\<close> | |
| 2937 | by (simp add: not_eq_complement sub_inc_One_eq add_One) | |
| 2938 | ||
| 2939 | lemma sub_one_eq_not_neg: | |
| 2940 | \<open>Num.sub n num.One = NOT (- numeral n)\<close> | |
| 2941 | by (simp add: not_eq_complement) | |
| 2942 | ||
| 2943 | lemma minus_numeral_eq_not_sub_one: | |
| 2944 | \<open>- numeral n = NOT (Num.sub n num.One)\<close> | |
| 2945 | by (simp add: not_eq_complement) | |
| 2946 | ||
| 74495 | 2947 | lemma not_numeral_eq [simp]: | 
| 74163 | 2948 | \<open>NOT (numeral n) = - numeral (Num.inc n)\<close> | 
| 2949 | by (simp add: minus_numeral_inc_eq) | |
| 2950 | ||
| 2951 | lemma not_minus_numeral_eq [simp]: | |
| 2952 | \<open>NOT (- numeral n) = Num.sub n num.One\<close> | |
| 2953 | by (simp add: sub_one_eq_not_neg) | |
| 2954 | ||
| 2955 | lemma minus_not_numeral_eq [simp]: | |
| 2956 | \<open>- (NOT (numeral n)) = numeral (Num.inc n)\<close> | |
| 74495 | 2957 | by simp | 
| 2958 | ||
| 2959 | lemma not_numeral_BitM_eq: | |
| 2960 | \<open>NOT (numeral (Num.BitM n)) = - numeral (num.Bit0 n)\<close> | |
| 2961 | by (simp add: inc_BitM_eq) | |
| 2962 | ||
| 2963 | lemma not_numeral_Bit0_eq: | |
| 2964 | \<open>NOT (numeral (Num.Bit0 n)) = - numeral (num.Bit1 n)\<close> | |
| 2965 | by simp | |
| 74163 | 2966 | |
| 2967 | end | |
| 2968 | ||
| 2969 | lemma bit_minus_numeral_int [simp]: | |
| 2970 | \<open>bit (- numeral (num.Bit0 w) :: int) (numeral n) \<longleftrightarrow> bit (- numeral w :: int) (pred_numeral n)\<close> | |
| 2971 | \<open>bit (- numeral (num.Bit1 w) :: int) (numeral n) \<longleftrightarrow> \<not> bit (numeral w :: int) (pred_numeral n)\<close> | |
| 2972 | by (simp_all add: bit_minus_iff bit_not_iff numeral_eq_Suc bit_Suc add_One sub_inc_One_eq) | |
| 2973 | ||
| 74592 | 2974 | lemma bit_minus_numeral_Bit0_Suc_iff [simp]: | 
| 2975 | \<open>bit (- numeral (num.Bit0 w) :: int) (Suc n) \<longleftrightarrow> bit (- numeral w :: int) n\<close> | |
| 2976 | by (simp add: bit_Suc) | |
| 2977 | ||
| 2978 | lemma bit_minus_numeral_Bit1_Suc_iff [simp]: | |
| 2979 | \<open>bit (- numeral (num.Bit1 w) :: int) (Suc n) \<longleftrightarrow> \<not> bit (numeral w :: int) n\<close> | |
| 2980 | by (simp add: bit_Suc add_One flip: bit_not_int_iff) | |
| 2981 | ||
| 74495 | 2982 | lemma and_not_numerals: | 
| 74163 | 2983 | \<open>1 AND NOT 1 = (0 :: int)\<close> | 
| 2984 | \<open>1 AND NOT (numeral (Num.Bit0 n)) = (1 :: int)\<close> | |
| 2985 | \<open>1 AND NOT (numeral (Num.Bit1 n)) = (0 :: int)\<close> | |
| 2986 | \<open>numeral (Num.Bit0 m) AND NOT (1 :: int) = numeral (Num.Bit0 m)\<close> | |
| 2987 | \<open>numeral (Num.Bit0 m) AND NOT (numeral (Num.Bit0 n)) = (2 :: int) * (numeral m AND NOT (numeral n))\<close> | |
| 2988 | \<open>numeral (Num.Bit0 m) AND NOT (numeral (Num.Bit1 n)) = (2 :: int) * (numeral m AND NOT (numeral n))\<close> | |
| 2989 | \<open>numeral (Num.Bit1 m) AND NOT (1 :: int) = numeral (Num.Bit0 m)\<close> | |
| 2990 | \<open>numeral (Num.Bit1 m) AND NOT (numeral (Num.Bit0 n)) = 1 + (2 :: int) * (numeral m AND NOT (numeral n))\<close> | |
| 2991 | \<open>numeral (Num.Bit1 m) AND NOT (numeral (Num.Bit1 n)) = (2 :: int) * (numeral m AND NOT (numeral n))\<close> | |
| 74495 | 2992 | by (simp_all add: bit_eq_iff) (auto simp add: bit_simps bit_Suc bit_numeral_rec BitM_inc_eq sub_inc_One_eq split: nat.split) | 
| 74163 | 2993 | |
| 2994 | fun and_not_num :: \<open>num \<Rightarrow> num \<Rightarrow> num option\<close> \<^marker>\<open>contributor \<open>Andreas Lochbihler\<close>\<close> | |
| 2995 | where | |
| 2996 | \<open>and_not_num num.One num.One = None\<close> | |
| 2997 | | \<open>and_not_num num.One (num.Bit0 n) = Some num.One\<close> | |
| 2998 | | \<open>and_not_num num.One (num.Bit1 n) = None\<close> | |
| 2999 | | \<open>and_not_num (num.Bit0 m) num.One = Some (num.Bit0 m)\<close> | |
| 3000 | | \<open>and_not_num (num.Bit0 m) (num.Bit0 n) = map_option num.Bit0 (and_not_num m n)\<close> | |
| 3001 | | \<open>and_not_num (num.Bit0 m) (num.Bit1 n) = map_option num.Bit0 (and_not_num m n)\<close> | |
| 3002 | | \<open>and_not_num (num.Bit1 m) num.One = Some (num.Bit0 m)\<close> | |
| 3003 | | \<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> | |
| 3004 | | \<open>and_not_num (num.Bit1 m) (num.Bit1 n) = map_option num.Bit0 (and_not_num m n)\<close> | |
| 3005 | ||
| 3006 | lemma int_numeral_and_not_num: | |
| 3007 | \<open>numeral m AND NOT (numeral n) = (case and_not_num m n of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close> | |
| 74495 | 3008 | by (induction m n rule: and_not_num.induct) (simp_all del: not_numeral_eq not_one_eq add: and_not_numerals split: option.splits) | 
| 74163 | 3009 | |
| 3010 | lemma int_numeral_not_and_num: | |
| 3011 | \<open>NOT (numeral m) AND numeral n = (case and_not_num n m of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close> | |
| 3012 | using int_numeral_and_not_num [of n m] by (simp add: ac_simps) | |
| 3013 | ||
| 3014 | lemma and_not_num_eq_None_iff: | |
| 3015 | \<open>and_not_num m n = None \<longleftrightarrow> numeral m AND NOT (numeral n) = (0 :: int)\<close> | |
| 74495 | 3016 | by (simp del: not_numeral_eq add: int_numeral_and_not_num split: option.split) | 
| 74163 | 3017 | |
| 3018 | lemma and_not_num_eq_Some_iff: | |
| 3019 | \<open>and_not_num m n = Some q \<longleftrightarrow> numeral m AND NOT (numeral n) = (numeral q :: int)\<close> | |
| 74495 | 3020 | by (simp del: not_numeral_eq add: int_numeral_and_not_num split: option.split) | 
| 3021 | ||
| 3022 | lemma and_minus_numerals [simp]: | |
| 3023 | \<open>1 AND - (numeral (num.Bit0 n)) = (0::int)\<close> | |
| 3024 | \<open>1 AND - (numeral (num.Bit1 n)) = (1::int)\<close> | |
| 3025 | \<open>numeral m AND - (numeral (num.Bit0 n)) = (case and_not_num m (Num.BitM n) of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close> | |
| 3026 | \<open>numeral m AND - (numeral (num.Bit1 n)) = (case and_not_num m (Num.Bit0 n) of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close> | |
| 3027 | \<open>- (numeral (num.Bit0 n)) AND 1 = (0::int)\<close> | |
| 3028 | \<open>- (numeral (num.Bit1 n)) AND 1 = (1::int)\<close> | |
| 3029 | \<open>- (numeral (num.Bit0 n)) AND numeral m = (case and_not_num m (Num.BitM n) of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close> | |
| 3030 | \<open>- (numeral (num.Bit1 n)) AND numeral m = (case and_not_num m (Num.Bit0 n) of None \<Rightarrow> 0 :: int | Some n' \<Rightarrow> numeral n')\<close> | |
| 3031 | by (simp_all del: not_numeral_eq add: ac_simps | |
| 3032 | and_not_numerals one_and_eq not_numeral_BitM_eq not_numeral_Bit0_eq and_not_num_eq_None_iff and_not_num_eq_Some_iff split: option.split) | |
| 3033 | ||
| 3034 | lemma and_minus_minus_numerals [simp]: | |
| 3035 | \<open>- (numeral m :: int) AND - (numeral n :: int) = NOT ((numeral m - 1) OR (numeral n - 1))\<close> | |
| 3036 | by (simp add: minus_numeral_eq_not_sub_one) | |
| 3037 | ||
| 3038 | lemma or_not_numerals: | |
| 74163 | 3039 | \<open>1 OR NOT 1 = NOT (0 :: int)\<close> | 
| 3040 | \<open>1 OR NOT (numeral (Num.Bit0 n)) = NOT (numeral (Num.Bit0 n) :: int)\<close> | |
| 3041 | \<open>1 OR NOT (numeral (Num.Bit1 n)) = NOT (numeral (Num.Bit0 n) :: int)\<close> | |
| 3042 | \<open>numeral (Num.Bit0 m) OR NOT (1 :: int) = NOT (1 :: int)\<close> | |
| 3043 | \<open>numeral (Num.Bit0 m) OR NOT (numeral (Num.Bit0 n)) = 1 + (2 :: int) * (numeral m OR NOT (numeral n))\<close> | |
| 3044 | \<open>numeral (Num.Bit0 m) OR NOT (numeral (Num.Bit1 n)) = (2 :: int) * (numeral m OR NOT (numeral n))\<close> | |
| 3045 | \<open>numeral (Num.Bit1 m) OR NOT (1 :: int) = NOT (0 :: int)\<close> | |
| 3046 | \<open>numeral (Num.Bit1 m) OR NOT (numeral (Num.Bit0 n)) = 1 + (2 :: int) * (numeral m OR NOT (numeral n))\<close> | |
| 3047 | \<open>numeral (Num.Bit1 m) OR NOT (numeral (Num.Bit1 n)) = 1 + (2 :: int) * (numeral m OR NOT (numeral n))\<close> | |
| 74495 | 3048 | by (simp_all add: bit_eq_iff) | 
| 3049 | (auto simp add: bit_simps bit_Suc bit_numeral_rec sub_inc_One_eq split: nat.split) | |
| 74163 | 3050 | |
| 3051 | fun or_not_num_neg :: \<open>num \<Rightarrow> num \<Rightarrow> num\<close> \<^marker>\<open>contributor \<open>Andreas Lochbihler\<close>\<close> | |
| 3052 | where | |
| 3053 | \<open>or_not_num_neg num.One num.One = num.One\<close> | |
| 3054 | | \<open>or_not_num_neg num.One (num.Bit0 m) = num.Bit1 m\<close> | |
| 3055 | | \<open>or_not_num_neg num.One (num.Bit1 m) = num.Bit1 m\<close> | |
| 3056 | | \<open>or_not_num_neg (num.Bit0 n) num.One = num.Bit0 num.One\<close> | |
| 3057 | | \<open>or_not_num_neg (num.Bit0 n) (num.Bit0 m) = Num.BitM (or_not_num_neg n m)\<close> | |
| 3058 | | \<open>or_not_num_neg (num.Bit0 n) (num.Bit1 m) = num.Bit0 (or_not_num_neg n m)\<close> | |
| 3059 | | \<open>or_not_num_neg (num.Bit1 n) num.One = num.One\<close> | |
| 3060 | | \<open>or_not_num_neg (num.Bit1 n) (num.Bit0 m) = Num.BitM (or_not_num_neg n m)\<close> | |
| 3061 | | \<open>or_not_num_neg (num.Bit1 n) (num.Bit1 m) = Num.BitM (or_not_num_neg n m)\<close> | |
| 3062 | ||
| 3063 | lemma int_numeral_or_not_num_neg: | |
| 3064 | \<open>numeral m OR NOT (numeral n :: int) = - numeral (or_not_num_neg m n)\<close> | |
| 74495 | 3065 | by (induction m n rule: or_not_num_neg.induct) (simp_all del: not_numeral_eq not_one_eq add: or_not_numerals, simp_all) | 
| 74163 | 3066 | |
| 3067 | lemma int_numeral_not_or_num_neg: | |
| 3068 | \<open>NOT (numeral m) OR (numeral n :: int) = - numeral (or_not_num_neg n m)\<close> | |
| 3069 | using int_numeral_or_not_num_neg [of n m] by (simp add: ac_simps) | |
| 3070 | ||
| 3071 | lemma numeral_or_not_num_eq: | |
| 3072 | \<open>numeral (or_not_num_neg m n) = - (numeral m OR NOT (numeral n :: int))\<close> | |
| 3073 | using int_numeral_or_not_num_neg [of m n] by simp | |
| 3074 | ||
| 74495 | 3075 | lemma or_minus_numerals [simp]: | 
| 3076 | \<open>1 OR - (numeral (num.Bit0 n)) = - (numeral (or_not_num_neg num.One (Num.BitM n)) :: int)\<close> | |
| 3077 | \<open>1 OR - (numeral (num.Bit1 n)) = - (numeral (num.Bit1 n) :: int)\<close> | |
| 3078 | \<open>numeral m OR - (numeral (num.Bit0 n)) = - (numeral (or_not_num_neg m (Num.BitM n)) :: int)\<close> | |
| 3079 | \<open>numeral m OR - (numeral (num.Bit1 n)) = - (numeral (or_not_num_neg m (Num.Bit0 n)) :: int)\<close> | |
| 3080 | \<open>- (numeral (num.Bit0 n)) OR 1 = - (numeral (or_not_num_neg num.One (Num.BitM n)) :: int)\<close> | |
| 3081 | \<open>- (numeral (num.Bit1 n)) OR 1 = - (numeral (num.Bit1 n) :: int)\<close> | |
| 3082 | \<open>- (numeral (num.Bit0 n)) OR numeral m = - (numeral (or_not_num_neg m (Num.BitM n)) :: int)\<close> | |
| 3083 | \<open>- (numeral (num.Bit1 n)) OR numeral m = - (numeral (or_not_num_neg m (Num.Bit0 n)) :: int)\<close> | |
| 3084 | by (simp_all only: or.commute [of _ 1] or.commute [of _ \<open>numeral m\<close>] | |
| 3085 | minus_numeral_eq_not_sub_one or_not_numerals | |
| 3086 | numeral_or_not_num_eq arith_simps minus_minus numeral_One) | |
| 3087 | ||
| 3088 | lemma or_minus_minus_numerals [simp]: | |
| 3089 | \<open>- (numeral m :: int) OR - (numeral n :: int) = NOT ((numeral m - 1) AND (numeral n - 1))\<close> | |
| 3090 | by (simp add: minus_numeral_eq_not_sub_one) | |
| 3091 | ||
| 74163 | 3092 | lemma xor_minus_numerals [simp]: | 
| 3093 | \<open>- numeral n XOR k = NOT (neg_numeral_class.sub n num.One XOR k)\<close> | |
| 3094 | \<open>k XOR - numeral n = NOT (k XOR (neg_numeral_class.sub n num.One))\<close> for k :: int | |
| 3095 | by (simp_all add: minus_numeral_eq_not_sub_one) | |
| 3096 | ||
| 74592 | 3097 | definition take_bit_num :: \<open>nat \<Rightarrow> num \<Rightarrow> num option\<close> | 
| 3098 | where \<open>take_bit_num n m = | |
| 3099 | (if take_bit n (numeral m ::nat) = 0 then None else Some (num_of_nat (take_bit n (numeral m ::nat))))\<close> | |
| 3100 | ||
| 74618 | 3101 | lemma take_bit_num_simps: | 
| 74592 | 3102 | \<open>take_bit_num 0 m = None\<close> | 
| 3103 | \<open>take_bit_num (Suc n) Num.One = | |
| 3104 | Some Num.One\<close> | |
| 3105 | \<open>take_bit_num (Suc n) (Num.Bit0 m) = | |
| 3106 | (case take_bit_num n m of None \<Rightarrow> None | Some q \<Rightarrow> Some (Num.Bit0 q))\<close> | |
| 3107 | \<open>take_bit_num (Suc n) (Num.Bit1 m) = | |
| 3108 | Some (case take_bit_num n m of None \<Rightarrow> Num.One | Some q \<Rightarrow> Num.Bit1 q)\<close> | |
| 74618 | 3109 | \<open>take_bit_num (numeral r) Num.One = | 
| 74592 | 3110 | Some Num.One\<close> | 
| 74618 | 3111 | \<open>take_bit_num (numeral r) (Num.Bit0 m) = | 
| 3112 | (case take_bit_num (pred_numeral r) m of None \<Rightarrow> None | Some q \<Rightarrow> Some (Num.Bit0 q))\<close> | |
| 3113 | \<open>take_bit_num (numeral r) (Num.Bit1 m) = | |
| 3114 | Some (case take_bit_num (pred_numeral r) m of None \<Rightarrow> Num.One | Some q \<Rightarrow> Num.Bit1 q)\<close> | |
| 74592 | 3115 | by (auto simp add: take_bit_num_def ac_simps mult_2 num_of_nat_double | 
| 74618 | 3116 | take_bit_Suc_bit0 take_bit_Suc_bit1 take_bit_numeral_bit0 take_bit_numeral_bit1) | 
| 3117 | ||
| 3118 | lemma take_bit_num_code [code]: | |
| 3119 | \<comment> \<open>Ocaml-style pattern matching is more robust wrt. different representations of \<^typ>\<open>nat\<close>\<close> | |
| 3120 | \<open>take_bit_num n m = (case (n, m) | |
| 3121 | of (0, _) \<Rightarrow> None | |
| 3122 | | (Suc n, Num.One) \<Rightarrow> Some Num.One | |
| 3123 | | (Suc n, Num.Bit0 m) \<Rightarrow> (case take_bit_num n m of None \<Rightarrow> None | Some q \<Rightarrow> Some (Num.Bit0 q)) | |
| 3124 | | (Suc n, Num.Bit1 m) \<Rightarrow> Some (case take_bit_num n m of None \<Rightarrow> Num.One | Some q \<Rightarrow> Num.Bit1 q))\<close> | |
| 3125 | by (cases n; cases m) (simp_all add: take_bit_num_simps) | |
| 74592 | 3126 | |
| 3127 | context semiring_bit_operations | |
| 3128 | begin | |
| 3129 | ||
| 3130 | lemma take_bit_num_eq_None_imp: | |
| 3131 | \<open>take_bit m (numeral n) = 0\<close> if \<open>take_bit_num m n = None\<close> | |
| 3132 | proof - | |
| 3133 | from that have \<open>take_bit m (numeral n :: nat) = 0\<close> | |
| 3134 | by (simp add: take_bit_num_def split: if_splits) | |
| 3135 | then have \<open>of_nat (take_bit m (numeral n)) = of_nat 0\<close> | |
| 3136 | by simp | |
| 3137 | then show ?thesis | |
| 3138 | by (simp add: of_nat_take_bit) | |
| 3139 | qed | |
| 3140 | ||
| 3141 | lemma take_bit_num_eq_Some_imp: | |
| 3142 | \<open>take_bit m (numeral n) = numeral q\<close> if \<open>take_bit_num m n = Some q\<close> | |
| 3143 | proof - | |
| 3144 | from that have \<open>take_bit m (numeral n :: nat) = numeral q\<close> | |
| 3145 | by (auto simp add: take_bit_num_def Num.numeral_num_of_nat_unfold split: if_splits) | |
| 3146 | then have \<open>of_nat (take_bit m (numeral n)) = of_nat (numeral q)\<close> | |
| 3147 | by simp | |
| 3148 | then show ?thesis | |
| 3149 | by (simp add: of_nat_take_bit) | |
| 3150 | qed | |
| 3151 | ||
| 3152 | lemma take_bit_numeral_numeral: | |
| 3153 | \<open>take_bit (numeral m) (numeral n) = | |
| 3154 | (case take_bit_num (numeral m) n of None \<Rightarrow> 0 | Some q \<Rightarrow> numeral q)\<close> | |
| 3155 | by (auto split: option.split dest: take_bit_num_eq_None_imp take_bit_num_eq_Some_imp) | |
| 3156 | ||
| 3157 | end | |
| 3158 | ||
| 3159 | lemma take_bit_numeral_minus_numeral_int: | |
| 3160 | \<open>take_bit (numeral m) (- numeral n :: int) = | |
| 3161 | (case take_bit_num (numeral m) n of None \<Rightarrow> 0 | Some q \<Rightarrow> take_bit (numeral m) (2 ^ numeral m - numeral q))\<close> (is \<open>?lhs = ?rhs\<close>) | |
| 3162 | proof (cases \<open>take_bit_num (numeral m) n\<close>) | |
| 3163 | case None | |
| 3164 | then show ?thesis | |
| 3165 | by (auto dest: take_bit_num_eq_None_imp [where ?'a = int] simp add: take_bit_eq_0_iff) | |
| 3166 | next | |
| 3167 | case (Some q) | |
| 3168 | then have q: \<open>take_bit (numeral m) (numeral n :: int) = numeral q\<close> | |
| 3169 | by (auto dest: take_bit_num_eq_Some_imp) | |
| 3170 | let ?T = \<open>take_bit (numeral m) :: int \<Rightarrow> int\<close> | |
| 3171 | have *: \<open>?T (2 ^ numeral m) = ?T (?T 0)\<close> | |
| 3172 | by (simp add: take_bit_eq_0_iff) | |
| 3173 | have \<open>?lhs = ?T (0 - numeral n)\<close> | |
| 3174 | by simp | |
| 3175 | also have \<open>\<dots> = ?T (?T (?T 0) - ?T (?T (numeral n)))\<close> | |
| 3176 | by (simp only: take_bit_diff) | |
| 3177 | also have \<open>\<dots> = ?T (2 ^ numeral m - ?T (numeral n))\<close> | |
| 3178 | by (simp only: take_bit_diff flip: *) | |
| 3179 | also have \<open>\<dots> = ?rhs\<close> | |
| 3180 | by (simp add: q Some) | |
| 3181 | finally show ?thesis . | |
| 3182 | qed | |
| 3183 | ||
| 74618 | 3184 | declare take_bit_num_simps [simp] | 
| 3185 | take_bit_numeral_numeral [simp] | |
| 74592 | 3186 | take_bit_numeral_minus_numeral_int [simp] | 
| 3187 | ||
| 74163 | 3188 | |
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 3189 | subsection \<open>More properties\<close> | 
| 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 3190 | |
| 72830 | 3191 | lemma take_bit_eq_mask_iff: | 
| 3192 | \<open>take_bit n k = mask n \<longleftrightarrow> take_bit n (k + 1) = 0\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) | |
| 3193 | for k :: int | |
| 3194 | proof | |
| 3195 | assume ?P | |
| 3196 | then have \<open>take_bit n (take_bit n k + take_bit n 1) = 0\<close> | |
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 3197 | by (simp add: mask_eq_exp_minus_1 take_bit_eq_0_iff) | 
| 72830 | 3198 | then show ?Q | 
| 3199 | by (simp only: take_bit_add) | |
| 3200 | next | |
| 3201 | assume ?Q | |
| 3202 | then have \<open>take_bit n (k + 1) - 1 = - 1\<close> | |
| 3203 | by simp | |
| 3204 | then have \<open>take_bit n (take_bit n (k + 1) - 1) = take_bit n (- 1)\<close> | |
| 3205 | by simp | |
| 3206 | moreover have \<open>take_bit n (take_bit n (k + 1) - 1) = take_bit n k\<close> | |
| 3207 | by (simp add: take_bit_eq_mod mod_simps) | |
| 3208 | ultimately show ?P | |
| 74592 | 3209 | by simp | 
| 72830 | 3210 | qed | 
| 3211 | ||
| 3212 | lemma take_bit_eq_mask_iff_exp_dvd: | |
| 3213 | \<open>take_bit n k = mask n \<longleftrightarrow> 2 ^ n dvd k + 1\<close> | |
| 3214 | for k :: int | |
| 3215 | by (simp add: take_bit_eq_mask_iff flip: take_bit_eq_0_iff) | |
| 3216 | ||
| 71442 | 3217 | |
| 72028 | 3218 | subsection \<open>Bit concatenation\<close> | 
| 3219 | ||
| 3220 | definition concat_bit :: \<open>nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> int\<close> | |
| 72227 | 3221 | where \<open>concat_bit n k l = take_bit n k OR push_bit n l\<close> | 
| 72028 | 3222 | |
| 72611 
c7bc3e70a8c7
official collection for bit projection simplifications
 haftmann parents: 
72512diff
changeset | 3223 | lemma bit_concat_bit_iff [bit_simps]: | 
| 72028 | 3224 | \<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> | 
| 72227 | 3225 | by (simp add: concat_bit_def bit_or_iff bit_and_iff bit_take_bit_iff bit_push_bit_iff ac_simps) | 
| 72028 | 3226 | |
| 3227 | lemma concat_bit_eq: | |
| 3228 | \<open>concat_bit n k l = take_bit n k + push_bit n l\<close> | |
| 3229 | by (simp add: concat_bit_def take_bit_eq_mask | |
| 3230 | bit_and_iff bit_mask_iff bit_push_bit_iff disjunctive_add) | |
| 3231 | ||
| 3232 | lemma concat_bit_0 [simp]: | |
| 3233 | \<open>concat_bit 0 k l = l\<close> | |
| 3234 | by (simp add: concat_bit_def) | |
| 3235 | ||
| 3236 | lemma concat_bit_Suc: | |
| 3237 | \<open>concat_bit (Suc n) k l = k mod 2 + 2 * concat_bit n (k div 2) l\<close> | |
| 3238 | by (simp add: concat_bit_eq take_bit_Suc push_bit_double) | |
| 3239 | ||
| 3240 | lemma concat_bit_of_zero_1 [simp]: | |
| 3241 | \<open>concat_bit n 0 l = push_bit n l\<close> | |
| 3242 | by (simp add: concat_bit_def) | |
| 3243 | ||
| 3244 | lemma concat_bit_of_zero_2 [simp]: | |
| 3245 | \<open>concat_bit n k 0 = take_bit n k\<close> | |
| 3246 | by (simp add: concat_bit_def take_bit_eq_mask) | |
| 3247 | ||
| 3248 | lemma concat_bit_nonnegative_iff [simp]: | |
| 3249 | \<open>concat_bit n k l \<ge> 0 \<longleftrightarrow> l \<ge> 0\<close> | |
| 3250 | by (simp add: concat_bit_def) | |
| 3251 | ||
| 3252 | lemma concat_bit_negative_iff [simp]: | |
| 3253 | \<open>concat_bit n k l < 0 \<longleftrightarrow> l < 0\<close> | |
| 3254 | by (simp add: concat_bit_def) | |
| 3255 | ||
| 3256 | lemma concat_bit_assoc: | |
| 3257 | \<open>concat_bit n k (concat_bit m l r) = concat_bit (m + n) (concat_bit n k l) r\<close> | |
| 3258 | by (rule bit_eqI) (auto simp add: bit_concat_bit_iff ac_simps) | |
| 3259 | ||
| 3260 | lemma concat_bit_assoc_sym: | |
| 3261 | \<open>concat_bit m (concat_bit n k l) r = concat_bit (min m n) k (concat_bit (m - n) l r)\<close> | |
| 3262 | by (rule bit_eqI) (auto simp add: bit_concat_bit_iff ac_simps min_def) | |
| 3263 | ||
| 72227 | 3264 | lemma concat_bit_eq_iff: | 
| 3265 | \<open>concat_bit n k l = concat_bit n r s | |
| 3266 | \<longleftrightarrow> take_bit n k = take_bit n r \<and> l = s\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) | |
| 3267 | proof | |
| 3268 | assume ?Q | |
| 3269 | then show ?P | |
| 3270 | by (simp add: concat_bit_def) | |
| 3271 | next | |
| 3272 | assume ?P | |
| 3273 | then have *: \<open>bit (concat_bit n k l) m = bit (concat_bit n r s) m\<close> for m | |
| 3274 | by (simp add: bit_eq_iff) | |
| 3275 | have \<open>take_bit n k = take_bit n r\<close> | |
| 3276 | proof (rule bit_eqI) | |
| 3277 | fix m | |
| 3278 | from * [of m] | |
| 3279 | show \<open>bit (take_bit n k) m \<longleftrightarrow> bit (take_bit n r) m\<close> | |
| 3280 | by (auto simp add: bit_take_bit_iff bit_concat_bit_iff) | |
| 3281 | qed | |
| 3282 | moreover have \<open>push_bit n l = push_bit n s\<close> | |
| 3283 | proof (rule bit_eqI) | |
| 3284 | fix m | |
| 3285 | from * [of m] | |
| 3286 | show \<open>bit (push_bit n l) m \<longleftrightarrow> bit (push_bit n s) m\<close> | |
| 3287 | by (auto simp add: bit_push_bit_iff bit_concat_bit_iff) | |
| 3288 | qed | |
| 3289 | then have \<open>l = s\<close> | |
| 3290 | by (simp add: push_bit_eq_mult) | |
| 3291 | ultimately show ?Q | |
| 3292 | by (simp add: concat_bit_def) | |
| 3293 | qed | |
| 3294 | ||
| 3295 | lemma take_bit_concat_bit_eq: | |
| 3296 | \<open>take_bit m (concat_bit n k l) = concat_bit (min m n) k (take_bit (m - n) l)\<close> | |
| 3297 | by (rule bit_eqI) | |
| 3298 | (auto simp add: bit_take_bit_iff bit_concat_bit_iff min_def) | |
| 3299 | ||
| 72488 | 3300 | lemma concat_bit_take_bit_eq: | 
| 3301 | \<open>concat_bit n (take_bit n b) = concat_bit n b\<close> | |
| 3302 | by (simp add: concat_bit_def [abs_def]) | |
| 3303 | ||
| 72028 | 3304 | |
| 72241 | 3305 | subsection \<open>Taking bits with sign propagation\<close> | 
| 72010 | 3306 | |
| 72241 | 3307 | context ring_bit_operations | 
| 3308 | begin | |
| 72010 | 3309 | |
| 72241 | 3310 | definition signed_take_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close> | 
| 3311 | where \<open>signed_take_bit n a = take_bit n a OR (of_bool (bit a n) * NOT (mask n))\<close> | |
| 72227 | 3312 | |
| 72241 | 3313 | lemma signed_take_bit_eq_if_positive: | 
| 3314 | \<open>signed_take_bit n a = take_bit n a\<close> if \<open>\<not> bit a n\<close> | |
| 72010 | 3315 | using that by (simp add: signed_take_bit_def) | 
| 3316 | ||
| 72241 | 3317 | lemma signed_take_bit_eq_if_negative: | 
| 3318 | \<open>signed_take_bit n a = take_bit n a OR NOT (mask n)\<close> if \<open>bit a n\<close> | |
| 3319 | using that by (simp add: signed_take_bit_def) | |
| 3320 | ||
| 3321 | lemma even_signed_take_bit_iff: | |
| 3322 | \<open>even (signed_take_bit m a) \<longleftrightarrow> even a\<close> | |
| 3323 | by (auto simp add: signed_take_bit_def even_or_iff even_mask_iff bit_double_iff) | |
| 3324 | ||
| 72611 
c7bc3e70a8c7
official collection for bit projection simplifications
 haftmann parents: 
72512diff
changeset | 3325 | lemma bit_signed_take_bit_iff [bit_simps]: | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 3326 |   \<open>bit (signed_take_bit m a) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> bit a (min m n)\<close>
 | 
| 72241 | 3327 | by (simp add: signed_take_bit_def bit_take_bit_iff bit_or_iff bit_not_iff bit_mask_iff min_def not_le) | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 3328 | (blast dest: bit_imp_possible_bit) | 
| 72010 | 3329 | |
| 3330 | lemma signed_take_bit_0 [simp]: | |
| 72241 | 3331 | \<open>signed_take_bit 0 a = - (a mod 2)\<close> | 
| 72010 | 3332 | by (simp add: signed_take_bit_def odd_iff_mod_2_eq_one) | 
| 3333 | ||
| 3334 | lemma signed_take_bit_Suc: | |
| 72241 | 3335 | \<open>signed_take_bit (Suc n) a = a mod 2 + 2 * signed_take_bit n (a div 2)\<close> | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 3336 | apply (simp add: bit_eq_iff bit_sum_mult_2_cases bit_simps bit_Suc[symmetric]) | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 3337 | apply (simp add: possible_bit_less_imp flip: min_Suc_Suc) | 
| 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 3338 | done | 
| 72010 | 3339 | |
| 72187 | 3340 | lemma signed_take_bit_of_0 [simp]: | 
| 3341 | \<open>signed_take_bit n 0 = 0\<close> | |
| 3342 | by (simp add: signed_take_bit_def) | |
| 3343 | ||
| 3344 | lemma signed_take_bit_of_minus_1 [simp]: | |
| 3345 | \<open>signed_take_bit n (- 1) = - 1\<close> | |
| 74592 | 3346 | by (simp add: signed_take_bit_def mask_eq_exp_minus_1 possible_bit_def) | 
| 72187 | 3347 | |
| 72241 | 3348 | lemma signed_take_bit_Suc_1 [simp]: | 
| 3349 | \<open>signed_take_bit (Suc n) 1 = 1\<close> | |
| 3350 | by (simp add: signed_take_bit_Suc) | |
| 3351 | ||
| 74497 | 3352 | lemma signed_take_bit_numeral_of_1 [simp]: | 
| 3353 | \<open>signed_take_bit (numeral k) 1 = 1\<close> | |
| 3354 | by (simp add: bit_1_iff signed_take_bit_eq_if_positive) | |
| 3355 | ||
| 72241 | 3356 | lemma signed_take_bit_rec: | 
| 3357 | \<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> | |
| 3358 | by (cases n) (simp_all add: signed_take_bit_Suc) | |
| 72187 | 3359 | |
| 3360 | lemma signed_take_bit_eq_iff_take_bit_eq: | |
| 72241 | 3361 | \<open>signed_take_bit n a = signed_take_bit n b \<longleftrightarrow> take_bit (Suc n) a = take_bit (Suc n) b\<close> | 
| 3362 | proof - | |
| 3363 | 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> | |
| 3364 | by (simp add: fun_eq_iff bit_signed_take_bit_iff bit_take_bit_iff not_le less_Suc_eq_le min_def) | |
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 3365 | (use bit_imp_possible_bit in fastforce) | 
| 72187 | 3366 | then show ?thesis | 
| 74309 
42523fbf643b
explicit predicate for confined bit range avoids cyclic rewriting in presence of extensionality rule for bit values (contributed by Thomas Sewell)
 haftmann parents: 
74163diff
changeset | 3367 | by (auto simp add: fun_eq_iff intro: bit_eqI) | 
| 72187 | 3368 | qed | 
| 3369 | ||
| 72241 | 3370 | lemma signed_take_bit_signed_take_bit [simp]: | 
| 3371 | \<open>signed_take_bit m (signed_take_bit n a) = signed_take_bit (min m n) a\<close> | |
| 74495 | 3372 | by (auto simp add: bit_eq_iff bit_simps ac_simps) | 
| 72241 | 3373 | |
| 3374 | lemma signed_take_bit_take_bit: | |
| 3375 | \<open>signed_take_bit m (take_bit n a) = (if n \<le> m then take_bit n else signed_take_bit m) a\<close> | |
| 3376 | by (rule bit_eqI) (auto simp add: bit_signed_take_bit_iff min_def bit_take_bit_iff) | |
| 3377 | ||
| 72187 | 3378 | lemma take_bit_signed_take_bit: | 
| 72241 | 3379 | \<open>take_bit m (signed_take_bit n a) = take_bit m a\<close> if \<open>m \<le> Suc n\<close> | 
| 72187 | 3380 | using that by (rule le_SucE; intro bit_eqI) | 
| 3381 | (auto simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def less_Suc_eq) | |
| 3382 | ||
| 72241 | 3383 | end | 
| 3384 | ||
| 3385 | text \<open>Modulus centered around 0\<close> | |
| 3386 | ||
| 3387 | lemma signed_take_bit_eq_concat_bit: | |
| 3388 | \<open>signed_take_bit n k = concat_bit n k (- of_bool (bit k n))\<close> | |
| 74592 | 3389 | by (simp add: concat_bit_def signed_take_bit_def) | 
| 72241 | 3390 | |
| 72187 | 3391 | lemma signed_take_bit_add: | 
| 3392 | \<open>signed_take_bit n (signed_take_bit n k + signed_take_bit n l) = signed_take_bit n (k + l)\<close> | |
| 72241 | 3393 | for k l :: int | 
| 72187 | 3394 | proof - | 
| 3395 | have \<open>take_bit (Suc n) | |
| 3396 | (take_bit (Suc n) (signed_take_bit n k) + | |
| 3397 | take_bit (Suc n) (signed_take_bit n l)) = | |
| 3398 | take_bit (Suc n) (k + l)\<close> | |
| 3399 | by (simp add: take_bit_signed_take_bit take_bit_add) | |
| 3400 | then show ?thesis | |
| 3401 | by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_add) | |
| 3402 | qed | |
| 3403 | ||
| 3404 | lemma signed_take_bit_diff: | |
| 3405 | \<open>signed_take_bit n (signed_take_bit n k - signed_take_bit n l) = signed_take_bit n (k - l)\<close> | |
| 72241 | 3406 | for k l :: int | 
| 72187 | 3407 | proof - | 
| 3408 | have \<open>take_bit (Suc n) | |
| 3409 | (take_bit (Suc n) (signed_take_bit n k) - | |
| 3410 | take_bit (Suc n) (signed_take_bit n l)) = | |
| 3411 | take_bit (Suc n) (k - l)\<close> | |
| 3412 | by (simp add: take_bit_signed_take_bit take_bit_diff) | |
| 3413 | then show ?thesis | |
| 3414 | by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_diff) | |
| 3415 | qed | |
| 3416 | ||
| 3417 | lemma signed_take_bit_minus: | |
| 3418 | \<open>signed_take_bit n (- signed_take_bit n k) = signed_take_bit n (- k)\<close> | |
| 72241 | 3419 | for k :: int | 
| 72187 | 3420 | proof - | 
| 3421 | have \<open>take_bit (Suc n) | |
| 3422 | (- take_bit (Suc n) (signed_take_bit n k)) = | |
| 3423 | take_bit (Suc n) (- k)\<close> | |
| 3424 | by (simp add: take_bit_signed_take_bit take_bit_minus) | |
| 3425 | then show ?thesis | |
| 3426 | by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_minus) | |
| 3427 | qed | |
| 3428 | ||
| 3429 | lemma signed_take_bit_mult: | |
| 3430 | \<open>signed_take_bit n (signed_take_bit n k * signed_take_bit n l) = signed_take_bit n (k * l)\<close> | |
| 72241 | 3431 | for k l :: int | 
| 72187 | 3432 | proof - | 
| 3433 | have \<open>take_bit (Suc n) | |
| 3434 | (take_bit (Suc n) (signed_take_bit n k) * | |
| 3435 | take_bit (Suc n) (signed_take_bit n l)) = | |
| 3436 | take_bit (Suc n) (k * l)\<close> | |
| 3437 | by (simp add: take_bit_signed_take_bit take_bit_mult) | |
| 3438 | then show ?thesis | |
| 3439 | by (simp only: signed_take_bit_eq_iff_take_bit_eq take_bit_mult) | |
| 3440 | qed | |
| 3441 | ||
| 72010 | 3442 | lemma signed_take_bit_eq_take_bit_minus: | 
| 3443 | \<open>signed_take_bit n k = take_bit (Suc n) k - 2 ^ Suc n * of_bool (bit k n)\<close> | |
| 72241 | 3444 | for k :: int | 
| 72010 | 3445 | proof (cases \<open>bit k n\<close>) | 
| 3446 | case True | |
| 3447 | have \<open>signed_take_bit n k = take_bit (Suc n) k OR NOT (mask (Suc n))\<close> | |
| 3448 | 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) | |
| 3449 | then have \<open>signed_take_bit n k = take_bit (Suc n) k + NOT (mask (Suc n))\<close> | |
| 3450 | by (simp add: disjunctive_add bit_take_bit_iff bit_not_iff bit_mask_iff) | |
| 3451 | with True show ?thesis | |
| 3452 | by (simp flip: minus_exp_eq_not_mask) | |
| 3453 | next | |
| 3454 | case False | |
| 72241 | 3455 | show ?thesis | 
| 3456 | by (rule bit_eqI) (simp add: False bit_signed_take_bit_iff bit_take_bit_iff min_def less_Suc_eq) | |
| 72010 | 3457 | qed | 
| 3458 | ||
| 3459 | lemma signed_take_bit_eq_take_bit_shift: | |
| 3460 | \<open>signed_take_bit n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n\<close> | |
| 72241 | 3461 | for k :: int | 
| 72010 | 3462 | proof - | 
| 3463 | have *: \<open>take_bit n k OR 2 ^ n = take_bit n k + 2 ^ n\<close> | |
| 3464 | by (simp add: disjunctive_add bit_exp_iff bit_take_bit_iff) | |
| 3465 | have \<open>take_bit n k - 2 ^ n = take_bit n k + NOT (mask n)\<close> | |
| 3466 | by (simp add: minus_exp_eq_not_mask) | |
| 3467 | also have \<open>\<dots> = take_bit n k OR NOT (mask n)\<close> | |
| 3468 | by (rule disjunctive_add) | |
| 3469 | (simp add: bit_exp_iff bit_take_bit_iff bit_not_iff bit_mask_iff) | |
| 3470 | finally have **: \<open>take_bit n k - 2 ^ n = take_bit n k OR NOT (mask n)\<close> . | |
| 3471 | 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> | |
| 3472 | by (simp only: take_bit_add) | |
| 3473 | also have \<open>take_bit (Suc n) k = 2 ^ n * of_bool (bit k n) + take_bit n k\<close> | |
| 3474 | by (simp add: take_bit_Suc_from_most) | |
| 3475 | 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> | |
| 3476 | by (simp add: ac_simps) | |
| 3477 | 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> | |
| 3478 | by (rule disjunctive_add) | |
| 3479 | (auto simp add: disjunctive_add bit_take_bit_iff bit_double_iff bit_exp_iff) | |
| 3480 | finally show ?thesis | |
| 72241 | 3481 | using * ** by (simp add: signed_take_bit_def concat_bit_Suc min_def ac_simps) | 
| 72010 | 3482 | qed | 
| 3483 | ||
| 3484 | lemma signed_take_bit_nonnegative_iff [simp]: | |
| 3485 | \<open>0 \<le> signed_take_bit n k \<longleftrightarrow> \<not> bit k n\<close> | |
| 72241 | 3486 | for k :: int | 
| 72028 | 3487 | by (simp add: signed_take_bit_def not_less concat_bit_def) | 
| 72010 | 3488 | |
| 3489 | lemma signed_take_bit_negative_iff [simp]: | |
| 3490 | \<open>signed_take_bit n k < 0 \<longleftrightarrow> bit k n\<close> | |
| 72241 | 3491 | for k :: int | 
| 72028 | 3492 | by (simp add: signed_take_bit_def not_less concat_bit_def) | 
| 72010 | 3493 | |
| 73868 | 3494 | lemma signed_take_bit_int_greater_eq_minus_exp [simp]: | 
| 3495 | \<open>- (2 ^ n) \<le> signed_take_bit n k\<close> | |
| 3496 | for k :: int | |
| 3497 | by (simp add: signed_take_bit_eq_take_bit_shift) | |
| 3498 | ||
| 3499 | lemma signed_take_bit_int_less_exp [simp]: | |
| 3500 | \<open>signed_take_bit n k < 2 ^ n\<close> | |
| 3501 | for k :: int | |
| 3502 | using take_bit_int_less_exp [of \<open>Suc n\<close>] | |
| 3503 | by (simp add: signed_take_bit_eq_take_bit_shift) | |
| 3504 | ||
| 72261 | 3505 | lemma signed_take_bit_int_eq_self_iff: | 
| 3506 | \<open>signed_take_bit n k = k \<longleftrightarrow> - (2 ^ n) \<le> k \<and> k < 2 ^ n\<close> | |
| 3507 | for k :: int | |
| 3508 | by (auto simp add: signed_take_bit_eq_take_bit_shift take_bit_int_eq_self_iff algebra_simps) | |
| 3509 | ||
| 72262 | 3510 | lemma signed_take_bit_int_eq_self: | 
| 3511 | \<open>signed_take_bit n k = k\<close> if \<open>- (2 ^ n) \<le> k\<close> \<open>k < 2 ^ n\<close> | |
| 3512 | for k :: int | |
| 3513 | using that by (simp add: signed_take_bit_int_eq_self_iff) | |
| 3514 | ||
| 72261 | 3515 | lemma signed_take_bit_int_less_eq_self_iff: | 
| 3516 | \<open>signed_take_bit n k \<le> k \<longleftrightarrow> - (2 ^ n) \<le> k\<close> | |
| 3517 | for k :: int | |
| 3518 | by (simp add: signed_take_bit_eq_take_bit_shift take_bit_int_less_eq_self_iff algebra_simps) | |
| 3519 | linarith | |
| 3520 | ||
| 3521 | lemma signed_take_bit_int_less_self_iff: | |
| 3522 | \<open>signed_take_bit n k < k \<longleftrightarrow> 2 ^ n \<le> k\<close> | |
| 3523 | for k :: int | |
| 3524 | by (simp add: signed_take_bit_eq_take_bit_shift take_bit_int_less_self_iff algebra_simps) | |
| 3525 | ||
| 3526 | lemma signed_take_bit_int_greater_self_iff: | |
| 3527 | \<open>k < signed_take_bit n k \<longleftrightarrow> k < - (2 ^ n)\<close> | |
| 3528 | for k :: int | |
| 3529 | by (simp add: signed_take_bit_eq_take_bit_shift take_bit_int_greater_self_iff algebra_simps) | |
| 3530 | linarith | |
| 3531 | ||
| 3532 | lemma signed_take_bit_int_greater_eq_self_iff: | |
| 3533 | \<open>k \<le> signed_take_bit n k \<longleftrightarrow> k < 2 ^ n\<close> | |
| 3534 | for k :: int | |
| 3535 | by (simp add: signed_take_bit_eq_take_bit_shift take_bit_int_greater_eq_self_iff algebra_simps) | |
| 3536 | ||
| 3537 | lemma signed_take_bit_int_greater_eq: | |
| 72010 | 3538 | \<open>k + 2 ^ Suc n \<le> signed_take_bit n k\<close> if \<open>k < - (2 ^ n)\<close> | 
| 72241 | 3539 | for k :: int | 
| 72262 | 3540 | using that take_bit_int_greater_eq [of \<open>k + 2 ^ n\<close> \<open>Suc n\<close>] | 
| 72010 | 3541 | by (simp add: signed_take_bit_eq_take_bit_shift) | 
| 3542 | ||
| 72261 | 3543 | lemma signed_take_bit_int_less_eq: | 
| 72010 | 3544 | \<open>signed_take_bit n k \<le> k - 2 ^ Suc n\<close> if \<open>k \<ge> 2 ^ n\<close> | 
| 72241 | 3545 | for k :: int | 
| 72262 | 3546 | using that take_bit_int_less_eq [of \<open>Suc n\<close> \<open>k + 2 ^ n\<close>] | 
| 72010 | 3547 | by (simp add: signed_take_bit_eq_take_bit_shift) | 
| 3548 | ||
| 3549 | lemma signed_take_bit_Suc_bit0 [simp]: | |
| 72241 | 3550 | \<open>signed_take_bit (Suc n) (numeral (Num.Bit0 k)) = signed_take_bit n (numeral k) * (2 :: int)\<close> | 
| 72010 | 3551 | by (simp add: signed_take_bit_Suc) | 
| 3552 | ||
| 3553 | lemma signed_take_bit_Suc_bit1 [simp]: | |
| 72241 | 3554 | \<open>signed_take_bit (Suc n) (numeral (Num.Bit1 k)) = signed_take_bit n (numeral k) * 2 + (1 :: int)\<close> | 
| 72010 | 3555 | by (simp add: signed_take_bit_Suc) | 
| 3556 | ||
| 3557 | lemma signed_take_bit_Suc_minus_bit0 [simp]: | |
| 72241 | 3558 | \<open>signed_take_bit (Suc n) (- numeral (Num.Bit0 k)) = signed_take_bit n (- numeral k) * (2 :: int)\<close> | 
| 72010 | 3559 | by (simp add: signed_take_bit_Suc) | 
| 3560 | ||
| 3561 | lemma signed_take_bit_Suc_minus_bit1 [simp]: | |
| 72241 | 3562 | \<open>signed_take_bit (Suc n) (- numeral (Num.Bit1 k)) = signed_take_bit n (- numeral k - 1) * 2 + (1 :: int)\<close> | 
| 72010 | 3563 | by (simp add: signed_take_bit_Suc) | 
| 3564 | ||
| 3565 | lemma signed_take_bit_numeral_bit0 [simp]: | |
| 72241 | 3566 | \<open>signed_take_bit (numeral l) (numeral (Num.Bit0 k)) = signed_take_bit (pred_numeral l) (numeral k) * (2 :: int)\<close> | 
| 72010 | 3567 | by (simp add: signed_take_bit_rec) | 
| 3568 | ||
| 3569 | lemma signed_take_bit_numeral_bit1 [simp]: | |
| 72241 | 3570 | \<open>signed_take_bit (numeral l) (numeral (Num.Bit1 k)) = signed_take_bit (pred_numeral l) (numeral k) * 2 + (1 :: int)\<close> | 
| 72010 | 3571 | by (simp add: signed_take_bit_rec) | 
| 3572 | ||
| 3573 | lemma signed_take_bit_numeral_minus_bit0 [simp]: | |
| 72241 | 3574 | \<open>signed_take_bit (numeral l) (- numeral (Num.Bit0 k)) = signed_take_bit (pred_numeral l) (- numeral k) * (2 :: int)\<close> | 
| 72010 | 3575 | by (simp add: signed_take_bit_rec) | 
| 3576 | ||
| 3577 | lemma signed_take_bit_numeral_minus_bit1 [simp]: | |
| 72241 | 3578 | \<open>signed_take_bit (numeral l) (- numeral (Num.Bit1 k)) = signed_take_bit (pred_numeral l) (- numeral k - 1) * 2 + (1 :: int)\<close> | 
| 72010 | 3579 | by (simp add: signed_take_bit_rec) | 
| 3580 | ||
| 3581 | lemma signed_take_bit_code [code]: | |
| 72241 | 3582 | \<open>signed_take_bit n a = | 
| 3583 | (let l = take_bit (Suc n) a | |
| 3584 | in if bit l n then l + push_bit (Suc n) (- 1) else l)\<close> | |
| 72010 | 3585 | proof - | 
| 72241 | 3586 | have *: \<open>take_bit (Suc n) a + push_bit n (- 2) = | 
| 3587 | take_bit (Suc n) a OR NOT (mask (Suc n))\<close> | |
| 3588 | by (auto simp add: bit_take_bit_iff bit_push_bit_iff bit_not_iff bit_mask_iff disjunctive_add | |
| 3589 | simp flip: push_bit_minus_one_eq_not_mask) | |
| 72010 | 3590 | show ?thesis | 
| 3591 | by (rule bit_eqI) | |
| 74592 | 3592 | (auto simp add: Let_def * bit_signed_take_bit_iff bit_take_bit_iff min_def less_Suc_eq bit_not_iff | 
| 3593 | bit_mask_iff bit_or_iff simp del: push_bit_minus_one_eq_not_mask) | |
| 72010 | 3594 | qed | 
| 3595 | ||
| 3596 | ||
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 3597 | subsection \<open>Horner sums\<close> | 
| 71804 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 3598 | |
| 72227 | 3599 | context semiring_bit_operations | 
| 3600 | begin | |
| 3601 | ||
| 74101 | 3602 | lemma horner_sum_bit_eq_take_bit: | 
| 3603 | \<open>horner_sum of_bool 2 (map (bit a) [0..<n]) = take_bit n a\<close> | |
| 3604 | proof (induction a arbitrary: n rule: bits_induct) | |
| 3605 | case (stable a) | |
| 3606 | moreover have \<open>bit a = (\<lambda>_. odd a)\<close> | |
| 3607 | using stable by (simp add: stable_imp_bit_iff_odd fun_eq_iff) | |
| 3608 |   moreover have \<open>{q. q < n} = {0..<n}\<close>
 | |
| 3609 | by auto | |
| 3610 | ultimately show ?case | |
| 3611 | by (simp add: stable_imp_take_bit_eq horner_sum_eq_sum mask_eq_sum_exp) | |
| 3612 | next | |
| 3613 | case (rec a b) | |
| 3614 | show ?case | |
| 3615 | proof (cases n) | |
| 3616 | case 0 | |
| 3617 | then show ?thesis | |
| 3618 | by simp | |
| 3619 | next | |
| 3620 | case (Suc m) | |
| 3621 | 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> | |
| 3622 | by (simp only: upt_conv_Cons) simp | |
| 3623 | also have \<open>\<dots> = b # map (bit a) [0..<m]\<close> | |
| 3624 | by (simp only: flip: map_Suc_upt) (simp add: bit_Suc rec.hyps) | |
| 3625 | finally show ?thesis | |
| 3626 | using Suc rec.IH [of m] by (simp add: take_bit_Suc rec.hyps) | |
| 3627 | (simp_all add: ac_simps mod_2_eq_odd) | |
| 3628 | qed | |
| 3629 | qed | |
| 3630 | ||
| 3631 | end | |
| 3632 | ||
| 74108 
3146646a43a7
simplified hierarchy of type classes for bit operations
 haftmann parents: 
74101diff
changeset | 3633 | context unique_euclidean_semiring_with_bit_operations | 
| 74101 | 3634 | begin | 
| 3635 | ||
| 3636 | lemma bit_horner_sum_bit_iff [bit_simps]: | |
| 3637 | \<open>bit (horner_sum of_bool 2 bs) n \<longleftrightarrow> n < length bs \<and> bs ! n\<close> | |
| 3638 | proof (induction bs arbitrary: n) | |
| 3639 | case Nil | |
| 3640 | then show ?case | |
| 3641 | by simp | |
| 3642 | next | |
| 3643 | case (Cons b bs) | |
| 3644 | show ?case | |
| 3645 | proof (cases n) | |
| 3646 | case 0 | |
| 3647 | then show ?thesis | |
| 3648 | by simp | |
| 3649 | next | |
| 3650 | case (Suc m) | |
| 3651 | with bit_rec [of _ n] Cons.prems Cons.IH [of m] | |
| 3652 | show ?thesis by simp | |
| 3653 | qed | |
| 3654 | qed | |
| 3655 | ||
| 3656 | lemma take_bit_horner_sum_bit_eq: | |
| 3657 | \<open>take_bit n (horner_sum of_bool 2 bs) = horner_sum of_bool 2 (take n bs)\<close> | |
| 3658 | by (auto simp add: bit_eq_iff bit_take_bit_iff bit_horner_sum_bit_iff) | |
| 3659 | ||
| 3660 | end | |
| 3661 | ||
| 3662 | lemma horner_sum_of_bool_2_less: | |
| 3663 | \<open>(horner_sum of_bool 2 bs :: int) < 2 ^ length bs\<close> | |
| 3664 | proof - | |
| 3665 | have \<open>(\<Sum>n = 0..<length bs. of_bool (bs ! n) * (2::int) ^ n) \<le> (\<Sum>n = 0..<length bs. 2 ^ n)\<close> | |
| 3666 | by (rule sum_mono) simp | |
| 3667 | also have \<open>\<dots> = 2 ^ length bs - 1\<close> | |
| 3668 | by (induction bs) simp_all | |
| 3669 | finally show ?thesis | |
| 3670 | by (simp add: horner_sum_eq_sum) | |
| 3671 | qed | |
| 3672 | ||
| 3673 | ||
| 71800 | 3674 | subsection \<open>Key ideas of bit operations\<close> | 
| 3675 | ||
| 3676 | text \<open> | |
| 3677 | When formalizing bit operations, it is tempting to represent | |
| 3678 | bit values as explicit lists over a binary type. This however | |
| 3679 | is a bad idea, mainly due to the inherent ambiguities in | |
| 3680 | representation concerning repeating leading bits. | |
| 3681 | ||
| 3682 | Hence this approach avoids such explicit lists altogether | |
| 3683 | following an algebraic path: | |
| 3684 | ||
| 3685 | \<^item> Bit values are represented by numeric types: idealized | |
| 3686 | unbounded bit values can be represented by type \<^typ>\<open>int\<close>, | |
| 3687 | bounded bit values by quotient types over \<^typ>\<open>int\<close>. | |
| 3688 | ||
| 3689 | \<^item> (A special case are idealized unbounded bit values ending | |
| 3690 |     in @{term [source] 0} which can be represented by type \<^typ>\<open>nat\<close> but
 | |
| 3691 | only support a restricted set of operations). | |
| 3692 | ||
| 3693 | \<^item> From this idea follows that | |
| 3694 | ||
| 3695 | \<^item> multiplication by \<^term>\<open>2 :: int\<close> is a bit shift to the left and | |
| 3696 | ||
| 3697 | \<^item> division by \<^term>\<open>2 :: int\<close> is a bit shift to the right. | |
| 3698 | ||
| 3699 | \<^item> Concerning bounded bit values, iterated shifts to the left | |
| 3700 | may result in eliminating all bits by shifting them all | |
| 3701 | beyond the boundary. The property \<^prop>\<open>(2 :: int) ^ n \<noteq> 0\<close> | |
| 3702 | represents that \<^term>\<open>n\<close> is \<^emph>\<open>not\<close> beyond that boundary. | |
| 3703 | ||
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71956diff
changeset | 3704 |   \<^item> The projection on a single bit is then @{thm bit_iff_odd [where ?'a = int, no_vars]}.
 | 
| 71800 | 3705 | |
| 3706 | \<^item> This leads to the most fundamental properties of bit values: | |
| 3707 | ||
| 3708 |       \<^item> Equality rule: @{thm bit_eqI [where ?'a = int, no_vars]}
 | |
| 3709 | ||
| 3710 |       \<^item> Induction rule: @{thm bits_induct [where ?'a = int, no_vars]}
 | |
| 3711 | ||
| 3712 | \<^item> Typical operations are characterized as follows: | |
| 3713 | ||
| 3714 | \<^item> Singleton \<^term>\<open>n\<close>th bit: \<^term>\<open>(2 :: int) ^ n\<close> | |
| 3715 | ||
| 71956 | 3716 |       \<^item> Bit mask upto bit \<^term>\<open>n\<close>: @{thm mask_eq_exp_minus_1 [where ?'a = int, no_vars]}
 | 
| 71800 | 3717 | |
| 3718 |       \<^item> Left shift: @{thm push_bit_eq_mult [where ?'a = int, no_vars]}
 | |
| 3719 | ||
| 3720 |       \<^item> Right shift: @{thm drop_bit_eq_div [where ?'a = int, no_vars]}
 | |
| 3721 | ||
| 3722 |       \<^item> Truncation: @{thm take_bit_eq_mod [where ?'a = int, no_vars]}
 | |
| 3723 | ||
| 3724 |       \<^item> Negation: @{thm bit_not_iff [where ?'a = int, no_vars]}
 | |
| 3725 | ||
| 3726 |       \<^item> And: @{thm bit_and_iff [where ?'a = int, no_vars]}
 | |
| 3727 | ||
| 3728 |       \<^item> Or: @{thm bit_or_iff [where ?'a = int, no_vars]}
 | |
| 3729 | ||
| 3730 |       \<^item> Xor: @{thm bit_xor_iff [where ?'a = int, no_vars]}
 | |
| 3731 | ||
| 3732 |       \<^item> Set a single bit: @{thm set_bit_def [where ?'a = int, no_vars]}
 | |
| 3733 | ||
| 3734 |       \<^item> Unset a single bit: @{thm unset_bit_def [where ?'a = int, no_vars]}
 | |
| 3735 | ||
| 3736 |       \<^item> Flip a single bit: @{thm flip_bit_def [where ?'a = int, no_vars]}
 | |
| 72028 | 3737 | |
| 72241 | 3738 |       \<^item> Signed truncation, or modulus centered around \<^term>\<open>0::int\<close>: @{thm signed_take_bit_def [no_vars]}
 | 
| 72028 | 3739 | |
| 72241 | 3740 |       \<^item> Bit concatenation: @{thm concat_bit_def [no_vars]}
 | 
| 72028 | 3741 | |
| 3742 |       \<^item> (Bounded) conversion from and to a list of bits: @{thm horner_sum_bit_eq_take_bit [where ?'a = int, no_vars]}
 | |
| 71800 | 3743 | \<close> | 
| 3744 | ||
| 74097 | 3745 | no_notation | 
| 74391 | 3746 | not (\<open>NOT\<close>) | 
| 74364 | 3747 | and "and" (infixr \<open>AND\<close> 64) | 
| 74097 | 3748 | and or (infixr \<open>OR\<close> 59) | 
| 3749 | and xor (infixr \<open>XOR\<close> 59) | |
| 3750 | ||
| 3751 | bundle bit_operations_syntax | |
| 74101 | 3752 | begin | 
| 74097 | 3753 | |
| 3754 | notation | |
| 74391 | 3755 | not (\<open>NOT\<close>) | 
| 74364 | 3756 | and "and" (infixr \<open>AND\<close> 64) | 
| 74097 | 3757 | and or (infixr \<open>OR\<close> 59) | 
| 3758 | and xor (infixr \<open>XOR\<close> 59) | |
| 3759 | ||
| 71442 | 3760 | end | 
| 74097 | 3761 | |
| 3762 | end |