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