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