| author | haftmann |
| Sat, 20 Jun 2020 05:56:28 +0000 | |
| changeset 71966 | e18e9ac8c205 |
| parent 71965 | d45f5d4c41bd |
| child 71986 | 76193dd4aec8 |
| 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 |
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
7 |
imports |
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
8 |
"HOL-Library.Boolean_Algebra" |
| 71095 | 9 |
Main |
|
71042
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
10 |
begin |
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
11 |
|
| 71956 | 12 |
subsection \<open>Bit operations\<close> |
|
71042
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
13 |
|
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
14 |
class semiring_bit_operations = semiring_bit_shifts + |
| 71426 | 15 |
fixes "and" :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close> (infixr \<open>AND\<close> 64) |
16 |
and or :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close> (infixr \<open>OR\<close> 59) |
|
17 |
and xor :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close> (infixr \<open>XOR\<close> 59) |
|
| 71186 | 18 |
assumes bit_and_iff: \<open>\<And>n. bit (a AND b) n \<longleftrightarrow> bit a n \<and> bit b n\<close> |
19 |
and bit_or_iff: \<open>\<And>n. bit (a OR b) n \<longleftrightarrow> bit a n \<or> bit b n\<close> |
|
20 |
and bit_xor_iff: \<open>\<And>n. bit (a XOR b) n \<longleftrightarrow> bit a n \<noteq> bit b 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
|
21 |
begin |
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
22 |
|
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
23 |
text \<open> |
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
24 |
We want the bitwise operations to bind slightly weaker |
| 71094 | 25 |
than \<open>+\<close> and \<open>-\<close>. |
|
71042
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
26 |
For the sake of code generation |
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
27 |
the operations \<^const>\<open>and\<close>, \<^const>\<open>or\<close> and \<^const>\<open>xor\<close> |
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
28 |
are specified as definitional class operations. |
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
29 |
\<close> |
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
30 |
|
| 71418 | 31 |
sublocale "and": semilattice \<open>(AND)\<close> |
32 |
by standard (auto simp add: bit_eq_iff bit_and_iff) |
|
33 |
||
34 |
sublocale or: semilattice_neutr \<open>(OR)\<close> 0 |
|
35 |
by standard (auto simp add: bit_eq_iff bit_or_iff) |
|
36 |
||
37 |
sublocale xor: comm_monoid \<open>(XOR)\<close> 0 |
|
38 |
by standard (auto simp add: bit_eq_iff bit_xor_iff) |
|
39 |
||
| 71823 | 40 |
lemma even_and_iff: |
41 |
\<open>even (a AND b) \<longleftrightarrow> even a \<or> even b\<close> |
|
42 |
using bit_and_iff [of a b 0] by auto |
|
43 |
||
44 |
lemma even_or_iff: |
|
45 |
\<open>even (a OR b) \<longleftrightarrow> even a \<and> even b\<close> |
|
46 |
using bit_or_iff [of a b 0] by auto |
|
47 |
||
48 |
lemma even_xor_iff: |
|
49 |
\<open>even (a XOR b) \<longleftrightarrow> (even a \<longleftrightarrow> even b)\<close> |
|
50 |
using bit_xor_iff [of a b 0] by auto |
|
51 |
||
| 71412 | 52 |
lemma zero_and_eq [simp]: |
53 |
"0 AND a = 0" |
|
54 |
by (simp add: bit_eq_iff bit_and_iff) |
|
55 |
||
56 |
lemma and_zero_eq [simp]: |
|
57 |
"a AND 0 = 0" |
|
58 |
by (simp add: bit_eq_iff bit_and_iff) |
|
59 |
||
| 71921 | 60 |
lemma one_and_eq: |
| 71822 | 61 |
"1 AND a = a mod 2" |
| 71418 | 62 |
by (simp add: bit_eq_iff bit_and_iff) (auto simp add: bit_1_iff) |
| 71412 | 63 |
|
| 71921 | 64 |
lemma and_one_eq: |
| 71822 | 65 |
"a AND 1 = a mod 2" |
| 71418 | 66 |
using one_and_eq [of a] by (simp add: ac_simps) |
67 |
||
| 71822 | 68 |
lemma one_or_eq: |
| 71418 | 69 |
"1 OR a = a + of_bool (even a)" |
70 |
by (simp add: bit_eq_iff bit_or_iff add.commute [of _ 1] even_bit_succ_iff) (auto simp add: bit_1_iff) |
|
| 71412 | 71 |
|
| 71822 | 72 |
lemma or_one_eq: |
| 71418 | 73 |
"a OR 1 = a + of_bool (even a)" |
74 |
using one_or_eq [of a] by (simp add: ac_simps) |
|
| 71412 | 75 |
|
| 71822 | 76 |
lemma one_xor_eq: |
| 71418 | 77 |
"1 XOR a = a + of_bool (even a) - of_bool (odd a)" |
78 |
by (simp add: bit_eq_iff bit_xor_iff add.commute [of _ 1] even_bit_succ_iff) (auto simp add: bit_1_iff odd_bit_iff_bit_pred elim: oddE) |
|
79 |
||
| 71822 | 80 |
lemma xor_one_eq: |
| 71418 | 81 |
"a XOR 1 = a + of_bool (even a) - of_bool (odd a)" |
82 |
using one_xor_eq [of a] by (simp add: ac_simps) |
|
| 71412 | 83 |
|
| 71409 | 84 |
lemma take_bit_and [simp]: |
85 |
\<open>take_bit n (a AND b) = take_bit n a AND take_bit n b\<close> |
|
86 |
by (auto simp add: bit_eq_iff bit_take_bit_iff bit_and_iff) |
|
87 |
||
88 |
lemma take_bit_or [simp]: |
|
89 |
\<open>take_bit n (a OR b) = take_bit n a OR take_bit n b\<close> |
|
90 |
by (auto simp add: bit_eq_iff bit_take_bit_iff bit_or_iff) |
|
91 |
||
92 |
lemma take_bit_xor [simp]: |
|
93 |
\<open>take_bit n (a XOR b) = take_bit n a XOR take_bit n b\<close> |
|
94 |
by (auto simp add: bit_eq_iff bit_take_bit_iff bit_xor_iff) |
|
95 |
||
| 71823 | 96 |
definition mask :: \<open>nat \<Rightarrow> 'a\<close> |
97 |
where mask_eq_exp_minus_1: \<open>mask n = 2 ^ n - 1\<close> |
|
98 |
||
99 |
lemma bit_mask_iff: |
|
100 |
\<open>bit (mask m) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n < m\<close> |
|
101 |
by (simp add: mask_eq_exp_minus_1 bit_mask_iff) |
|
102 |
||
103 |
lemma even_mask_iff: |
|
104 |
\<open>even (mask n) \<longleftrightarrow> n = 0\<close> |
|
105 |
using bit_mask_iff [of n 0] by auto |
|
106 |
||
107 |
lemma mask_0 [simp, code]: |
|
108 |
\<open>mask 0 = 0\<close> |
|
109 |
by (simp add: mask_eq_exp_minus_1) |
|
110 |
||
111 |
lemma mask_Suc_exp [code]: |
|
112 |
\<open>mask (Suc n) = 2 ^ n OR mask n\<close> |
|
113 |
by (rule bit_eqI) |
|
114 |
(auto simp add: bit_or_iff bit_mask_iff bit_exp_iff not_less le_less_Suc_eq) |
|
115 |
||
116 |
lemma mask_Suc_double: |
|
117 |
\<open>mask (Suc n) = 2 * mask n OR 1\<close> |
|
118 |
proof (rule bit_eqI) |
|
119 |
fix q |
|
120 |
assume \<open>2 ^ q \<noteq> 0\<close> |
|
121 |
show \<open>bit (mask (Suc n)) q \<longleftrightarrow> bit (2 * mask n OR 1) q\<close> |
|
122 |
by (cases q) |
|
123 |
(simp_all add: even_mask_iff even_or_iff bit_or_iff bit_mask_iff bit_exp_iff bit_double_iff not_less le_less_Suc_eq bit_1_iff, auto simp add: mult_2) |
|
124 |
qed |
|
125 |
||
|
71965
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71956
diff
changeset
|
126 |
lemma take_bit_eq_mask: |
| 71823 | 127 |
\<open>take_bit n a = a AND mask n\<close> |
128 |
by (rule bit_eqI) |
|
129 |
(auto simp add: bit_take_bit_iff bit_and_iff bit_mask_iff) |
|
130 |
||
|
71042
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
131 |
end |
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
132 |
|
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
133 |
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
|
134 |
fixes not :: \<open>'a \<Rightarrow> 'a\<close> (\<open>NOT\<close>) |
| 71186 | 135 |
assumes bit_not_iff: \<open>\<And>n. bit (NOT a) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> \<not> bit a n\<close> |
| 71409 | 136 |
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
|
137 |
begin |
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
138 |
|
| 71409 | 139 |
text \<open> |
140 |
For the sake of code generation \<^const>\<open>not\<close> is specified as |
|
141 |
definitional class operation. Note that \<^const>\<open>not\<close> has no |
|
142 |
sensible definition for unlimited but only positive bit strings |
|
143 |
(type \<^typ>\<open>nat\<close>). |
|
144 |
\<close> |
|
145 |
||
| 71186 | 146 |
lemma bits_minus_1_mod_2_eq [simp]: |
147 |
\<open>(- 1) mod 2 = 1\<close> |
|
148 |
by (simp add: mod_2_eq_odd) |
|
149 |
||
| 71409 | 150 |
lemma not_eq_complement: |
151 |
\<open>NOT a = - a - 1\<close> |
|
152 |
using minus_eq_not_minus_1 [of \<open>a + 1\<close>] by simp |
|
153 |
||
154 |
lemma minus_eq_not_plus_1: |
|
155 |
\<open>- a = NOT a + 1\<close> |
|
156 |
using not_eq_complement [of a] by simp |
|
157 |
||
158 |
lemma bit_minus_iff: |
|
159 |
\<open>bit (- a) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> \<not> bit (a - 1) n\<close> |
|
160 |
by (simp add: minus_eq_not_minus_1 bit_not_iff) |
|
161 |
||
| 71418 | 162 |
lemma even_not_iff [simp]: |
163 |
"even (NOT a) \<longleftrightarrow> odd a" |
|
164 |
using bit_not_iff [of a 0] by auto |
|
165 |
||
| 71409 | 166 |
lemma bit_not_exp_iff: |
167 |
\<open>bit (NOT (2 ^ m)) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n \<noteq> m\<close> |
|
168 |
by (auto simp add: bit_not_iff bit_exp_iff) |
|
169 |
||
| 71186 | 170 |
lemma bit_minus_1_iff [simp]: |
171 |
\<open>bit (- 1) n \<longleftrightarrow> 2 ^ n \<noteq> 0\<close> |
|
| 71409 | 172 |
by (simp add: bit_minus_iff) |
173 |
||
174 |
lemma bit_minus_exp_iff: |
|
175 |
\<open>bit (- (2 ^ m)) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n \<ge> m\<close> |
|
176 |
oops |
|
177 |
||
178 |
lemma bit_minus_2_iff [simp]: |
|
179 |
\<open>bit (- 2) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n > 0\<close> |
|
180 |
by (simp add: bit_minus_iff bit_1_iff) |
|
| 71186 | 181 |
|
| 71418 | 182 |
lemma not_one [simp]: |
183 |
"NOT 1 = - 2" |
|
184 |
by (simp add: bit_eq_iff bit_not_iff) (simp add: bit_1_iff) |
|
185 |
||
186 |
sublocale "and": semilattice_neutr \<open>(AND)\<close> \<open>- 1\<close> |
|
187 |
apply standard |
|
| 71426 | 188 |
apply (simp add: bit_eq_iff bit_and_iff) |
189 |
apply (auto simp add: exp_eq_0_imp_not_bit bit_exp_iff) |
|
| 71418 | 190 |
done |
191 |
||
|
71042
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
192 |
sublocale bit: boolean_algebra \<open>(AND)\<close> \<open>(OR)\<close> NOT 0 \<open>- 1\<close> |
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
193 |
rewrites \<open>bit.xor = (XOR)\<close> |
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
194 |
proof - |
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
195 |
interpret bit: boolean_algebra \<open>(AND)\<close> \<open>(OR)\<close> NOT 0 \<open>- 1\<close> |
| 71186 | 196 |
apply standard |
|
71965
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71956
diff
changeset
|
197 |
apply (simp_all add: bit_eq_iff) |
|
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71956
diff
changeset
|
198 |
apply (auto simp add: bit_and_iff bit_or_iff bit_not_iff bit_exp_iff exp_eq_0_imp_not_bit) |
| 71186 | 199 |
done |
|
71042
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
200 |
show \<open>boolean_algebra (AND) (OR) NOT 0 (- 1)\<close> |
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
201 |
by standard |
| 71426 | 202 |
show \<open>boolean_algebra.xor (AND) (OR) NOT = (XOR)\<close> |
|
71965
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71956
diff
changeset
|
203 |
apply (simp add: fun_eq_iff bit_eq_iff bit.xor_def) |
|
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71956
diff
changeset
|
204 |
apply (auto simp add: bit_and_iff bit_or_iff bit_not_iff bit_xor_iff exp_eq_0_imp_not_bit) |
| 71186 | 205 |
done |
|
71042
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
206 |
qed |
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
207 |
|
| 71802 | 208 |
lemma and_eq_not_not_or: |
209 |
\<open>a AND b = NOT (NOT a OR NOT b)\<close> |
|
210 |
by simp |
|
211 |
||
212 |
lemma or_eq_not_not_and: |
|
213 |
\<open>a OR b = NOT (NOT a AND NOT b)\<close> |
|
214 |
by simp |
|
215 |
||
| 71412 | 216 |
lemma push_bit_minus: |
217 |
\<open>push_bit n (- a) = - push_bit n a\<close> |
|
218 |
by (simp add: push_bit_eq_mult) |
|
219 |
||
| 71409 | 220 |
lemma take_bit_not_take_bit: |
221 |
\<open>take_bit n (NOT (take_bit n a)) = take_bit n (NOT a)\<close> |
|
222 |
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
|
223 |
|
| 71418 | 224 |
lemma take_bit_not_iff: |
225 |
"take_bit n (NOT a) = take_bit n (NOT b) \<longleftrightarrow> take_bit n a = take_bit n b" |
|
226 |
apply (simp add: bit_eq_iff bit_not_iff bit_take_bit_iff) |
|
227 |
apply (simp add: bit_exp_iff) |
|
228 |
apply (use local.exp_eq_0_imp_not_bit in blast) |
|
229 |
done |
|
230 |
||
| 71922 | 231 |
lemma take_bit_minus_one_eq_mask: |
232 |
\<open>take_bit n (- 1) = mask n\<close> |
|
233 |
by (simp add: bit_eq_iff bit_mask_iff bit_take_bit_iff conj_commute) |
|
234 |
||
235 |
lemma push_bit_minus_one_eq_not_mask: |
|
236 |
\<open>push_bit n (- 1) = NOT (mask n)\<close> |
|
237 |
proof (rule bit_eqI) |
|
238 |
fix m |
|
239 |
assume \<open>2 ^ m \<noteq> 0\<close> |
|
240 |
show \<open>bit (push_bit n (- 1)) m \<longleftrightarrow> bit (NOT (mask n)) m\<close> |
|
241 |
proof (cases \<open>n \<le> m\<close>) |
|
242 |
case True |
|
243 |
moreover define q where \<open>q = m - n\<close> |
|
244 |
ultimately have \<open>m = n + q\<close> \<open>m - n = q\<close> |
|
245 |
by simp_all |
|
246 |
with \<open>2 ^ m \<noteq> 0\<close> have \<open>2 ^ n * 2 ^ q \<noteq> 0\<close> |
|
247 |
by (simp add: power_add) |
|
248 |
then have \<open>2 ^ q \<noteq> 0\<close> |
|
249 |
using mult_not_zero by blast |
|
250 |
with \<open>m - n = q\<close> show ?thesis |
|
251 |
by (auto simp add: bit_not_iff bit_mask_iff bit_push_bit_iff not_less) |
|
252 |
next |
|
253 |
case False |
|
254 |
then show ?thesis |
|
255 |
by (simp add: bit_not_iff bit_mask_iff bit_push_bit_iff not_le) |
|
256 |
qed |
|
257 |
qed |
|
258 |
||
| 71426 | 259 |
definition set_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close> |
260 |
where \<open>set_bit n a = a OR 2 ^ n\<close> |
|
261 |
||
262 |
definition unset_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close> |
|
263 |
where \<open>unset_bit n a = a AND NOT (2 ^ n)\<close> |
|
264 |
||
265 |
definition flip_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close> |
|
266 |
where \<open>flip_bit n a = a XOR 2 ^ n\<close> |
|
267 |
||
268 |
lemma bit_set_bit_iff: |
|
269 |
\<open>bit (set_bit m a) n \<longleftrightarrow> bit a n \<or> (m = n \<and> 2 ^ n \<noteq> 0)\<close> |
|
270 |
by (auto simp add: set_bit_def bit_or_iff bit_exp_iff) |
|
271 |
||
272 |
lemma even_set_bit_iff: |
|
273 |
\<open>even (set_bit m a) \<longleftrightarrow> even a \<and> m \<noteq> 0\<close> |
|
274 |
using bit_set_bit_iff [of m a 0] by auto |
|
275 |
||
276 |
lemma bit_unset_bit_iff: |
|
277 |
\<open>bit (unset_bit m a) n \<longleftrightarrow> bit a n \<and> m \<noteq> n\<close> |
|
278 |
by (auto simp add: unset_bit_def bit_and_iff bit_not_iff bit_exp_iff exp_eq_0_imp_not_bit) |
|
279 |
||
280 |
lemma even_unset_bit_iff: |
|
281 |
\<open>even (unset_bit m a) \<longleftrightarrow> even a \<or> m = 0\<close> |
|
282 |
using bit_unset_bit_iff [of m a 0] by auto |
|
283 |
||
284 |
lemma bit_flip_bit_iff: |
|
285 |
\<open>bit (flip_bit m a) n \<longleftrightarrow> (m = n \<longleftrightarrow> \<not> bit a n) \<and> 2 ^ n \<noteq> 0\<close> |
|
286 |
by (auto simp add: flip_bit_def bit_xor_iff bit_exp_iff exp_eq_0_imp_not_bit) |
|
287 |
||
288 |
lemma even_flip_bit_iff: |
|
289 |
\<open>even (flip_bit m a) \<longleftrightarrow> \<not> (even a \<longleftrightarrow> m = 0)\<close> |
|
290 |
using bit_flip_bit_iff [of m a 0] by auto |
|
291 |
||
292 |
lemma set_bit_0 [simp]: |
|
293 |
\<open>set_bit 0 a = 1 + 2 * (a div 2)\<close> |
|
294 |
proof (rule bit_eqI) |
|
295 |
fix m |
|
296 |
assume *: \<open>2 ^ m \<noteq> 0\<close> |
|
297 |
then show \<open>bit (set_bit 0 a) m = bit (1 + 2 * (a div 2)) m\<close> |
|
298 |
by (simp add: bit_set_bit_iff bit_double_iff even_bit_succ_iff) |
|
|
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71442
diff
changeset
|
299 |
(cases m, simp_all add: bit_Suc) |
| 71426 | 300 |
qed |
301 |
||
| 71821 | 302 |
lemma set_bit_Suc: |
| 71426 | 303 |
\<open>set_bit (Suc n) a = a mod 2 + 2 * set_bit n (a div 2)\<close> |
304 |
proof (rule bit_eqI) |
|
305 |
fix m |
|
306 |
assume *: \<open>2 ^ m \<noteq> 0\<close> |
|
307 |
show \<open>bit (set_bit (Suc n) a) m \<longleftrightarrow> bit (a mod 2 + 2 * set_bit n (a div 2)) m\<close> |
|
308 |
proof (cases m) |
|
309 |
case 0 |
|
310 |
then show ?thesis |
|
311 |
by (simp add: even_set_bit_iff) |
|
312 |
next |
|
313 |
case (Suc m) |
|
314 |
with * have \<open>2 ^ m \<noteq> 0\<close> |
|
315 |
using mult_2 by auto |
|
316 |
show ?thesis |
|
317 |
by (cases a rule: parity_cases) |
|
318 |
(simp_all add: bit_set_bit_iff bit_double_iff even_bit_succ_iff *, |
|
|
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71442
diff
changeset
|
319 |
simp_all add: Suc \<open>2 ^ m \<noteq> 0\<close> bit_Suc) |
| 71426 | 320 |
qed |
321 |
qed |
|
322 |
||
323 |
lemma unset_bit_0 [simp]: |
|
324 |
\<open>unset_bit 0 a = 2 * (a div 2)\<close> |
|
325 |
proof (rule bit_eqI) |
|
326 |
fix m |
|
327 |
assume *: \<open>2 ^ m \<noteq> 0\<close> |
|
328 |
then show \<open>bit (unset_bit 0 a) m = bit (2 * (a div 2)) m\<close> |
|
329 |
by (simp add: bit_unset_bit_iff bit_double_iff) |
|
|
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71442
diff
changeset
|
330 |
(cases m, simp_all add: bit_Suc) |
| 71426 | 331 |
qed |
332 |
||
| 71821 | 333 |
lemma unset_bit_Suc: |
| 71426 | 334 |
\<open>unset_bit (Suc n) a = a mod 2 + 2 * unset_bit n (a div 2)\<close> |
335 |
proof (rule bit_eqI) |
|
336 |
fix m |
|
337 |
assume *: \<open>2 ^ m \<noteq> 0\<close> |
|
338 |
then show \<open>bit (unset_bit (Suc n) a) m \<longleftrightarrow> bit (a mod 2 + 2 * unset_bit n (a div 2)) m\<close> |
|
339 |
proof (cases m) |
|
340 |
case 0 |
|
341 |
then show ?thesis |
|
342 |
by (simp add: even_unset_bit_iff) |
|
343 |
next |
|
344 |
case (Suc m) |
|
345 |
show ?thesis |
|
346 |
by (cases a rule: parity_cases) |
|
347 |
(simp_all add: bit_unset_bit_iff bit_double_iff even_bit_succ_iff *, |
|
|
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71442
diff
changeset
|
348 |
simp_all add: Suc bit_Suc) |
| 71426 | 349 |
qed |
350 |
qed |
|
351 |
||
352 |
lemma flip_bit_0 [simp]: |
|
353 |
\<open>flip_bit 0 a = of_bool (even a) + 2 * (a div 2)\<close> |
|
354 |
proof (rule bit_eqI) |
|
355 |
fix m |
|
356 |
assume *: \<open>2 ^ m \<noteq> 0\<close> |
|
357 |
then show \<open>bit (flip_bit 0 a) m = bit (of_bool (even a) + 2 * (a div 2)) m\<close> |
|
358 |
by (simp add: bit_flip_bit_iff bit_double_iff even_bit_succ_iff) |
|
|
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71442
diff
changeset
|
359 |
(cases m, simp_all add: bit_Suc) |
| 71426 | 360 |
qed |
361 |
||
| 71821 | 362 |
lemma flip_bit_Suc: |
| 71426 | 363 |
\<open>flip_bit (Suc n) a = a mod 2 + 2 * flip_bit n (a div 2)\<close> |
364 |
proof (rule bit_eqI) |
|
365 |
fix m |
|
366 |
assume *: \<open>2 ^ m \<noteq> 0\<close> |
|
367 |
show \<open>bit (flip_bit (Suc n) a) m \<longleftrightarrow> bit (a mod 2 + 2 * flip_bit n (a div 2)) m\<close> |
|
368 |
proof (cases m) |
|
369 |
case 0 |
|
370 |
then show ?thesis |
|
371 |
by (simp add: even_flip_bit_iff) |
|
372 |
next |
|
373 |
case (Suc m) |
|
374 |
with * have \<open>2 ^ m \<noteq> 0\<close> |
|
375 |
using mult_2 by auto |
|
376 |
show ?thesis |
|
377 |
by (cases a rule: parity_cases) |
|
378 |
(simp_all add: bit_flip_bit_iff bit_double_iff even_bit_succ_iff, |
|
|
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71442
diff
changeset
|
379 |
simp_all add: Suc \<open>2 ^ m \<noteq> 0\<close> bit_Suc) |
| 71426 | 380 |
qed |
381 |
qed |
|
382 |
||
|
71042
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
383 |
end |
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
384 |
|
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
385 |
|
| 71956 | 386 |
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
|
387 |
|
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
388 |
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
|
389 |
begin |
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
390 |
|
| 71420 | 391 |
definition not_int :: \<open>int \<Rightarrow> int\<close> |
392 |
where \<open>not_int k = - k - 1\<close> |
|
393 |
||
394 |
lemma not_int_rec: |
|
395 |
"NOT k = of_bool (even k) + 2 * NOT (k div 2)" for k :: int |
|
396 |
by (auto simp add: not_int_def elim: oddE) |
|
397 |
||
398 |
lemma even_not_iff_int: |
|
399 |
\<open>even (NOT k) \<longleftrightarrow> odd k\<close> for k :: int |
|
400 |
by (simp add: not_int_def) |
|
401 |
||
402 |
lemma not_int_div_2: |
|
403 |
\<open>NOT k div 2 = NOT (k div 2)\<close> for k :: int |
|
404 |
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
|
405 |
|
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
406 |
lemma bit_not_int_iff: |
| 71186 | 407 |
\<open>bit (NOT k) n \<longleftrightarrow> \<not> bit k n\<close> |
408 |
for k :: int |
|
|
71535
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
haftmann
parents:
71442
diff
changeset
|
409 |
by (induction n arbitrary: k) (simp_all add: not_int_div_2 even_not_iff_int bit_Suc) |
| 71186 | 410 |
|
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
411 |
function and_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close> |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
412 |
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
|
413 |
then - of_bool (odd k \<and> odd l) |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
414 |
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
|
415 |
by auto |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
416 |
|
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
417 |
termination |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
418 |
by (relation \<open>measure (\<lambda>(k, l). nat (\<bar>k\<bar> + \<bar>l\<bar>))\<close>) auto |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
419 |
|
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
420 |
declare and_int.simps [simp del] |
| 71802 | 421 |
|
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
422 |
lemma and_int_rec: |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
423 |
\<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
|
424 |
for k l :: int |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
425 |
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
|
426 |
case True |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
427 |
then show ?thesis |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
428 |
by auto (simp_all add: and_int.simps) |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
429 |
next |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
430 |
case False |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
431 |
then show ?thesis |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
432 |
by (auto simp add: ac_simps and_int.simps [of k l]) |
| 71802 | 433 |
qed |
434 |
||
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
435 |
lemma bit_and_int_iff: |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
436 |
\<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
|
437 |
proof (induction n arbitrary: k l) |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
438 |
case 0 |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
439 |
then show ?case |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
440 |
by (simp add: and_int_rec [of k l]) |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
441 |
next |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
442 |
case (Suc n) |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
443 |
then show ?case |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
444 |
by (simp add: and_int_rec [of k l] bit_Suc) |
| 71802 | 445 |
qed |
446 |
||
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
447 |
lemma even_and_iff_int: |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
448 |
\<open>even (k AND l) \<longleftrightarrow> even k \<or> even l\<close> for k l :: int |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
449 |
using bit_and_int_iff [of k l 0] by auto |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
450 |
|
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
451 |
definition or_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close> |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
452 |
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
|
453 |
|
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
454 |
lemma or_int_rec: |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
455 |
\<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
|
456 |
for k l :: int |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
457 |
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
|
458 |
by (simp add: or_int_def even_not_iff_int not_int_div_2) |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
459 |
(simp add: not_int_def) |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
460 |
|
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
461 |
lemma bit_or_int_iff: |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
462 |
\<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
|
463 |
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
|
464 |
|
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
465 |
definition xor_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close> |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
466 |
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
|
467 |
|
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
468 |
lemma xor_int_rec: |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
469 |
\<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
|
470 |
for k l :: int |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
471 |
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
|
472 |
(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
|
473 |
|
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
474 |
lemma bit_xor_int_iff: |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
475 |
\<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
|
476 |
by (auto simp add: xor_int_def bit_or_int_iff bit_and_int_iff bit_not_int_iff) |
| 71802 | 477 |
|
|
71042
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
478 |
instance proof |
| 71186 | 479 |
fix k l :: int and n :: nat |
| 71409 | 480 |
show \<open>- k = NOT (k - 1)\<close> |
481 |
by (simp add: not_int_def) |
|
| 71186 | 482 |
show \<open>bit (k AND l) n \<longleftrightarrow> bit k n \<and> bit l n\<close> |
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
483 |
by (fact bit_and_int_iff) |
| 71186 | 484 |
show \<open>bit (k OR l) n \<longleftrightarrow> bit k n \<or> bit l n\<close> |
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
485 |
by (fact bit_or_int_iff) |
| 71186 | 486 |
show \<open>bit (k XOR l) n \<longleftrightarrow> bit k n \<noteq> bit l n\<close> |
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
487 |
by (fact bit_xor_int_iff) |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
488 |
qed (simp_all add: bit_not_int_iff) |
|
71042
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
489 |
|
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
490 |
end |
|
400e9512f1d3
proof-of-concept theory for bit operations without a constructivistic representation and a minimal common logical foundation
haftmann
parents:
diff
changeset
|
491 |
|
| 71802 | 492 |
lemma not_nonnegative_int_iff [simp]: |
493 |
\<open>NOT k \<ge> 0 \<longleftrightarrow> k < 0\<close> for k :: int |
|
494 |
by (simp add: not_int_def) |
|
495 |
||
496 |
lemma not_negative_int_iff [simp]: |
|
497 |
\<open>NOT k < 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int |
|
498 |
by (subst Not_eq_iff [symmetric]) (simp add: not_less not_le) |
|
499 |
||
500 |
lemma and_nonnegative_int_iff [simp]: |
|
501 |
\<open>k AND l \<ge> 0 \<longleftrightarrow> k \<ge> 0 \<or> l \<ge> 0\<close> for k l :: int |
|
502 |
proof (induction k arbitrary: l rule: int_bit_induct) |
|
503 |
case zero |
|
504 |
then show ?case |
|
505 |
by simp |
|
506 |
next |
|
507 |
case minus |
|
508 |
then show ?case |
|
509 |
by simp |
|
510 |
next |
|
511 |
case (even k) |
|
512 |
then show ?case |
|
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
513 |
using and_int_rec [of \<open>k * 2\<close> l] by (simp add: pos_imp_zdiv_nonneg_iff) |
| 71802 | 514 |
next |
515 |
case (odd k) |
|
516 |
from odd have \<open>0 \<le> k AND l div 2 \<longleftrightarrow> 0 \<le> k \<or> 0 \<le> l div 2\<close> |
|
517 |
by simp |
|
518 |
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> |
|
519 |
by simp |
|
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
520 |
with and_int_rec [of \<open>1 + k * 2\<close> l] |
| 71802 | 521 |
show ?case |
522 |
by auto |
|
523 |
qed |
|
524 |
||
525 |
lemma and_negative_int_iff [simp]: |
|
526 |
\<open>k AND l < 0 \<longleftrightarrow> k < 0 \<and> l < 0\<close> for k l :: int |
|
527 |
by (subst Not_eq_iff [symmetric]) (simp add: not_less) |
|
528 |
||
529 |
lemma or_nonnegative_int_iff [simp]: |
|
530 |
\<open>k OR l \<ge> 0 \<longleftrightarrow> k \<ge> 0 \<and> l \<ge> 0\<close> for k l :: int |
|
531 |
by (simp only: or_eq_not_not_and not_nonnegative_int_iff) simp |
|
532 |
||
533 |
lemma or_negative_int_iff [simp]: |
|
534 |
\<open>k OR l < 0 \<longleftrightarrow> k < 0 \<or> l < 0\<close> for k l :: int |
|
535 |
by (subst Not_eq_iff [symmetric]) (simp add: not_less) |
|
536 |
||
537 |
lemma xor_nonnegative_int_iff [simp]: |
|
538 |
\<open>k XOR l \<ge> 0 \<longleftrightarrow> (k \<ge> 0 \<longleftrightarrow> l \<ge> 0)\<close> for k l :: int |
|
539 |
by (simp only: bit.xor_def or_nonnegative_int_iff) auto |
|
540 |
||
541 |
lemma xor_negative_int_iff [simp]: |
|
542 |
\<open>k XOR l < 0 \<longleftrightarrow> (k < 0) \<noteq> (l < 0)\<close> for k l :: int |
|
543 |
by (subst Not_eq_iff [symmetric]) (auto simp add: not_less) |
|
544 |
||
545 |
lemma set_bit_nonnegative_int_iff [simp]: |
|
546 |
\<open>set_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int |
|
547 |
by (simp add: set_bit_def) |
|
548 |
||
549 |
lemma set_bit_negative_int_iff [simp]: |
|
550 |
\<open>set_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int |
|
551 |
by (simp add: set_bit_def) |
|
552 |
||
553 |
lemma unset_bit_nonnegative_int_iff [simp]: |
|
554 |
\<open>unset_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int |
|
555 |
by (simp add: unset_bit_def) |
|
556 |
||
557 |
lemma unset_bit_negative_int_iff [simp]: |
|
558 |
\<open>unset_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int |
|
559 |
by (simp add: unset_bit_def) |
|
560 |
||
561 |
lemma flip_bit_nonnegative_int_iff [simp]: |
|
562 |
\<open>flip_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int |
|
563 |
by (simp add: flip_bit_def) |
|
564 |
||
565 |
lemma flip_bit_negative_int_iff [simp]: |
|
566 |
\<open>flip_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int |
|
567 |
by (simp add: flip_bit_def) |
|
568 |
||
| 71442 | 569 |
|
| 71956 | 570 |
subsection \<open>Instance \<^typ>\<open>nat\<close>\<close> |
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
571 |
|
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
572 |
instantiation nat :: semiring_bit_operations |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
573 |
begin |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
574 |
|
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
575 |
definition and_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close> |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
576 |
where \<open>m AND n = nat (int m AND int n)\<close> for m n :: nat |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
577 |
|
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
578 |
definition or_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close> |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
579 |
where \<open>m OR n = nat (int m OR int n)\<close> for m n :: nat |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
580 |
|
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
581 |
definition xor_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close> |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
582 |
where \<open>m XOR n = nat (int m XOR int n)\<close> for m n :: nat |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
583 |
|
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
584 |
instance proof |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
585 |
fix m n q :: nat |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
586 |
show \<open>bit (m AND n) q \<longleftrightarrow> bit m q \<and> bit n q\<close> |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
587 |
by (auto simp add: and_nat_def bit_and_iff less_le bit_eq_iff) |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
588 |
show \<open>bit (m OR n) q \<longleftrightarrow> bit m q \<or> bit n q\<close> |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
589 |
by (auto simp add: or_nat_def bit_or_iff less_le bit_eq_iff) |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
590 |
show \<open>bit (m XOR n) q \<longleftrightarrow> bit m q \<noteq> bit n q\<close> |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
591 |
by (auto simp add: xor_nat_def bit_xor_iff less_le bit_eq_iff) |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
592 |
qed |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
593 |
|
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
594 |
end |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
595 |
|
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
596 |
lemma and_nat_rec: |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
597 |
\<open>m AND n = of_bool (odd m \<and> odd n) + 2 * ((m div 2) AND (n div 2))\<close> for m n :: nat |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
598 |
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) |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
599 |
|
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
600 |
lemma or_nat_rec: |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
601 |
\<open>m OR n = of_bool (odd m \<or> odd n) + 2 * ((m div 2) OR (n div 2))\<close> for m n :: nat |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
602 |
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) |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
603 |
|
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
604 |
lemma xor_nat_rec: |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
605 |
\<open>m XOR n = of_bool (odd m \<noteq> odd n) + 2 * ((m div 2) XOR (n div 2))\<close> for m n :: nat |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
606 |
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) |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
607 |
|
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
608 |
lemma Suc_0_and_eq [simp]: |
| 71822 | 609 |
\<open>Suc 0 AND n = n mod 2\<close> |
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
610 |
using one_and_eq [of n] by simp |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
611 |
|
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
612 |
lemma and_Suc_0_eq [simp]: |
| 71822 | 613 |
\<open>n AND Suc 0 = n mod 2\<close> |
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
614 |
using and_one_eq [of n] by simp |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
615 |
|
| 71822 | 616 |
lemma Suc_0_or_eq: |
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
617 |
\<open>Suc 0 OR n = n + of_bool (even n)\<close> |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
618 |
using one_or_eq [of n] by simp |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
619 |
|
| 71822 | 620 |
lemma or_Suc_0_eq: |
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
621 |
\<open>n OR Suc 0 = n + of_bool (even n)\<close> |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
622 |
using or_one_eq [of n] by simp |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
623 |
|
| 71822 | 624 |
lemma Suc_0_xor_eq: |
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
625 |
\<open>Suc 0 XOR n = n + of_bool (even n) - of_bool (odd n)\<close> |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
626 |
using one_xor_eq [of n] by simp |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
627 |
|
| 71822 | 628 |
lemma xor_Suc_0_eq: |
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
629 |
\<open>n XOR Suc 0 = n + of_bool (even n) - of_bool (odd n)\<close> |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
630 |
using xor_one_eq [of n] by simp |
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
631 |
|
|
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
632 |
|
| 71956 | 633 |
subsection \<open>Instances for \<^typ>\<open>integer\<close> and \<^typ>\<open>natural\<close>\<close> |
| 71442 | 634 |
|
635 |
unbundle integer.lifting natural.lifting |
|
636 |
||
637 |
instantiation integer :: ring_bit_operations |
|
638 |
begin |
|
639 |
||
640 |
lift_definition not_integer :: \<open>integer \<Rightarrow> integer\<close> |
|
641 |
is not . |
|
642 |
||
643 |
lift_definition and_integer :: \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close> |
|
644 |
is \<open>and\<close> . |
|
645 |
||
646 |
lift_definition or_integer :: \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close> |
|
647 |
is or . |
|
648 |
||
649 |
lift_definition xor_integer :: \<open>integer \<Rightarrow> integer \<Rightarrow> integer\<close> |
|
650 |
is xor . |
|
651 |
||
652 |
instance proof |
|
653 |
fix k l :: \<open>integer\<close> and n :: nat |
|
654 |
show \<open>- k = NOT (k - 1)\<close> |
|
655 |
by transfer (simp add: minus_eq_not_minus_1) |
|
656 |
show \<open>bit (NOT k) n \<longleftrightarrow> (2 :: integer) ^ n \<noteq> 0 \<and> \<not> bit k n\<close> |
|
657 |
by transfer (fact bit_not_iff) |
|
658 |
show \<open>bit (k AND l) n \<longleftrightarrow> bit k n \<and> bit l n\<close> |
|
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
659 |
by transfer (fact bit_and_iff) |
| 71442 | 660 |
show \<open>bit (k OR l) n \<longleftrightarrow> bit k n \<or> bit l n\<close> |
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
661 |
by transfer (fact bit_or_iff) |
| 71442 | 662 |
show \<open>bit (k XOR l) n \<longleftrightarrow> bit k n \<noteq> bit l n\<close> |
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
663 |
by transfer (fact bit_xor_iff) |
| 71442 | 664 |
qed |
665 |
||
666 |
end |
|
667 |
||
668 |
instantiation natural :: semiring_bit_operations |
|
669 |
begin |
|
670 |
||
671 |
lift_definition and_natural :: \<open>natural \<Rightarrow> natural \<Rightarrow> natural\<close> |
|
672 |
is \<open>and\<close> . |
|
673 |
||
674 |
lift_definition or_natural :: \<open>natural \<Rightarrow> natural \<Rightarrow> natural\<close> |
|
675 |
is or . |
|
676 |
||
677 |
lift_definition xor_natural :: \<open>natural \<Rightarrow> natural \<Rightarrow> natural\<close> |
|
678 |
is xor . |
|
679 |
||
680 |
instance proof |
|
681 |
fix m n :: \<open>natural\<close> and q :: nat |
|
682 |
show \<open>bit (m AND n) q \<longleftrightarrow> bit m q \<and> bit n q\<close> |
|
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
683 |
by transfer (fact bit_and_iff) |
| 71442 | 684 |
show \<open>bit (m OR n) q \<longleftrightarrow> bit m q \<or> bit n q\<close> |
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
685 |
by transfer (fact bit_or_iff) |
| 71442 | 686 |
show \<open>bit (m XOR n) q \<longleftrightarrow> bit m q \<noteq> bit n q\<close> |
|
71804
6fd70ed18199
simplified construction of binary bit operations
haftmann
parents:
71802
diff
changeset
|
687 |
by transfer (fact bit_xor_iff) |
| 71442 | 688 |
qed |
689 |
||
690 |
end |
|
691 |
||
692 |
lifting_update integer.lifting |
|
693 |
lifting_forget integer.lifting |
|
694 |
||
695 |
lifting_update natural.lifting |
|
696 |
lifting_forget natural.lifting |
|
697 |
||
| 71800 | 698 |
|
699 |
subsection \<open>Key ideas of bit operations\<close> |
|
700 |
||
701 |
text \<open> |
|
702 |
When formalizing bit operations, it is tempting to represent |
|
703 |
bit values as explicit lists over a binary type. This however |
|
704 |
is a bad idea, mainly due to the inherent ambiguities in |
|
705 |
representation concerning repeating leading bits. |
|
706 |
||
707 |
Hence this approach avoids such explicit lists altogether |
|
708 |
following an algebraic path: |
|
709 |
||
710 |
\<^item> Bit values are represented by numeric types: idealized |
|
711 |
unbounded bit values can be represented by type \<^typ>\<open>int\<close>, |
|
712 |
bounded bit values by quotient types over \<^typ>\<open>int\<close>. |
|
713 |
||
714 |
\<^item> (A special case are idealized unbounded bit values ending |
|
715 |
in @{term [source] 0} which can be represented by type \<^typ>\<open>nat\<close> but
|
|
716 |
only support a restricted set of operations). |
|
717 |
||
718 |
\<^item> From this idea follows that |
|
719 |
||
720 |
\<^item> multiplication by \<^term>\<open>2 :: int\<close> is a bit shift to the left and |
|
721 |
||
722 |
\<^item> division by \<^term>\<open>2 :: int\<close> is a bit shift to the right. |
|
723 |
||
724 |
\<^item> Concerning bounded bit values, iterated shifts to the left |
|
725 |
may result in eliminating all bits by shifting them all |
|
726 |
beyond the boundary. The property \<^prop>\<open>(2 :: int) ^ n \<noteq> 0\<close> |
|
727 |
represents that \<^term>\<open>n\<close> is \<^emph>\<open>not\<close> beyond that boundary. |
|
728 |
||
|
71965
d45f5d4c41bd
more class operations for the sake of efficient generated code
haftmann
parents:
71956
diff
changeset
|
729 |
\<^item> The projection on a single bit is then @{thm bit_iff_odd [where ?'a = int, no_vars]}.
|
| 71800 | 730 |
|
731 |
\<^item> This leads to the most fundamental properties of bit values: |
|
732 |
||
733 |
\<^item> Equality rule: @{thm bit_eqI [where ?'a = int, no_vars]}
|
|
734 |
||
735 |
\<^item> Induction rule: @{thm bits_induct [where ?'a = int, no_vars]}
|
|
736 |
||
737 |
\<^item> Typical operations are characterized as follows: |
|
738 |
||
739 |
\<^item> Singleton \<^term>\<open>n\<close>th bit: \<^term>\<open>(2 :: int) ^ n\<close> |
|
740 |
||
| 71956 | 741 |
\<^item> Bit mask upto bit \<^term>\<open>n\<close>: @{thm mask_eq_exp_minus_1 [where ?'a = int, no_vars]}
|
| 71800 | 742 |
|
743 |
\<^item> Left shift: @{thm push_bit_eq_mult [where ?'a = int, no_vars]}
|
|
744 |
||
745 |
\<^item> Right shift: @{thm drop_bit_eq_div [where ?'a = int, no_vars]}
|
|
746 |
||
747 |
\<^item> Truncation: @{thm take_bit_eq_mod [where ?'a = int, no_vars]}
|
|
748 |
||
749 |
\<^item> Negation: @{thm bit_not_iff [where ?'a = int, no_vars]}
|
|
750 |
||
751 |
\<^item> And: @{thm bit_and_iff [where ?'a = int, no_vars]}
|
|
752 |
||
753 |
\<^item> Or: @{thm bit_or_iff [where ?'a = int, no_vars]}
|
|
754 |
||
755 |
\<^item> Xor: @{thm bit_xor_iff [where ?'a = int, no_vars]}
|
|
756 |
||
757 |
\<^item> Set a single bit: @{thm set_bit_def [where ?'a = int, no_vars]}
|
|
758 |
||
759 |
\<^item> Unset a single bit: @{thm unset_bit_def [where ?'a = int, no_vars]}
|
|
760 |
||
761 |
\<^item> Flip a single bit: @{thm flip_bit_def [where ?'a = int, no_vars]}
|
|
762 |
\<close> |
|
763 |
||
| 71442 | 764 |
end |