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