| author | haftmann | 
| Thu, 16 Oct 2025 11:03:48 +0200 | |
| changeset 83274 | 0bd014c32479 | 
| parent 82687 | 97b9ac57751e | 
| child 83279 | 28541336560a | 
| permissions | -rw-r--r-- | 
| 
72515
 
c7038c397ae3
moved most material from session HOL-Word to Word_Lib in the AFP
 
haftmann 
parents: 
72512 
diff
changeset
 | 
1  | 
(* Title: HOL/Library/Word.thy  | 
| 
 
c7038c397ae3
moved most material from session HOL-Word to Word_Lib in the AFP
 
haftmann 
parents: 
72512 
diff
changeset
 | 
2  | 
Author: Jeremy Dawson and Gerwin Klein, NICTA, et. al.  | 
| 24333 | 3  | 
*)  | 
4  | 
||
| 61799 | 5  | 
section \<open>A type of finite bit strings\<close>  | 
| 24350 | 6  | 
|
| 29628 | 7  | 
theory Word  | 
| 
41413
 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 
wenzelm 
parents: 
41060 
diff
changeset
 | 
8  | 
imports  | 
| 
66453
 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 
wenzelm 
parents: 
65363 
diff
changeset
 | 
9  | 
"HOL-Library.Type_Length"  | 
| 37660 | 10  | 
begin  | 
11  | 
||
| 72243 | 12  | 
subsection \<open>Preliminaries\<close>  | 
13  | 
||
14  | 
lemma signed_take_bit_decr_length_iff:  | 
|
15  | 
  \<open>signed_take_bit (LENGTH('a::len) - Suc 0) k = signed_take_bit (LENGTH('a) - Suc 0) l
 | 
|
16  | 
    \<longleftrightarrow> take_bit LENGTH('a) k = take_bit LENGTH('a) l\<close>
 | 
|
| 80777 | 17  | 
by (simp add: signed_take_bit_eq_iff_take_bit_eq)  | 
| 72243 | 18  | 
|
19  | 
||
| 72244 | 20  | 
subsection \<open>Fundamentals\<close>  | 
21  | 
||
22  | 
subsubsection \<open>Type definition\<close>  | 
|
| 37660 | 23  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
24  | 
quotient_type (overloaded) 'a word = int / \<open>\<lambda>k l. take_bit LENGTH('a) k = take_bit LENGTH('a::len) l\<close>
 | 
| 72243 | 25  | 
morphisms rep Word by (auto intro!: equivpI reflpI sympI transpI)  | 
26  | 
||
27  | 
hide_const (open) rep \<comment> \<open>only for foundational purpose\<close>  | 
|
| 
72130
 
9e5862223442
dedicated symbols for code generation, to pave way for generic conversions from and to word
 
haftmann 
parents: 
72128 
diff
changeset
 | 
28  | 
hide_const (open) Word \<comment> \<open>only for code generation\<close>  | 
| 
 
9e5862223442
dedicated symbols for code generation, to pave way for generic conversions from and to word
 
haftmann 
parents: 
72128 
diff
changeset
 | 
29  | 
|
| 72244 | 30  | 
|
31  | 
subsubsection \<open>Basic arithmetic\<close>  | 
|
32  | 
||
| 72243 | 33  | 
instantiation word :: (len) comm_ring_1  | 
34  | 
begin  | 
|
35  | 
||
36  | 
lift_definition zero_word :: \<open>'a word\<close>  | 
|
37  | 
is 0 .  | 
|
38  | 
||
39  | 
lift_definition one_word :: \<open>'a word\<close>  | 
|
40  | 
is 1 .  | 
|
41  | 
||
42  | 
lift_definition plus_word :: \<open>'a word \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>  | 
|
43  | 
is \<open>(+)\<close>  | 
|
| 80777 | 44  | 
by (auto simp: take_bit_eq_mod intro: mod_add_cong)  | 
| 72243 | 45  | 
|
46  | 
lift_definition minus_word :: \<open>'a word \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>  | 
|
47  | 
is \<open>(-)\<close>  | 
|
| 80777 | 48  | 
by (auto simp: take_bit_eq_mod intro: mod_diff_cong)  | 
| 72243 | 49  | 
|
50  | 
lift_definition uminus_word :: \<open>'a word \<Rightarrow> 'a word\<close>  | 
|
51  | 
is uminus  | 
|
| 80777 | 52  | 
by (auto simp: take_bit_eq_mod intro: mod_minus_cong)  | 
| 72243 | 53  | 
|
54  | 
lift_definition times_word :: \<open>'a word \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>  | 
|
55  | 
is \<open>(*)\<close>  | 
|
| 80777 | 56  | 
by (auto simp: take_bit_eq_mod intro: mod_mult_cong)  | 
| 72243 | 57  | 
|
58  | 
instance  | 
|
59  | 
by (standard; transfer) (simp_all add: algebra_simps)  | 
|
60  | 
||
61  | 
end  | 
|
62  | 
||
63  | 
context  | 
|
64  | 
includes lifting_syntax  | 
|
| 72244 | 65  | 
notes  | 
66  | 
power_transfer [transfer_rule]  | 
|
67  | 
transfer_rule_of_bool [transfer_rule]  | 
|
68  | 
transfer_rule_numeral [transfer_rule]  | 
|
69  | 
transfer_rule_of_nat [transfer_rule]  | 
|
70  | 
transfer_rule_of_int [transfer_rule]  | 
|
| 72243 | 71  | 
begin  | 
72  | 
||
73  | 
lemma power_transfer_word [transfer_rule]:  | 
|
74  | 
\<open>(pcr_word ===> (=) ===> pcr_word) (^) (^)\<close>  | 
|
75  | 
by transfer_prover  | 
|
76  | 
||
| 72244 | 77  | 
lemma [transfer_rule]:  | 
78  | 
\<open>((=) ===> pcr_word) of_bool of_bool\<close>  | 
|
79  | 
by transfer_prover  | 
|
80  | 
||
81  | 
lemma [transfer_rule]:  | 
|
82  | 
\<open>((=) ===> pcr_word) numeral numeral\<close>  | 
|
83  | 
by transfer_prover  | 
|
84  | 
||
85  | 
lemma [transfer_rule]:  | 
|
86  | 
\<open>((=) ===> pcr_word) int of_nat\<close>  | 
|
87  | 
by transfer_prover  | 
|
88  | 
||
89  | 
lemma [transfer_rule]:  | 
|
90  | 
\<open>((=) ===> pcr_word) (\<lambda>k. k) of_int\<close>  | 
|
91  | 
proof -  | 
|
92  | 
have \<open>((=) ===> pcr_word) of_int of_int\<close>  | 
|
93  | 
by transfer_prover  | 
|
94  | 
then show ?thesis by (simp add: id_def)  | 
|
95  | 
qed  | 
|
96  | 
||
97  | 
lemma [transfer_rule]:  | 
|
98  | 
\<open>(pcr_word ===> (\<longleftrightarrow>)) even ((dvd) 2 :: 'a::len word \<Rightarrow> bool)\<close>  | 
|
99  | 
proof -  | 
|
100  | 
  have even_word_unfold: "even k \<longleftrightarrow> (\<exists>l. take_bit LENGTH('a) k = take_bit LENGTH('a) (2 * l))" (is "?P \<longleftrightarrow> ?Q")
 | 
|
101  | 
for k :: int  | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
102  | 
by (metis dvd_triv_left evenE even_take_bit_eq len_not_eq_0)  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
103  | 
show ?thesis  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
104  | 
unfolding even_word_unfold [abs_def] dvd_def [where ?'a = "'a word", abs_def]  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
105  | 
by transfer_prover  | 
| 72244 | 106  | 
qed  | 
107  | 
||
| 72243 | 108  | 
end  | 
109  | 
||
| 72512 | 110  | 
lemma exp_eq_zero_iff [simp]:  | 
111  | 
  \<open>2 ^ n = (0 :: 'a::len word) \<longleftrightarrow> n \<ge> LENGTH('a)\<close>
 | 
|
| 73535 | 112  | 
by transfer auto  | 
| 72512 | 113  | 
|
| 72244 | 114  | 
lemma word_exp_length_eq_0 [simp]:  | 
115  | 
  \<open>(2 :: 'a::len word) ^ LENGTH('a) = 0\<close>
 | 
|
| 72512 | 116  | 
by simp  | 
| 72262 | 117  | 
|
| 72244 | 118  | 
|
| 72489 | 119  | 
subsubsection \<open>Basic tool setup\<close>  | 
120  | 
||
121  | 
ML_file \<open>Tools/word_lib.ML\<close>  | 
|
122  | 
||
123  | 
||
| 72244 | 124  | 
subsubsection \<open>Basic code generation setup\<close>  | 
| 
71948
 
6ede899d26d3
fundamental construction of word type following existing transfer rules
 
haftmann 
parents: 
71947 
diff
changeset
 | 
125  | 
|
| 72262 | 126  | 
context  | 
127  | 
begin  | 
|
128  | 
||
129  | 
qualified lift_definition the_int :: \<open>'a::len word \<Rightarrow> int\<close>  | 
|
| 
71948
 
6ede899d26d3
fundamental construction of word type following existing transfer rules
 
haftmann 
parents: 
71947 
diff
changeset
 | 
130  | 
  is \<open>take_bit LENGTH('a)\<close> .
 | 
| 37660 | 131  | 
|
| 72262 | 132  | 
end  | 
133  | 
||
| 72243 | 134  | 
lemma [code abstype]:  | 
| 72262 | 135  | 
\<open>Word.Word (Word.the_int w) = w\<close>  | 
| 72243 | 136  | 
by transfer simp  | 
137  | 
||
| 72262 | 138  | 
lemma Word_eq_word_of_int [code_post, simp]:  | 
139  | 
\<open>Word.Word = of_int\<close>  | 
|
140  | 
by (rule; transfer) simp  | 
|
141  | 
||
| 72243 | 142  | 
quickcheck_generator word  | 
143  | 
constructors:  | 
|
144  | 
\<open>0 :: 'a::len word\<close>,  | 
|
145  | 
\<open>numeral :: num \<Rightarrow> 'a::len word\<close>  | 
|
146  | 
||
147  | 
instantiation word :: (len) equal  | 
|
148  | 
begin  | 
|
149  | 
||
150  | 
lift_definition equal_word :: \<open>'a word \<Rightarrow> 'a word \<Rightarrow> bool\<close>  | 
|
151  | 
  is \<open>\<lambda>k l. take_bit LENGTH('a) k = take_bit LENGTH('a) l\<close>
 | 
|
152  | 
by simp  | 
|
153  | 
||
154  | 
instance  | 
|
155  | 
by (standard; transfer) rule  | 
|
156  | 
||
157  | 
end  | 
|
158  | 
||
159  | 
lemma [code]:  | 
|
| 72262 | 160  | 
\<open>HOL.equal v w \<longleftrightarrow> HOL.equal (Word.the_int v) (Word.the_int w)\<close>  | 
| 72243 | 161  | 
by transfer (simp add: equal)  | 
162  | 
||
163  | 
lemma [code]:  | 
|
| 72262 | 164  | 
\<open>Word.the_int 0 = 0\<close>  | 
| 72243 | 165  | 
by transfer simp  | 
166  | 
||
167  | 
lemma [code]:  | 
|
| 72262 | 168  | 
\<open>Word.the_int 1 = 1\<close>  | 
| 72243 | 169  | 
by transfer simp  | 
170  | 
||
171  | 
lemma [code]:  | 
|
| 72262 | 172  | 
  \<open>Word.the_int (v + w) = take_bit LENGTH('a) (Word.the_int v + Word.the_int w)\<close>
 | 
| 72243 | 173  | 
for v w :: \<open>'a::len word\<close>  | 
174  | 
by transfer (simp add: take_bit_add)  | 
|
175  | 
||
176  | 
lemma [code]:  | 
|
| 72262 | 177  | 
  \<open>Word.the_int (- w) = (let k = Word.the_int w in if w = 0 then 0 else 2 ^ LENGTH('a) - k)\<close>
 | 
| 72243 | 178  | 
for w :: \<open>'a::len word\<close>  | 
| 80777 | 179  | 
by transfer (auto simp: take_bit_eq_mod zmod_zminus1_eq_if)  | 
| 72243 | 180  | 
|
181  | 
lemma [code]:  | 
|
| 72262 | 182  | 
  \<open>Word.the_int (v - w) = take_bit LENGTH('a) (Word.the_int v - Word.the_int w)\<close>
 | 
| 72243 | 183  | 
for v w :: \<open>'a::len word\<close>  | 
184  | 
by transfer (simp add: take_bit_diff)  | 
|
185  | 
||
186  | 
lemma [code]:  | 
|
| 72262 | 187  | 
  \<open>Word.the_int (v * w) = take_bit LENGTH('a) (Word.the_int v * Word.the_int w)\<close>
 | 
| 72243 | 188  | 
for v w :: \<open>'a::len word\<close>  | 
189  | 
by transfer (simp add: take_bit_mult)  | 
|
190  | 
||
191  | 
||
| 72244 | 192  | 
subsubsection \<open>Basic conversions\<close>  | 
| 70185 | 193  | 
|
| 72262 | 194  | 
abbreviation word_of_nat :: \<open>nat \<Rightarrow> 'a::len word\<close>  | 
195  | 
where \<open>word_of_nat \<equiv> of_nat\<close>  | 
|
196  | 
||
197  | 
abbreviation word_of_int :: \<open>int \<Rightarrow> 'a::len word\<close>  | 
|
198  | 
where \<open>word_of_int \<equiv> of_int\<close>  | 
|
199  | 
||
200  | 
lemma word_of_nat_eq_iff:  | 
|
201  | 
  \<open>word_of_nat m = (word_of_nat n :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) m = take_bit LENGTH('a) n\<close>
 | 
|
202  | 
by transfer (simp add: take_bit_of_nat)  | 
|
203  | 
||
204  | 
lemma word_of_int_eq_iff:  | 
|
205  | 
  \<open>word_of_int k = (word_of_int l :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) k = take_bit LENGTH('a) l\<close>
 | 
|
206  | 
by transfer rule  | 
|
207  | 
||
| 74496 | 208  | 
lemma word_of_nat_eq_0_iff:  | 
| 72262 | 209  | 
  \<open>word_of_nat n = (0 :: 'a::len word) \<longleftrightarrow> 2 ^ LENGTH('a) dvd n\<close>
 | 
210  | 
using word_of_nat_eq_iff [where ?'a = 'a, of n 0] by (simp add: take_bit_eq_0_iff)  | 
|
211  | 
||
| 74496 | 212  | 
lemma word_of_int_eq_0_iff:  | 
| 72262 | 213  | 
  \<open>word_of_int k = (0 :: 'a::len word) \<longleftrightarrow> 2 ^ LENGTH('a) dvd k\<close>
 | 
214  | 
using word_of_int_eq_iff [where ?'a = 'a, of k 0] by (simp add: take_bit_eq_0_iff)  | 
|
215  | 
||
216  | 
context semiring_1  | 
|
217  | 
begin  | 
|
218  | 
||
219  | 
lift_definition unsigned :: \<open>'b::len word \<Rightarrow> 'a\<close>  | 
|
220  | 
  is \<open>of_nat \<circ> nat \<circ> take_bit LENGTH('b)\<close>
 | 
|
| 72244 | 221  | 
by simp  | 
| 
55816
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55415 
diff
changeset
 | 
222  | 
|
| 72262 | 223  | 
lemma unsigned_0 [simp]:  | 
224  | 
\<open>unsigned 0 = 0\<close>  | 
|
225  | 
by transfer simp  | 
|
226  | 
||
227  | 
lemma unsigned_1 [simp]:  | 
|
228  | 
\<open>unsigned 1 = 1\<close>  | 
|
229  | 
by transfer simp  | 
|
230  | 
||
231  | 
lemma unsigned_numeral [simp]:  | 
|
232  | 
  \<open>unsigned (numeral n :: 'b::len word) = of_nat (take_bit LENGTH('b) (numeral n))\<close>
 | 
|
233  | 
by transfer (simp add: nat_take_bit_eq)  | 
|
234  | 
||
235  | 
lemma unsigned_neg_numeral [simp]:  | 
|
236  | 
  \<open>unsigned (- numeral n :: 'b::len word) = of_nat (nat (take_bit LENGTH('b) (- numeral n)))\<close>
 | 
|
237  | 
by transfer simp  | 
|
238  | 
||
239  | 
end  | 
|
240  | 
||
241  | 
context semiring_1  | 
|
242  | 
begin  | 
|
243  | 
||
| 74496 | 244  | 
lemma unsigned_of_nat:  | 
| 72262 | 245  | 
  \<open>unsigned (word_of_nat n :: 'b::len word) = of_nat (take_bit LENGTH('b) n)\<close>
 | 
246  | 
by transfer (simp add: nat_eq_iff take_bit_of_nat)  | 
|
247  | 
||
| 74496 | 248  | 
lemma unsigned_of_int:  | 
| 72262 | 249  | 
  \<open>unsigned (word_of_int k :: 'b::len word) = of_nat (nat (take_bit LENGTH('b) k))\<close>
 | 
250  | 
by transfer simp  | 
|
251  | 
||
252  | 
end  | 
|
253  | 
||
254  | 
context semiring_char_0  | 
|
255  | 
begin  | 
|
256  | 
||
257  | 
lemma unsigned_word_eqI:  | 
|
258  | 
\<open>v = w\<close> if \<open>unsigned v = unsigned w\<close>  | 
|
259  | 
using that by transfer (simp add: eq_nat_nat_iff)  | 
|
260  | 
||
261  | 
lemma word_eq_iff_unsigned:  | 
|
262  | 
\<open>v = w \<longleftrightarrow> unsigned v = unsigned w\<close>  | 
|
263  | 
by (auto intro: unsigned_word_eqI)  | 
|
264  | 
||
| 72292 | 265  | 
lemma inj_unsigned [simp]:  | 
266  | 
\<open>inj unsigned\<close>  | 
|
267  | 
by (rule injI) (simp add: unsigned_word_eqI)  | 
|
268  | 
||
269  | 
lemma unsigned_eq_0_iff:  | 
|
| 
72281
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
270  | 
\<open>unsigned w = 0 \<longleftrightarrow> w = 0\<close>  | 
| 
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
271  | 
using word_eq_iff_unsigned [of w 0] by simp  | 
| 
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
272  | 
|
| 72262 | 273  | 
end  | 
274  | 
||
275  | 
context ring_1  | 
|
276  | 
begin  | 
|
277  | 
||
278  | 
lift_definition signed :: \<open>'b::len word \<Rightarrow> 'a\<close>  | 
|
279  | 
  is \<open>of_int \<circ> signed_take_bit (LENGTH('b) - Suc 0)\<close>
 | 
|
280  | 
by (simp flip: signed_take_bit_decr_length_iff)  | 
|
281  | 
||
282  | 
lemma signed_0 [simp]:  | 
|
283  | 
\<open>signed 0 = 0\<close>  | 
|
284  | 
by transfer simp  | 
|
285  | 
||
286  | 
lemma signed_1 [simp]:  | 
|
287  | 
  \<open>signed (1 :: 'b::len word) = (if LENGTH('b) = 1 then - 1 else 1)\<close>
 | 
|
| 72488 | 288  | 
  by (transfer fixing: uminus; cases \<open>LENGTH('b)\<close>) (auto dest: gr0_implies_Suc)
 | 
| 72262 | 289  | 
|
290  | 
lemma signed_minus_1 [simp]:  | 
|
291  | 
\<open>signed (- 1 :: 'b::len word) = - 1\<close>  | 
|
292  | 
by (transfer fixing: uminus) simp  | 
|
293  | 
||
294  | 
lemma signed_numeral [simp]:  | 
|
295  | 
  \<open>signed (numeral n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - 1) (numeral n))\<close>
 | 
|
296  | 
by transfer simp  | 
|
297  | 
||
298  | 
lemma signed_neg_numeral [simp]:  | 
|
299  | 
  \<open>signed (- numeral n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - 1) (- numeral n))\<close>
 | 
|
300  | 
by transfer simp  | 
|
301  | 
||
| 74496 | 302  | 
lemma signed_of_nat:  | 
| 72262 | 303  | 
  \<open>signed (word_of_nat n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - Suc 0) (int n))\<close>
 | 
304  | 
by transfer simp  | 
|
305  | 
||
| 74496 | 306  | 
lemma signed_of_int:  | 
| 72262 | 307  | 
  \<open>signed (word_of_int n :: 'b::len word) = of_int (signed_take_bit (LENGTH('b) - Suc 0) n)\<close>
 | 
308  | 
by transfer simp  | 
|
309  | 
||
310  | 
end  | 
|
311  | 
||
312  | 
context ring_char_0  | 
|
313  | 
begin  | 
|
314  | 
||
315  | 
lemma signed_word_eqI:  | 
|
316  | 
\<open>v = w\<close> if \<open>signed v = signed w\<close>  | 
|
317  | 
using that by transfer (simp flip: signed_take_bit_decr_length_iff)  | 
|
318  | 
||
319  | 
lemma word_eq_iff_signed:  | 
|
320  | 
\<open>v = w \<longleftrightarrow> signed v = signed w\<close>  | 
|
321  | 
by (auto intro: signed_word_eqI)  | 
|
322  | 
||
| 72292 | 323  | 
lemma inj_signed [simp]:  | 
324  | 
\<open>inj signed\<close>  | 
|
325  | 
by (rule injI) (simp add: signed_word_eqI)  | 
|
326  | 
||
| 
72281
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
327  | 
lemma signed_eq_0_iff:  | 
| 
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
328  | 
\<open>signed w = 0 \<longleftrightarrow> w = 0\<close>  | 
| 
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
329  | 
using word_eq_iff_signed [of w 0] by simp  | 
| 
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
330  | 
|
| 72262 | 331  | 
end  | 
332  | 
||
333  | 
abbreviation unat :: \<open>'a::len word \<Rightarrow> nat\<close>  | 
|
334  | 
where \<open>unat \<equiv> unsigned\<close>  | 
|
335  | 
||
336  | 
abbreviation uint :: \<open>'a::len word \<Rightarrow> int\<close>  | 
|
337  | 
where \<open>uint \<equiv> unsigned\<close>  | 
|
338  | 
||
339  | 
abbreviation sint :: \<open>'a::len word \<Rightarrow> int\<close>  | 
|
340  | 
where \<open>sint \<equiv> signed\<close>  | 
|
341  | 
||
342  | 
abbreviation ucast :: \<open>'a::len word \<Rightarrow> 'b::len word\<close>  | 
|
343  | 
where \<open>ucast \<equiv> unsigned\<close>  | 
|
344  | 
||
345  | 
abbreviation scast :: \<open>'a::len word \<Rightarrow> 'b::len word\<close>  | 
|
346  | 
where \<open>scast \<equiv> signed\<close>  | 
|
347  | 
||
348  | 
context  | 
|
349  | 
includes lifting_syntax  | 
|
350  | 
begin  | 
|
351  | 
||
352  | 
lemma [transfer_rule]:  | 
|
353  | 
  \<open>(pcr_word ===> (=)) (nat \<circ> take_bit LENGTH('a)) (unat :: 'a::len word \<Rightarrow> nat)\<close>
 | 
|
354  | 
using unsigned.transfer [where ?'a = nat] by simp  | 
|
355  | 
||
356  | 
lemma [transfer_rule]:  | 
|
357  | 
  \<open>(pcr_word ===> (=)) (take_bit LENGTH('a)) (uint :: 'a::len word \<Rightarrow> int)\<close>
 | 
|
358  | 
using unsigned.transfer [where ?'a = int] by (simp add: comp_def)  | 
|
359  | 
||
360  | 
lemma [transfer_rule]:  | 
|
361  | 
  \<open>(pcr_word ===> (=)) (signed_take_bit (LENGTH('a) - Suc 0)) (sint :: 'a::len word \<Rightarrow> int)\<close>
 | 
|
362  | 
using signed.transfer [where ?'a = int] by simp  | 
|
363  | 
||
364  | 
lemma [transfer_rule]:  | 
|
365  | 
  \<open>(pcr_word ===> pcr_word) (take_bit LENGTH('a)) (ucast :: 'a::len word \<Rightarrow> 'b::len word)\<close>
 | 
|
366  | 
proof (rule rel_funI)  | 
|
367  | 
fix k :: int and w :: \<open>'a word\<close>  | 
|
368  | 
assume \<open>pcr_word k w\<close>  | 
|
369  | 
then have \<open>w = word_of_int k\<close>  | 
|
370  | 
by (simp add: pcr_word_def cr_word_def relcompp_apply)  | 
|
371  | 
  moreover have \<open>pcr_word (take_bit LENGTH('a) k) (ucast (word_of_int k :: 'a word))\<close>
 | 
|
372  | 
by transfer (simp add: pcr_word_def cr_word_def relcompp_apply)  | 
|
373  | 
  ultimately show \<open>pcr_word (take_bit LENGTH('a) k) (ucast w)\<close>
 | 
|
374  | 
by simp  | 
|
375  | 
qed  | 
|
376  | 
||
377  | 
lemma [transfer_rule]:  | 
|
378  | 
  \<open>(pcr_word ===> pcr_word) (signed_take_bit (LENGTH('a) - Suc 0)) (scast :: 'a::len word \<Rightarrow> 'b::len word)\<close>
 | 
|
379  | 
proof (rule rel_funI)  | 
|
380  | 
fix k :: int and w :: \<open>'a word\<close>  | 
|
381  | 
assume \<open>pcr_word k w\<close>  | 
|
382  | 
then have \<open>w = word_of_int k\<close>  | 
|
383  | 
by (simp add: pcr_word_def cr_word_def relcompp_apply)  | 
|
384  | 
  moreover have \<open>pcr_word (signed_take_bit (LENGTH('a) - Suc 0) k) (scast (word_of_int k :: 'a word))\<close>
 | 
|
385  | 
by transfer (simp add: pcr_word_def cr_word_def relcompp_apply)  | 
|
386  | 
  ultimately show \<open>pcr_word (signed_take_bit (LENGTH('a) - Suc 0) k) (scast w)\<close>
 | 
|
387  | 
by simp  | 
|
388  | 
qed  | 
|
389  | 
||
390  | 
end  | 
|
391  | 
||
392  | 
lemma of_nat_unat [simp]:  | 
|
393  | 
\<open>of_nat (unat w) = unsigned w\<close>  | 
|
394  | 
by transfer simp  | 
|
395  | 
||
396  | 
lemma of_int_uint [simp]:  | 
|
397  | 
\<open>of_int (uint w) = unsigned w\<close>  | 
|
398  | 
by transfer simp  | 
|
399  | 
||
400  | 
lemma of_int_sint [simp]:  | 
|
401  | 
\<open>of_int (sint a) = signed a\<close>  | 
|
402  | 
by transfer (simp_all add: take_bit_signed_take_bit)  | 
|
| 72079 | 403  | 
|
404  | 
lemma nat_uint_eq [simp]:  | 
|
405  | 
\<open>nat (uint w) = unat w\<close>  | 
|
406  | 
by transfer simp  | 
|
407  | 
||
| 81609 | 408  | 
lemma nat_of_natural_unsigned_eq [simp]:  | 
409  | 
\<open>nat_of_natural (unsigned w) = unat w\<close>  | 
|
410  | 
by transfer simp  | 
|
411  | 
||
412  | 
lemma int_of_integer_unsigned_eq [simp]:  | 
|
413  | 
\<open>int_of_integer (unsigned w) = uint w\<close>  | 
|
414  | 
by transfer simp  | 
|
415  | 
||
416  | 
lemma int_of_integer_signed_eq [simp]:  | 
|
417  | 
\<open>int_of_integer (signed w) = sint w\<close>  | 
|
418  | 
by transfer simp  | 
|
419  | 
||
| 
72281
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
420  | 
lemma sgn_uint_eq [simp]:  | 
| 
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
421  | 
\<open>sgn (uint w) = of_bool (w \<noteq> 0)\<close>  | 
| 
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
422  | 
by transfer (simp add: less_le)  | 
| 
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
423  | 
|
| 72262 | 424  | 
text \<open>Aliasses only for code generation\<close>  | 
425  | 
||
426  | 
context  | 
|
427  | 
begin  | 
|
428  | 
||
429  | 
qualified lift_definition of_int :: \<open>int \<Rightarrow> 'a::len word\<close>  | 
|
430  | 
  is \<open>take_bit LENGTH('a)\<close> .
 | 
|
431  | 
||
432  | 
qualified lift_definition of_nat :: \<open>nat \<Rightarrow> 'a::len word\<close>  | 
|
433  | 
  is \<open>int \<circ> take_bit LENGTH('a)\<close> .
 | 
|
434  | 
||
435  | 
qualified lift_definition the_nat :: \<open>'a::len word \<Rightarrow> nat\<close>  | 
|
436  | 
  is \<open>nat \<circ> take_bit LENGTH('a)\<close> by simp
 | 
|
437  | 
||
438  | 
qualified lift_definition the_signed_int :: \<open>'a::len word \<Rightarrow> int\<close>  | 
|
439  | 
  is \<open>signed_take_bit (LENGTH('a) - Suc 0)\<close> by (simp add: signed_take_bit_decr_length_iff)
 | 
|
440  | 
||
441  | 
qualified lift_definition cast :: \<open>'a::len word \<Rightarrow> 'b::len word\<close>  | 
|
442  | 
  is \<open>take_bit LENGTH('a)\<close> by simp
 | 
|
443  | 
||
444  | 
qualified lift_definition signed_cast :: \<open>'a::len word \<Rightarrow> 'b::len word\<close>  | 
|
445  | 
  is \<open>signed_take_bit (LENGTH('a) - Suc 0)\<close> by (metis signed_take_bit_decr_length_iff)
 | 
|
446  | 
||
447  | 
end  | 
|
448  | 
||
449  | 
lemma [code_abbrev, simp]:  | 
|
450  | 
\<open>Word.the_int = uint\<close>  | 
|
451  | 
by transfer rule  | 
|
452  | 
||
453  | 
lemma [code]:  | 
|
454  | 
  \<open>Word.the_int (Word.of_int k :: 'a::len word) = take_bit LENGTH('a) k\<close>
 | 
|
455  | 
by transfer simp  | 
|
456  | 
||
457  | 
lemma [code_abbrev, simp]:  | 
|
458  | 
\<open>Word.of_int = word_of_int\<close>  | 
|
459  | 
by (rule; transfer) simp  | 
|
460  | 
||
461  | 
lemma [code]:  | 
|
462  | 
  \<open>Word.the_int (Word.of_nat n :: 'a::len word) = take_bit LENGTH('a) (int n)\<close>
 | 
|
| 72244 | 463  | 
by transfer (simp add: take_bit_of_nat)  | 
464  | 
||
| 72262 | 465  | 
lemma [code_abbrev, simp]:  | 
466  | 
\<open>Word.of_nat = word_of_nat\<close>  | 
|
467  | 
by (rule; transfer) (simp add: take_bit_of_nat)  | 
|
468  | 
||
469  | 
lemma [code]:  | 
|
470  | 
\<open>Word.the_nat w = nat (Word.the_int w)\<close>  | 
|
471  | 
by transfer simp  | 
|
472  | 
||
473  | 
lemma [code_abbrev, simp]:  | 
|
474  | 
\<open>Word.the_nat = unat\<close>  | 
|
475  | 
by (rule; transfer) simp  | 
|
476  | 
||
477  | 
lemma [code]:  | 
|
| 82367 | 478  | 
\<open>Word.the_signed_int w = (let k = Word.the_int w  | 
479  | 
    in if bit k (LENGTH('a) - Suc 0) then k + push_bit LENGTH('a) (- 1) else k)\<close>
 | 
|
| 72262 | 480  | 
for w :: \<open>'a::len word\<close>  | 
| 82367 | 481  | 
by transfer (simp add: bit_simps signed_take_bit_eq_take_bit_add)  | 
| 72262 | 482  | 
|
483  | 
lemma [code_abbrev, simp]:  | 
|
484  | 
\<open>Word.the_signed_int = sint\<close>  | 
|
485  | 
by (rule; transfer) simp  | 
|
486  | 
||
487  | 
lemma [code]:  | 
|
488  | 
  \<open>Word.the_int (Word.cast w :: 'b::len word) = take_bit LENGTH('b) (Word.the_int w)\<close>
 | 
|
489  | 
for w :: \<open>'a::len word\<close>  | 
|
490  | 
by transfer simp  | 
|
491  | 
||
492  | 
lemma [code_abbrev, simp]:  | 
|
493  | 
\<open>Word.cast = ucast\<close>  | 
|
494  | 
by (rule; transfer) simp  | 
|
495  | 
||
496  | 
lemma [code]:  | 
|
497  | 
  \<open>Word.the_int (Word.signed_cast w :: 'b::len word) = take_bit LENGTH('b) (Word.the_signed_int w)\<close>
 | 
|
498  | 
for w :: \<open>'a::len word\<close>  | 
|
499  | 
by transfer simp  | 
|
500  | 
||
501  | 
lemma [code_abbrev, simp]:  | 
|
502  | 
\<open>Word.signed_cast = scast\<close>  | 
|
503  | 
by (rule; transfer) simp  | 
|
504  | 
||
505  | 
lemma [code]:  | 
|
506  | 
\<open>unsigned w = of_nat (nat (Word.the_int w))\<close>  | 
|
507  | 
by transfer simp  | 
|
508  | 
||
509  | 
lemma [code]:  | 
|
510  | 
\<open>signed w = of_int (Word.the_signed_int w)\<close>  | 
|
511  | 
by transfer simp  | 
|
| 72244 | 512  | 
|
513  | 
||
| 82526 | 514  | 
subsection \<open>Elementary case distinctions\<close>  | 
515  | 
||
516  | 
lemma word_length_one [case_names zero minus_one length_beyond]:  | 
|
517  | 
fixes w :: \<open>'a::len word\<close>  | 
|
518  | 
  obtains (zero) \<open>LENGTH('a) = Suc 0\<close> \<open>w = 0\<close>
 | 
|
519  | 
  | (minus_one) \<open>LENGTH('a) = Suc 0\<close> \<open>w = - 1\<close>
 | 
|
520  | 
  | (length_beyond) \<open>2 \<le> LENGTH('a)\<close>
 | 
|
521  | 
proof (cases \<open>2 \<le> LENGTH('a)\<close>)
 | 
|
522  | 
case True  | 
|
523  | 
with length_beyond show ?thesis .  | 
|
524  | 
next  | 
|
525  | 
case False  | 
|
526  | 
  then have \<open>LENGTH('a) = Suc 0\<close>
 | 
|
527  | 
by simp  | 
|
528  | 
then have \<open>w = 0 \<or> w = - 1\<close>  | 
|
529  | 
by transfer auto  | 
|
530  | 
  with \<open>LENGTH('a) = Suc 0\<close> zero minus_one show ?thesis
 | 
|
531  | 
by blast  | 
|
532  | 
qed  | 
|
533  | 
||
534  | 
||
| 72244 | 535  | 
subsubsection \<open>Basic ordering\<close>  | 
| 45547 | 536  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
537  | 
instantiation word :: (len) linorder  | 
| 45547 | 538  | 
begin  | 
539  | 
||
| 71950 | 540  | 
lift_definition less_eq_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool"  | 
541  | 
  is "\<lambda>a b. take_bit LENGTH('a) a \<le> take_bit LENGTH('a) b"
 | 
|
542  | 
by simp  | 
|
543  | 
||
544  | 
lift_definition less_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool"  | 
|
545  | 
  is "\<lambda>a b. take_bit LENGTH('a) a < take_bit LENGTH('a) b"
 | 
|
546  | 
by simp  | 
|
| 37660 | 547  | 
|
| 45547 | 548  | 
instance  | 
| 71950 | 549  | 
by (standard; transfer) auto  | 
| 45547 | 550  | 
|
551  | 
end  | 
|
552  | 
||
| 
71957
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
553  | 
interpretation word_order: ordering_top \<open>(\<le>)\<close> \<open>(<)\<close> \<open>- 1 :: 'a::len word\<close>  | 
| 
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
554  | 
by (standard; transfer) (simp add: take_bit_eq_mod zmod_minus1)  | 
| 
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
555  | 
|
| 
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
556  | 
interpretation word_coorder: ordering_top \<open>(\<ge>)\<close> \<open>(>)\<close> \<open>0 :: 'a::len word\<close>  | 
| 
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
557  | 
by (standard; transfer) simp  | 
| 
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
558  | 
|
| 72262 | 559  | 
lemma word_of_nat_less_eq_iff:  | 
560  | 
  \<open>word_of_nat m \<le> (word_of_nat n :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) m \<le> take_bit LENGTH('a) n\<close>
 | 
|
561  | 
by transfer (simp add: take_bit_of_nat)  | 
|
562  | 
||
563  | 
lemma word_of_int_less_eq_iff:  | 
|
564  | 
  \<open>word_of_int k \<le> (word_of_int l :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) k \<le> take_bit LENGTH('a) l\<close>
 | 
|
565  | 
by transfer rule  | 
|
566  | 
||
567  | 
lemma word_of_nat_less_iff:  | 
|
568  | 
  \<open>word_of_nat m < (word_of_nat n :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) m < take_bit LENGTH('a) n\<close>
 | 
|
569  | 
by transfer (simp add: take_bit_of_nat)  | 
|
570  | 
||
571  | 
lemma word_of_int_less_iff:  | 
|
572  | 
  \<open>word_of_int k < (word_of_int l :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) k < take_bit LENGTH('a) l\<close>
 | 
|
573  | 
by transfer rule  | 
|
574  | 
||
| 71950 | 575  | 
lemma word_le_def [code]:  | 
576  | 
"a \<le> b \<longleftrightarrow> uint a \<le> uint b"  | 
|
577  | 
by transfer rule  | 
|
578  | 
||
579  | 
lemma word_less_def [code]:  | 
|
580  | 
"a < b \<longleftrightarrow> uint a < uint b"  | 
|
581  | 
by transfer rule  | 
|
582  | 
||
| 71951 | 583  | 
lemma word_greater_zero_iff:  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
584  | 
\<open>a > 0 \<longleftrightarrow> a \<noteq> 0\<close> for a :: \<open>'a::len word\<close>  | 
| 71951 | 585  | 
by transfer (simp add: less_le)  | 
586  | 
||
587  | 
lemma of_nat_word_less_eq_iff:  | 
|
588  | 
  \<open>of_nat m \<le> (of_nat n :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) m \<le> take_bit LENGTH('a) n\<close>
 | 
|
589  | 
by transfer (simp add: take_bit_of_nat)  | 
|
590  | 
||
591  | 
lemma of_nat_word_less_iff:  | 
|
592  | 
  \<open>of_nat m < (of_nat n :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) m < take_bit LENGTH('a) n\<close>
 | 
|
593  | 
by transfer (simp add: take_bit_of_nat)  | 
|
594  | 
||
595  | 
lemma of_int_word_less_eq_iff:  | 
|
596  | 
  \<open>of_int k \<le> (of_int l :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) k \<le> take_bit LENGTH('a) l\<close>
 | 
|
597  | 
by transfer rule  | 
|
598  | 
||
599  | 
lemma of_int_word_less_iff:  | 
|
600  | 
  \<open>of_int k < (of_int l :: 'a::len word) \<longleftrightarrow> take_bit LENGTH('a) k < take_bit LENGTH('a) l\<close>
 | 
|
601  | 
by transfer rule  | 
|
602  | 
||
| 37660 | 603  | 
|
| 72280 | 604  | 
|
605  | 
subsection \<open>Enumeration\<close>  | 
|
606  | 
||
607  | 
lemma inj_on_word_of_nat:  | 
|
608  | 
  \<open>inj_on (word_of_nat :: nat \<Rightarrow> 'a::len word) {0..<2 ^ LENGTH('a)}\<close>
 | 
|
609  | 
by (rule inj_onI; transfer) (simp_all add: take_bit_int_eq_self)  | 
|
610  | 
||
611  | 
lemma UNIV_word_eq_word_of_nat:  | 
|
612  | 
  \<open>(UNIV :: 'a::len word set) = word_of_nat ` {0..<2 ^ LENGTH('a)}\<close> (is \<open>_ = ?A\<close>)
 | 
|
613  | 
proof  | 
|
614  | 
  show \<open>word_of_nat ` {0..<2 ^ LENGTH('a)} \<subseteq> UNIV\<close>
 | 
|
615  | 
by simp  | 
|
616  | 
show \<open>UNIV \<subseteq> ?A\<close>  | 
|
617  | 
proof  | 
|
618  | 
fix w :: \<open>'a word\<close>  | 
|
619  | 
    show \<open>w \<in> (word_of_nat ` {0..<2 ^ LENGTH('a)} :: 'a word set)\<close>
 | 
|
620  | 
by (rule image_eqI [of _ _ \<open>unat w\<close>]; transfer) simp_all  | 
|
621  | 
qed  | 
|
622  | 
qed  | 
|
623  | 
||
624  | 
instantiation word :: (len) enum  | 
|
625  | 
begin  | 
|
626  | 
||
627  | 
definition enum_word :: \<open>'a word list\<close>  | 
|
628  | 
  where \<open>enum_word = map word_of_nat [0..<2 ^ LENGTH('a)]\<close>
 | 
|
629  | 
||
630  | 
definition enum_all_word :: \<open>('a word \<Rightarrow> bool) \<Rightarrow> bool\<close>
 | 
|
| 77225 | 631  | 
where \<open>enum_all_word = All\<close>  | 
| 72280 | 632  | 
|
633  | 
definition enum_ex_word :: \<open>('a word \<Rightarrow> bool) \<Rightarrow> bool\<close>
 | 
|
| 77225 | 634  | 
where \<open>enum_ex_word = Ex\<close>  | 
635  | 
||
636  | 
instance  | 
|
637  | 
by standard  | 
|
638  | 
(simp_all add: enum_all_word_def enum_ex_word_def enum_word_def distinct_map inj_on_word_of_nat flip: UNIV_word_eq_word_of_nat)  | 
|
639  | 
||
640  | 
end  | 
|
| 72280 | 641  | 
|
642  | 
lemma [code]:  | 
|
| 77225 | 643  | 
\<open>Enum.enum_all P \<longleftrightarrow> list_all P Enum.enum\<close>  | 
644  | 
\<open>Enum.enum_ex P \<longleftrightarrow> list_ex P Enum.enum\<close> for P :: \<open>'a::len word \<Rightarrow> bool\<close>  | 
|
645  | 
by (simp_all add: enum_all_word_def enum_ex_word_def enum_UNIV list_all_iff list_ex_iff)  | 
|
| 72280 | 646  | 
|
647  | 
||
| 61799 | 648  | 
subsection \<open>Bit-wise operations\<close>  | 
| 37660 | 649  | 
|
| 77812 | 650  | 
text \<open>  | 
651  | 
The following specification of word division just lifts the pre-existing  | 
|
| 79950 | 652  | 
division on integers named ``F-Division'' in \<^cite>\<open>"leijen01"\<close>.  | 
| 77812 | 653  | 
\<close>  | 
654  | 
||
| 72244 | 655  | 
instantiation word :: (len) semiring_modulo  | 
656  | 
begin  | 
|
657  | 
||
658  | 
lift_definition divide_word :: \<open>'a word \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>  | 
|
659  | 
  is \<open>\<lambda>a b. take_bit LENGTH('a) a div take_bit LENGTH('a) b\<close>
 | 
|
660  | 
by simp  | 
|
661  | 
||
662  | 
lift_definition modulo_word :: \<open>'a word \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>  | 
|
663  | 
  is \<open>\<lambda>a b. take_bit LENGTH('a) a mod take_bit LENGTH('a) b\<close>
 | 
|
664  | 
by simp  | 
|
665  | 
||
666  | 
instance proof  | 
|
667  | 
show "a div b * b + a mod b = a" for a b :: "'a word"  | 
|
668  | 
proof transfer  | 
|
669  | 
fix k l :: int  | 
|
670  | 
    define r :: int where "r = 2 ^ LENGTH('a)"
 | 
|
671  | 
    then have r: "take_bit LENGTH('a) k = k mod r" for k
 | 
|
672  | 
by (simp add: take_bit_eq_mod)  | 
|
673  | 
have "k mod r = ((k mod r) div (l mod r) * (l mod r)  | 
|
674  | 
+ (k mod r) mod (l mod r)) mod r"  | 
|
675  | 
by (simp add: div_mult_mod_eq)  | 
|
| 80777 | 676  | 
also have "\<dots> = (((k mod r) div (l mod r) * (l mod r)) mod r  | 
| 72244 | 677  | 
+ (k mod r) mod (l mod r)) mod r"  | 
678  | 
by (simp add: mod_add_left_eq)  | 
|
| 80777 | 679  | 
also have "\<dots> = (((k mod r) div (l mod r) * l) mod r  | 
| 72244 | 680  | 
+ (k mod r) mod (l mod r)) mod r"  | 
681  | 
by (simp add: mod_mult_right_eq)  | 
|
682  | 
finally have "k mod r = ((k mod r) div (l mod r) * l  | 
|
683  | 
+ (k mod r) mod (l mod r)) mod r"  | 
|
684  | 
by (simp add: mod_simps)  | 
|
685  | 
    with r show "take_bit LENGTH('a) (take_bit LENGTH('a) k div take_bit LENGTH('a) l * l
 | 
|
686  | 
      + take_bit LENGTH('a) k mod take_bit LENGTH('a) l) = take_bit LENGTH('a) k"
 | 
|
687  | 
by simp  | 
|
688  | 
qed  | 
|
689  | 
qed  | 
|
690  | 
||
691  | 
end  | 
|
692  | 
||
| 
79673
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
693  | 
lemma unat_div_distrib:  | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
694  | 
\<open>unat (v div w) = unat v div unat w\<close>  | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
695  | 
proof transfer  | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
696  | 
fix k l  | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
697  | 
  have \<open>nat (take_bit LENGTH('a) k) div nat (take_bit LENGTH('a) l) \<le> nat (take_bit LENGTH('a) k)\<close>
 | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
698  | 
by (rule div_le_dividend)  | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
699  | 
  also have \<open>nat (take_bit LENGTH('a) k) < 2 ^ LENGTH('a)\<close>
 | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
700  | 
by (simp add: nat_less_iff)  | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
701  | 
  finally show \<open>(nat \<circ> take_bit LENGTH('a)) (take_bit LENGTH('a) k div take_bit LENGTH('a) l) =
 | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
702  | 
    (nat \<circ> take_bit LENGTH('a)) k div (nat \<circ> take_bit LENGTH('a)) l\<close>
 | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
703  | 
by (simp add: nat_take_bit_eq div_int_pos_iff nat_div_distrib take_bit_nat_eq_self_iff)  | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
704  | 
qed  | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
705  | 
|
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
706  | 
lemma unat_mod_distrib:  | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
707  | 
\<open>unat (v mod w) = unat v mod unat w\<close>  | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
708  | 
proof transfer  | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
709  | 
fix k l  | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
710  | 
  have \<open>nat (take_bit LENGTH('a) k) mod nat (take_bit LENGTH('a) l) \<le> nat (take_bit LENGTH('a) k)\<close>
 | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
711  | 
by (rule mod_less_eq_dividend)  | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
712  | 
  also have \<open>nat (take_bit LENGTH('a) k) < 2 ^ LENGTH('a)\<close>
 | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
713  | 
by (simp add: nat_less_iff)  | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
714  | 
  finally show \<open>(nat \<circ> take_bit LENGTH('a)) (take_bit LENGTH('a) k mod take_bit LENGTH('a) l) =
 | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
715  | 
    (nat \<circ> take_bit LENGTH('a)) k mod (nat \<circ> take_bit LENGTH('a)) l\<close>
 | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
716  | 
by (simp add: nat_take_bit_eq mod_int_pos_iff less_le nat_mod_distrib take_bit_nat_eq_self_iff)  | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
717  | 
qed  | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
718  | 
|
| 72244 | 719  | 
instance word :: (len) semiring_parity  | 
| 79555 | 720  | 
by (standard; transfer)  | 
| 80777 | 721  | 
(auto simp: mod_2_eq_odd take_bit_Suc elim: evenE dest: le_Suc_ex)  | 
| 72244 | 722  | 
|
| 71951 | 723  | 
lemma word_bit_induct [case_names zero even odd]:  | 
724  | 
\<open>P a\<close> if word_zero: \<open>P 0\<close>  | 
|
| 72262 | 725  | 
    and word_even: \<open>\<And>a. P a \<Longrightarrow> 0 < a \<Longrightarrow> a < 2 ^ (LENGTH('a) - Suc 0) \<Longrightarrow> P (2 * a)\<close>
 | 
726  | 
    and word_odd: \<open>\<And>a. P a \<Longrightarrow> a < 2 ^ (LENGTH('a) - Suc 0) \<Longrightarrow> P (1 + 2 * a)\<close>
 | 
|
| 71951 | 727  | 
for P and a :: \<open>'a::len word\<close>  | 
728  | 
proof -  | 
|
| 72262 | 729  | 
  define m :: nat where \<open>m = LENGTH('a) - Suc 0\<close>
 | 
| 71951 | 730  | 
  then have l: \<open>LENGTH('a) = Suc m\<close>
 | 
731  | 
by simp  | 
|
732  | 
define n :: nat where \<open>n = unat a\<close>  | 
|
733  | 
  then have \<open>n < 2 ^ LENGTH('a)\<close>
 | 
|
| 72262 | 734  | 
by transfer (simp add: take_bit_eq_mod)  | 
| 71951 | 735  | 
then have \<open>n < 2 * 2 ^ m\<close>  | 
736  | 
by (simp add: l)  | 
|
737  | 
then have \<open>P (of_nat n)\<close>  | 
|
738  | 
proof (induction n rule: nat_bit_induct)  | 
|
739  | 
case zero  | 
|
740  | 
show ?case  | 
|
741  | 
by simp (rule word_zero)  | 
|
742  | 
next  | 
|
743  | 
case (even n)  | 
|
744  | 
then have \<open>n < 2 ^ m\<close>  | 
|
745  | 
by simp  | 
|
746  | 
with even.IH have \<open>P (of_nat n)\<close>  | 
|
747  | 
by simp  | 
|
748  | 
moreover from \<open>n < 2 ^ m\<close> even.hyps have \<open>0 < (of_nat n :: 'a word)\<close>  | 
|
| 80777 | 749  | 
by (auto simp: word_greater_zero_iff l word_of_nat_eq_0_iff)  | 
| 72262 | 750  | 
    moreover from \<open>n < 2 ^ m\<close> have \<open>(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - Suc 0)\<close>
 | 
| 71951 | 751  | 
using of_nat_word_less_iff [where ?'a = 'a, of n \<open>2 ^ m\<close>]  | 
| 72261 | 752  | 
by (simp add: l take_bit_eq_mod)  | 
| 71951 | 753  | 
ultimately have \<open>P (2 * of_nat n)\<close>  | 
754  | 
by (rule word_even)  | 
|
755  | 
then show ?case  | 
|
756  | 
by simp  | 
|
757  | 
next  | 
|
758  | 
case (odd n)  | 
|
759  | 
then have \<open>Suc n \<le> 2 ^ m\<close>  | 
|
760  | 
by simp  | 
|
761  | 
with odd.IH have \<open>P (of_nat n)\<close>  | 
|
762  | 
by simp  | 
|
| 72262 | 763  | 
    moreover from \<open>Suc n \<le> 2 ^ m\<close> have \<open>(of_nat n :: 'a word) < 2 ^ (LENGTH('a) - Suc 0)\<close>
 | 
| 71951 | 764  | 
using of_nat_word_less_iff [where ?'a = 'a, of n \<open>2 ^ m\<close>]  | 
| 72261 | 765  | 
by (simp add: l take_bit_eq_mod)  | 
| 71951 | 766  | 
ultimately have \<open>P (1 + 2 * of_nat n)\<close>  | 
767  | 
by (rule word_odd)  | 
|
768  | 
then show ?case  | 
|
769  | 
by simp  | 
|
770  | 
qed  | 
|
771  | 
moreover have \<open>of_nat (nat (uint a)) = a\<close>  | 
|
772  | 
by transfer simp  | 
|
773  | 
ultimately show ?thesis  | 
|
| 72079 | 774  | 
by (simp add: n_def)  | 
| 71951 | 775  | 
qed  | 
776  | 
||
777  | 
lemma bit_word_half_eq:  | 
|
778  | 
\<open>(of_bool b + a * 2) div 2 = a\<close>  | 
|
779  | 
    if \<open>a < 2 ^ (LENGTH('a) - Suc 0)\<close>
 | 
|
780  | 
for a :: \<open>'a::len word\<close>  | 
|
781  | 
proof (cases \<open>2 \<le> LENGTH('a::len)\<close>)
 | 
|
782  | 
case False  | 
|
| 82524 | 783  | 
with that show ?thesis  | 
784  | 
by transfer simp  | 
|
| 71951 | 785  | 
next  | 
786  | 
case True  | 
|
787  | 
  obtain n where length: \<open>LENGTH('a) = Suc n\<close>
 | 
|
788  | 
    by (cases \<open>LENGTH('a)\<close>) simp_all
 | 
|
789  | 
show ?thesis proof (cases b)  | 
|
790  | 
case False  | 
|
791  | 
moreover have \<open>a * 2 div 2 = a\<close>  | 
|
792  | 
using that proof transfer  | 
|
793  | 
fix k :: int  | 
|
794  | 
      from length have \<open>k * 2 mod 2 ^ LENGTH('a) = (k mod 2 ^ n) * 2\<close>
 | 
|
795  | 
by simp  | 
|
796  | 
      moreover assume \<open>take_bit LENGTH('a) k < take_bit LENGTH('a) (2 ^ (LENGTH('a) - Suc 0))\<close>
 | 
|
| 76231 | 797  | 
      with \<open>LENGTH('a) = Suc n\<close> have \<open>take_bit LENGTH('a) k = take_bit n k\<close>
 | 
| 80777 | 798  | 
by (auto simp: take_bit_Suc_from_most)  | 
| 71951 | 799  | 
      ultimately have \<open>take_bit LENGTH('a) (k * 2) = take_bit LENGTH('a) k * 2\<close>
 | 
800  | 
by (simp add: take_bit_eq_mod)  | 
|
801  | 
      with True show \<open>take_bit LENGTH('a) (take_bit LENGTH('a) (k * 2) div take_bit LENGTH('a) 2)
 | 
|
802  | 
        = take_bit LENGTH('a) k\<close>
 | 
|
803  | 
by simp  | 
|
804  | 
qed  | 
|
805  | 
ultimately show ?thesis  | 
|
806  | 
by simp  | 
|
807  | 
next  | 
|
808  | 
case True  | 
|
809  | 
moreover have \<open>(1 + a * 2) div 2 = a\<close>  | 
|
810  | 
using that proof transfer  | 
|
811  | 
fix k :: int  | 
|
812  | 
      from length have \<open>(1 + k * 2) mod 2 ^ LENGTH('a) = 1 + (k mod 2 ^ n) * 2\<close>
 | 
|
813  | 
using pos_zmod_mult_2 [of \<open>2 ^ n\<close> k] by (simp add: ac_simps)  | 
|
814  | 
      moreover assume \<open>take_bit LENGTH('a) k < take_bit LENGTH('a) (2 ^ (LENGTH('a) - Suc 0))\<close>
 | 
|
| 76231 | 815  | 
      with \<open>LENGTH('a) = Suc n\<close> have \<open>take_bit LENGTH('a) k = take_bit n k\<close>
 | 
| 80777 | 816  | 
by (auto simp: take_bit_Suc_from_most)  | 
| 71951 | 817  | 
      ultimately have \<open>take_bit LENGTH('a) (1 + k * 2) = 1 + take_bit LENGTH('a) k * 2\<close>
 | 
818  | 
by (simp add: take_bit_eq_mod)  | 
|
819  | 
      with True show \<open>take_bit LENGTH('a) (take_bit LENGTH('a) (1 + k * 2) div take_bit LENGTH('a) 2)
 | 
|
820  | 
        = take_bit LENGTH('a) k\<close>
 | 
|
| 80777 | 821  | 
by (auto simp: take_bit_Suc)  | 
| 71951 | 822  | 
qed  | 
823  | 
ultimately show ?thesis  | 
|
824  | 
by simp  | 
|
825  | 
qed  | 
|
826  | 
qed  | 
|
827  | 
||
828  | 
lemma even_mult_exp_div_word_iff:  | 
|
829  | 
\<open>even (a * 2 ^ m div 2 ^ n) \<longleftrightarrow> \<not> (  | 
|
830  | 
m \<le> n \<and>  | 
|
831  | 
    n < LENGTH('a) \<and> odd (a div 2 ^ (n - m)))\<close> for a :: \<open>'a::len word\<close>
 | 
|
832  | 
by transfer  | 
|
833  | 
(auto simp flip: drop_bit_eq_div simp add: even_drop_bit_iff_not_bit bit_take_bit_iff,  | 
|
834  | 
simp_all flip: push_bit_eq_mult add: bit_push_bit_iff_int)  | 
|
835  | 
||
| 
71965
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
836  | 
instantiation word :: (len) semiring_bits  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
837  | 
begin  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
838  | 
|
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
839  | 
lift_definition bit_word :: \<open>'a word \<Rightarrow> nat \<Rightarrow> bool\<close>  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
840  | 
  is \<open>\<lambda>k n. n < LENGTH('a) \<and> bit k n\<close>
 | 
| 71951 | 841  | 
proof  | 
| 
71965
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
842  | 
fix k l :: int and n :: nat  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
843  | 
  assume *: \<open>take_bit LENGTH('a) k = take_bit LENGTH('a) l\<close>
 | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
844  | 
  show \<open>n < LENGTH('a) \<and> bit k n \<longleftrightarrow> n < LENGTH('a) \<and> bit l n\<close>
 | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
845  | 
  proof (cases \<open>n < LENGTH('a)\<close>)
 | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
846  | 
case True  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
847  | 
    from * have \<open>bit (take_bit LENGTH('a) k) n \<longleftrightarrow> bit (take_bit LENGTH('a) l) n\<close>
 | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
848  | 
by simp  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
849  | 
then show ?thesis  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
850  | 
by (simp add: bit_take_bit_iff)  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
851  | 
next  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
852  | 
case False  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
853  | 
then show ?thesis  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
854  | 
by simp  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
855  | 
qed  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
856  | 
qed  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
857  | 
|
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
858  | 
instance proof  | 
| 71951 | 859  | 
show \<open>P a\<close> if stable: \<open>\<And>a. a div 2 = a \<Longrightarrow> P a\<close>  | 
860  | 
and rec: \<open>\<And>a b. P a \<Longrightarrow> (of_bool b + 2 * a) div 2 = a \<Longrightarrow> P (of_bool b + 2 * a)\<close>  | 
|
861  | 
for P and a :: \<open>'a word\<close>  | 
|
862  | 
proof (induction a rule: word_bit_induct)  | 
|
863  | 
case zero  | 
|
864  | 
have \<open>0 div 2 = (0::'a word)\<close>  | 
|
865  | 
by transfer simp  | 
|
866  | 
with stable [of 0] show ?case  | 
|
867  | 
by simp  | 
|
868  | 
next  | 
|
869  | 
case (even a)  | 
|
870  | 
with rec [of a False] show ?case  | 
|
871  | 
using bit_word_half_eq [of a False] by (simp add: ac_simps)  | 
|
872  | 
next  | 
|
873  | 
case (odd a)  | 
|
874  | 
with rec [of a True] show ?case  | 
|
875  | 
using bit_word_half_eq [of a True] by (simp add: ac_simps)  | 
|
876  | 
qed  | 
|
| 
71965
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
877  | 
show \<open>bit a n \<longleftrightarrow> odd (a div 2 ^ n)\<close> for a :: \<open>'a word\<close> and n  | 
| 
 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 
haftmann 
parents: 
71958 
diff
changeset
 | 
878  | 
by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit bit_iff_odd_drop_bit)  | 
| 79481 | 879  | 
show \<open>a div 0 = 0\<close>  | 
880  | 
for a :: \<open>'a word\<close>  | 
|
881  | 
by transfer simp  | 
|
| 71951 | 882  | 
show \<open>a div 1 = a\<close>  | 
883  | 
for a :: \<open>'a word\<close>  | 
|
884  | 
by transfer simp  | 
|
| 
79531
 
22a137a6de44
rearranged and reformulated abstract classes for bit structures and operations
 
haftmann 
parents: 
79489 
diff
changeset
 | 
885  | 
show \<open>0 div a = 0\<close>  | 
| 
 
22a137a6de44
rearranged and reformulated abstract classes for bit structures and operations
 
haftmann 
parents: 
79489 
diff
changeset
 | 
886  | 
for a :: \<open>'a word\<close>  | 
| 
 
22a137a6de44
rearranged and reformulated abstract classes for bit structures and operations
 
haftmann 
parents: 
79489 
diff
changeset
 | 
887  | 
by transfer simp  | 
| 
79673
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
888  | 
show \<open>a mod b div b = 0\<close>  | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
889  | 
for a b :: \<open>'a word\<close>  | 
| 
 
c172eecba85d
simplified specification of type class semiring_bits
 
haftmann 
parents: 
79590 
diff
changeset
 | 
890  | 
by (simp add: word_eq_iff_unsigned [where ?'a = nat] unat_div_distrib unat_mod_distrib)  | 
| 
79531
 
22a137a6de44
rearranged and reformulated abstract classes for bit structures and operations
 
haftmann 
parents: 
79489 
diff
changeset
 | 
891  | 
show \<open>a div 2 div 2 ^ n = a div 2 ^ Suc n\<close>  | 
| 
 
22a137a6de44
rearranged and reformulated abstract classes for bit structures and operations
 
haftmann 
parents: 
79489 
diff
changeset
 | 
892  | 
for a :: \<open>'a word\<close> and m n :: nat  | 
| 71951 | 893  | 
apply transfer  | 
| 
79531
 
22a137a6de44
rearranged and reformulated abstract classes for bit structures and operations
 
haftmann 
parents: 
79489 
diff
changeset
 | 
894  | 
using drop_bit_eq_div [symmetric, where ?'a = int,of _ 1]  | 
| 80777 | 895  | 
apply (auto simp: not_less take_bit_drop_bit ac_simps simp flip: drop_bit_eq_div simp del: power.simps)  | 
| 71951 | 896  | 
apply (simp add: drop_bit_take_bit)  | 
897  | 
done  | 
|
| 
79531
 
22a137a6de44
rearranged and reformulated abstract classes for bit structures and operations
 
haftmann 
parents: 
79489 
diff
changeset
 | 
898  | 
show \<open>even (2 * a div 2 ^ Suc n) \<longleftrightarrow> even (a div 2 ^ n)\<close> if \<open>2 ^ Suc n \<noteq> (0::'a word)\<close>  | 
| 
 
22a137a6de44
rearranged and reformulated abstract classes for bit structures and operations
 
haftmann 
parents: 
79489 
diff
changeset
 | 
899  | 
for a :: \<open>'a word\<close> and n :: nat  | 
| 
 
22a137a6de44
rearranged and reformulated abstract classes for bit structures and operations
 
haftmann 
parents: 
79489 
diff
changeset
 | 
900  | 
using that by transfer  | 
| 
 
22a137a6de44
rearranged and reformulated abstract classes for bit structures and operations
 
haftmann 
parents: 
79489 
diff
changeset
 | 
901  | 
(simp add: even_drop_bit_iff_not_bit bit_simps flip: drop_bit_eq_div del: power.simps)  | 
| 71951 | 902  | 
qed  | 
903  | 
||
904  | 
end  | 
|
905  | 
||
| 72262 | 906  | 
lemma bit_word_eqI:  | 
907  | 
  \<open>a = b\<close> if \<open>\<And>n. n < LENGTH('a) \<Longrightarrow> bit a n \<longleftrightarrow> bit b n\<close>
 | 
|
908  | 
for a b :: \<open>'a::len word\<close>  | 
|
| 80777 | 909  | 
using that by transfer (auto simp: nat_less_le bit_eq_iff bit_take_bit_iff)  | 
| 72262 | 910  | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
911  | 
lemma bit_imp_le_length: \<open>n < LENGTH('a)\<close> if \<open>bit w n\<close> for w :: \<open>'a::len word\<close>
 | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
912  | 
by (meson bit_word.rep_eq that)  | 
| 72262 | 913  | 
|
914  | 
lemma not_bit_length [simp]:  | 
|
915  | 
  \<open>\<not> bit w LENGTH('a)\<close> for w :: \<open>'a::len word\<close>
 | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
916  | 
using bit_imp_le_length by blast  | 
| 72262 | 917  | 
|
| 72830 | 918  | 
lemma finite_bit_word [simp]:  | 
919  | 
  \<open>finite {n. bit w n}\<close>
 | 
|
920  | 
for w :: \<open>'a::len word\<close>  | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
921  | 
by (metis bit_imp_le_length bounded_nat_set_is_finite mem_Collect_eq)  | 
| 72830 | 922  | 
|
| 73789 | 923  | 
lemma bit_numeral_word_iff [simp]:  | 
924  | 
\<open>bit (numeral w :: 'a::len word) n  | 
|
925  | 
    \<longleftrightarrow> n < LENGTH('a) \<and> bit (numeral w :: int) n\<close>
 | 
|
926  | 
by transfer simp  | 
|
927  | 
||
928  | 
lemma bit_neg_numeral_word_iff [simp]:  | 
|
929  | 
\<open>bit (- numeral w :: 'a::len word) n  | 
|
930  | 
    \<longleftrightarrow> n < LENGTH('a) \<and> bit (- numeral w :: int) n\<close>
 | 
|
931  | 
by transfer simp  | 
|
932  | 
||
| 72262 | 933  | 
instantiation word :: (len) ring_bit_operations  | 
934  | 
begin  | 
|
935  | 
||
936  | 
lift_definition not_word :: \<open>'a word \<Rightarrow> 'a word\<close>  | 
|
937  | 
is not  | 
|
938  | 
by (simp add: take_bit_not_iff)  | 
|
939  | 
||
940  | 
lift_definition and_word :: \<open>'a word \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>  | 
|
941  | 
is \<open>and\<close>  | 
|
942  | 
by simp  | 
|
943  | 
||
944  | 
lift_definition or_word :: \<open>'a word \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>  | 
|
945  | 
is or  | 
|
946  | 
by simp  | 
|
947  | 
||
948  | 
lift_definition xor_word :: \<open>'a word \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>  | 
|
949  | 
is xor  | 
|
950  | 
by simp  | 
|
951  | 
||
952  | 
lift_definition mask_word :: \<open>nat \<Rightarrow> 'a word\<close>  | 
|
953  | 
is mask  | 
|
954  | 
.  | 
|
955  | 
||
| 
73682
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
956  | 
lift_definition set_bit_word :: \<open>nat \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>  | 
| 
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
957  | 
is set_bit  | 
| 
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
958  | 
by (simp add: set_bit_def)  | 
| 
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
959  | 
|
| 
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
960  | 
lift_definition unset_bit_word :: \<open>nat \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>  | 
| 
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
961  | 
is unset_bit  | 
| 
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
962  | 
by (simp add: unset_bit_def)  | 
| 
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
963  | 
|
| 
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
964  | 
lift_definition flip_bit_word :: \<open>nat \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>  | 
| 
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
965  | 
is flip_bit  | 
| 
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
966  | 
by (simp add: flip_bit_def)  | 
| 
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
967  | 
|
| 
74108
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
968  | 
lift_definition push_bit_word :: \<open>nat \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
969  | 
is push_bit  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
970  | 
proof -  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
971  | 
  show \<open>take_bit LENGTH('a) (push_bit n k) = take_bit LENGTH('a) (push_bit n l)\<close>
 | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
972  | 
    if \<open>take_bit LENGTH('a) k = take_bit LENGTH('a) l\<close> for k l :: int and n :: nat
 | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
973  | 
by (metis le_add2 push_bit_take_bit take_bit_tightened that)  | 
| 
74108
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
974  | 
qed  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
975  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
976  | 
lift_definition drop_bit_word :: \<open>nat \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
977  | 
  is \<open>\<lambda>n. drop_bit n \<circ> take_bit LENGTH('a)\<close>
 | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
978  | 
by (simp add: take_bit_eq_mod)  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
979  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
980  | 
lift_definition take_bit_word :: \<open>nat \<Rightarrow> 'a word \<Rightarrow> 'a word\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
981  | 
  is \<open>\<lambda>n. take_bit (min LENGTH('a) n)\<close>
 | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
982  | 
by (simp add: ac_simps) (simp only: flip: take_bit_take_bit)  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
983  | 
|
| 
79008
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
984  | 
context  | 
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
985  | 
includes bit_operations_syntax  | 
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
986  | 
begin  | 
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
987  | 
|
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
988  | 
instance proof  | 
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
989  | 
fix v w :: \<open>'a word\<close> and n m :: nat  | 
| 
79072
 
a91050cd5c93
de-duplicated specification of class ring_bit_operations
 
haftmann 
parents: 
79031 
diff
changeset
 | 
990  | 
show \<open>NOT v = - v - 1\<close>  | 
| 
 
a91050cd5c93
de-duplicated specification of class ring_bit_operations
 
haftmann 
parents: 
79031 
diff
changeset
 | 
991  | 
by transfer (simp add: not_eq_complement)  | 
| 
79008
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
992  | 
show \<open>v AND w = of_bool (odd v \<and> odd w) + 2 * (v div 2 AND w div 2)\<close>  | 
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
993  | 
apply transfer  | 
| 80777 | 994  | 
by (rule bit_eqI) (auto simp: even_bit_succ_iff bit_simps bit_0 simp flip: bit_Suc)  | 
| 
79008
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
995  | 
show \<open>v OR w = of_bool (odd v \<or> odd w) + 2 * (v div 2 OR w div 2)\<close>  | 
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
996  | 
apply transfer  | 
| 80777 | 997  | 
by (rule bit_eqI) (auto simp: even_bit_succ_iff bit_simps bit_0 simp flip: bit_Suc)  | 
| 
79008
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
998  | 
show \<open>v XOR w = of_bool (odd v \<noteq> odd w) + 2 * (v div 2 XOR w div 2)\<close>  | 
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
999  | 
apply transfer  | 
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
1000  | 
apply (rule bit_eqI)  | 
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
1001  | 
subgoal for k l n  | 
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
1002  | 
apply (cases n)  | 
| 80777 | 1003  | 
apply (auto simp: even_bit_succ_iff bit_simps bit_0 even_xor_iff simp flip: bit_Suc)  | 
| 
79008
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
1004  | 
done  | 
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
1005  | 
done  | 
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
1006  | 
show \<open>mask n = 2 ^ n - (1 :: 'a word)\<close>  | 
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
1007  | 
by transfer (simp flip: mask_eq_exp_minus_1)  | 
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
1008  | 
show \<open>set_bit n v = v OR push_bit n 1\<close>  | 
| 79489 | 1009  | 
by transfer (simp add: set_bit_eq_or)  | 
1010  | 
show \<open>unset_bit n v = (v OR push_bit n 1) XOR push_bit n 1\<close>  | 
|
1011  | 
by transfer (simp add: unset_bit_eq_or_xor)  | 
|
| 
79008
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
1012  | 
show \<open>flip_bit n v = v XOR push_bit n 1\<close>  | 
| 79489 | 1013  | 
by transfer (simp add: flip_bit_eq_xor)  | 
| 
79008
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
1014  | 
show \<open>push_bit n v = v * 2 ^ n\<close>  | 
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
1015  | 
by transfer (simp add: push_bit_eq_mult)  | 
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
1016  | 
show \<open>drop_bit n v = v div 2 ^ n\<close>  | 
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
1017  | 
by transfer (simp add: drop_bit_take_bit flip: drop_bit_eq_div)  | 
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
1018  | 
show \<open>take_bit n v = v mod 2 ^ n\<close>  | 
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
1019  | 
by transfer (simp flip: take_bit_eq_mod)  | 
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
1020  | 
qed  | 
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
1021  | 
|
| 
 
74a4776f7a22
operations AND, OR, XOR are specified by characteristic recursive equation
 
haftmann 
parents: 
78955 
diff
changeset
 | 
1022  | 
end  | 
| 72262 | 1023  | 
|
1024  | 
end  | 
|
1025  | 
||
| 
74108
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1026  | 
lemma [code]:  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1027  | 
\<open>push_bit n w = w * 2 ^ n\<close> for w :: \<open>'a::len word\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1028  | 
by (fact push_bit_eq_mult)  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1029  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1030  | 
lemma [code]:  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1031  | 
\<open>Word.the_int (drop_bit n w) = drop_bit n (Word.the_int w)\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1032  | 
by transfer (simp add: drop_bit_take_bit min_def le_less less_diff_conv)  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1033  | 
|
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1034  | 
lemma [code]:  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1035  | 
  \<open>Word.the_int (take_bit n w) = (if n < LENGTH('a::len) then take_bit n (Word.the_int w) else Word.the_int w)\<close>
 | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1036  | 
for w :: \<open>'a::len word\<close>  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1037  | 
by transfer (simp add: not_le not_less ac_simps min_absorb2)  | 
| 
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1038  | 
|
| 72262 | 1039  | 
lemma [code_abbrev]:  | 
1040  | 
\<open>push_bit n 1 = (2 :: 'a::len word) ^ n\<close>  | 
|
1041  | 
by (fact push_bit_of_1)  | 
|
1042  | 
||
| 74391 | 1043  | 
context  | 
1044  | 
includes bit_operations_syntax  | 
|
1045  | 
begin  | 
|
1046  | 
||
| 72262 | 1047  | 
lemma [code]:  | 
1048  | 
\<open>NOT w = Word.of_int (NOT (Word.the_int w))\<close>  | 
|
1049  | 
for w :: \<open>'a::len word\<close>  | 
|
1050  | 
by transfer (simp add: take_bit_not_take_bit)  | 
|
1051  | 
||
1052  | 
lemma [code]:  | 
|
1053  | 
\<open>Word.the_int (v AND w) = Word.the_int v AND Word.the_int w\<close>  | 
|
| 71990 | 1054  | 
by transfer simp  | 
1055  | 
||
| 72262 | 1056  | 
lemma [code]:  | 
1057  | 
\<open>Word.the_int (v OR w) = Word.the_int v OR Word.the_int w\<close>  | 
|
1058  | 
by transfer simp  | 
|
1059  | 
||
1060  | 
lemma [code]:  | 
|
1061  | 
\<open>Word.the_int (v XOR w) = Word.the_int v XOR Word.the_int w\<close>  | 
|
1062  | 
by transfer simp  | 
|
1063  | 
||
1064  | 
lemma [code]:  | 
|
1065  | 
  \<open>Word.the_int (mask n :: 'a::len word) = mask (min LENGTH('a) n)\<close>
 | 
|
1066  | 
by transfer simp  | 
|
1067  | 
||
| 
73682
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
1068  | 
lemma [code]:  | 
| 
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
1069  | 
\<open>set_bit n w = w OR push_bit n 1\<close> for w :: \<open>'a::len word\<close>  | 
| 
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
1070  | 
by (fact set_bit_eq_or)  | 
| 
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
1071  | 
|
| 
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
1072  | 
lemma [code]:  | 
| 
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
1073  | 
\<open>unset_bit n w = w AND NOT (push_bit n 1)\<close> for w :: \<open>'a::len word\<close>  | 
| 
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
1074  | 
by (fact unset_bit_eq_and_not)  | 
| 
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
1075  | 
|
| 
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
1076  | 
lemma [code]:  | 
| 
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
1077  | 
\<open>flip_bit n w = w XOR push_bit n 1\<close> for w :: \<open>'a::len word\<close>  | 
| 
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
1078  | 
by (fact flip_bit_eq_xor)  | 
| 
 
78044b2f001c
explicit type class operations for type-specific implementations
 
haftmann 
parents: 
73535 
diff
changeset
 | 
1079  | 
|
| 72262 | 1080  | 
context  | 
1081  | 
includes lifting_syntax  | 
|
1082  | 
begin  | 
|
1083  | 
||
1084  | 
lemma set_bit_word_transfer [transfer_rule]:  | 
|
1085  | 
\<open>((=) ===> pcr_word ===> pcr_word) set_bit set_bit\<close>  | 
|
1086  | 
by (unfold set_bit_def) transfer_prover  | 
|
1087  | 
||
1088  | 
lemma unset_bit_word_transfer [transfer_rule]:  | 
|
1089  | 
\<open>((=) ===> pcr_word ===> pcr_word) unset_bit unset_bit\<close>  | 
|
1090  | 
by (unfold unset_bit_def) transfer_prover  | 
|
1091  | 
||
1092  | 
lemma flip_bit_word_transfer [transfer_rule]:  | 
|
1093  | 
\<open>((=) ===> pcr_word ===> pcr_word) flip_bit flip_bit\<close>  | 
|
1094  | 
by (unfold flip_bit_def) transfer_prover  | 
|
1095  | 
||
1096  | 
lemma signed_take_bit_word_transfer [transfer_rule]:  | 
|
1097  | 
\<open>((=) ===> pcr_word ===> pcr_word)  | 
|
1098  | 
    (\<lambda>n k. signed_take_bit n (take_bit LENGTH('a::len) k))
 | 
|
1099  | 
(signed_take_bit :: nat \<Rightarrow> 'a word \<Rightarrow> 'a word)\<close>  | 
|
1100  | 
proof -  | 
|
1101  | 
  let ?K = \<open>\<lambda>n (k :: int). take_bit (min LENGTH('a) n) k OR of_bool (n < LENGTH('a) \<and> bit k n) * NOT (mask n)\<close>
 | 
|
1102  | 
let ?W = \<open>\<lambda>n (w :: 'a word). take_bit n w OR of_bool (bit w n) * NOT (mask n)\<close>  | 
|
1103  | 
have \<open>((=) ===> pcr_word ===> pcr_word) ?K ?W\<close>  | 
|
1104  | 
by transfer_prover  | 
|
1105  | 
  also have \<open>?K = (\<lambda>n k. signed_take_bit n (take_bit LENGTH('a::len) k))\<close>
 | 
|
1106  | 
by (simp add: fun_eq_iff signed_take_bit_def bit_take_bit_iff ac_simps)  | 
|
1107  | 
also have \<open>?W = signed_take_bit\<close>  | 
|
1108  | 
by (simp add: fun_eq_iff signed_take_bit_def)  | 
|
1109  | 
finally show ?thesis .  | 
|
1110  | 
qed  | 
|
1111  | 
||
1112  | 
end  | 
|
1113  | 
||
| 74097 | 1114  | 
end  | 
1115  | 
||
| 72244 | 1116  | 
|
1117  | 
subsection \<open>Conversions including casts\<close>  | 
|
1118  | 
||
| 72262 | 1119  | 
subsubsection \<open>Generic unsigned conversion\<close>  | 
1120  | 
||
1121  | 
context semiring_bits  | 
|
1122  | 
begin  | 
|
1123  | 
||
| 
72611
 
c7bc3e70a8c7
official collection for bit projection simplifications
 
haftmann 
parents: 
72515 
diff
changeset
 | 
1124  | 
lemma bit_unsigned_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
 | 
1125  | 
  \<open>bit (unsigned w) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> bit w n\<close>
 | 
| 72262 | 1126  | 
for w :: \<open>'b::len word\<close>  | 
1127  | 
by (transfer fixing: bit) (simp add: bit_of_nat_iff bit_nat_iff bit_take_bit_iff)  | 
|
1128  | 
||
1129  | 
end  | 
|
1130  | 
||
| 
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
 | 
1131  | 
lemma possible_bit_word[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
 | 
1132  | 
  \<open>possible_bit TYPE(('a :: len) word) m \<longleftrightarrow> m < LENGTH('a)\<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
 | 
1133  | 
by (simp add: possible_bit_def linorder_not_le)  | 
| 
 
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
 | 
1134  | 
|
| 
74108
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1135  | 
context semiring_bit_operations  | 
| 72262 | 1136  | 
begin  | 
1137  | 
||
| 74592 | 1138  | 
lemma unsigned_minus_1_eq_mask:  | 
1139  | 
  \<open>unsigned (- 1 :: 'b::len word) = mask LENGTH('b)\<close>
 | 
|
1140  | 
by (transfer fixing: mask) (simp add: nat_mask_eq of_nat_mask_eq)  | 
|
1141  | 
||
| 72262 | 1142  | 
lemma unsigned_push_bit_eq:  | 
1143  | 
  \<open>unsigned (push_bit n w) = take_bit LENGTH('b) (push_bit n (unsigned w))\<close>
 | 
|
1144  | 
for w :: \<open>'b::len word\<close>  | 
|
1145  | 
proof (rule bit_eqI)  | 
|
1146  | 
fix m  | 
|
| 
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
 | 
1147  | 
  assume \<open>possible_bit TYPE('a) m\<close>
 | 
| 72262 | 1148  | 
  show \<open>bit (unsigned (push_bit n w)) m = bit (take_bit LENGTH('b) (push_bit n (unsigned w))) m\<close>
 | 
1149  | 
proof (cases \<open>n \<le> m\<close>)  | 
|
1150  | 
case True  | 
|
| 
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
 | 
1151  | 
    with \<open>possible_bit TYPE('a) m\<close> have \<open>possible_bit TYPE('a) (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
 | 
1152  | 
by (simp add: possible_bit_less_imp)  | 
| 72262 | 1153  | 
with True show ?thesis  | 
| 74101 | 1154  | 
by (simp add: bit_unsigned_iff bit_push_bit_iff Bit_Operations.bit_push_bit_iff bit_take_bit_iff not_le ac_simps)  | 
| 72262 | 1155  | 
next  | 
1156  | 
case False  | 
|
1157  | 
then show ?thesis  | 
|
| 74101 | 1158  | 
by (simp add: not_le bit_unsigned_iff bit_push_bit_iff Bit_Operations.bit_push_bit_iff bit_take_bit_iff)  | 
| 72262 | 1159  | 
qed  | 
1160  | 
qed  | 
|
1161  | 
||
1162  | 
lemma unsigned_take_bit_eq:  | 
|
1163  | 
\<open>unsigned (take_bit n w) = take_bit n (unsigned w)\<close>  | 
|
1164  | 
for w :: \<open>'b::len word\<close>  | 
|
| 74101 | 1165  | 
by (rule bit_eqI) (simp add: bit_unsigned_iff bit_take_bit_iff Bit_Operations.bit_take_bit_iff)  | 
| 72262 | 1166  | 
|
1167  | 
end  | 
|
1168  | 
||
| 78955 | 1169  | 
context linordered_euclidean_semiring_bit_operations  | 
| 72512 | 1170  | 
begin  | 
1171  | 
||
1172  | 
lemma unsigned_drop_bit_eq:  | 
|
1173  | 
  \<open>unsigned (drop_bit n w) = drop_bit n (take_bit LENGTH('b) (unsigned w))\<close>
 | 
|
1174  | 
for w :: \<open>'b::len word\<close>  | 
|
| 80777 | 1175  | 
by (rule bit_eqI) (auto simp: bit_unsigned_iff bit_take_bit_iff bit_drop_bit_eq Bit_Operations.bit_drop_bit_eq possible_bit_def dest: bit_imp_le_length)  | 
| 72512 | 1176  | 
|
1177  | 
end  | 
|
1178  | 
||
| 73853 | 1179  | 
lemma ucast_drop_bit_eq:  | 
1180  | 
\<open>ucast (drop_bit n w) = drop_bit n (ucast w :: 'b::len word)\<close>  | 
|
1181  | 
  if \<open>LENGTH('a) \<le> LENGTH('b)\<close> for w :: \<open>'a::len word\<close>
 | 
|
| 80777 | 1182  | 
by (rule bit_word_eqI) (use that in \<open>auto simp: bit_unsigned_iff bit_drop_bit_eq dest: bit_imp_le_length\<close>)  | 
| 73853 | 1183  | 
|
| 72262 | 1184  | 
context semiring_bit_operations  | 
1185  | 
begin  | 
|
1186  | 
||
| 74097 | 1187  | 
context  | 
1188  | 
includes bit_operations_syntax  | 
|
1189  | 
begin  | 
|
1190  | 
||
| 72262 | 1191  | 
lemma unsigned_and_eq:  | 
1192  | 
\<open>unsigned (v AND w) = unsigned v AND unsigned w\<close>  | 
|
1193  | 
for v w :: \<open>'b::len word\<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
 | 
1194  | 
by (simp add: bit_eq_iff bit_simps)  | 
| 72262 | 1195  | 
|
1196  | 
lemma unsigned_or_eq:  | 
|
1197  | 
\<open>unsigned (v OR w) = unsigned v OR unsigned w\<close>  | 
|
1198  | 
for v w :: \<open>'b::len word\<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
 | 
1199  | 
by (simp add: bit_eq_iff bit_simps)  | 
| 72262 | 1200  | 
|
1201  | 
lemma unsigned_xor_eq:  | 
|
1202  | 
\<open>unsigned (v XOR w) = unsigned v XOR unsigned w\<close>  | 
|
1203  | 
for v w :: \<open>'b::len word\<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
 | 
1204  | 
by (simp add: bit_eq_iff bit_simps)  | 
| 72262 | 1205  | 
|
1206  | 
end  | 
|
1207  | 
||
| 74097 | 1208  | 
end  | 
1209  | 
||
| 72262 | 1210  | 
context ring_bit_operations  | 
1211  | 
begin  | 
|
1212  | 
||
| 74097 | 1213  | 
context  | 
1214  | 
includes bit_operations_syntax  | 
|
1215  | 
begin  | 
|
1216  | 
||
| 72262 | 1217  | 
lemma unsigned_not_eq:  | 
1218  | 
  \<open>unsigned (NOT w) = take_bit LENGTH('b) (NOT (unsigned w))\<close>
 | 
|
1219  | 
for w :: \<open>'b::len word\<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
 | 
1220  | 
by (simp add: bit_eq_iff bit_simps)  | 
| 72262 | 1221  | 
|
1222  | 
end  | 
|
1223  | 
||
| 74097 | 1224  | 
end  | 
1225  | 
||
| 
80401
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
1226  | 
context linordered_euclidean_semiring  | 
| 72262 | 1227  | 
begin  | 
1228  | 
||
| 72292 | 1229  | 
lemma unsigned_greater_eq [simp]:  | 
| 72262 | 1230  | 
\<open>0 \<le> unsigned w\<close> for w :: \<open>'b::len word\<close>  | 
1231  | 
by (transfer fixing: less_eq) simp  | 
|
1232  | 
||
| 72292 | 1233  | 
lemma unsigned_less [simp]:  | 
| 72262 | 1234  | 
  \<open>unsigned w < 2 ^ LENGTH('b)\<close> for w :: \<open>'b::len word\<close>
 | 
1235  | 
by (transfer fixing: less) simp  | 
|
1236  | 
||
1237  | 
end  | 
|
1238  | 
||
1239  | 
context linordered_semidom  | 
|
1240  | 
begin  | 
|
1241  | 
||
1242  | 
lemma word_less_eq_iff_unsigned:  | 
|
1243  | 
"a \<le> b \<longleftrightarrow> unsigned a \<le> unsigned b"  | 
|
1244  | 
by (transfer fixing: less_eq) (simp add: nat_le_eq_zle)  | 
|
1245  | 
||
1246  | 
lemma word_less_iff_unsigned:  | 
|
1247  | 
"a < b \<longleftrightarrow> unsigned a < unsigned b"  | 
|
1248  | 
by (transfer fixing: less) (auto dest: preorder_class.le_less_trans [OF take_bit_nonnegative])  | 
|
1249  | 
||
1250  | 
end  | 
|
1251  | 
||
1252  | 
||
1253  | 
subsubsection \<open>Generic signed conversion\<close>  | 
|
1254  | 
||
1255  | 
context ring_bit_operations  | 
|
1256  | 
begin  | 
|
1257  | 
||
| 
72611
 
c7bc3e70a8c7
official collection for bit projection simplifications
 
haftmann 
parents: 
72515 
diff
changeset
 | 
1258  | 
lemma bit_signed_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
 | 
1259  | 
  \<open>bit (signed w) n \<longleftrightarrow> possible_bit TYPE('a) n \<and> bit w (min (LENGTH('b) - Suc 0) n)\<close>
 | 
| 72262 | 1260  | 
for w :: \<open>'b::len word\<close>  | 
1261  | 
by (transfer fixing: bit)  | 
|
| 80777 | 1262  | 
(auto simp: bit_of_int_iff Bit_Operations.bit_signed_take_bit_iff min_def)  | 
| 72262 | 1263  | 
|
1264  | 
lemma signed_push_bit_eq:  | 
|
1265  | 
  \<open>signed (push_bit n w) = signed_take_bit (LENGTH('b) - Suc 0) (push_bit n (signed w :: 'a))\<close>
 | 
|
1266  | 
for w :: \<open>'b::len word\<close>  | 
|
| 82524 | 1267  | 
by (simp add: bit_eq_iff bit_simps possible_bit_less_imp min_less_iff_disj) auto  | 
| 72262 | 1268  | 
|
1269  | 
lemma signed_take_bit_eq:  | 
|
1270  | 
  \<open>signed (take_bit n w) = (if n < LENGTH('b) then take_bit n (signed w) else signed w)\<close>
 | 
|
1271  | 
for w :: \<open>'b::len word\<close>  | 
|
1272  | 
  by (transfer fixing: take_bit; cases \<open>LENGTH('b)\<close>)
 | 
|
| 80777 | 1273  | 
(auto simp: Bit_Operations.signed_take_bit_take_bit Bit_Operations.take_bit_signed_take_bit take_bit_of_int min_def less_Suc_eq)  | 
| 72262 | 1274  | 
|
| 74391 | 1275  | 
context  | 
1276  | 
includes bit_operations_syntax  | 
|
1277  | 
begin  | 
|
1278  | 
||
| 72262 | 1279  | 
lemma signed_not_eq:  | 
1280  | 
  \<open>signed (NOT w) = signed_take_bit LENGTH('b) (NOT (signed w))\<close>
 | 
|
1281  | 
for w :: \<open>'b::len word\<close>  | 
|
| 74592 | 1282  | 
by (simp add: bit_eq_iff bit_simps possible_bit_less_imp min_less_iff_disj)  | 
| 
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
 | 
1283  | 
(auto simp: min_def)  | 
| 72262 | 1284  | 
|
1285  | 
lemma signed_and_eq:  | 
|
1286  | 
\<open>signed (v AND w) = signed v AND signed w\<close>  | 
|
1287  | 
for v w :: \<open>'b::len word\<close>  | 
|
1288  | 
by (rule bit_eqI) (simp add: bit_signed_iff bit_and_iff Bit_Operations.bit_and_iff)  | 
|
1289  | 
||
1290  | 
lemma signed_or_eq:  | 
|
1291  | 
\<open>signed (v OR w) = signed v OR signed w\<close>  | 
|
1292  | 
for v w :: \<open>'b::len word\<close>  | 
|
1293  | 
by (rule bit_eqI) (simp add: bit_signed_iff bit_or_iff Bit_Operations.bit_or_iff)  | 
|
1294  | 
||
1295  | 
lemma signed_xor_eq:  | 
|
1296  | 
\<open>signed (v XOR w) = signed v XOR signed w\<close>  | 
|
1297  | 
for v w :: \<open>'b::len word\<close>  | 
|
1298  | 
by (rule bit_eqI) (simp add: bit_signed_iff bit_xor_iff Bit_Operations.bit_xor_iff)  | 
|
1299  | 
||
1300  | 
end  | 
|
1301  | 
||
| 74097 | 1302  | 
end  | 
1303  | 
||
| 72262 | 1304  | 
|
1305  | 
subsubsection \<open>More\<close>  | 
|
1306  | 
||
1307  | 
lemma sint_greater_eq:  | 
|
1308  | 
  \<open>- (2 ^ (LENGTH('a) - Suc 0)) \<le> sint w\<close> for w :: \<open>'a::len word\<close>
 | 
|
1309  | 
proof (cases \<open>bit w (LENGTH('a) - Suc 0)\<close>)
 | 
|
1310  | 
case True  | 
|
1311  | 
then show ?thesis  | 
|
1312  | 
by transfer (simp add: signed_take_bit_eq_if_negative minus_exp_eq_not_mask or_greater_eq ac_simps)  | 
|
1313  | 
next  | 
|
1314  | 
  have *: \<open>- (2 ^ (LENGTH('a) - Suc 0)) \<le> (0::int)\<close>
 | 
|
1315  | 
by simp  | 
|
1316  | 
case False  | 
|
1317  | 
then show ?thesis  | 
|
| 80777 | 1318  | 
by transfer (auto simp: signed_take_bit_eq intro: order_trans *)  | 
| 72262 | 1319  | 
qed  | 
1320  | 
||
1321  | 
lemma sint_less:  | 
|
1322  | 
  \<open>sint w < 2 ^ (LENGTH('a) - Suc 0)\<close> for w :: \<open>'a::len word\<close>
 | 
|
1323  | 
  by (cases \<open>bit w (LENGTH('a) - Suc 0)\<close>; transfer)
 | 
|
1324  | 
(simp_all add: signed_take_bit_eq signed_take_bit_def not_eq_complement mask_eq_exp_minus_1 OR_upper)  | 
|
1325  | 
||
1326  | 
lemma uint_div_distrib:  | 
|
1327  | 
\<open>uint (v div w) = uint v div uint w\<close>  | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
1328  | 
by (metis int_ops(8) of_nat_unat unat_div_distrib)  | 
| 72262 | 1329  | 
|
| 72388 | 1330  | 
lemma unat_drop_bit_eq:  | 
1331  | 
\<open>unat (drop_bit n w) = drop_bit n (unat w)\<close>  | 
|
1332  | 
by (rule bit_eqI) (simp add: bit_unsigned_iff bit_drop_bit_eq)  | 
|
1333  | 
||
| 72262 | 1334  | 
lemma uint_mod_distrib:  | 
1335  | 
\<open>uint (v mod w) = uint v mod uint w\<close>  | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
1336  | 
by (metis int_ops(9) of_nat_unat unat_mod_distrib)  | 
| 72262 | 1337  | 
|
| 
74108
 
3146646a43a7
simplified hierarchy of type classes for bit operations
 
haftmann 
parents: 
74101 
diff
changeset
 | 
1338  | 
context semiring_bit_operations  | 
| 72262 | 1339  | 
begin  | 
1340  | 
||
1341  | 
lemma unsigned_ucast_eq:  | 
|
1342  | 
  \<open>unsigned (ucast w :: 'c::len word) = take_bit LENGTH('c) (unsigned w)\<close>
 | 
|
1343  | 
for w :: \<open>'b::len word\<close>  | 
|
| 74101 | 1344  | 
by (rule bit_eqI) (simp add: bit_unsigned_iff Word.bit_unsigned_iff bit_take_bit_iff not_le)  | 
| 72262 | 1345  | 
|
1346  | 
end  | 
|
1347  | 
||
1348  | 
context ring_bit_operations  | 
|
1349  | 
begin  | 
|
1350  | 
||
1351  | 
lemma signed_ucast_eq:  | 
|
1352  | 
  \<open>signed (ucast w :: 'c::len word) = signed_take_bit (LENGTH('c) - Suc 0) (unsigned w)\<close>
 | 
|
1353  | 
for w :: \<open>'b::len word\<close>  | 
|
| 74592 | 1354  | 
by (simp add: bit_eq_iff bit_simps min_less_iff_disj)  | 
| 72262 | 1355  | 
|
1356  | 
lemma signed_scast_eq:  | 
|
1357  | 
  \<open>signed (scast w :: 'c::len word) = signed_take_bit (LENGTH('c) - Suc 0) (signed w)\<close>
 | 
|
1358  | 
for w :: \<open>'b::len word\<close>  | 
|
| 74496 | 1359  | 
by (simp add: bit_eq_iff bit_simps min_less_iff_disj)  | 
| 72262 | 1360  | 
|
1361  | 
end  | 
|
1362  | 
||
| 72244 | 1363  | 
lemma uint_nonnegative: "0 \<le> uint w"  | 
| 72262 | 1364  | 
by (fact unsigned_greater_eq)  | 
| 72244 | 1365  | 
|
1366  | 
lemma uint_bounded: "uint w < 2 ^ LENGTH('a)"
 | 
|
1367  | 
for w :: "'a::len word"  | 
|
| 72262 | 1368  | 
by (fact unsigned_less)  | 
| 72244 | 1369  | 
|
1370  | 
lemma uint_idem: "uint w mod 2 ^ LENGTH('a) = uint w"
 | 
|
1371  | 
for w :: "'a::len word"  | 
|
| 72262 | 1372  | 
by transfer (simp add: take_bit_eq_mod)  | 
| 72244 | 1373  | 
|
1374  | 
lemma word_uint_eqI: "uint a = uint b \<Longrightarrow> a = b"  | 
|
| 72262 | 1375  | 
by (fact unsigned_word_eqI)  | 
| 72244 | 1376  | 
|
1377  | 
lemma word_uint_eq_iff: "a = b \<longleftrightarrow> uint a = uint b"  | 
|
| 72262 | 1378  | 
by (fact word_eq_iff_unsigned)  | 
1379  | 
||
1380  | 
lemma uint_word_of_int_eq:  | 
|
| 72244 | 1381  | 
  \<open>uint (word_of_int k :: 'a::len word) = take_bit LENGTH('a) k\<close>
 | 
1382  | 
by transfer rule  | 
|
1383  | 
||
1384  | 
lemma uint_word_of_int: "uint (word_of_int k :: 'a::len word) = k mod 2 ^ LENGTH('a)"
 | 
|
1385  | 
by (simp add: uint_word_of_int_eq take_bit_eq_mod)  | 
|
1386  | 
||
1387  | 
lemma word_of_int_uint: "word_of_int (uint w) = w"  | 
|
1388  | 
by transfer simp  | 
|
1389  | 
||
1390  | 
lemma word_div_def [code]:  | 
|
1391  | 
"a div b = word_of_int (uint a div uint b)"  | 
|
1392  | 
by transfer rule  | 
|
1393  | 
||
1394  | 
lemma word_mod_def [code]:  | 
|
1395  | 
"a mod b = word_of_int (uint a mod uint b)"  | 
|
1396  | 
by transfer rule  | 
|
1397  | 
||
1398  | 
lemma split_word_all: "(\<And>x::'a::len word. PROP P x) \<equiv> (\<And>x. PROP P (word_of_int x))"  | 
|
1399  | 
proof  | 
|
1400  | 
fix x :: "'a word"  | 
|
1401  | 
assume "\<And>x. PROP P (word_of_int x)"  | 
|
1402  | 
then have "PROP P (word_of_int (uint x))" .  | 
|
| 72262 | 1403  | 
then show "PROP P x"  | 
1404  | 
by (simp only: word_of_int_uint)  | 
|
| 72244 | 1405  | 
qed  | 
1406  | 
||
| 72262 | 1407  | 
lemma sint_uint:  | 
1408  | 
  \<open>sint w = signed_take_bit (LENGTH('a) - Suc 0) (uint w)\<close>
 | 
|
| 72244 | 1409  | 
for w :: \<open>'a::len word\<close>  | 
1410  | 
  by (cases \<open>LENGTH('a)\<close>; transfer) (simp_all add: signed_take_bit_take_bit)
 | 
|
1411  | 
||
| 72262 | 1412  | 
lemma unat_eq_nat_uint:  | 
| 72244 | 1413  | 
\<open>unat w = nat (uint w)\<close>  | 
1414  | 
by simp  | 
|
1415  | 
||
| 72262 | 1416  | 
lemma ucast_eq:  | 
| 72244 | 1417  | 
\<open>ucast w = word_of_int (uint w)\<close>  | 
1418  | 
by transfer simp  | 
|
1419  | 
||
| 72262 | 1420  | 
lemma scast_eq:  | 
| 72244 | 1421  | 
\<open>scast w = word_of_int (sint w)\<close>  | 
1422  | 
by transfer simp  | 
|
1423  | 
||
| 72262 | 1424  | 
lemma uint_0_eq:  | 
| 72244 | 1425  | 
\<open>uint 0 = 0\<close>  | 
| 72262 | 1426  | 
by (fact unsigned_0)  | 
1427  | 
||
1428  | 
lemma uint_1_eq:  | 
|
| 72244 | 1429  | 
\<open>uint 1 = 1\<close>  | 
| 72262 | 1430  | 
by (fact unsigned_1)  | 
| 72244 | 1431  | 
|
1432  | 
lemma word_m1_wi: "- 1 = word_of_int (- 1)"  | 
|
| 72262 | 1433  | 
by simp  | 
| 72244 | 1434  | 
|
1435  | 
lemma uint_0_iff: "uint x = 0 \<longleftrightarrow> x = 0"  | 
|
| 80777 | 1436  | 
by (auto simp: unsigned_word_eqI)  | 
| 72244 | 1437  | 
|
1438  | 
lemma unat_0_iff: "unat x = 0 \<longleftrightarrow> x = 0"  | 
|
| 80777 | 1439  | 
by (auto simp: unsigned_word_eqI)  | 
| 72262 | 1440  | 
|
1441  | 
lemma unat_0: "unat 0 = 0"  | 
|
1442  | 
by (fact unsigned_0)  | 
|
| 72244 | 1443  | 
|
1444  | 
lemma unat_gt_0: "0 < unat x \<longleftrightarrow> x \<noteq> 0"  | 
|
1445  | 
by (auto simp: unat_0_iff [symmetric])  | 
|
1446  | 
||
| 72262 | 1447  | 
lemma ucast_0: "ucast 0 = 0"  | 
1448  | 
by (fact unsigned_0)  | 
|
1449  | 
||
1450  | 
lemma sint_0: "sint 0 = 0"  | 
|
1451  | 
by (fact signed_0)  | 
|
1452  | 
||
1453  | 
lemma scast_0: "scast 0 = 0"  | 
|
1454  | 
by (fact signed_0)  | 
|
1455  | 
||
1456  | 
lemma sint_n1: "sint (- 1) = - 1"  | 
|
1457  | 
by (fact signed_minus_1)  | 
|
1458  | 
||
1459  | 
lemma scast_n1: "scast (- 1) = - 1"  | 
|
1460  | 
by (fact signed_minus_1)  | 
|
| 72244 | 1461  | 
|
1462  | 
lemma uint_1: "uint (1::'a::len word) = 1"  | 
|
1463  | 
by (fact uint_1_eq)  | 
|
1464  | 
||
| 72262 | 1465  | 
lemma unat_1: "unat (1::'a::len word) = 1"  | 
1466  | 
by (fact unsigned_1)  | 
|
1467  | 
||
1468  | 
lemma ucast_1: "ucast (1::'a::len word) = 1"  | 
|
1469  | 
by (fact unsigned_1)  | 
|
| 72244 | 1470  | 
|
1471  | 
instantiation word :: (len) size  | 
|
1472  | 
begin  | 
|
1473  | 
||
1474  | 
lift_definition size_word :: \<open>'a word \<Rightarrow> nat\<close>  | 
|
1475  | 
  is \<open>\<lambda>_. LENGTH('a)\<close> ..
 | 
|
1476  | 
||
1477  | 
instance ..  | 
|
1478  | 
||
1479  | 
end  | 
|
1480  | 
||
1481  | 
lemma word_size [code]:  | 
|
1482  | 
  \<open>size w = LENGTH('a)\<close> for w :: \<open>'a::len word\<close>
 | 
|
1483  | 
by (fact size_word.rep_eq)  | 
|
1484  | 
||
1485  | 
lemma word_size_gt_0 [iff]: "0 < size w"  | 
|
1486  | 
for w :: "'a::len word"  | 
|
1487  | 
by (simp add: word_size)  | 
|
1488  | 
||
1489  | 
lemmas lens_gt_0 = word_size_gt_0 len_gt_0  | 
|
1490  | 
||
1491  | 
lemma lens_not_0 [iff]:  | 
|
1492  | 
\<open>size w \<noteq> 0\<close> for w :: \<open>'a::len word\<close>  | 
|
1493  | 
by auto  | 
|
1494  | 
||
1495  | 
lift_definition source_size :: \<open>('a::len word \<Rightarrow> 'b) \<Rightarrow> nat\<close>
 | 
|
1496  | 
  is \<open>\<lambda>_. LENGTH('a)\<close> .
 | 
|
1497  | 
||
1498  | 
lift_definition target_size :: \<open>('a \<Rightarrow> 'b::len word) \<Rightarrow> nat\<close>
 | 
|
1499  | 
  is \<open>\<lambda>_. LENGTH('b)\<close> ..
 | 
|
1500  | 
||
1501  | 
lift_definition is_up :: \<open>('a::len word \<Rightarrow> 'b::len word) \<Rightarrow> bool\<close>
 | 
|
1502  | 
  is \<open>\<lambda>_. LENGTH('a) \<le> LENGTH('b)\<close> ..
 | 
|
1503  | 
||
1504  | 
lift_definition is_down :: \<open>('a::len word \<Rightarrow> 'b::len word) \<Rightarrow> bool\<close>
 | 
|
1505  | 
  is \<open>\<lambda>_. LENGTH('a) \<ge> LENGTH('b)\<close> ..
 | 
|
1506  | 
||
1507  | 
lemma is_up_eq:  | 
|
1508  | 
\<open>is_up f \<longleftrightarrow> source_size f \<le> target_size f\<close>  | 
|
1509  | 
for f :: \<open>'a::len word \<Rightarrow> 'b::len word\<close>  | 
|
1510  | 
by (simp add: source_size.rep_eq target_size.rep_eq is_up.rep_eq)  | 
|
1511  | 
||
1512  | 
lemma is_down_eq:  | 
|
1513  | 
\<open>is_down f \<longleftrightarrow> target_size f \<le> source_size f\<close>  | 
|
1514  | 
for f :: \<open>'a::len word \<Rightarrow> 'b::len word\<close>  | 
|
1515  | 
by (simp add: source_size.rep_eq target_size.rep_eq is_down.rep_eq)  | 
|
1516  | 
||
1517  | 
lift_definition word_int_case :: \<open>(int \<Rightarrow> 'b) \<Rightarrow> 'a::len word \<Rightarrow> 'b\<close>  | 
|
1518  | 
  is \<open>\<lambda>f. f \<circ> take_bit LENGTH('a)\<close> by simp
 | 
|
1519  | 
||
1520  | 
lemma word_int_case_eq_uint [code]:  | 
|
1521  | 
\<open>word_int_case f w = f (uint w)\<close>  | 
|
1522  | 
by transfer simp  | 
|
1523  | 
||
1524  | 
translations  | 
|
1525  | 
"case x of XCONST of_int y \<Rightarrow> b" \<rightleftharpoons> "CONST word_int_case (\<lambda>y. b) x"  | 
|
1526  | 
"case x of (XCONST of_int :: 'a) y \<Rightarrow> b" \<rightharpoonup> "CONST word_int_case (\<lambda>y. b) x"  | 
|
1527  | 
||
1528  | 
||
1529  | 
subsection \<open>Arithmetic operations\<close>  | 
|
1530  | 
||
| 74592 | 1531  | 
lemma div_word_self:  | 
1532  | 
\<open>w div w = 1\<close> if \<open>w \<noteq> 0\<close> for w :: \<open>'a::len word\<close>  | 
|
1533  | 
using that by transfer simp  | 
|
1534  | 
||
1535  | 
lemma mod_word_self [simp]:  | 
|
1536  | 
\<open>w mod w = 0\<close> for w :: \<open>'a::len word\<close>  | 
|
| 80777 | 1537  | 
by (simp add: word_mod_def)  | 
| 74592 | 1538  | 
|
1539  | 
lemma div_word_less:  | 
|
1540  | 
\<open>w div v = 0\<close> if \<open>w < v\<close> for w v :: \<open>'a::len word\<close>  | 
|
1541  | 
using that by transfer simp  | 
|
1542  | 
||
1543  | 
lemma mod_word_less:  | 
|
1544  | 
\<open>w mod v = w\<close> if \<open>w < v\<close> for w v :: \<open>'a::len word\<close>  | 
|
1545  | 
using div_mult_mod_eq [of w v] using that by (simp add: div_word_less)  | 
|
1546  | 
||
1547  | 
lemma div_word_one [simp]:  | 
|
1548  | 
\<open>1 div w = of_bool (w = 1)\<close> for w :: \<open>'a::len word\<close>  | 
|
1549  | 
proof transfer  | 
|
1550  | 
fix k :: int  | 
|
1551  | 
  show \<open>take_bit LENGTH('a) (take_bit LENGTH('a) 1 div take_bit LENGTH('a) k) =
 | 
|
1552  | 
         take_bit LENGTH('a) (of_bool (take_bit LENGTH('a) k = take_bit LENGTH('a) 1))\<close>
 | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
1553  | 
    using take_bit_nonnegative [of \<open>LENGTH('a)\<close> k]
 | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
1554  | 
by (smt (verit, best) div_by_1 of_bool_eq take_bit_of_0 take_bit_of_1 zdiv_eq_0_iff)  | 
| 74592 | 1555  | 
qed  | 
1556  | 
||
1557  | 
lemma mod_word_one [simp]:  | 
|
1558  | 
\<open>1 mod w = 1 - w * of_bool (w = 1)\<close> for w :: \<open>'a::len word\<close>  | 
|
| 75087 | 1559  | 
using div_mult_mod_eq [of 1 w] by auto  | 
| 74592 | 1560  | 
|
1561  | 
lemma div_word_by_minus_1_eq [simp]:  | 
|
1562  | 
\<open>w div - 1 = of_bool (w = - 1)\<close> for w :: \<open>'a::len word\<close>  | 
|
1563  | 
by (auto intro: div_word_less simp add: div_word_self word_order.not_eq_extremum)  | 
|
1564  | 
||
1565  | 
lemma mod_word_by_minus_1_eq [simp]:  | 
|
1566  | 
\<open>w mod - 1 = w * of_bool (w < - 1)\<close> for w :: \<open>'a::len word\<close>  | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
1567  | 
using mod_word_less word_order.not_eq_extremum by fastforce  | 
| 74592 | 1568  | 
|
| 72244 | 1569  | 
text \<open>Legacy theorems:\<close>  | 
1570  | 
||
1571  | 
lemma word_add_def [code]:  | 
|
1572  | 
"a + b = word_of_int (uint a + uint b)"  | 
|
1573  | 
by transfer (simp add: take_bit_add)  | 
|
1574  | 
||
1575  | 
lemma word_sub_wi [code]:  | 
|
1576  | 
"a - b = word_of_int (uint a - uint b)"  | 
|
1577  | 
by transfer (simp add: take_bit_diff)  | 
|
1578  | 
||
1579  | 
lemma word_mult_def [code]:  | 
|
1580  | 
"a * b = word_of_int (uint a * uint b)"  | 
|
1581  | 
by transfer (simp add: take_bit_eq_mod mod_simps)  | 
|
1582  | 
||
1583  | 
lemma word_minus_def [code]:  | 
|
1584  | 
"- a = word_of_int (- uint a)"  | 
|
1585  | 
by transfer (simp add: take_bit_minus)  | 
|
1586  | 
||
1587  | 
lemma word_0_wi:  | 
|
1588  | 
"0 = word_of_int 0"  | 
|
1589  | 
by transfer simp  | 
|
1590  | 
||
1591  | 
lemma word_1_wi:  | 
|
1592  | 
"1 = word_of_int 1"  | 
|
1593  | 
by transfer simp  | 
|
1594  | 
||
1595  | 
lift_definition word_succ :: "'a::len word \<Rightarrow> 'a word" is "\<lambda>x. x + 1"  | 
|
| 80777 | 1596  | 
by (auto simp: take_bit_eq_mod intro: mod_add_cong)  | 
| 72244 | 1597  | 
|
1598  | 
lift_definition word_pred :: "'a::len word \<Rightarrow> 'a word" is "\<lambda>x. x - 1"  | 
|
| 80777 | 1599  | 
by (auto simp: take_bit_eq_mod intro: mod_diff_cong)  | 
| 72244 | 1600  | 
|
1601  | 
lemma word_succ_alt [code]:  | 
|
1602  | 
"word_succ a = word_of_int (uint a + 1)"  | 
|
1603  | 
by transfer (simp add: take_bit_eq_mod mod_simps)  | 
|
1604  | 
||
1605  | 
lemma word_pred_alt [code]:  | 
|
1606  | 
"word_pred a = word_of_int (uint a - 1)"  | 
|
1607  | 
by transfer (simp add: take_bit_eq_mod mod_simps)  | 
|
1608  | 
||
1609  | 
lemmas word_arith_wis =  | 
|
1610  | 
word_add_def word_sub_wi word_mult_def  | 
|
1611  | 
word_minus_def word_succ_alt word_pred_alt  | 
|
1612  | 
word_0_wi word_1_wi  | 
|
1613  | 
||
1614  | 
lemma wi_homs:  | 
|
1615  | 
shows wi_hom_add: "word_of_int a + word_of_int b = word_of_int (a + b)"  | 
|
1616  | 
and wi_hom_sub: "word_of_int a - word_of_int b = word_of_int (a - b)"  | 
|
1617  | 
and wi_hom_mult: "word_of_int a * word_of_int b = word_of_int (a * b)"  | 
|
1618  | 
and wi_hom_neg: "- word_of_int a = word_of_int (- a)"  | 
|
1619  | 
and wi_hom_succ: "word_succ (word_of_int a) = word_of_int (a + 1)"  | 
|
1620  | 
and wi_hom_pred: "word_pred (word_of_int a) = word_of_int (a - 1)"  | 
|
1621  | 
by (transfer, simp)+  | 
|
1622  | 
||
1623  | 
lemmas wi_hom_syms = wi_homs [symmetric]  | 
|
1624  | 
||
1625  | 
lemmas word_of_int_homs = wi_homs word_0_wi word_1_wi  | 
|
1626  | 
||
1627  | 
lemmas word_of_int_hom_syms = word_of_int_homs [symmetric]  | 
|
1628  | 
||
1629  | 
lemma double_eq_zero_iff:  | 
|
1630  | 
  \<open>2 * a = 0 \<longleftrightarrow> a = 0 \<or> a = 2 ^ (LENGTH('a) - Suc 0)\<close>
 | 
|
1631  | 
for a :: \<open>'a::len word\<close>  | 
|
1632  | 
proof -  | 
|
1633  | 
  define n where \<open>n = LENGTH('a) - Suc 0\<close>
 | 
|
1634  | 
  then have *: \<open>LENGTH('a) = Suc n\<close>
 | 
|
1635  | 
by simp  | 
|
1636  | 
  have \<open>a = 0\<close> if \<open>2 * a = 0\<close> and \<open>a \<noteq> 2 ^ (LENGTH('a) - Suc 0)\<close>
 | 
|
1637  | 
using that by transfer  | 
|
| 80777 | 1638  | 
(auto simp: take_bit_eq_0_iff take_bit_eq_mod *)  | 
| 72244 | 1639  | 
  moreover have \<open>2 ^ LENGTH('a) = (0 :: 'a word)\<close>
 | 
1640  | 
by transfer simp  | 
|
1641  | 
  then have \<open>2 * 2 ^ (LENGTH('a) - Suc 0) = (0 :: 'a word)\<close>
 | 
|
1642  | 
by (simp add: *)  | 
|
1643  | 
ultimately show ?thesis  | 
|
1644  | 
by auto  | 
|
1645  | 
qed  | 
|
1646  | 
||
1647  | 
||
1648  | 
subsection \<open>Ordering\<close>  | 
|
1649  | 
||
| 82376 | 1650  | 
instance word :: (len) wellorder  | 
1651  | 
proof  | 
|
1652  | 
fix P :: "'a word \<Rightarrow> bool" and a  | 
|
1653  | 
assume *: "(\<And>b. (\<And>a. a < b \<Longrightarrow> P a) \<Longrightarrow> P b)"  | 
|
1654  | 
have "wf (measure unat)" ..  | 
|
1655  | 
  moreover have "{(a, b :: ('a::len) word). a < b} \<subseteq> measure unat"
 | 
|
1656  | 
by (auto simp add: word_less_iff_unsigned [where ?'a = nat])  | 
|
1657  | 
  ultimately have "wf {(a, b :: ('a::len) word). a < b}"
 | 
|
1658  | 
by (rule wf_subset)  | 
|
1659  | 
then show "P a" using *  | 
|
1660  | 
by induction blast  | 
|
1661  | 
qed  | 
|
1662  | 
||
1663  | 
lemma word_m1_ge [simp]: (* FIXME: delete *)  | 
|
1664  | 
"word_pred 0 \<ge> y"  | 
|
1665  | 
by transfer (simp add: mask_eq_exp_minus_1)  | 
|
1666  | 
||
1667  | 
lemma word_less_alt:  | 
|
1668  | 
"a < b \<longleftrightarrow> uint a < uint b"  | 
|
1669  | 
by (fact word_less_def)  | 
|
1670  | 
||
1671  | 
lemma word_zero_le [simp]:  | 
|
1672  | 
"0 \<le> y" for y :: "'a::len word"  | 
|
1673  | 
by (fact word_coorder.extremum)  | 
|
1674  | 
||
1675  | 
lemma word_n1_ge [simp]:  | 
|
1676  | 
"y \<le> - 1" for y :: "'a::len word"  | 
|
1677  | 
by (fact word_order.extremum)  | 
|
1678  | 
||
1679  | 
lemmas word_not_simps [simp] =  | 
|
1680  | 
word_zero_le [THEN leD] word_m1_ge [THEN leD] word_n1_ge [THEN leD]  | 
|
1681  | 
||
1682  | 
lemma word_gt_0:  | 
|
1683  | 
"0 < y \<longleftrightarrow> 0 \<noteq> y"  | 
|
1684  | 
for y :: "'a::len word"  | 
|
1685  | 
by (simp add: less_le)  | 
|
1686  | 
||
1687  | 
lemma word_gt_0_no [simp]:  | 
|
1688  | 
\<open>(0 :: 'a::len word) < numeral y \<longleftrightarrow> (0 :: 'a::len word) \<noteq> numeral y\<close>  | 
|
1689  | 
by (fact word_gt_0)  | 
|
1690  | 
||
1691  | 
lemma word_le_nat_alt:  | 
|
1692  | 
"a \<le> b \<longleftrightarrow> unat a \<le> unat b"  | 
|
1693  | 
by transfer (simp add: nat_le_eq_zle)  | 
|
1694  | 
||
1695  | 
lemma word_less_nat_alt:  | 
|
1696  | 
"a < b \<longleftrightarrow> unat a < unat b"  | 
|
1697  | 
by transfer (auto simp: less_le [of 0])  | 
|
1698  | 
||
1699  | 
lemmas unat_mono = word_less_nat_alt [THEN iffD1]  | 
|
1700  | 
||
1701  | 
lemma wi_less:  | 
|
1702  | 
"(word_of_int n < (word_of_int m :: 'a::len word)) =  | 
|
1703  | 
    (n mod 2 ^ LENGTH('a) < m mod 2 ^ LENGTH('a))"
 | 
|
1704  | 
by (simp add: uint_word_of_int word_less_def)  | 
|
1705  | 
||
1706  | 
lemma wi_le:  | 
|
1707  | 
"(word_of_int n \<le> (word_of_int m :: 'a::len word)) =  | 
|
1708  | 
    (n mod 2 ^ LENGTH('a) \<le> m mod 2 ^ LENGTH('a))"
 | 
|
1709  | 
by (simp add: uint_word_of_int word_le_def)  | 
|
1710  | 
||
| 72388 | 1711  | 
lift_definition word_sle :: \<open>'a::len word \<Rightarrow> 'a word \<Rightarrow> bool\<close>  | 
1712  | 
  is \<open>\<lambda>k l. signed_take_bit (LENGTH('a) - Suc 0) k \<le> signed_take_bit (LENGTH('a) - Suc 0) l\<close>
 | 
|
1713  | 
by (simp flip: signed_take_bit_decr_length_iff)  | 
|
1714  | 
||
1715  | 
lift_definition word_sless :: \<open>'a::len word \<Rightarrow> 'a word \<Rightarrow> bool\<close>  | 
|
1716  | 
  is \<open>\<lambda>k l. signed_take_bit (LENGTH('a) - Suc 0) k < signed_take_bit (LENGTH('a) - Suc 0) l\<close>
 | 
|
| 72244 | 1717  | 
by (simp flip: signed_take_bit_decr_length_iff)  | 
1718  | 
||
| 72388 | 1719  | 
notation  | 
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
80777 
diff
changeset
 | 
1720  | 
word_sle (\<open>'(\<le>s')\<close>) and  | 
| 81142 | 1721  | 
word_sle (\<open>(\<open>notation=\<open>infix \<le>s\<close>\<close>_/ \<le>s _)\<close> [51, 51] 50) and  | 
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
80777 
diff
changeset
 | 
1722  | 
word_sless (\<open>'(<s')\<close>) and  | 
| 81142 | 1723  | 
word_sless (\<open>(\<open>notation=\<open>infix <s\<close>\<close>_/ <s _)\<close> [51, 51] 50)  | 
| 72388 | 1724  | 
|
1725  | 
notation (input)  | 
|
| 81142 | 1726  | 
word_sle (\<open>(\<open>notation=\<open>infix <=s\<close>\<close>_/ <=s _)\<close> [51, 51] 50)  | 
| 72388 | 1727  | 
|
| 72244 | 1728  | 
lemma word_sle_eq [code]:  | 
1729  | 
\<open>a <=s b \<longleftrightarrow> sint a \<le> sint b\<close>  | 
|
1730  | 
by transfer simp  | 
|
1731  | 
||
| 82376 | 1732  | 
lemma word_sless_alt [code]:  | 
1733  | 
"a <s b \<longleftrightarrow> sint a < sint b"  | 
|
| 72244 | 1734  | 
by transfer simp  | 
1735  | 
||
| 72388 | 1736  | 
lemma signed_ordering: \<open>ordering word_sle word_sless\<close>  | 
| 82376 | 1737  | 
by (standard; transfer) (auto simp flip: signed_take_bit_decr_length_iff)  | 
| 72388 | 1738  | 
|
1739  | 
lemma signed_linorder: \<open>class.linorder word_sle word_sless\<close>  | 
|
| 80777 | 1740  | 
by (standard; transfer) (auto simp: signed_take_bit_decr_length_iff)  | 
| 72388 | 1741  | 
|
1742  | 
interpretation signed: linorder word_sle word_sless  | 
|
1743  | 
by (fact signed_linorder)  | 
|
1744  | 
||
1745  | 
lemma word_sless_eq:  | 
|
1746  | 
\<open>x <s y \<longleftrightarrow> x <=s y \<and> x \<noteq> y\<close>  | 
|
1747  | 
by (fact signed.less_le)  | 
|
1748  | 
||
| 82524 | 1749  | 
lemma minus_1_sless_0 [simp]:  | 
1750  | 
\<open>- 1 <s 0\<close>  | 
|
1751  | 
by transfer simp  | 
|
1752  | 
||
| 82525 | 1753  | 
lemma not_0_sless_minus_1 [simp]:  | 
1754  | 
\<open>\<not> 0 <s - 1\<close>  | 
|
1755  | 
by transfer simp  | 
|
1756  | 
||
| 82524 | 1757  | 
lemma minus_1_sless_eq_0 [simp]:  | 
1758  | 
\<open>- 1 \<le>s 0\<close>  | 
|
1759  | 
by transfer simp  | 
|
1760  | 
||
| 82525 | 1761  | 
lemma not_0_sless_eq_minus_1 [simp]:  | 
1762  | 
\<open>\<not> 0 \<le>s - 1\<close>  | 
|
1763  | 
by transfer simp  | 
|
1764  | 
||
| 72244 | 1765  | 
|
1766  | 
subsection \<open>Bit-wise operations\<close>  | 
|
1767  | 
||
| 74097 | 1768  | 
context  | 
1769  | 
includes bit_operations_syntax  | 
|
1770  | 
begin  | 
|
1771  | 
||
| 72227 | 1772  | 
lemma take_bit_word_eq_self:  | 
1773  | 
  \<open>take_bit n w = w\<close> if \<open>LENGTH('a) \<le> n\<close> for w :: \<open>'a::len word\<close>
 | 
|
1774  | 
using that by transfer simp  | 
|
1775  | 
||
| 
72027
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
1776  | 
lemma take_bit_length_eq [simp]:  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
1777  | 
  \<open>take_bit LENGTH('a) w = w\<close> for w :: \<open>'a::len word\<close>
 | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
1778  | 
by (simp add: nat_le_linear take_bit_word_eq_self)  | 
| 
72027
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
1779  | 
|
| 71990 | 1780  | 
lemma bit_word_of_int_iff:  | 
1781  | 
  \<open>bit (word_of_int k :: 'a::len word) n \<longleftrightarrow> n < LENGTH('a) \<and> bit k n\<close>
 | 
|
| 82525 | 1782  | 
by (simp add: bit_simps)  | 
| 71990 | 1783  | 
|
1784  | 
lemma bit_uint_iff:  | 
|
1785  | 
  \<open>bit (uint w) n \<longleftrightarrow> n < LENGTH('a) \<and> bit w n\<close>
 | 
|
1786  | 
for w :: \<open>'a::len word\<close>  | 
|
1787  | 
by transfer (simp add: bit_take_bit_iff)  | 
|
1788  | 
||
1789  | 
lemma bit_sint_iff:  | 
|
1790  | 
  \<open>bit (sint w) n \<longleftrightarrow> n \<ge> LENGTH('a) \<and> bit w (LENGTH('a) - 1) \<or> bit w n\<close>
 | 
|
1791  | 
for w :: \<open>'a::len word\<close>  | 
|
| 80777 | 1792  | 
by transfer (auto simp: bit_signed_take_bit_iff min_def le_less not_less)  | 
| 71990 | 1793  | 
|
1794  | 
lemma bit_word_ucast_iff:  | 
|
1795  | 
  \<open>bit (ucast w :: 'b::len word) n \<longleftrightarrow> n < LENGTH('a) \<and> n < LENGTH('b) \<and> bit w n\<close>
 | 
|
1796  | 
for w :: \<open>'a::len word\<close>  | 
|
| 82525 | 1797  | 
by (auto simp add: bit_unsigned_iff bit_imp_le_length)  | 
| 71990 | 1798  | 
|
1799  | 
lemma bit_word_scast_iff:  | 
|
1800  | 
\<open>bit (scast w :: 'b::len word) n \<longleftrightarrow>  | 
|
1801  | 
    n < LENGTH('b) \<and> (bit w n \<or> LENGTH('a) \<le> n \<and> bit w (LENGTH('a) - Suc 0))\<close>
 | 
|
1802  | 
for w :: \<open>'a::len word\<close>  | 
|
| 82525 | 1803  | 
by (auto simp add: bit_signed_iff bit_imp_le_length min_def le_less dest: bit_imp_possible_bit)  | 
| 72079 | 1804  | 
|
| 
72088
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
1805  | 
lemma bit_word_iff_drop_bit_and [code]:  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
1806  | 
\<open>bit a n \<longleftrightarrow> drop_bit n a AND 1 = 1\<close> for a :: \<open>'a::len word\<close>  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
1807  | 
by (simp add: bit_iff_odd_drop_bit odd_iff_mod_2_eq_one and_one_eq)  | 
| 72079 | 1808  | 
|
1809  | 
lemma  | 
|
| 72262 | 1810  | 
word_not_def: "NOT (a::'a::len word) = word_of_int (NOT (uint a))"  | 
| 65268 | 1811  | 
and word_and_def: "(a::'a word) AND b = word_of_int (uint a AND uint b)"  | 
1812  | 
and word_or_def: "(a::'a word) OR b = word_of_int (uint a OR uint b)"  | 
|
1813  | 
and word_xor_def: "(a::'a word) XOR b = word_of_int (uint a XOR uint b)"  | 
|
| 
71957
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
1814  | 
by (transfer, simp add: take_bit_not_take_bit)+  | 
| 
47374
 
9475d524bafb
set up and use lift_definition for word operations
 
huffman 
parents: 
47372 
diff
changeset
 | 
1815  | 
|
| 
71957
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
1816  | 
definition even_word :: \<open>'a::len word \<Rightarrow> bool\<close>  | 
| 
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
1817  | 
where [code_abbrev]: \<open>even_word = even\<close>  | 
| 
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
1818  | 
|
| 
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
1819  | 
lemma even_word_iff [code]:  | 
| 
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
1820  | 
\<open>even_word a \<longleftrightarrow> a AND 1 = 0\<close>  | 
| 
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
1821  | 
by (simp add: and_one_eq even_iff_mod_2_eq_zero even_word_def)  | 
| 
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
1822  | 
|
| 72079 | 1823  | 
lemma map_bit_range_eq_if_take_bit_eq:  | 
1824  | 
\<open>map (bit k) [0..<n] = map (bit l) [0..<n]\<close>  | 
|
1825  | 
if \<open>take_bit n k = take_bit n l\<close> for k l :: int  | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
1826  | 
using that  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
1827  | 
proof (induction n arbitrary: k l)  | 
| 72079 | 1828  | 
case 0  | 
1829  | 
then show ?case  | 
|
1830  | 
by simp  | 
|
1831  | 
next  | 
|
1832  | 
case (Suc n)  | 
|
1833  | 
from Suc.prems have \<open>take_bit n (k div 2) = take_bit n (l div 2)\<close>  | 
|
1834  | 
by (simp add: take_bit_Suc)  | 
|
1835  | 
then have \<open>map (bit (k div 2)) [0..<n] = map (bit (l div 2)) [0..<n]\<close>  | 
|
1836  | 
by (rule Suc.IH)  | 
|
1837  | 
moreover have \<open>bit (r div 2) = bit r \<circ> Suc\<close> for r :: int  | 
|
1838  | 
by (simp add: fun_eq_iff bit_Suc)  | 
|
1839  | 
moreover from Suc.prems have \<open>even k \<longleftrightarrow> even l\<close>  | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
1840  | 
by (metis Zero_neq_Suc even_take_bit_eq)  | 
| 72079 | 1841  | 
ultimately show ?case  | 
| 75085 | 1842  | 
by (simp only: map_Suc_upt upt_conv_Cons flip: list.map_comp) (simp add: bit_0)  | 
| 72079 | 1843  | 
qed  | 
1844  | 
||
| 72262 | 1845  | 
lemma  | 
1846  | 
take_bit_word_Bit0_eq [simp]: \<open>take_bit (numeral n) (numeral (num.Bit0 m) :: 'a::len word)  | 
|
1847  | 
= 2 * take_bit (pred_numeral n) (numeral m)\<close> (is ?P)  | 
|
1848  | 
and take_bit_word_Bit1_eq [simp]: \<open>take_bit (numeral n) (numeral (num.Bit1 m) :: 'a::len word)  | 
|
1849  | 
= 1 + 2 * take_bit (pred_numeral n) (numeral m)\<close> (is ?Q)  | 
|
1850  | 
and take_bit_word_minus_Bit0_eq [simp]: \<open>take_bit (numeral n) (- numeral (num.Bit0 m) :: 'a::len word)  | 
|
1851  | 
= 2 * take_bit (pred_numeral n) (- numeral m)\<close> (is ?R)  | 
|
1852  | 
and take_bit_word_minus_Bit1_eq [simp]: \<open>take_bit (numeral n) (- numeral (num.Bit1 m) :: 'a::len word)  | 
|
1853  | 
= 1 + 2 * take_bit (pred_numeral n) (- numeral (Num.inc m))\<close> (is ?S)  | 
|
1854  | 
proof -  | 
|
1855  | 
define w :: \<open>'a::len word\<close>  | 
|
1856  | 
where \<open>w = numeral m\<close>  | 
|
1857  | 
moreover define q :: nat  | 
|
1858  | 
where \<open>q = pred_numeral n\<close>  | 
|
1859  | 
ultimately have num:  | 
|
1860  | 
\<open>numeral m = w\<close>  | 
|
1861  | 
\<open>numeral (num.Bit0 m) = 2 * w\<close>  | 
|
1862  | 
\<open>numeral (num.Bit1 m) = 1 + 2 * w\<close>  | 
|
1863  | 
\<open>numeral (Num.inc m) = 1 + w\<close>  | 
|
1864  | 
\<open>pred_numeral n = q\<close>  | 
|
1865  | 
\<open>numeral n = Suc q\<close>  | 
|
1866  | 
by (simp_all only: w_def q_def numeral_Bit0 [of m] numeral_Bit1 [of m] ac_simps  | 
|
1867  | 
numeral_inc numeral_eq_Suc flip: mult_2)  | 
|
1868  | 
have even: \<open>take_bit (Suc q) (2 * w) = 2 * take_bit q w\<close> for w :: \<open>'a::len word\<close>  | 
|
1869  | 
by (rule bit_word_eqI)  | 
|
| 80777 | 1870  | 
(auto simp: bit_take_bit_iff bit_double_iff)  | 
| 72262 | 1871  | 
have odd: \<open>take_bit (Suc q) (1 + 2 * w) = 1 + 2 * take_bit q w\<close> for w :: \<open>'a::len word\<close>  | 
1872  | 
by (rule bit_eqI)  | 
|
| 80777 | 1873  | 
(auto simp: bit_take_bit_iff bit_double_iff even_bit_succ_iff)  | 
| 72262 | 1874  | 
show ?P  | 
1875  | 
using even [of w] by (simp add: num)  | 
|
1876  | 
show ?Q  | 
|
1877  | 
using odd [of w] by (simp add: num)  | 
|
1878  | 
show ?R  | 
|
1879  | 
using even [of \<open>- w\<close>] by (simp add: num)  | 
|
1880  | 
show ?S  | 
|
1881  | 
using odd [of \<open>- (1 + w)\<close>] by (simp add: num)  | 
|
1882  | 
qed  | 
|
1883  | 
||
| 72079 | 1884  | 
|
1885  | 
subsection \<open>More shift operations\<close>  | 
|
1886  | 
||
| 72388 | 1887  | 
lift_definition signed_drop_bit :: \<open>nat \<Rightarrow> 'a word \<Rightarrow> 'a::len word\<close>  | 
1888  | 
  is \<open>\<lambda>n. drop_bit n \<circ> signed_take_bit (LENGTH('a) - Suc 0)\<close>
 | 
|
1889  | 
using signed_take_bit_decr_length_iff  | 
|
1890  | 
by (simp add: take_bit_drop_bit) force  | 
|
1891  | 
||
| 
72611
 
c7bc3e70a8c7
official collection for bit projection simplifications
 
haftmann 
parents: 
72515 
diff
changeset
 | 
1892  | 
lemma bit_signed_drop_bit_iff [bit_simps]:  | 
| 72388 | 1893  | 
  \<open>bit (signed_drop_bit m w) n \<longleftrightarrow> bit w (if LENGTH('a) - m \<le> n \<and> n < LENGTH('a) then LENGTH('a) - 1 else m + n)\<close>
 | 
1894  | 
for w :: \<open>'a::len word\<close>  | 
|
| 82525 | 1895  | 
by transfer (simp add: bit_drop_bit_eq bit_signed_take_bit_iff min_def le_less not_less, auto)  | 
| 72388 | 1896  | 
|
| 72508 | 1897  | 
lemma [code]:  | 
1898  | 
  \<open>Word.the_int (signed_drop_bit n w) = take_bit LENGTH('a) (drop_bit n (Word.the_signed_int w))\<close>
 | 
|
1899  | 
for w :: \<open>'a::len word\<close>  | 
|
1900  | 
by transfer simp  | 
|
1901  | 
||
| 73816 | 1902  | 
lemma signed_drop_bit_of_0 [simp]:  | 
1903  | 
\<open>signed_drop_bit n 0 = 0\<close>  | 
|
1904  | 
by transfer simp  | 
|
1905  | 
||
1906  | 
lemma signed_drop_bit_of_minus_1 [simp]:  | 
|
1907  | 
\<open>signed_drop_bit n (- 1) = - 1\<close>  | 
|
1908  | 
by transfer simp  | 
|
1909  | 
||
| 72488 | 1910  | 
lemma signed_drop_bit_signed_drop_bit [simp]:  | 
1911  | 
\<open>signed_drop_bit m (signed_drop_bit n w) = signed_drop_bit (m + n) w\<close>  | 
|
1912  | 
for w :: \<open>'a::len word\<close>  | 
|
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
1913  | 
proof (cases \<open>LENGTH('a)\<close>)
 | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
1914  | 
case 0  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
1915  | 
then show ?thesis  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
1916  | 
using len_not_eq_0 by blast  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
1917  | 
next  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
1918  | 
case (Suc n)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
1919  | 
then show ?thesis  | 
| 80777 | 1920  | 
by (force simp: bit_signed_drop_bit_iff not_le less_diff_conv ac_simps intro!: bit_word_eqI)  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
1921  | 
qed  | 
| 72488 | 1922  | 
|
| 72388 | 1923  | 
lemma signed_drop_bit_0 [simp]:  | 
1924  | 
\<open>signed_drop_bit 0 w = w\<close>  | 
|
| 72488 | 1925  | 
by transfer (simp add: take_bit_signed_take_bit)  | 
| 72388 | 1926  | 
|
1927  | 
lemma sint_signed_drop_bit_eq:  | 
|
1928  | 
\<open>sint (signed_drop_bit n w) = drop_bit n (sint w)\<close>  | 
|
| 82525 | 1929  | 
by (rule bit_eqI; cases n) (auto simp add: bit_simps not_less)  | 
| 72388 | 1930  | 
|
| 37660 | 1931  | 
|
| 75623 | 1932  | 
subsection \<open>Single-bit operations\<close>  | 
1933  | 
||
1934  | 
lemma set_bit_eq_idem_iff:  | 
|
| 82525 | 1935  | 
  \<open>set_bit n w = w \<longleftrightarrow> bit w n \<or> n \<ge> LENGTH('a)\<close>
 | 
| 75623 | 1936  | 
for w :: \<open>'a::len word\<close>  | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
1937  | 
unfolding bit_eq_iff  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
1938  | 
by (auto simp: bit_simps not_le)  | 
| 75623 | 1939  | 
|
1940  | 
lemma unset_bit_eq_idem_iff:  | 
|
1941  | 
  \<open>unset_bit n w = w \<longleftrightarrow> bit w n \<longrightarrow> n \<ge> LENGTH('a)\<close>
 | 
|
1942  | 
for w :: \<open>'a::len word\<close>  | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
1943  | 
unfolding bit_eq_iff  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
1944  | 
by (auto simp: bit_simps dest: bit_imp_le_length)  | 
| 75623 | 1945  | 
|
1946  | 
lemma flip_bit_eq_idem_iff:  | 
|
1947  | 
  \<open>flip_bit n w = w \<longleftrightarrow> n \<ge> LENGTH('a)\<close>
 | 
|
1948  | 
for w :: \<open>'a::len word\<close>  | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
1949  | 
by (simp add: flip_bit_eq_if set_bit_eq_idem_iff unset_bit_eq_idem_iff)  | 
| 75623 | 1950  | 
|
1951  | 
||
| 61799 | 1952  | 
subsection \<open>Rotation\<close>  | 
| 37660 | 1953  | 
|
| 72079 | 1954  | 
lift_definition word_rotr :: \<open>nat \<Rightarrow> 'a::len word \<Rightarrow> 'a::len word\<close>  | 
1955  | 
  is \<open>\<lambda>n k. concat_bit (LENGTH('a) - n mod LENGTH('a))
 | 
|
1956  | 
    (drop_bit (n mod LENGTH('a)) (take_bit LENGTH('a) k))
 | 
|
1957  | 
    (take_bit (n mod LENGTH('a)) k)\<close>
 | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
1958  | 
using take_bit_tightened by fastforce  | 
| 72079 | 1959  | 
|
1960  | 
lift_definition word_rotl :: \<open>nat \<Rightarrow> 'a::len word \<Rightarrow> 'a::len word\<close>  | 
|
1961  | 
  is \<open>\<lambda>n k. concat_bit (n mod LENGTH('a))
 | 
|
1962  | 
    (drop_bit (LENGTH('a) - n mod LENGTH('a)) (take_bit LENGTH('a) k))
 | 
|
1963  | 
    (take_bit (LENGTH('a) - n mod LENGTH('a)) k)\<close>
 | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
1964  | 
using take_bit_tightened by fastforce  | 
| 72079 | 1965  | 
|
1966  | 
lift_definition word_roti :: \<open>int \<Rightarrow> 'a::len word \<Rightarrow> 'a::len word\<close>  | 
|
1967  | 
  is \<open>\<lambda>r k. concat_bit (LENGTH('a) - nat (r mod int LENGTH('a)))
 | 
|
1968  | 
    (drop_bit (nat (r mod int LENGTH('a))) (take_bit LENGTH('a) k))
 | 
|
1969  | 
    (take_bit (nat (r mod int LENGTH('a))) k)\<close>
 | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
1970  | 
by (smt (verit, best) len_gt_0 nat_le_iff of_nat_0_less_iff pos_mod_bound  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
1971  | 
take_bit_tightened)  | 
| 72079 | 1972  | 
|
1973  | 
lemma word_rotl_eq_word_rotr [code]:  | 
|
1974  | 
  \<open>word_rotl n = (word_rotr (LENGTH('a) - n mod LENGTH('a)) :: 'a::len word \<Rightarrow> 'a word)\<close>
 | 
|
1975  | 
  by (rule ext, cases \<open>n mod LENGTH('a) = 0\<close>; transfer) simp_all
 | 
|
1976  | 
||
1977  | 
lemma word_roti_eq_word_rotr_word_rotl [code]:  | 
|
1978  | 
\<open>word_roti i w =  | 
|
1979  | 
(if i \<ge> 0 then word_rotr (nat i) w else word_rotl (nat (- i)) w)\<close>  | 
|
1980  | 
proof (cases \<open>i \<ge> 0\<close>)  | 
|
1981  | 
case True  | 
|
1982  | 
moreover define n where \<open>n = nat i\<close>  | 
|
1983  | 
ultimately have \<open>i = int n\<close>  | 
|
1984  | 
by simp  | 
|
1985  | 
moreover have \<open>word_roti (int n) = (word_rotr n :: _ \<Rightarrow> 'a word)\<close>  | 
|
1986  | 
by (rule ext, transfer) (simp add: nat_mod_distrib)  | 
|
1987  | 
ultimately show ?thesis  | 
|
1988  | 
by simp  | 
|
1989  | 
next  | 
|
1990  | 
case False  | 
|
1991  | 
moreover define n where \<open>n = nat (- i)\<close>  | 
|
1992  | 
ultimately have \<open>i = - int n\<close> \<open>n > 0\<close>  | 
|
1993  | 
by simp_all  | 
|
1994  | 
moreover have \<open>word_roti (- int n) = (word_rotl n :: _ \<Rightarrow> 'a word)\<close>  | 
|
1995  | 
by (rule ext, transfer)  | 
|
1996  | 
(simp add: zmod_zminus1_eq_if flip: of_nat_mod of_nat_diff)  | 
|
1997  | 
ultimately show ?thesis  | 
|
1998  | 
by simp  | 
|
1999  | 
qed  | 
|
2000  | 
||
| 
72611
 
c7bc3e70a8c7
official collection for bit projection simplifications
 
haftmann 
parents: 
72515 
diff
changeset
 | 
2001  | 
lemma bit_word_rotr_iff [bit_simps]:  | 
| 72079 | 2002  | 
\<open>bit (word_rotr m w) n \<longleftrightarrow>  | 
2003  | 
    n < LENGTH('a) \<and> bit w ((n + m) mod LENGTH('a))\<close>
 | 
|
2004  | 
for w :: \<open>'a::len word\<close>  | 
|
2005  | 
proof transfer  | 
|
2006  | 
fix k :: int and m n :: nat  | 
|
2007  | 
  define q where \<open>q = m mod LENGTH('a)\<close>
 | 
|
2008  | 
  have \<open>q < LENGTH('a)\<close> 
 | 
|
2009  | 
by (simp add: q_def)  | 
|
2010  | 
  then have \<open>q \<le> LENGTH('a)\<close>
 | 
|
2011  | 
by simp  | 
|
2012  | 
  have \<open>m mod LENGTH('a) = q\<close>
 | 
|
2013  | 
by (simp add: q_def)  | 
|
2014  | 
  moreover have \<open>(n + m) mod LENGTH('a) = (n + q) mod LENGTH('a)\<close>
 | 
|
2015  | 
    by (subst mod_add_right_eq [symmetric]) (simp add: \<open>m mod LENGTH('a) = q\<close>)
 | 
|
2016  | 
  moreover have \<open>n < LENGTH('a) \<and>
 | 
|
2017  | 
    bit (concat_bit (LENGTH('a) - q) (drop_bit q (take_bit LENGTH('a) k)) (take_bit q k)) n \<longleftrightarrow>
 | 
|
2018  | 
    n < LENGTH('a) \<and> bit k ((n + q) mod LENGTH('a))\<close>
 | 
|
2019  | 
    using \<open>q < LENGTH('a)\<close>
 | 
|
2020  | 
    by (cases \<open>q + n \<ge> LENGTH('a)\<close>)
 | 
|
| 80777 | 2021  | 
(auto simp: bit_concat_bit_iff bit_drop_bit_eq  | 
| 72079 | 2022  | 
bit_take_bit_iff le_mod_geq ac_simps)  | 
2023  | 
  ultimately show \<open>n < LENGTH('a) \<and>
 | 
|
2024  | 
    bit (concat_bit (LENGTH('a) - m mod LENGTH('a))
 | 
|
2025  | 
      (drop_bit (m mod LENGTH('a)) (take_bit LENGTH('a) k))
 | 
|
2026  | 
      (take_bit (m mod LENGTH('a)) k)) n
 | 
|
2027  | 
    \<longleftrightarrow> n < LENGTH('a) \<and>
 | 
|
2028  | 
      (n + m) mod LENGTH('a) < LENGTH('a) \<and>
 | 
|
2029  | 
      bit k ((n + m) mod LENGTH('a))\<close>
 | 
|
2030  | 
by simp  | 
|
2031  | 
qed  | 
|
2032  | 
||
| 
72611
 
c7bc3e70a8c7
official collection for bit projection simplifications
 
haftmann 
parents: 
72515 
diff
changeset
 | 
2033  | 
lemma bit_word_rotl_iff [bit_simps]:  | 
| 72079 | 2034  | 
\<open>bit (word_rotl m w) n \<longleftrightarrow>  | 
2035  | 
    n < LENGTH('a) \<and> bit w ((n + (LENGTH('a) - m mod LENGTH('a))) mod LENGTH('a))\<close>
 | 
|
2036  | 
for w :: \<open>'a::len word\<close>  | 
|
2037  | 
by (simp add: word_rotl_eq_word_rotr bit_word_rotr_iff)  | 
|
2038  | 
||
| 
72611
 
c7bc3e70a8c7
official collection for bit projection simplifications
 
haftmann 
parents: 
72515 
diff
changeset
 | 
2039  | 
lemma bit_word_roti_iff [bit_simps]:  | 
| 72079 | 2040  | 
\<open>bit (word_roti k w) n \<longleftrightarrow>  | 
2041  | 
    n < LENGTH('a) \<and> bit w (nat ((int n + k) mod int LENGTH('a)))\<close>
 | 
|
2042  | 
for w :: \<open>'a::len word\<close>  | 
|
2043  | 
proof transfer  | 
|
2044  | 
fix k l :: int and n :: nat  | 
|
2045  | 
  define m where \<open>m = nat (k mod int LENGTH('a))\<close>
 | 
|
2046  | 
  have \<open>m < LENGTH('a)\<close> 
 | 
|
2047  | 
by (simp add: nat_less_iff m_def)  | 
|
2048  | 
  then have \<open>m \<le> LENGTH('a)\<close>
 | 
|
2049  | 
by simp  | 
|
2050  | 
  have \<open>k mod int LENGTH('a) = int m\<close>
 | 
|
2051  | 
by (simp add: nat_less_iff m_def)  | 
|
2052  | 
  moreover have \<open>(int n + k) mod int LENGTH('a) = int ((n + m) mod LENGTH('a))\<close>
 | 
|
2053  | 
    by (subst mod_add_right_eq [symmetric]) (simp add: of_nat_mod \<open>k mod int LENGTH('a) = int m\<close>)
 | 
|
2054  | 
  moreover have \<open>n < LENGTH('a) \<and>
 | 
|
2055  | 
    bit (concat_bit (LENGTH('a) - m) (drop_bit m (take_bit LENGTH('a) l)) (take_bit m l)) n \<longleftrightarrow>
 | 
|
2056  | 
    n < LENGTH('a) \<and> bit l ((n + m) mod LENGTH('a))\<close>
 | 
|
2057  | 
    using \<open>m < LENGTH('a)\<close>
 | 
|
2058  | 
    by (cases \<open>m + n \<ge> LENGTH('a)\<close>)
 | 
|
| 80777 | 2059  | 
(auto simp: bit_concat_bit_iff bit_drop_bit_eq  | 
| 72079 | 2060  | 
bit_take_bit_iff nat_less_iff not_le not_less ac_simps  | 
2061  | 
le_diff_conv le_mod_geq)  | 
|
2062  | 
  ultimately show \<open>n < LENGTH('a)
 | 
|
2063  | 
    \<and> bit (concat_bit (LENGTH('a) - nat (k mod int LENGTH('a)))
 | 
|
2064  | 
             (drop_bit (nat (k mod int LENGTH('a))) (take_bit LENGTH('a) l))
 | 
|
2065  | 
             (take_bit (nat (k mod int LENGTH('a))) l)) n \<longleftrightarrow>
 | 
|
2066  | 
       n < LENGTH('a) 
 | 
|
2067  | 
    \<and> nat ((int n + k) mod int LENGTH('a)) < LENGTH('a)
 | 
|
2068  | 
    \<and> bit l (nat ((int n + k) mod int LENGTH('a)))\<close>
 | 
|
2069  | 
by simp  | 
|
2070  | 
qed  | 
|
2071  | 
||
| 72262 | 2072  | 
lemma uint_word_rotr_eq:  | 
| 72079 | 2073  | 
  \<open>uint (word_rotr n w) = concat_bit (LENGTH('a) - n mod LENGTH('a))
 | 
2074  | 
    (drop_bit (n mod LENGTH('a)) (uint w))
 | 
|
2075  | 
    (uint (take_bit (n mod LENGTH('a)) w))\<close>
 | 
|
2076  | 
for w :: \<open>'a::len word\<close>  | 
|
| 74101 | 2077  | 
by transfer (simp add: take_bit_concat_bit_eq)  | 
| 72079 | 2078  | 
|
| 72262 | 2079  | 
lemma [code]:  | 
2080  | 
  \<open>Word.the_int (word_rotr n w) = concat_bit (LENGTH('a) - n mod LENGTH('a))
 | 
|
2081  | 
    (drop_bit (n mod LENGTH('a)) (Word.the_int w))
 | 
|
2082  | 
    (Word.the_int (take_bit (n mod LENGTH('a)) w))\<close>
 | 
|
2083  | 
for w :: \<open>'a::len word\<close>  | 
|
2084  | 
using uint_word_rotr_eq [of n w] by simp  | 
|
2085  | 
||
| 72079 | 2086  | 
|
| 61799 | 2087  | 
subsection \<open>Split and cat operations\<close>  | 
| 37660 | 2088  | 
|
| 72079 | 2089  | 
lift_definition word_cat :: \<open>'a::len word \<Rightarrow> 'b::len word \<Rightarrow> 'c::len word\<close>  | 
2090  | 
  is \<open>\<lambda>k l. concat_bit LENGTH('b) l (take_bit LENGTH('a) k)\<close>
 | 
|
2091  | 
by (simp add: bit_eq_iff bit_concat_bit_iff bit_take_bit_iff)  | 
|
| 65268 | 2092  | 
|
| 71990 | 2093  | 
lemma word_cat_eq:  | 
2094  | 
  \<open>(word_cat v w :: 'c::len word) = push_bit LENGTH('b) (ucast v) + ucast w\<close>
 | 
|
2095  | 
for v :: \<open>'a::len word\<close> and w :: \<open>'b::len word\<close>  | 
|
| 72128 | 2096  | 
by transfer (simp add: concat_bit_eq ac_simps)  | 
| 72079 | 2097  | 
|
2098  | 
lemma word_cat_eq' [code]:  | 
|
2099  | 
  \<open>word_cat a b = word_of_int (concat_bit LENGTH('b) (uint b) (uint a))\<close>
 | 
|
2100  | 
for a :: \<open>'a::len word\<close> and b :: \<open>'b::len word\<close>  | 
|
| 72488 | 2101  | 
by transfer (simp add: concat_bit_take_bit_eq)  | 
| 71990 | 2102  | 
|
| 
72611
 
c7bc3e70a8c7
official collection for bit projection simplifications
 
haftmann 
parents: 
72515 
diff
changeset
 | 
2103  | 
lemma bit_word_cat_iff [bit_simps]:  | 
| 71990 | 2104  | 
  \<open>bit (word_cat v w :: 'c::len word) n \<longleftrightarrow> n < LENGTH('c) \<and> (if n < LENGTH('b) then bit w n else bit v (n - LENGTH('b)))\<close> 
 | 
2105  | 
for v :: \<open>'a::len word\<close> and w :: \<open>'b::len word\<close>  | 
|
| 72079 | 2106  | 
by transfer (simp add: bit_concat_bit_iff bit_take_bit_iff)  | 
| 71990 | 2107  | 
|
| 72488 | 2108  | 
definition word_split :: \<open>'a::len word \<Rightarrow> 'b::len word \<times> 'c::len word\<close>  | 
2109  | 
where \<open>word_split w =  | 
|
2110  | 
    (ucast (drop_bit LENGTH('c) w) :: 'b::len word, ucast w :: 'c::len word)\<close>
 | 
|
| 65268 | 2111  | 
|
| 
72088
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
2112  | 
definition word_rcat :: \<open>'a::len word list \<Rightarrow> 'b::len word\<close>  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
2113  | 
  where \<open>word_rcat = word_of_int \<circ> horner_sum uint (2 ^ LENGTH('a)) \<circ> rev\<close>
 | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
2114  | 
|
| 37660 | 2115  | 
|
| 72292 | 2116  | 
subsection \<open>More on conversions\<close>  | 
2117  | 
||
2118  | 
lemma int_word_sint:  | 
|
2119  | 
  \<open>sint (word_of_int x :: 'a::len word) = (x + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - 2 ^ (LENGTH('a) - 1)\<close>
 | 
|
| 72488 | 2120  | 
by transfer (simp flip: take_bit_eq_mod add: signed_take_bit_eq_take_bit_shift)  | 
| 46010 | 2121  | 
|
| 72128 | 2122  | 
lemma sint_sbintrunc': "sint (word_of_int bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) bin"
 | 
| 74496 | 2123  | 
by (simp add: signed_of_int)  | 
| 65268 | 2124  | 
|
| 72488 | 2125  | 
lemma uint_sint: "uint w = take_bit LENGTH('a) (sint w)"
 | 
| 65328 | 2126  | 
for w :: "'a::len word"  | 
| 72488 | 2127  | 
by transfer (simp add: take_bit_signed_take_bit)  | 
| 65268 | 2128  | 
|
| 72128 | 2129  | 
lemma bintr_uint: "LENGTH('a) \<le> n \<Longrightarrow> take_bit n (uint w) = uint w"
 | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2130  | 
for w :: "'a::len word"  | 
| 72292 | 2131  | 
by transfer (simp add: min_def)  | 
| 37660 | 2132  | 
|
| 46057 | 2133  | 
lemma wi_bintr:  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2134  | 
  "LENGTH('a::len) \<le> n \<Longrightarrow>
 | 
| 72128 | 2135  | 
word_of_int (take_bit n w) = (word_of_int w :: 'a word)"  | 
| 72292 | 2136  | 
by transfer simp  | 
| 45805 | 2137  | 
|
| 65268 | 2138  | 
lemma word_numeral_alt: "numeral b = word_of_int (numeral b)"  | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2139  | 
by simp  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
2140  | 
|
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
2141  | 
declare word_numeral_alt [symmetric, code_abbrev]  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
2142  | 
|
| 65268 | 2143  | 
lemma word_neg_numeral_alt: "- numeral b = word_of_int (- numeral b)"  | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2144  | 
by simp  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
2145  | 
|
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
2146  | 
declare word_neg_numeral_alt [symmetric, code_abbrev]  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
2147  | 
|
| 45805 | 2148  | 
lemma uint_bintrunc [simp]:  | 
| 82523 | 2149  | 
  "uint (numeral bin :: 'a word) = take_bit LENGTH('a::len) (numeral bin)"
 | 
| 72292 | 2150  | 
by transfer rule  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
2151  | 
|
| 65268 | 2152  | 
lemma uint_bintrunc_neg [simp]:  | 
| 82523 | 2153  | 
  "uint (- numeral bin :: 'a word) = take_bit LENGTH('a::len) (- numeral bin)"
 | 
| 72292 | 2154  | 
by transfer rule  | 
| 37660 | 2155  | 
|
| 45805 | 2156  | 
lemma sint_sbintrunc [simp]:  | 
| 72128 | 2157  | 
  "sint (numeral bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) (numeral bin)"
 | 
| 72292 | 2158  | 
by transfer simp  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
2159  | 
|
| 65268 | 2160  | 
lemma sint_sbintrunc_neg [simp]:  | 
| 72128 | 2161  | 
  "sint (- numeral bin :: 'a word) = signed_take_bit (LENGTH('a::len) - 1) (- numeral bin)"
 | 
| 72292 | 2162  | 
by transfer simp  | 
| 37660 | 2163  | 
|
| 45805 | 2164  | 
lemma unat_bintrunc [simp]:  | 
| 82523 | 2165  | 
  "unat (numeral bin :: 'a::len word) = take_bit LENGTH('a) (numeral bin)"
 | 
| 72079 | 2166  | 
by transfer simp  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
2167  | 
|
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
2168  | 
lemma unat_bintrunc_neg [simp]:  | 
| 82523 | 2169  | 
  "unat (- numeral bin :: 'a::len word) = nat (take_bit LENGTH('a) (- numeral bin))"
 | 
| 72079 | 2170  | 
by transfer simp  | 
| 37660 | 2171  | 
|
| 65328 | 2172  | 
lemma size_0_eq: "size w = 0 \<Longrightarrow> v = w"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2173  | 
for v w :: "'a::len word"  | 
| 72292 | 2174  | 
by transfer simp  | 
| 37660 | 2175  | 
|
| 65268 | 2176  | 
lemma uint_ge_0 [iff]: "0 \<le> uint x"  | 
| 72292 | 2177  | 
by (fact unsigned_greater_eq)  | 
| 45805 | 2178  | 
|
| 70185 | 2179  | 
lemma uint_lt2p [iff]: "uint x < 2 ^ LENGTH('a)"
 | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2180  | 
for x :: "'a::len word"  | 
| 72292 | 2181  | 
by (fact unsigned_less)  | 
| 45805 | 2182  | 
|
| 70185 | 2183  | 
lemma sint_ge: "- (2 ^ (LENGTH('a) - 1)) \<le> sint x"
 | 
| 65268 | 2184  | 
for x :: "'a::len word"  | 
| 72292 | 2185  | 
using sint_greater_eq [of x] by simp  | 
| 45805 | 2186  | 
|
| 70185 | 2187  | 
lemma sint_lt: "sint x < 2 ^ (LENGTH('a) - 1)"
 | 
| 65268 | 2188  | 
for x :: "'a::len word"  | 
| 72292 | 2189  | 
using sint_less [of x] by simp  | 
| 37660 | 2190  | 
|
| 70185 | 2191  | 
lemma uint_m2p_neg: "uint x - 2 ^ LENGTH('a) < 0"
 | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2192  | 
for x :: "'a::len word"  | 
| 45805 | 2193  | 
by (simp only: diff_less_0_iff_less uint_lt2p)  | 
2194  | 
||
| 70185 | 2195  | 
lemma uint_m2p_not_non_neg: "\<not> 0 \<le> uint x - 2 ^ LENGTH('a)"
 | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2196  | 
for x :: "'a::len word"  | 
| 45805 | 2197  | 
by (simp only: not_le uint_m2p_neg)  | 
| 37660 | 2198  | 
|
| 70185 | 2199  | 
lemma lt2p_lem: "LENGTH('a) \<le> n \<Longrightarrow> uint w < 2 ^ n"
 | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2200  | 
for w :: "'a::len word"  | 
| 72488 | 2201  | 
using uint_bounded [of w] by (rule less_le_trans) simp  | 
| 37660 | 2202  | 
|
| 45805 | 2203  | 
lemma uint_le_0_iff [simp]: "uint x \<le> 0 \<longleftrightarrow> uint x = 0"  | 
| 
70749
 
5d06b7bb9d22
More type class generalisations. Note that linorder_antisym_conv1 and linorder_antisym_conv2 no longer exist.
 
paulson <lp15@cam.ac.uk> 
parents: 
70342 
diff
changeset
 | 
2204  | 
by (fact uint_ge_0 [THEN leD, THEN antisym_conv1])  | 
| 37660 | 2205  | 
|
| 
40827
 
abbc05c20e24
code preprocessor setup for numerals on word type;
 
haftmann 
parents: 
39910 
diff
changeset
 | 
2206  | 
lemma uint_nat: "uint w = int (unat w)"  | 
| 72079 | 2207  | 
by transfer simp  | 
| 65268 | 2208  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2209  | 
lemma uint_numeral: "uint (numeral b :: 'a::len word) = numeral b mod 2 ^ LENGTH('a)"
 | 
| 72292 | 2210  | 
by (simp flip: take_bit_eq_mod add: of_nat_take_bit)  | 
| 65268 | 2211  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2212  | 
lemma uint_neg_numeral: "uint (- numeral b :: 'a::len word) = - numeral b mod 2 ^ LENGTH('a)"
 | 
| 72292 | 2213  | 
by (simp flip: take_bit_eq_mod add: of_nat_take_bit)  | 
| 65268 | 2214  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2215  | 
lemma unat_numeral: "unat (numeral b :: 'a::len word) = numeral b mod 2 ^ LENGTH('a)"
 | 
| 72079 | 2216  | 
by transfer (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq)  | 
| 37660 | 2217  | 
|
| 65268 | 2218  | 
lemma sint_numeral:  | 
2219  | 
"sint (numeral b :: 'a::len word) =  | 
|
| 72292 | 2220  | 
    (numeral b + 2 ^ (LENGTH('a) - 1)) mod 2 ^ LENGTH('a) - 2 ^ (LENGTH('a) - 1)"
 | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
2221  | 
by (metis int_word_sint word_numeral_alt)  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
2222  | 
|
| 65268 | 2223  | 
lemma word_of_int_0 [simp, code_post]: "word_of_int 0 = 0"  | 
| 72292 | 2224  | 
by (fact of_int_0)  | 
| 45958 | 2225  | 
|
| 65268 | 2226  | 
lemma word_of_int_1 [simp, code_post]: "word_of_int 1 = 1"  | 
| 72292 | 2227  | 
by (fact of_int_1)  | 
| 45958 | 2228  | 
|
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54225 
diff
changeset
 | 
2229  | 
lemma word_of_int_neg_1 [simp]: "word_of_int (- 1) = - 1"  | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2230  | 
by simp  | 
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54225 
diff
changeset
 | 
2231  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2232  | 
lemma word_of_int_numeral [simp] : "(word_of_int (numeral bin) :: 'a::len word) = numeral bin"  | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2233  | 
by simp  | 
| 65268 | 2234  | 
|
2235  | 
lemma word_int_case_wi:  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2236  | 
  "word_int_case f (word_of_int i :: 'b word) = f (i mod 2 ^ LENGTH('b::len))"
 | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2237  | 
by (simp add: uint_word_of_int word_int_case_eq_uint)  | 
| 65268 | 2238  | 
|
2239  | 
lemma word_int_split:  | 
|
2240  | 
"P (word_int_case f x) =  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2241  | 
    (\<forall>i. x = (word_of_int i :: 'b::len word) \<and> 0 \<le> i \<and> i < 2 ^ LENGTH('b) \<longrightarrow> P (f i))"
 | 
| 80777 | 2242  | 
by transfer (auto simp: take_bit_eq_mod)  | 
| 65268 | 2243  | 
|
2244  | 
lemma word_int_split_asm:  | 
|
2245  | 
"P (word_int_case f x) =  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2246  | 
    (\<nexists>n. x = (word_of_int n :: 'b::len word) \<and> 0 \<le> n \<and> n < 2 ^ LENGTH('b::len) \<and> \<not> P (f n))"
 | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2247  | 
using word_int_split by auto  | 
| 45805 | 2248  | 
|
| 65268 | 2249  | 
lemma uint_range_size: "0 \<le> uint w \<and> uint w < 2 ^ size w"  | 
| 72292 | 2250  | 
by transfer simp  | 
| 37660 | 2251  | 
|
| 65268 | 2252  | 
lemma sint_range_size: "- (2 ^ (size w - Suc 0)) \<le> sint w \<and> sint w < 2 ^ (size w - Suc 0)"  | 
| 72488 | 2253  | 
by (simp add: word_size sint_greater_eq sint_less)  | 
| 37660 | 2254  | 
|
| 65268 | 2255  | 
lemma sint_above_size: "2 ^ (size w - 1) \<le> x \<Longrightarrow> sint w < x"  | 
2256  | 
for w :: "'a::len word"  | 
|
| 45805 | 2257  | 
unfolding word_size by (rule less_le_trans [OF sint_lt])  | 
2258  | 
||
| 65268 | 2259  | 
lemma sint_below_size: "x \<le> - (2 ^ (size w - 1)) \<Longrightarrow> x \<le> sint w"  | 
2260  | 
for w :: "'a::len word"  | 
|
| 45805 | 2261  | 
unfolding word_size by (rule order_trans [OF _ sint_ge])  | 
| 37660 | 2262  | 
|
| 74592 | 2263  | 
lemma word_unat_eq_iff:  | 
2264  | 
\<open>v = w \<longleftrightarrow> unat v = unat w\<close>  | 
|
2265  | 
for v w :: \<open>'a::len word\<close>  | 
|
2266  | 
by (fact word_eq_iff_unsigned)  | 
|
2267  | 
||
| 
55816
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55415 
diff
changeset
 | 
2268  | 
|
| 61799 | 2269  | 
subsection \<open>Testing bits\<close>  | 
| 46010 | 2270  | 
|
| 72488 | 2271  | 
lemma bin_nth_uint_imp: "bit (uint w) n \<Longrightarrow> n < LENGTH('a)"
 | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2272  | 
for w :: "'a::len word"  | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2273  | 
by (simp add: bit_uint_iff)  | 
| 37660 | 2274  | 
|
| 46057 | 2275  | 
lemma bin_nth_sint:  | 
| 70185 | 2276  | 
  "LENGTH('a) \<le> n \<Longrightarrow>
 | 
| 72488 | 2277  | 
    bit (sint w) n = bit (sint w) (LENGTH('a) - 1)"
 | 
| 65268 | 2278  | 
for w :: "'a::len word"  | 
| 72292 | 2279  | 
by (transfer fixing: n) (simp add: bit_signed_take_bit_iff le_diff_conv min_def)  | 
| 37660 | 2280  | 
|
2281  | 
lemma num_of_bintr':  | 
|
| 72128 | 2282  | 
  "take_bit (LENGTH('a::len)) (numeral a :: int) = (numeral b) \<Longrightarrow>
 | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
2283  | 
numeral a = (numeral b :: 'a word)"  | 
| 72292 | 2284  | 
proof (transfer fixing: a b)  | 
2285  | 
  assume \<open>take_bit LENGTH('a) (numeral a :: int) = numeral b\<close>
 | 
|
2286  | 
  then have \<open>take_bit LENGTH('a) (take_bit LENGTH('a) (numeral a :: int)) = take_bit LENGTH('a) (numeral b)\<close>
 | 
|
2287  | 
by simp  | 
|
2288  | 
  then show \<open>take_bit LENGTH('a) (numeral a :: int) = take_bit LENGTH('a) (numeral b)\<close>
 | 
|
2289  | 
by simp  | 
|
2290  | 
qed  | 
|
| 37660 | 2291  | 
|
2292  | 
lemma num_of_sbintr':  | 
|
| 72241 | 2293  | 
  "signed_take_bit (LENGTH('a::len) - 1) (numeral a :: int) = (numeral b) \<Longrightarrow>
 | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
2294  | 
numeral a = (numeral b :: 'a word)"  | 
| 72292 | 2295  | 
proof (transfer fixing: a b)  | 
2296  | 
  assume \<open>signed_take_bit (LENGTH('a) - 1) (numeral a :: int) = numeral b\<close>
 | 
|
2297  | 
  then have \<open>take_bit LENGTH('a) (signed_take_bit (LENGTH('a) - 1) (numeral a :: int)) = take_bit LENGTH('a) (numeral b)\<close>
 | 
|
2298  | 
by simp  | 
|
2299  | 
  then show \<open>take_bit LENGTH('a) (numeral a :: int) = take_bit LENGTH('a) (numeral b)\<close>
 | 
|
| 72488 | 2300  | 
by (simp add: take_bit_signed_take_bit)  | 
| 72292 | 2301  | 
qed  | 
2302  | 
||
| 
46962
 
5bdcdb28be83
make more word theorems respect int/bin distinction
 
huffman 
parents: 
46656 
diff
changeset
 | 
2303  | 
lemma num_abs_bintr:  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
2304  | 
"(numeral x :: 'a word) =  | 
| 72128 | 2305  | 
    word_of_int (take_bit (LENGTH('a::len)) (numeral x))"
 | 
| 72292 | 2306  | 
by transfer simp  | 
| 
46962
 
5bdcdb28be83
make more word theorems respect int/bin distinction
 
huffman 
parents: 
46656 
diff
changeset
 | 
2307  | 
|
| 
 
5bdcdb28be83
make more word theorems respect int/bin distinction
 
huffman 
parents: 
46656 
diff
changeset
 | 
2308  | 
lemma num_abs_sbintr:  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
2309  | 
"(numeral x :: 'a word) =  | 
| 72128 | 2310  | 
    word_of_int (signed_take_bit (LENGTH('a::len) - 1) (numeral x))"
 | 
| 72488 | 2311  | 
by transfer (simp add: take_bit_signed_take_bit)  | 
| 
46962
 
5bdcdb28be83
make more word theorems respect int/bin distinction
 
huffman 
parents: 
46656 
diff
changeset
 | 
2312  | 
|
| 67408 | 2313  | 
text \<open>  | 
2314  | 
\<open>cast\<close> -- note, no arg for new length, as it's determined by type of result,  | 
|
2315  | 
thus in \<open>cast w = w\<close>, the type means cast to length of \<open>w\<close>!  | 
|
2316  | 
\<close>  | 
|
| 37660 | 2317  | 
|
| 
71957
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
2318  | 
lemma bit_ucast_iff:  | 
| 74101 | 2319  | 
  \<open>bit (ucast a :: 'a::len word) n \<longleftrightarrow> n < LENGTH('a::len) \<and> bit a n\<close>
 | 
| 72079 | 2320  | 
by transfer (simp add: bit_take_bit_iff)  | 
2321  | 
||
2322  | 
lemma ucast_id [simp]: "ucast w = w"  | 
|
2323  | 
by transfer simp  | 
|
2324  | 
||
2325  | 
lemma scast_id [simp]: "scast w = w"  | 
|
| 72488 | 2326  | 
by transfer (simp add: take_bit_signed_take_bit)  | 
| 37660 | 2327  | 
|
| 
71957
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
2328  | 
lemma ucast_mask_eq:  | 
| 72082 | 2329  | 
  \<open>ucast (mask n :: 'b word) = mask (min LENGTH('b::len) n)\<close>
 | 
| 80777 | 2330  | 
by (simp add: bit_eq_iff) (auto simp: bit_mask_iff bit_ucast_iff)  | 
| 
71957
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
2331  | 
|
| 67408 | 2332  | 
\<comment> \<open>literal u(s)cast\<close>  | 
| 
46001
 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 
huffman 
parents: 
46000 
diff
changeset
 | 
2333  | 
lemma ucast_bintr [simp]:  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2334  | 
"ucast (numeral w :: 'a::len word) =  | 
| 72128 | 2335  | 
    word_of_int (take_bit (LENGTH('a)) (numeral w))"
 | 
| 72079 | 2336  | 
by transfer simp  | 
| 65268 | 2337  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
2338  | 
(* TODO: neg_numeral *)  | 
| 37660 | 2339  | 
|
| 
46001
 
0b562d564d5f
redefine some binary operations on integers work on abstract numerals instead of Int.Pls and Int.Min
 
huffman 
parents: 
46000 
diff
changeset
 | 
2340  | 
lemma scast_sbintr [simp]:  | 
| 65268 | 2341  | 
"scast (numeral w ::'a::len word) =  | 
| 72128 | 2342  | 
    word_of_int (signed_take_bit (LENGTH('a) - Suc 0) (numeral w))"
 | 
| 72079 | 2343  | 
by transfer simp  | 
| 37660 | 2344  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2345  | 
lemma source_size: "source_size (c::'a::len word \<Rightarrow> _) = LENGTH('a)"
 | 
| 72079 | 2346  | 
by transfer simp  | 
| 46011 | 2347  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2348  | 
lemma target_size: "target_size (c::_ \<Rightarrow> 'b::len word) = LENGTH('b)"
 | 
| 72079 | 2349  | 
by transfer simp  | 
| 46011 | 2350  | 
|
| 70185 | 2351  | 
lemma is_down: "is_down c \<longleftrightarrow> LENGTH('b) \<le> LENGTH('a)"
 | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2352  | 
for c :: "'a::len word \<Rightarrow> 'b::len word"  | 
| 72079 | 2353  | 
by transfer simp  | 
| 65268 | 2354  | 
|
| 70185 | 2355  | 
lemma is_up: "is_up c \<longleftrightarrow> LENGTH('a) \<le> LENGTH('b)"
 | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2356  | 
for c :: "'a::len word \<Rightarrow> 'b::len word"  | 
| 72079 | 2357  | 
by transfer simp  | 
2358  | 
||
2359  | 
lemma is_up_down:  | 
|
2360  | 
\<open>is_up c \<longleftrightarrow> is_down d\<close>  | 
|
2361  | 
for c :: \<open>'a::len word \<Rightarrow> 'b::len word\<close>  | 
|
2362  | 
and d :: \<open>'b::len word \<Rightarrow> 'a::len word\<close>  | 
|
2363  | 
by transfer simp  | 
|
2364  | 
||
2365  | 
context  | 
|
2366  | 
fixes dummy_types :: \<open>'a::len \<times> 'b::len\<close>  | 
|
2367  | 
begin  | 
|
2368  | 
||
2369  | 
private abbreviation (input) UCAST :: \<open>'a::len word \<Rightarrow> 'b::len word\<close>  | 
|
2370  | 
where \<open>UCAST == ucast\<close>  | 
|
2371  | 
||
2372  | 
private abbreviation (input) SCAST :: \<open>'a::len word \<Rightarrow> 'b::len word\<close>  | 
|
2373  | 
where \<open>SCAST == scast\<close>  | 
|
2374  | 
||
2375  | 
lemma down_cast_same:  | 
|
2376  | 
\<open>UCAST = scast\<close> if \<open>is_down UCAST\<close>  | 
|
2377  | 
by (rule ext, use that in transfer) (simp add: take_bit_signed_take_bit)  | 
|
2378  | 
||
2379  | 
lemma sint_up_scast:  | 
|
2380  | 
\<open>sint (SCAST w) = sint w\<close> if \<open>is_up SCAST\<close>  | 
|
2381  | 
using that by transfer (simp add: min_def Suc_leI le_diff_iff)  | 
|
2382  | 
||
2383  | 
lemma uint_up_ucast:  | 
|
2384  | 
\<open>uint (UCAST w) = uint w\<close> if \<open>is_up UCAST\<close>  | 
|
2385  | 
using that by transfer (simp add: min_def)  | 
|
2386  | 
||
2387  | 
lemma ucast_up_ucast:  | 
|
2388  | 
\<open>ucast (UCAST w) = ucast w\<close> if \<open>is_up UCAST\<close>  | 
|
2389  | 
using that by transfer (simp add: ac_simps)  | 
|
2390  | 
||
2391  | 
lemma ucast_up_ucast_id:  | 
|
2392  | 
\<open>ucast (UCAST w) = w\<close> if \<open>is_up UCAST\<close>  | 
|
2393  | 
using that by (simp add: ucast_up_ucast)  | 
|
2394  | 
||
2395  | 
lemma scast_up_scast:  | 
|
2396  | 
\<open>scast (SCAST w) = scast w\<close> if \<open>is_up SCAST\<close>  | 
|
2397  | 
using that by transfer (simp add: ac_simps)  | 
|
2398  | 
||
2399  | 
lemma scast_up_scast_id:  | 
|
2400  | 
\<open>scast (SCAST w) = w\<close> if \<open>is_up SCAST\<close>  | 
|
2401  | 
using that by (simp add: scast_up_scast)  | 
|
2402  | 
||
2403  | 
lemma isduu:  | 
|
2404  | 
\<open>is_up UCAST\<close> if \<open>is_down d\<close>  | 
|
2405  | 
for d :: \<open>'b word \<Rightarrow> 'a word\<close>  | 
|
2406  | 
using that is_up_down [of UCAST d] by simp  | 
|
2407  | 
||
2408  | 
lemma isdus:  | 
|
2409  | 
\<open>is_up SCAST\<close> if \<open>is_down d\<close>  | 
|
2410  | 
for d :: \<open>'b word \<Rightarrow> 'a word\<close>  | 
|
2411  | 
using that is_up_down [of SCAST d] by simp  | 
|
2412  | 
||
| 37660 | 2413  | 
lemmas ucast_down_ucast_id = isduu [THEN ucast_up_ucast_id]  | 
| 72079 | 2414  | 
lemmas scast_down_scast_id = isdus [THEN scast_up_scast_id]  | 
| 37660 | 2415  | 
|
2416  | 
lemma up_ucast_surj:  | 
|
| 72079 | 2417  | 
\<open>surj (ucast :: 'b word \<Rightarrow> 'a word)\<close> if \<open>is_up UCAST\<close>  | 
2418  | 
by (rule surjI) (use that in \<open>rule ucast_up_ucast_id\<close>)  | 
|
| 37660 | 2419  | 
|
2420  | 
lemma up_scast_surj:  | 
|
| 72079 | 2421  | 
\<open>surj (scast :: 'b word \<Rightarrow> 'a word)\<close> if \<open>is_up SCAST\<close>  | 
2422  | 
by (rule surjI) (use that in \<open>rule scast_up_scast_id\<close>)  | 
|
| 37660 | 2423  | 
|
2424  | 
lemma down_ucast_inj:  | 
|
| 72079 | 2425  | 
\<open>inj_on UCAST A\<close> if \<open>is_down (ucast :: 'b word \<Rightarrow> 'a word)\<close>  | 
2426  | 
by (rule inj_on_inverseI) (use that in \<open>rule ucast_down_ucast_id\<close>)  | 
|
2427  | 
||
2428  | 
lemma down_scast_inj:  | 
|
2429  | 
\<open>inj_on SCAST A\<close> if \<open>is_down (scast :: 'b word \<Rightarrow> 'a word)\<close>  | 
|
2430  | 
by (rule inj_on_inverseI) (use that in \<open>rule scast_down_scast_id\<close>)  | 
|
2431  | 
||
2432  | 
lemma ucast_down_wi:  | 
|
2433  | 
\<open>UCAST (word_of_int x) = word_of_int x\<close> if \<open>is_down UCAST\<close>  | 
|
2434  | 
using that by transfer simp  | 
|
2435  | 
||
2436  | 
lemma ucast_down_no:  | 
|
2437  | 
\<open>UCAST (numeral bin) = numeral bin\<close> if \<open>is_down UCAST\<close>  | 
|
2438  | 
using that by transfer simp  | 
|
2439  | 
||
2440  | 
end  | 
|
| 37660 | 2441  | 
|
2442  | 
lemmas word_log_defs = word_and_def word_or_def word_xor_def word_not_def  | 
|
2443  | 
||
| 72000 | 2444  | 
lemma bit_last_iff:  | 
2445  | 
  \<open>bit w (LENGTH('a) - Suc 0) \<longleftrightarrow> sint w < 0\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)
 | 
|
2446  | 
for w :: \<open>'a::len word\<close>  | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2447  | 
by (simp add: bit_unsigned_iff sint_uint)  | 
| 72000 | 2448  | 
|
2449  | 
lemma drop_bit_eq_zero_iff_not_bit_last:  | 
|
2450  | 
  \<open>drop_bit (LENGTH('a) - Suc 0) w = 0 \<longleftrightarrow> \<not> bit w (LENGTH('a) - Suc 0)\<close>
 | 
|
2451  | 
for w :: "'a::len word"  | 
|
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
2452  | 
proof (cases \<open>LENGTH('a)\<close>)
 | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
2453  | 
case (Suc n)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
2454  | 
then show ?thesis  | 
| 72000 | 2455  | 
apply transfer  | 
2456  | 
apply (simp add: take_bit_drop_bit)  | 
|
| 74101 | 2457  | 
by (simp add: bit_iff_odd_drop_bit drop_bit_take_bit odd_iff_mod_2_eq_one)  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
2458  | 
qed auto  | 
| 72000 | 2459  | 
|
| 74592 | 2460  | 
lemma unat_div:  | 
2461  | 
\<open>unat (x div y) = unat x div unat y\<close>  | 
|
2462  | 
by (fact unat_div_distrib)  | 
|
2463  | 
||
2464  | 
lemma unat_mod:  | 
|
2465  | 
\<open>unat (x mod y) = unat x mod unat y\<close>  | 
|
2466  | 
by (fact unat_mod_distrib)  | 
|
2467  | 
||
| 37660 | 2468  | 
|
| 61799 | 2469  | 
subsection \<open>Word Arithmetic\<close>  | 
| 37660 | 2470  | 
|
| 74592 | 2471  | 
lemmas less_eq_word_numeral_numeral [simp] =  | 
2472  | 
word_le_def [of \<open>numeral a\<close> \<open>numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]  | 
|
2473  | 
for a b  | 
|
2474  | 
lemmas less_word_numeral_numeral [simp] =  | 
|
2475  | 
word_less_def [of \<open>numeral a\<close> \<open>numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]  | 
|
2476  | 
for a b  | 
|
2477  | 
lemmas less_eq_word_minus_numeral_numeral [simp] =  | 
|
2478  | 
word_le_def [of \<open>- numeral a\<close> \<open>numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]  | 
|
2479  | 
for a b  | 
|
2480  | 
lemmas less_word_minus_numeral_numeral [simp] =  | 
|
2481  | 
word_less_def [of \<open>- numeral a\<close> \<open>numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]  | 
|
2482  | 
for a b  | 
|
2483  | 
lemmas less_eq_word_numeral_minus_numeral [simp] =  | 
|
2484  | 
word_le_def [of \<open>numeral a\<close> \<open>- numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]  | 
|
2485  | 
for a b  | 
|
2486  | 
lemmas less_word_numeral_minus_numeral [simp] =  | 
|
2487  | 
word_less_def [of \<open>numeral a\<close> \<open>- numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]  | 
|
2488  | 
for a b  | 
|
2489  | 
lemmas less_eq_word_minus_numeral_minus_numeral [simp] =  | 
|
2490  | 
word_le_def [of \<open>- numeral a\<close> \<open>- numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]  | 
|
2491  | 
for a b  | 
|
2492  | 
lemmas less_word_minus_numeral_minus_numeral [simp] =  | 
|
2493  | 
word_less_def [of \<open>- numeral a\<close> \<open>- numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]  | 
|
2494  | 
for a b  | 
|
2495  | 
lemmas less_word_numeral_minus_1 [simp] =  | 
|
2496  | 
word_less_def [of \<open>numeral a\<close> \<open>- 1\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]  | 
|
2497  | 
for a b  | 
|
2498  | 
lemmas less_word_minus_numeral_minus_1 [simp] =  | 
|
2499  | 
word_less_def [of \<open>- numeral a\<close> \<open>- 1\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]  | 
|
2500  | 
for a b  | 
|
2501  | 
||
2502  | 
lemmas sless_eq_word_numeral_numeral [simp] =  | 
|
2503  | 
word_sle_eq [of \<open>numeral a\<close> \<open>numeral b\<close>, simplified sint_sbintrunc sint_sbintrunc_neg]  | 
|
2504  | 
for a b  | 
|
2505  | 
lemmas sless_word_numeral_numeral [simp] =  | 
|
2506  | 
word_sless_alt [of \<open>numeral a\<close> \<open>numeral b\<close>, simplified sint_sbintrunc sint_sbintrunc_neg]  | 
|
2507  | 
for a b  | 
|
2508  | 
lemmas sless_eq_word_minus_numeral_numeral [simp] =  | 
|
2509  | 
word_sle_eq [of \<open>- numeral a\<close> \<open>numeral b\<close>, simplified sint_sbintrunc sint_sbintrunc_neg]  | 
|
2510  | 
for a b  | 
|
2511  | 
lemmas sless_word_minus_numeral_numeral [simp] =  | 
|
2512  | 
word_sless_alt [of \<open>- numeral a\<close> \<open>numeral b\<close>, simplified sint_sbintrunc sint_sbintrunc_neg]  | 
|
2513  | 
for a b  | 
|
2514  | 
lemmas sless_eq_word_numeral_minus_numeral [simp] =  | 
|
2515  | 
word_sle_eq [of \<open>numeral a\<close> \<open>- numeral b\<close>, simplified sint_sbintrunc sint_sbintrunc_neg]  | 
|
2516  | 
for a b  | 
|
2517  | 
lemmas sless_word_numeral_minus_numeral [simp] =  | 
|
2518  | 
word_sless_alt [of \<open>numeral a\<close> \<open>- numeral b\<close>, simplified sint_sbintrunc sint_sbintrunc_neg]  | 
|
2519  | 
for a b  | 
|
2520  | 
lemmas sless_eq_word_minus_numeral_minus_numeral [simp] =  | 
|
2521  | 
word_sle_eq [of \<open>- numeral a\<close> \<open>- numeral b\<close>, simplified sint_sbintrunc sint_sbintrunc_neg]  | 
|
2522  | 
for a b  | 
|
2523  | 
lemmas sless_word_minus_numeral_minus_numeral [simp] =  | 
|
2524  | 
word_sless_alt [of \<open>- numeral a\<close> \<open>- numeral b\<close>, simplified sint_sbintrunc sint_sbintrunc_neg]  | 
|
2525  | 
for a b  | 
|
2526  | 
||
2527  | 
lemmas div_word_numeral_numeral [simp] =  | 
|
2528  | 
word_div_def [of \<open>numeral a\<close> \<open>numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]  | 
|
2529  | 
for a b  | 
|
2530  | 
lemmas div_word_minus_numeral_numeral [simp] =  | 
|
2531  | 
word_div_def [of \<open>- numeral a\<close> \<open>numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]  | 
|
2532  | 
for a b  | 
|
2533  | 
lemmas div_word_numeral_minus_numeral [simp] =  | 
|
2534  | 
word_div_def [of \<open>numeral a\<close> \<open>- numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]  | 
|
2535  | 
for a b  | 
|
2536  | 
lemmas div_word_minus_numeral_minus_numeral [simp] =  | 
|
2537  | 
word_div_def [of \<open>- numeral a\<close> \<open>- numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]  | 
|
2538  | 
for a b  | 
|
2539  | 
lemmas div_word_minus_1_numeral [simp] =  | 
|
2540  | 
word_div_def [of \<open>- 1\<close> \<open>numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]  | 
|
2541  | 
for a b  | 
|
2542  | 
lemmas div_word_minus_1_minus_numeral [simp] =  | 
|
2543  | 
word_div_def [of \<open>- 1\<close> \<open>- numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]  | 
|
2544  | 
for a b  | 
|
2545  | 
||
2546  | 
lemmas mod_word_numeral_numeral [simp] =  | 
|
2547  | 
word_mod_def [of \<open>numeral a\<close> \<open>numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]  | 
|
2548  | 
for a b  | 
|
2549  | 
lemmas mod_word_minus_numeral_numeral [simp] =  | 
|
2550  | 
word_mod_def [of \<open>- numeral a\<close> \<open>numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]  | 
|
2551  | 
for a b  | 
|
2552  | 
lemmas mod_word_numeral_minus_numeral [simp] =  | 
|
2553  | 
word_mod_def [of \<open>numeral a\<close> \<open>- numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]  | 
|
2554  | 
for a b  | 
|
2555  | 
lemmas mod_word_minus_numeral_minus_numeral [simp] =  | 
|
2556  | 
word_mod_def [of \<open>- numeral a\<close> \<open>- numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]  | 
|
2557  | 
for a b  | 
|
2558  | 
lemmas mod_word_minus_1_numeral [simp] =  | 
|
2559  | 
word_mod_def [of \<open>- 1\<close> \<open>numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]  | 
|
2560  | 
for a b  | 
|
2561  | 
lemmas mod_word_minus_1_minus_numeral [simp] =  | 
|
2562  | 
word_mod_def [of \<open>- 1\<close> \<open>- numeral b\<close>, simplified uint_bintrunc uint_bintrunc_neg unsigned_minus_1_eq_mask mask_eq_exp_minus_1]  | 
|
2563  | 
for a b  | 
|
2564  | 
||
2565  | 
lemma signed_drop_bit_of_1 [simp]:  | 
|
2566  | 
  \<open>signed_drop_bit n (1 :: 'a::len word) = of_bool (LENGTH('a) = 1 \<or> n = 0)\<close>
 | 
|
2567  | 
apply (transfer fixing: n)  | 
|
2568  | 
  apply (cases \<open>LENGTH('a)\<close>)
 | 
|
| 80777 | 2569  | 
apply (auto simp: take_bit_signed_take_bit)  | 
2570  | 
apply (auto simp: take_bit_drop_bit gr0_conv_Suc simp flip: take_bit_eq_self_iff_drop_bit_eq_0)  | 
|
| 74592 | 2571  | 
done  | 
2572  | 
||
2573  | 
lemma take_bit_word_beyond_length_eq:  | 
|
2574  | 
  \<open>take_bit n w = w\<close> if \<open>LENGTH('a) \<le> n\<close> for w :: \<open>'a::len word\<close>
 | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2575  | 
by (simp add: take_bit_word_eq_self that)  | 
| 74592 | 2576  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
2577  | 
lemmas word_div_no [simp] = word_div_def [of "numeral a" "numeral b"] for a b  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
2578  | 
lemmas word_mod_no [simp] = word_mod_def [of "numeral a" "numeral b"] for a b  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
2579  | 
lemmas word_less_no [simp] = word_less_def [of "numeral a" "numeral b"] for a b  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
2580  | 
lemmas word_le_no [simp] = word_le_def [of "numeral a" "numeral b"] for a b  | 
| 72079 | 2581  | 
lemmas word_sless_no [simp] = word_sless_eq [of "numeral a" "numeral b"] for a b  | 
2582  | 
lemmas word_sle_no [simp] = word_sle_eq [of "numeral a" "numeral b"] for a b  | 
|
| 37660 | 2583  | 
|
| 65268 | 2584  | 
lemma size_0_same': "size w = 0 \<Longrightarrow> w = v"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2585  | 
for v w :: "'a::len word"  | 
| 72079 | 2586  | 
by (unfold word_size) simp  | 
| 37660 | 2587  | 
|
| 
45816
 
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
 
huffman 
parents: 
45811 
diff
changeset
 | 
2588  | 
lemmas size_0_same = size_0_same' [unfolded word_size]  | 
| 37660 | 2589  | 
|
2590  | 
lemmas unat_eq_0 = unat_0_iff  | 
|
2591  | 
lemmas unat_eq_zero = unat_0_iff  | 
|
2592  | 
||
| 74592 | 2593  | 
lemma mask_1: "mask 1 = 1"  | 
2594  | 
by simp  | 
|
2595  | 
||
2596  | 
lemma mask_Suc_0: "mask (Suc 0) = 1"  | 
|
2597  | 
by simp  | 
|
2598  | 
||
2599  | 
lemma bin_last_bintrunc: "odd (take_bit l n) \<longleftrightarrow> l > 0 \<and> odd n"  | 
|
2600  | 
by simp  | 
|
2601  | 
||
2602  | 
lemma push_bit_word_beyond [simp]:  | 
|
2603  | 
  \<open>push_bit n w = 0\<close> if \<open>LENGTH('a) \<le> n\<close> for w :: \<open>'a::len word\<close>
 | 
|
2604  | 
using that by (transfer fixing: n) (simp add: take_bit_push_bit)  | 
|
2605  | 
||
2606  | 
lemma drop_bit_word_beyond [simp]:  | 
|
2607  | 
  \<open>drop_bit n w = 0\<close> if \<open>LENGTH('a) \<le> n\<close> for w :: \<open>'a::len word\<close>
 | 
|
2608  | 
using that by (transfer fixing: n) (simp add: drop_bit_take_bit)  | 
|
2609  | 
||
2610  | 
lemma signed_drop_bit_beyond:  | 
|
2611  | 
  \<open>signed_drop_bit n w = (if bit w (LENGTH('a) - Suc 0) then - 1 else 0)\<close>
 | 
|
2612  | 
  if \<open>LENGTH('a) \<le> n\<close> for w :: \<open>'a::len word\<close>
 | 
|
2613  | 
by (rule bit_word_eqI) (simp add: bit_signed_drop_bit_iff that)  | 
|
2614  | 
||
2615  | 
lemma take_bit_numeral_minus_numeral_word [simp]:  | 
|
2616  | 
\<open>take_bit (numeral m) (- numeral n :: 'a::len word) =  | 
|
2617  | 
(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>)  | 
|
2618  | 
proof (cases \<open>LENGTH('a) \<le> numeral m\<close>)
 | 
|
2619  | 
case True  | 
|
2620  | 
then have *: \<open>(take_bit (numeral m) :: 'a word \<Rightarrow> 'a word) = id\<close>  | 
|
2621  | 
by (simp add: fun_eq_iff take_bit_word_eq_self)  | 
|
2622  | 
have **: \<open>2 ^ numeral m = (0 :: 'a word)\<close>  | 
|
2623  | 
using True by (simp flip: exp_eq_zero_iff)  | 
|
2624  | 
show ?thesis  | 
|
2625  | 
by (auto simp only: * ** split: option.split  | 
|
2626  | 
dest!: take_bit_num_eq_None_imp [where ?'a = \<open>'a word\<close>] take_bit_num_eq_Some_imp [where ?'a = \<open>'a word\<close>])  | 
|
2627  | 
simp_all  | 
|
2628  | 
next  | 
|
2629  | 
case False  | 
|
2630  | 
then show ?thesis  | 
|
2631  | 
by (transfer fixing: m n) simp  | 
|
2632  | 
qed  | 
|
2633  | 
||
2634  | 
lemma of_nat_inverse:  | 
|
2635  | 
  \<open>word_of_nat r = a \<Longrightarrow> r < 2 ^ LENGTH('a) \<Longrightarrow> unat a = r\<close>
 | 
|
2636  | 
for a :: \<open>'a::len word\<close>  | 
|
2637  | 
by (metis id_apply of_nat_eq_id take_bit_nat_eq_self_iff unsigned_of_nat)  | 
|
2638  | 
||
| 
55816
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55415 
diff
changeset
 | 
2639  | 
|
| 61799 | 2640  | 
subsection \<open>Transferring goals from words to ints\<close>  | 
| 37660 | 2641  | 
|
| 65268 | 2642  | 
lemma word_ths:  | 
2643  | 
shows word_succ_p1: "word_succ a = a + 1"  | 
|
2644  | 
and word_pred_m1: "word_pred a = a - 1"  | 
|
2645  | 
and word_pred_succ: "word_pred (word_succ a) = a"  | 
|
2646  | 
and word_succ_pred: "word_succ (word_pred a) = a"  | 
|
2647  | 
and word_mult_succ: "word_succ a * b = b + a * b"  | 
|
| 
47374
 
9475d524bafb
set up and use lift_definition for word operations
 
huffman 
parents: 
47372 
diff
changeset
 | 
2648  | 
by (transfer, simp add: algebra_simps)+  | 
| 37660 | 2649  | 
|
| 
45816
 
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
 
huffman 
parents: 
45811 
diff
changeset
 | 
2650  | 
lemma uint_cong: "x = y \<Longrightarrow> uint x = uint y"  | 
| 
 
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
 
huffman 
parents: 
45811 
diff
changeset
 | 
2651  | 
by simp  | 
| 37660 | 2652  | 
|
| 55818 | 2653  | 
lemma uint_word_ariths:  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2654  | 
fixes a b :: "'a::len word"  | 
| 
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2655  | 
  shows "uint (a + b) = (uint a + uint b) mod 2 ^ LENGTH('a::len)"
 | 
| 70185 | 2656  | 
    and "uint (a - b) = (uint a - uint b) mod 2 ^ LENGTH('a)"
 | 
2657  | 
    and "uint (a * b) = uint a * uint b mod 2 ^ LENGTH('a)"
 | 
|
2658  | 
    and "uint (- a) = - uint a mod 2 ^ LENGTH('a)"
 | 
|
2659  | 
    and "uint (word_succ a) = (uint a + 1) mod 2 ^ LENGTH('a)"
 | 
|
2660  | 
    and "uint (word_pred a) = (uint a - 1) mod 2 ^ LENGTH('a)"
 | 
|
2661  | 
    and "uint (0 :: 'a word) = 0 mod 2 ^ LENGTH('a)"
 | 
|
2662  | 
    and "uint (1 :: 'a word) = 1 mod 2 ^ LENGTH('a)"
 | 
|
| 72262 | 2663  | 
by (simp_all only: word_arith_wis uint_word_of_int_eq flip: take_bit_eq_mod)  | 
| 55818 | 2664  | 
|
2665  | 
lemma uint_word_arith_bintrs:  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2666  | 
fixes a b :: "'a::len word"  | 
| 72128 | 2667  | 
  shows "uint (a + b) = take_bit (LENGTH('a)) (uint a + uint b)"
 | 
2668  | 
    and "uint (a - b) = take_bit (LENGTH('a)) (uint a - uint b)"
 | 
|
2669  | 
    and "uint (a * b) = take_bit (LENGTH('a)) (uint a * uint b)"
 | 
|
2670  | 
    and "uint (- a) = take_bit (LENGTH('a)) (- uint a)"
 | 
|
2671  | 
    and "uint (word_succ a) = take_bit (LENGTH('a)) (uint a + 1)"
 | 
|
2672  | 
    and "uint (word_pred a) = take_bit (LENGTH('a)) (uint a - 1)"
 | 
|
2673  | 
    and "uint (0 :: 'a word) = take_bit (LENGTH('a)) 0"
 | 
|
2674  | 
    and "uint (1 :: 'a word) = take_bit (LENGTH('a)) 1"
 | 
|
2675  | 
by (simp_all add: uint_word_ariths take_bit_eq_mod)  | 
|
| 55818 | 2676  | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2677  | 
context  | 
| 55818 | 2678  | 
fixes a b :: "'a::len word"  | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2679  | 
begin  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2680  | 
|
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2681  | 
lemma sint_word_add: "sint (a + b) = signed_take_bit (LENGTH('a) - 1) (sint a + sint b)"
 | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2682  | 
by transfer (simp add: signed_take_bit_add)  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2683  | 
|
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2684  | 
lemma sint_word_diff: "sint (a - b) = signed_take_bit (LENGTH('a) - 1) (sint a - sint b)"
 | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2685  | 
by transfer (simp add: signed_take_bit_diff)  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2686  | 
|
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2687  | 
lemma sint_word_mult: "sint (a * b) = signed_take_bit (LENGTH('a) - 1) (sint a * sint b)"
 | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2688  | 
by transfer (simp add: signed_take_bit_mult)  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2689  | 
|
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2690  | 
lemma sint_word_minus: "sint (- a) = signed_take_bit (LENGTH('a) - 1) (- sint a)"
 | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2691  | 
by transfer (simp add: signed_take_bit_minus)  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2692  | 
|
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2693  | 
lemma sint_word_succ: "sint (word_succ a) = signed_take_bit (LENGTH('a) - 1) (sint a + 1)"
 | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2694  | 
by (metis of_int_sint scast_id sint_sbintrunc' wi_hom_succ)  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2695  | 
|
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2696  | 
lemma sint_word_pred: "sint (word_pred a) = signed_take_bit (LENGTH('a) - 1) (sint a - 1)"
 | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2697  | 
by (metis of_int_sint scast_id sint_sbintrunc' wi_hom_pred)  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2698  | 
|
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2699  | 
lemma sint_word_01:  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2700  | 
  "sint (0 :: 'a word) = signed_take_bit (LENGTH('a) - 1) 0"
 | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2701  | 
  "sint (1 :: 'a word) = signed_take_bit (LENGTH('a) - 1) 1"
 | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2702  | 
by (simp_all add: sint_uint)  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2703  | 
|
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2704  | 
end  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2705  | 
|
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2706  | 
|
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2707  | 
lemmas sint_word_ariths =  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2708  | 
sint_word_add sint_word_diff sint_word_mult sint_word_minus  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2709  | 
sint_word_succ sint_word_pred sint_word_01  | 
| 45604 | 2710  | 
|
| 
58410
 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 
haftmann 
parents: 
58061 
diff
changeset
 | 
2711  | 
lemma word_pred_0_n1: "word_pred 0 = word_of_int (- 1)"  | 
| 
47374
 
9475d524bafb
set up and use lift_definition for word operations
 
huffman 
parents: 
47372 
diff
changeset
 | 
2712  | 
unfolding word_pred_m1 by simp  | 
| 37660 | 2713  | 
|
2714  | 
lemma succ_pred_no [simp]:  | 
|
| 65268 | 2715  | 
"word_succ (numeral w) = numeral w + 1"  | 
2716  | 
"word_pred (numeral w) = numeral w - 1"  | 
|
2717  | 
"word_succ (- numeral w) = - numeral w + 1"  | 
|
2718  | 
"word_pred (- numeral w) = - numeral w - 1"  | 
|
2719  | 
by (simp_all add: word_succ_p1 word_pred_m1)  | 
|
2720  | 
||
2721  | 
lemma word_sp_01 [simp]:  | 
|
2722  | 
"word_succ (- 1) = 0 \<and> word_succ 0 = 1 \<and> word_pred 0 = - 1 \<and> word_pred 1 = 0"  | 
|
2723  | 
by (simp_all add: word_succ_p1 word_pred_m1)  | 
|
| 37660 | 2724  | 
|
| 67408 | 2725  | 
\<comment> \<open>alternative approach to lifting arithmetic equalities\<close>  | 
| 65268 | 2726  | 
lemma word_of_int_Ex: "\<exists>y. x = word_of_int y"  | 
| 37660 | 2727  | 
by (rule_tac x="uint x" in exI) simp  | 
2728  | 
||
2729  | 
||
| 61799 | 2730  | 
subsection \<open>Order on fixed-length words\<close>  | 
| 37660 | 2731  | 
|
| 72262 | 2732  | 
lift_definition udvd :: \<open>'a::len word \<Rightarrow> 'a::len word \<Rightarrow> bool\<close> (infixl \<open>udvd\<close> 50)  | 
2733  | 
  is \<open>\<lambda>k l. take_bit LENGTH('a) k dvd take_bit LENGTH('a) l\<close> by simp
 | 
|
2734  | 
||
2735  | 
lemma udvd_iff_dvd:  | 
|
2736  | 
\<open>x udvd y \<longleftrightarrow> unat x dvd unat y\<close>  | 
|
2737  | 
by transfer (simp add: nat_dvd_iff)  | 
|
2738  | 
||
2739  | 
lemma udvd_iff_dvd_int:  | 
|
2740  | 
\<open>v udvd w \<longleftrightarrow> uint v dvd uint w\<close>  | 
|
2741  | 
by transfer rule  | 
|
2742  | 
||
2743  | 
lemma udvdI [intro]:  | 
|
2744  | 
\<open>v udvd w\<close> if \<open>unat w = unat v * unat u\<close>  | 
|
2745  | 
proof -  | 
|
2746  | 
from that have \<open>unat v dvd unat w\<close> ..  | 
|
2747  | 
then show ?thesis  | 
|
2748  | 
by (simp add: udvd_iff_dvd)  | 
|
2749  | 
qed  | 
|
2750  | 
||
2751  | 
lemma udvdE [elim]:  | 
|
2752  | 
fixes v w :: \<open>'a::len word\<close>  | 
|
2753  | 
assumes \<open>v udvd w\<close>  | 
|
2754  | 
obtains u :: \<open>'a word\<close> where \<open>unat w = unat v * unat u\<close>  | 
|
2755  | 
proof (cases \<open>v = 0\<close>)  | 
|
2756  | 
case True  | 
|
2757  | 
moreover from True \<open>v udvd w\<close> have \<open>w = 0\<close>  | 
|
2758  | 
by transfer simp  | 
|
2759  | 
ultimately show thesis  | 
|
2760  | 
using that by simp  | 
|
2761  | 
next  | 
|
2762  | 
case False  | 
|
2763  | 
then have \<open>unat v > 0\<close>  | 
|
2764  | 
by (simp add: unat_gt_0)  | 
|
2765  | 
from \<open>v udvd w\<close> have \<open>unat v dvd unat w\<close>  | 
|
2766  | 
by (simp add: udvd_iff_dvd)  | 
|
2767  | 
then obtain n where \<open>unat w = unat v * n\<close> ..  | 
|
2768  | 
  moreover have \<open>n < 2 ^ LENGTH('a)\<close>
 | 
|
2769  | 
proof (rule ccontr)  | 
|
2770  | 
    assume \<open>\<not> n < 2 ^ LENGTH('a)\<close>
 | 
|
2771  | 
    then have \<open>n \<ge> 2 ^ LENGTH('a)\<close>
 | 
|
2772  | 
by (simp add: not_le)  | 
|
2773  | 
    then have \<open>unat v * n \<ge> 2 ^ LENGTH('a)\<close>
 | 
|
2774  | 
      using \<open>unat v > 0\<close> mult_le_mono [of 1 \<open>unat v\<close> \<open>2 ^ LENGTH('a)\<close> n]
 | 
|
2775  | 
by simp  | 
|
| 72292 | 2776  | 
with \<open>unat w = unat v * n\<close>  | 
2777  | 
    have \<open>unat w \<ge> 2 ^ LENGTH('a)\<close>
 | 
|
| 72262 | 2778  | 
by simp  | 
| 72292 | 2779  | 
with unsigned_less [of w, where ?'a = nat] show False  | 
2780  | 
by linarith  | 
|
| 72262 | 2781  | 
qed  | 
2782  | 
ultimately have \<open>unat w = unat v * unat (word_of_nat n :: 'a word)\<close>  | 
|
| 80777 | 2783  | 
by (auto simp: take_bit_nat_eq_self_iff unsigned_of_nat intro: sym)  | 
| 72262 | 2784  | 
with that show thesis .  | 
2785  | 
qed  | 
|
2786  | 
||
2787  | 
lemma udvd_imp_mod_eq_0:  | 
|
2788  | 
\<open>w mod v = 0\<close> if \<open>v udvd w\<close>  | 
|
2789  | 
using that by transfer simp  | 
|
2790  | 
||
2791  | 
lemma mod_eq_0_imp_udvd [intro?]:  | 
|
2792  | 
\<open>v udvd w\<close> if \<open>w mod v = 0\<close>  | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2793  | 
by (metis mod_0_imp_dvd that udvd_iff_dvd unat_0 unat_mod_distrib)  | 
| 72262 | 2794  | 
|
| 72280 | 2795  | 
lemma udvd_imp_dvd:  | 
2796  | 
\<open>v dvd w\<close> if \<open>v udvd w\<close> for v w :: \<open>'a::len word\<close>  | 
|
2797  | 
proof -  | 
|
| 
72281
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
2798  | 
from that obtain u :: \<open>'a word\<close> where \<open>unat w = unat v * unat u\<close> ..  | 
| 72280 | 2799  | 
then have \<open>w = v * u\<close>  | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2800  | 
by (metis of_nat_mult of_nat_unat word_mult_def word_of_int_uint)  | 
| 72280 | 2801  | 
then show \<open>v dvd w\<close> ..  | 
2802  | 
qed  | 
|
2803  | 
||
| 
72281
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
2804  | 
lemma exp_dvd_iff_exp_udvd:  | 
| 
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
2805  | 
\<open>2 ^ n dvd w \<longleftrightarrow> 2 ^ n udvd w\<close> for v w :: \<open>'a::len word\<close>  | 
| 
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
2806  | 
proof  | 
| 
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
2807  | 
assume \<open>2 ^ n udvd w\<close> then show \<open>2 ^ n dvd w\<close>  | 
| 
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
2808  | 
by (rule udvd_imp_dvd)  | 
| 
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
2809  | 
next  | 
| 
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
2810  | 
assume \<open>2 ^ n dvd w\<close>  | 
| 
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
2811  | 
then obtain u :: \<open>'a word\<close> where \<open>w = 2 ^ n * u\<close> ..  | 
| 
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
2812  | 
then have \<open>w = push_bit n u\<close>  | 
| 
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
2813  | 
by (simp add: push_bit_eq_mult)  | 
| 
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
2814  | 
then show \<open>2 ^ n udvd w\<close>  | 
| 
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
2815  | 
by transfer (simp add: take_bit_push_bit dvd_eq_mod_eq_0 flip: take_bit_eq_mod)  | 
| 
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
2816  | 
qed  | 
| 
 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 
haftmann 
parents: 
72280 
diff
changeset
 | 
2817  | 
|
| 72262 | 2818  | 
lemma udvd_nat_alt:  | 
2819  | 
\<open>a udvd b \<longleftrightarrow> (\<exists>n. unat b = n * unat a)\<close>  | 
|
| 80777 | 2820  | 
by (auto simp: udvd_iff_dvd)  | 
| 72262 | 2821  | 
|
2822  | 
lemma udvd_unfold_int:  | 
|
2823  | 
\<open>a udvd b \<longleftrightarrow> (\<exists>n\<ge>0. uint b = n * uint a)\<close>  | 
|
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
2824  | 
unfolding udvd_iff_dvd_int  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
2825  | 
by (metis dvd_div_mult_self dvd_triv_right uint_div_distrib uint_ge_0)  | 
| 37660 | 2826  | 
|
| 
55816
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55415 
diff
changeset
 | 
2827  | 
lemma unat_minus_one:  | 
| 72079 | 2828  | 
\<open>unat (w - 1) = unat w - 1\<close> if \<open>w \<noteq> 0\<close>  | 
| 
55816
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55415 
diff
changeset
 | 
2829  | 
proof -  | 
| 
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55415 
diff
changeset
 | 
2830  | 
have "0 \<le> uint w" by (fact uint_nonnegative)  | 
| 72079 | 2831  | 
moreover from that have "0 \<noteq> uint w"  | 
| 65328 | 2832  | 
by (simp add: uint_0_iff)  | 
2833  | 
ultimately have "1 \<le> uint w"  | 
|
2834  | 
by arith  | 
|
| 70185 | 2835  | 
  from uint_lt2p [of w] have "uint w - 1 < 2 ^ LENGTH('a)"
 | 
| 65328 | 2836  | 
by arith  | 
| 70185 | 2837  | 
  with \<open>1 \<le> uint w\<close> have "(uint w - 1) mod 2 ^ LENGTH('a) = uint w - 1"
 | 
| 
55816
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55415 
diff
changeset
 | 
2838  | 
by (auto intro: mod_pos_pos_trivial)  | 
| 70185 | 2839  | 
  with \<open>1 \<le> uint w\<close> have "nat ((uint w - 1) mod 2 ^ LENGTH('a)) = nat (uint w) - 1"
 | 
| 72079 | 2840  | 
by (auto simp del: nat_uint_eq)  | 
| 
55816
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55415 
diff
changeset
 | 
2841  | 
then show ?thesis  | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2842  | 
by (metis uint_word_ariths(6) unat_eq_nat_uint word_pred_m1)  | 
| 
55816
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55415 
diff
changeset
 | 
2843  | 
qed  | 
| 
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55415 
diff
changeset
 | 
2844  | 
|
| 65328 | 2845  | 
lemma measure_unat: "p \<noteq> 0 \<Longrightarrow> unat (p - 1) < unat p"  | 
| 37660 | 2846  | 
by (simp add: unat_minus_one) (simp add: unat_0_iff [symmetric])  | 
| 65268 | 2847  | 
|
| 45604 | 2848  | 
lemmas uint_add_ge0 [simp] = add_nonneg_nonneg [OF uint_ge_0 uint_ge_0]  | 
2849  | 
lemmas uint_mult_ge0 [simp] = mult_nonneg_nonneg [OF uint_ge_0 uint_ge_0]  | 
|
| 37660 | 2850  | 
|
| 70185 | 2851  | 
lemma uint_sub_lt2p [simp]: "uint x - uint y < 2 ^ LENGTH('a)"
 | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2852  | 
for x :: "'a::len word" and y :: "'b::len word"  | 
| 37660 | 2853  | 
using uint_ge_0 [of y] uint_lt2p [of x] by arith  | 
2854  | 
||
2855  | 
||
| 61799 | 2856  | 
subsection \<open>Conditions for the addition (etc) of two words to overflow\<close>  | 
| 37660 | 2857  | 
|
| 65268 | 2858  | 
lemma uint_add_lem:  | 
| 70185 | 2859  | 
  "(uint x + uint y < 2 ^ LENGTH('a)) =
 | 
| 65328 | 2860  | 
(uint (x + y) = uint x + uint y)"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2861  | 
for x y :: "'a::len word"  | 
| 71997 | 2862  | 
by (metis add.right_neutral add_mono_thms_linordered_semiring(1) mod_pos_pos_trivial of_nat_0_le_iff uint_lt2p uint_nat uint_word_ariths(1))  | 
| 37660 | 2863  | 
|
| 65268 | 2864  | 
lemma uint_mult_lem:  | 
| 70185 | 2865  | 
  "(uint x * uint y < 2 ^ LENGTH('a)) =
 | 
| 65328 | 2866  | 
(uint (x * y) = uint x * uint y)"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2867  | 
for x y :: "'a::len word"  | 
| 71997 | 2868  | 
by (metis mod_pos_pos_trivial uint_lt2p uint_mult_ge0 uint_word_ariths(3))  | 
| 37660 | 2869  | 
|
| 65328 | 2870  | 
lemma uint_sub_lem: "uint x \<ge> uint y \<longleftrightarrow> uint (x - y) = uint x - uint y"  | 
| 
80401
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
2871  | 
by (simp add: uint_word_arith_bintrs take_bit_int_eq_self_iff)  | 
| 65328 | 2872  | 
|
2873  | 
lemma uint_add_le: "uint (x + y) \<le> uint x + uint y"  | 
|
| 71997 | 2874  | 
unfolding uint_word_ariths by (simp add: zmod_le_nonneg_dividend)  | 
| 37660 | 2875  | 
|
| 65328 | 2876  | 
lemma uint_sub_ge: "uint (x - y) \<ge> uint x - uint y"  | 
| 80777 | 2877  | 
by (smt (verit, ccfv_SIG) uint_nonnegative uint_sub_lem)  | 
| 72488 | 2878  | 
|
2879  | 
lemma int_mod_ge: \<open>a \<le> a mod n\<close> if \<open>a < n\<close> \<open>0 < n\<close>  | 
|
2880  | 
for a n :: int  | 
|
| 76231 | 2881  | 
using that order.trans [of a 0 \<open>a mod n\<close>] by (cases \<open>a < 0\<close>) auto  | 
| 72488 | 2882  | 
|
| 
55816
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55415 
diff
changeset
 | 
2883  | 
lemma mod_add_if_z:  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
2884  | 
"\<lbrakk>x < z; y < z; 0 \<le> y; 0 \<le> x; 0 \<le> z\<rbrakk> \<Longrightarrow>  | 
| 65328 | 2885  | 
(x + y) mod z = (if x + y < z then x + y else x + y - z)"  | 
2886  | 
for x y z :: int  | 
|
| 80777 | 2887  | 
by (smt (verit, best) minus_mod_self2 mod_pos_pos_trivial)  | 
| 
55816
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55415 
diff
changeset
 | 
2888  | 
|
| 
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55415 
diff
changeset
 | 
2889  | 
lemma uint_plus_if':  | 
| 65328 | 2890  | 
"uint (a + b) =  | 
| 70185 | 2891  | 
    (if uint a + uint b < 2 ^ LENGTH('a) then uint a + uint b
 | 
2892  | 
     else uint a + uint b - 2 ^ LENGTH('a))"
 | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2893  | 
for a b :: "'a::len word"  | 
| 
55816
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55415 
diff
changeset
 | 
2894  | 
using mod_add_if_z [of "uint a" _ "uint b"] by (simp add: uint_word_ariths)  | 
| 
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55415 
diff
changeset
 | 
2895  | 
|
| 
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55415 
diff
changeset
 | 
2896  | 
lemma mod_sub_if_z:  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
2897  | 
"\<lbrakk>x < z; y < z; 0 \<le> y; 0 \<le> x; 0 \<le> z\<rbrakk> \<Longrightarrow>  | 
| 65328 | 2898  | 
(x - y) mod z = (if y \<le> x then x - y else x - y + z)"  | 
2899  | 
for x y z :: int  | 
|
| 80777 | 2900  | 
using mod_pos_pos_trivial [of "x - y + z" z] by (auto simp: not_le)  | 
| 
55816
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55415 
diff
changeset
 | 
2901  | 
|
| 
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55415 
diff
changeset
 | 
2902  | 
lemma uint_sub_if':  | 
| 65328 | 2903  | 
"uint (a - b) =  | 
2904  | 
(if uint b \<le> uint a then uint a - uint b  | 
|
| 70185 | 2905  | 
     else uint a - uint b + 2 ^ LENGTH('a))"
 | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2906  | 
for a b :: "'a::len word"  | 
| 
55816
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55415 
diff
changeset
 | 
2907  | 
using mod_sub_if_z [of "uint a" _ "uint b"] by (simp add: uint_word_ariths)  | 
| 
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55415 
diff
changeset
 | 
2908  | 
|
| 37660 | 2909  | 
lemma word_of_int_inverse:  | 
| 70185 | 2910  | 
  "word_of_int r = a \<Longrightarrow> 0 \<le> r \<Longrightarrow> r < 2 ^ LENGTH('a) \<Longrightarrow> uint a = r"
 | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2911  | 
for a :: "'a::len word"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
2912  | 
by transfer (simp add: take_bit_int_eq_self)  | 
| 37660 | 2913  | 
|
| 74592 | 2914  | 
lemma unat_split: "P (unat x) \<longleftrightarrow> (\<forall>n. of_nat n = x \<and> n < 2^LENGTH('a) \<longrightarrow> P n)"
 | 
2915  | 
for x :: "'a::len word"  | 
|
| 80777 | 2916  | 
by (auto simp: unsigned_of_nat take_bit_nat_eq_self)  | 
| 74592 | 2917  | 
|
2918  | 
lemma unat_split_asm: "P (unat x) \<longleftrightarrow> (\<nexists>n. of_nat n = x \<and> n < 2^LENGTH('a) \<and> \<not> P n)"
 | 
|
2919  | 
for x :: "'a::len word"  | 
|
| 80777 | 2920  | 
using unat_split by auto  | 
| 74592 | 2921  | 
|
2922  | 
lemma un_ui_le:  | 
|
2923  | 
\<open>unat a \<le> unat b \<longleftrightarrow> uint a \<le> uint b\<close>  | 
|
2924  | 
by transfer (simp add: nat_le_iff)  | 
|
2925  | 
||
2926  | 
lemma unat_plus_if':  | 
|
2927  | 
\<open>unat (a + b) =  | 
|
2928  | 
    (if unat a + unat b < 2 ^ LENGTH('a)
 | 
|
2929  | 
then unat a + unat b  | 
|
2930  | 
    else unat a + unat b - 2 ^ LENGTH('a))\<close> for a b :: \<open>'a::len word\<close>
 | 
|
| 80777 | 2931  | 
apply (auto simp: not_less le_iff_add)  | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2932  | 
using of_nat_inverse apply force  | 
| 80777 | 2933  | 
by (smt (verit, ccfv_SIG) numeral_Bit0 numerals(1) of_nat_0_le_iff of_nat_1 of_nat_add of_nat_eq_iff of_nat_power of_nat_unat uint_plus_if')  | 
| 74592 | 2934  | 
|
2935  | 
lemma unat_sub_if_size:  | 
|
2936  | 
"unat (x - y) =  | 
|
2937  | 
(if unat y \<le> unat x  | 
|
2938  | 
then unat x - unat y  | 
|
2939  | 
else unat x + 2 ^ size x - unat y)"  | 
|
2940  | 
proof -  | 
|
2941  | 
  { assume xy: "\<not> uint y \<le> uint x"
 | 
|
2942  | 
    have "nat (uint x - uint y + 2 ^ LENGTH('a)) = nat (uint x + 2 ^ LENGTH('a) - uint y)"
 | 
|
2943  | 
by simp  | 
|
| 80777 | 2944  | 
    also have "\<dots> = nat (uint x + 2 ^ LENGTH('a)) - nat (uint y)"
 | 
| 74592 | 2945  | 
by (simp add: nat_diff_distrib')  | 
| 80777 | 2946  | 
    also have "\<dots> = nat (uint x) + 2 ^ LENGTH('a) - nat (uint y)"
 | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2947  | 
by (simp add: nat_add_distrib nat_power_eq)  | 
| 74592 | 2948  | 
    finally have "nat (uint x - uint y + 2 ^ LENGTH('a)) = nat (uint x) + 2 ^ LENGTH('a) - nat (uint y)" .
 | 
2949  | 
}  | 
|
2950  | 
then show ?thesis  | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2951  | 
by (metis nat_diff_distrib' uint_range_size uint_sub_if' un_ui_le unat_eq_nat_uint  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
2952  | 
word_size)  | 
| 74592 | 2953  | 
qed  | 
2954  | 
||
2955  | 
lemmas unat_sub_if' = unat_sub_if_size [unfolded word_size]  | 
|
2956  | 
||
| 37660 | 2957  | 
lemma uint_split:  | 
| 70185 | 2958  | 
  "P (uint x) = (\<forall>i. word_of_int i = x \<and> 0 \<le> i \<and> i < 2^LENGTH('a) \<longrightarrow> P i)"
 | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2959  | 
for x :: "'a::len word"  | 
| 80777 | 2960  | 
by transfer (auto simp: take_bit_eq_mod)  | 
| 37660 | 2961  | 
|
2962  | 
lemma uint_split_asm:  | 
|
| 70185 | 2963  | 
  "P (uint x) = (\<nexists>i. word_of_int i = x \<and> 0 \<le> i \<and> i < 2^LENGTH('a) \<and> \<not> P i)"
 | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
2964  | 
for x :: "'a::len word"  | 
| 80777 | 2965  | 
by (auto simp: unsigned_of_int take_bit_int_eq_self)  | 
| 37660 | 2966  | 
|
| 74592 | 2967  | 
|
2968  | 
subsection \<open>Some proof tool support\<close>  | 
|
2969  | 
||
2970  | 
\<comment> \<open>use this to stop, eg. \<open>2 ^ LENGTH(32)\<close> being simplified\<close>  | 
|
2971  | 
lemma power_False_cong: "False \<Longrightarrow> a ^ b = c ^ d"  | 
|
2972  | 
by auto  | 
|
2973  | 
||
2974  | 
lemmas unat_splits = unat_split unat_split_asm  | 
|
2975  | 
||
2976  | 
lemmas unat_arith_simps =  | 
|
2977  | 
word_le_nat_alt word_less_nat_alt  | 
|
2978  | 
word_unat_eq_iff  | 
|
2979  | 
unat_sub_if' unat_plus_if' unat_div unat_mod  | 
|
2980  | 
||
| 37660 | 2981  | 
lemmas uint_splits = uint_split uint_split_asm  | 
2982  | 
||
| 65268 | 2983  | 
lemmas uint_arith_simps =  | 
| 37660 | 2984  | 
word_le_def word_less_alt  | 
| 72292 | 2985  | 
word_uint_eq_iff  | 
| 37660 | 2986  | 
uint_sub_if' uint_plus_if'  | 
2987  | 
||
| 74592 | 2988  | 
\<comment> \<open>\<open>unat_arith_tac\<close>: tactic to reduce word arithmetic to \<open>nat\<close>, try to solve via \<open>arith\<close>\<close>  | 
2989  | 
ML \<open>  | 
|
2990  | 
val unat_arith_simpset =  | 
|
2991  | 
  @{context} (* TODO: completely explicitly determined simpset *)
 | 
|
2992  | 
  |> fold Simplifier.add_simp @{thms unat_arith_simps}
 | 
|
2993  | 
  |> fold Splitter.add_split @{thms if_split_asm}
 | 
|
2994  | 
  |> fold Simplifier.add_cong @{thms power_False_cong}
 | 
|
2995  | 
|> simpset_of  | 
|
2996  | 
||
2997  | 
fun unat_arith_tacs ctxt =  | 
|
2998  | 
let  | 
|
2999  | 
fun arith_tac' n t =  | 
|
3000  | 
Arith_Data.arith_tac ctxt n t  | 
|
3001  | 
handle Cooper.COOPER _ => Seq.empty;  | 
|
3002  | 
in  | 
|
3003  | 
[ clarify_tac ctxt 1,  | 
|
3004  | 
full_simp_tac (put_simpset unat_arith_simpset ctxt) 1,  | 
|
3005  | 
ALLGOALS (full_simp_tac  | 
|
3006  | 
(put_simpset HOL_ss ctxt  | 
|
3007  | 
          |> fold Splitter.add_split @{thms unat_splits}
 | 
|
3008  | 
          |> fold Simplifier.add_cong @{thms power_False_cong})),
 | 
|
3009  | 
      rewrite_goals_tac ctxt @{thms word_size},
 | 
|
3010  | 
ALLGOALS (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN  | 
|
3011  | 
REPEAT (eresolve_tac ctxt [conjE] n) THEN  | 
|
3012  | 
                         REPEAT (dresolve_tac ctxt @{thms of_nat_inverse} n THEN assume_tac ctxt n)),
 | 
|
3013  | 
TRYALL arith_tac' ]  | 
|
3014  | 
end  | 
|
3015  | 
||
3016  | 
fun unat_arith_tac ctxt = SELECT_GOAL (EVERY (unat_arith_tacs ctxt))  | 
|
3017  | 
\<close>  | 
|
3018  | 
||
3019  | 
method_setup unat_arith =  | 
|
3020  | 
\<open>Scan.succeed (SIMPLE_METHOD' o unat_arith_tac)\<close>  | 
|
3021  | 
"solving word arithmetic via natural numbers and arith"  | 
|
| 37660 | 3022  | 
|
| 67408 | 3023  | 
\<comment> \<open>\<open>uint_arith_tac\<close>: reduce to arithmetic on int, try to solve by arith\<close>  | 
| 61799 | 3024  | 
ML \<open>  | 
| 72292 | 3025  | 
val uint_arith_simpset =  | 
| 74592 | 3026  | 
  @{context} (* TODO: completely explicitly determined simpset *)
 | 
| 72292 | 3027  | 
  |> fold Simplifier.add_simp @{thms uint_arith_simps}
 | 
3028  | 
  |> fold Splitter.add_split @{thms if_split_asm}
 | 
|
3029  | 
  |> fold Simplifier.add_cong @{thms power_False_cong}
 | 
|
3030  | 
|> simpset_of;  | 
|
3031  | 
||
| 65268 | 3032  | 
fun uint_arith_tacs ctxt =  | 
| 37660 | 3033  | 
let  | 
3034  | 
fun arith_tac' n t =  | 
|
| 
59657
 
2441a80fb6c1
eliminated unused arith "verbose" flag -- tools that need options can use the context;
 
wenzelm 
parents: 
59498 
diff
changeset
 | 
3035  | 
Arith_Data.arith_tac ctxt n t  | 
| 37660 | 3036  | 
handle Cooper.COOPER _ => Seq.empty;  | 
| 65268 | 3037  | 
in  | 
| 42793 | 3038  | 
[ clarify_tac ctxt 1,  | 
| 72292 | 3039  | 
full_simp_tac (put_simpset uint_arith_simpset ctxt) 1,  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51375 
diff
changeset
 | 
3040  | 
ALLGOALS (full_simp_tac  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51375 
diff
changeset
 | 
3041  | 
(put_simpset HOL_ss ctxt  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51375 
diff
changeset
 | 
3042  | 
          |> fold Splitter.add_split @{thms uint_splits}
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51375 
diff
changeset
 | 
3043  | 
          |> fold Simplifier.add_cong @{thms power_False_cong})),
 | 
| 65268 | 3044  | 
      rewrite_goals_tac ctxt @{thms word_size},
 | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59487 
diff
changeset
 | 
3045  | 
ALLGOALS (fn n => REPEAT (resolve_tac ctxt [allI, impI] n) THEN  | 
| 60754 | 3046  | 
REPEAT (eresolve_tac ctxt [conjE] n) THEN  | 
| 65268 | 3047  | 
                         REPEAT (dresolve_tac ctxt @{thms word_of_int_inverse} n
 | 
3048  | 
THEN assume_tac ctxt n  | 
|
| 
58963
 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 
wenzelm 
parents: 
58874 
diff
changeset
 | 
3049  | 
THEN assume_tac ctxt n)),  | 
| 37660 | 3050  | 
TRYALL arith_tac' ]  | 
3051  | 
end  | 
|
3052  | 
||
3053  | 
fun uint_arith_tac ctxt = SELECT_GOAL (EVERY (uint_arith_tacs ctxt))  | 
|
| 61799 | 3054  | 
\<close>  | 
| 37660 | 3055  | 
|
| 65268 | 3056  | 
method_setup uint_arith =  | 
| 61799 | 3057  | 
\<open>Scan.succeed (SIMPLE_METHOD' o uint_arith_tac)\<close>  | 
| 37660 | 3058  | 
"solving word arithmetic via integers and arith"  | 
3059  | 
||
3060  | 
||
| 61799 | 3061  | 
subsection \<open>More on overflows and monotonicity\<close>  | 
| 37660 | 3062  | 
|
| 65328 | 3063  | 
lemma no_plus_overflow_uint_size: "x \<le> x + y \<longleftrightarrow> uint x + uint y < 2 ^ size x"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3064  | 
for x y :: "'a::len word"  | 
| 80777 | 3065  | 
by (auto simp: word_size word_le_def uint_add_lem uint_sub_lem)  | 
| 37660 | 3066  | 
|
3067  | 
lemmas no_olen_add = no_plus_overflow_uint_size [unfolded word_size]  | 
|
3068  | 
||
| 65328 | 3069  | 
lemma no_ulen_sub: "x \<ge> x - y \<longleftrightarrow> uint y \<le> uint x"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3070  | 
for x y :: "'a::len word"  | 
| 80777 | 3071  | 
by (auto simp: word_size word_le_def uint_add_lem uint_sub_lem)  | 
| 37660 | 3072  | 
|
| 70185 | 3073  | 
lemma no_olen_add': "x \<le> y + x \<longleftrightarrow> uint y + uint x < 2 ^ LENGTH('a)"
 | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3074  | 
for x y :: "'a::len word"  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
57512 
diff
changeset
 | 
3075  | 
by (simp add: ac_simps no_olen_add)  | 
| 37660 | 3076  | 
|
| 45604 | 3077  | 
lemmas olen_add_eqv = trans [OF no_olen_add no_olen_add' [symmetric]]  | 
3078  | 
||
3079  | 
lemmas uint_plus_simple_iff = trans [OF no_olen_add uint_add_lem]  | 
|
3080  | 
lemmas uint_plus_simple = uint_plus_simple_iff [THEN iffD1]  | 
|
3081  | 
lemmas uint_minus_simple_iff = trans [OF no_ulen_sub uint_sub_lem]  | 
|
| 37660 | 3082  | 
lemmas uint_minus_simple_alt = uint_sub_lem [folded word_le_def]  | 
3083  | 
lemmas word_sub_le_iff = no_ulen_sub [folded word_le_def]  | 
|
| 45604 | 3084  | 
lemmas word_sub_le = word_sub_le_iff [THEN iffD2]  | 
| 37660 | 3085  | 
|
| 65328 | 3086  | 
lemma word_less_sub1: "x \<noteq> 0 \<Longrightarrow> 1 < x \<longleftrightarrow> 0 < x - 1"  | 
3087  | 
for x :: "'a::len word"  | 
|
| 74592 | 3088  | 
by transfer (simp add: take_bit_decr_eq)  | 
| 37660 | 3089  | 
|
| 65328 | 3090  | 
lemma word_le_sub1: "x \<noteq> 0 \<Longrightarrow> 1 \<le> x \<longleftrightarrow> 0 \<le> x - 1"  | 
3091  | 
for x :: "'a::len word"  | 
|
| 74592 | 3092  | 
by transfer (simp add: int_one_le_iff_zero_less less_le)  | 
| 37660 | 3093  | 
|
| 65328 | 3094  | 
lemma sub_wrap_lt: "x < x - z \<longleftrightarrow> x < z"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3095  | 
for x z :: "'a::len word"  | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
3096  | 
by (meson linorder_not_le word_sub_le_iff)  | 
| 74592 | 3097  | 
|
| 65328 | 3098  | 
lemma sub_wrap: "x \<le> x - z \<longleftrightarrow> z = 0 \<or> x < z"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3099  | 
for x z :: "'a::len word"  | 
| 74592 | 3100  | 
by (simp add: le_less sub_wrap_lt ac_simps)  | 
| 37660 | 3101  | 
|
| 65328 | 3102  | 
lemma plus_minus_not_NULL_ab: "x \<le> ab - c \<Longrightarrow> c \<le> ab \<Longrightarrow> c \<noteq> 0 \<Longrightarrow> x + c \<noteq> 0"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3103  | 
for x ab c :: "'a::len word"  | 
| 37660 | 3104  | 
by uint_arith  | 
3105  | 
||
| 65328 | 3106  | 
lemma plus_minus_no_overflow_ab: "x \<le> ab - c \<Longrightarrow> c \<le> ab \<Longrightarrow> x \<le> x + c"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3107  | 
for x ab c :: "'a::len word"  | 
| 37660 | 3108  | 
by uint_arith  | 
3109  | 
||
| 65328 | 3110  | 
lemma le_minus': "a + c \<le> b \<Longrightarrow> a \<le> a + c \<Longrightarrow> c \<le> b - a"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3111  | 
for a b c :: "'a::len word"  | 
| 37660 | 3112  | 
by uint_arith  | 
3113  | 
||
| 65328 | 3114  | 
lemma le_plus': "a \<le> b \<Longrightarrow> c \<le> b - a \<Longrightarrow> a + c \<le> b"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3115  | 
for a b c :: "'a::len word"  | 
| 37660 | 3116  | 
by uint_arith  | 
3117  | 
||
3118  | 
lemmas le_plus = le_plus' [rotated]  | 
|
3119  | 
||
| 46011 | 3120  | 
lemmas le_minus = leD [THEN thin_rl, THEN le_minus'] (* FIXME *)  | 
| 37660 | 3121  | 
|
| 65328 | 3122  | 
lemma word_plus_mono_right: "y \<le> z \<Longrightarrow> x \<le> x + z \<Longrightarrow> x + y \<le> x + z"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3123  | 
for x y z :: "'a::len word"  | 
| 37660 | 3124  | 
by uint_arith  | 
3125  | 
||
| 65328 | 3126  | 
lemma word_less_minus_cancel: "y - x < z - x \<Longrightarrow> x \<le> z \<Longrightarrow> y < z"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3127  | 
for x y z :: "'a::len word"  | 
| 37660 | 3128  | 
by uint_arith  | 
3129  | 
||
| 65328 | 3130  | 
lemma word_less_minus_mono_left: "y < z \<Longrightarrow> x \<le> y \<Longrightarrow> y - x < z - x"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3131  | 
for x y z :: "'a::len word"  | 
| 37660 | 3132  | 
by uint_arith  | 
3133  | 
||
| 65328 | 3134  | 
lemma word_less_minus_mono: "a < c \<Longrightarrow> d < b \<Longrightarrow> a - b < a \<Longrightarrow> c - d < c \<Longrightarrow> a - b < c - d"  | 
3135  | 
for a b c d :: "'a::len word"  | 
|
| 37660 | 3136  | 
by uint_arith  | 
3137  | 
||
| 65328 | 3138  | 
lemma word_le_minus_cancel: "y - x \<le> z - x \<Longrightarrow> x \<le> z \<Longrightarrow> y \<le> z"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3139  | 
for x y z :: "'a::len word"  | 
| 37660 | 3140  | 
by uint_arith  | 
3141  | 
||
| 65328 | 3142  | 
lemma word_le_minus_mono_left: "y \<le> z \<Longrightarrow> x \<le> y \<Longrightarrow> y - x \<le> z - x"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3143  | 
for x y z :: "'a::len word"  | 
| 37660 | 3144  | 
by uint_arith  | 
3145  | 
||
| 65268 | 3146  | 
lemma word_le_minus_mono:  | 
| 65328 | 3147  | 
"a \<le> c \<Longrightarrow> d \<le> b \<Longrightarrow> a - b \<le> a \<Longrightarrow> c - d \<le> c \<Longrightarrow> a - b \<le> c - d"  | 
3148  | 
for a b c d :: "'a::len word"  | 
|
| 37660 | 3149  | 
by uint_arith  | 
3150  | 
||
| 65328 | 3151  | 
lemma plus_le_left_cancel_wrap: "x + y' < x \<Longrightarrow> x + y < x \<Longrightarrow> x + y' < x + y \<longleftrightarrow> y' < y"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3152  | 
for x y y' :: "'a::len word"  | 
| 37660 | 3153  | 
by uint_arith  | 
3154  | 
||
| 65328 | 3155  | 
lemma plus_le_left_cancel_nowrap: "x \<le> x + y' \<Longrightarrow> x \<le> x + y \<Longrightarrow> x + y' < x + y \<longleftrightarrow> y' < y"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3156  | 
for x y y' :: "'a::len word"  | 
| 37660 | 3157  | 
by uint_arith  | 
3158  | 
||
| 65328 | 3159  | 
lemma word_plus_mono_right2: "a \<le> a + b \<Longrightarrow> c \<le> b \<Longrightarrow> a \<le> a + c"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3160  | 
for a b c :: "'a::len word"  | 
| 65328 | 3161  | 
by uint_arith  | 
3162  | 
||
3163  | 
lemma word_less_add_right: "x < y - z \<Longrightarrow> z \<le> y \<Longrightarrow> x + z < y"  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3164  | 
for x y z :: "'a::len word"  | 
| 37660 | 3165  | 
by uint_arith  | 
3166  | 
||
| 65328 | 3167  | 
lemma word_less_sub_right: "x < y + z \<Longrightarrow> y \<le> x \<Longrightarrow> x - y < z"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3168  | 
for x y z :: "'a::len word"  | 
| 37660 | 3169  | 
by uint_arith  | 
3170  | 
||
| 65328 | 3171  | 
lemma word_le_plus_either: "x \<le> y \<or> x \<le> z \<Longrightarrow> y \<le> y + z \<Longrightarrow> x \<le> y + z"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3172  | 
for x y z :: "'a::len word"  | 
| 37660 | 3173  | 
by uint_arith  | 
3174  | 
||
| 65328 | 3175  | 
lemma word_less_nowrapI: "x < z - k \<Longrightarrow> k \<le> z \<Longrightarrow> 0 < k \<Longrightarrow> x < x + k"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3176  | 
for x z k :: "'a::len word"  | 
| 37660 | 3177  | 
by uint_arith  | 
3178  | 
||
| 65328 | 3179  | 
lemma inc_le: "i < m \<Longrightarrow> i + 1 \<le> m"  | 
3180  | 
for i m :: "'a::len word"  | 
|
| 37660 | 3181  | 
by uint_arith  | 
3182  | 
||
| 82687 | 3183  | 
lemma less_imp_less_eq_dec:  | 
3184  | 
\<open>v \<le> w - 1\<close> if \<open>v < w\<close> for v w :: \<open>'a::len word\<close>  | 
|
3185  | 
using that proof transfer  | 
|
3186  | 
  show \<open>take_bit LENGTH('a) k \<le> take_bit LENGTH('a) (l - 1)\<close>
 | 
|
3187  | 
    if \<open>take_bit LENGTH('a) k < take_bit LENGTH('a) l\<close>
 | 
|
3188  | 
for k l :: int  | 
|
3189  | 
    using that by (cases \<open>take_bit LENGTH('a) l = 0\<close>)
 | 
|
3190  | 
(auto simp add: take_bit_decr_eq)  | 
|
3191  | 
qed  | 
|
3192  | 
||
3193  | 
lemma inc_less_eq_triv_imp:  | 
|
3194  | 
\<open>w = - 1\<close> if \<open>w + 1 \<le> w\<close> for w :: \<open>'a::len word\<close>  | 
|
3195  | 
proof (rule ccontr)  | 
|
3196  | 
assume \<open>w \<noteq> - 1\<close>  | 
|
3197  | 
with that show False  | 
|
3198  | 
by transfer (auto simp add: take_bit_eq_mask_iff dest: take_bit_decr_eq)  | 
|
3199  | 
qed  | 
|
3200  | 
||
3201  | 
lemma less_eq_dec_triv_imp:  | 
|
3202  | 
\<open>w = 0\<close> if \<open>w \<le> w - 1\<close> for w :: \<open>'a::len word\<close>  | 
|
3203  | 
proof (rule ccontr)  | 
|
3204  | 
assume \<open>w \<noteq> 0\<close>  | 
|
3205  | 
with that show False  | 
|
3206  | 
by transfer (auto simp add: take_bit_eq_mask_iff dest: take_bit_decr_eq)  | 
|
3207  | 
qed  | 
|
3208  | 
||
3209  | 
lemma inc_less_eq_iff:  | 
|
3210  | 
\<open>v + 1 \<le> w \<longleftrightarrow> v = - 1 \<or> v < w\<close> for v w :: \<open>'a::len word\<close>  | 
|
3211  | 
by (auto intro: inc_less_eq_triv_imp inc_le)  | 
|
3212  | 
||
3213  | 
lemma less_eq_dec_iff:  | 
|
3214  | 
\<open>v \<le> w - 1 \<longleftrightarrow> w = 0 \<or> v < w\<close> for v w :: \<open>'a::len word\<close>  | 
|
3215  | 
by (auto intro: less_eq_dec_triv_imp less_imp_less_eq_dec)  | 
|
3216  | 
||
| 65328 | 3217  | 
lemma inc_i: "1 \<le> i \<Longrightarrow> i < m \<Longrightarrow> 1 \<le> i + 1 \<and> i + 1 \<le> m"  | 
3218  | 
for i m :: "'a::len word"  | 
|
| 37660 | 3219  | 
by uint_arith  | 
3220  | 
||
3221  | 
lemma udvd_incr_lem:  | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
3222  | 
"\<lbrakk>up < uq; up = ua + n * uint K; uq = ua + n' * uint K\<rbrakk>  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
3223  | 
\<Longrightarrow> up + uint K \<le> uq"  | 
| 71997 | 3224  | 
by auto (metis int_distrib(1) linorder_not_less mult.left_neutral mult_right_mono uint_nonnegative zless_imp_add1_zle)  | 
| 37660 | 3225  | 
|
| 65268 | 3226  | 
lemma udvd_incr':  | 
3227  | 
"p < q \<Longrightarrow> uint p = ua + n * uint K \<Longrightarrow>  | 
|
| 65328 | 3228  | 
uint q = ua + n' * uint K \<Longrightarrow> p + K \<le> q"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3229  | 
unfolding word_less_alt word_le_def  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3230  | 
by (metis (full_types) order_trans udvd_incr_lem uint_add_le)  | 
| 37660 | 3231  | 
|
| 65268 | 3232  | 
lemma udvd_decr':  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3233  | 
assumes "p < q" "uint p = ua + n * uint K" "uint q = ua + n' * uint K"  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3234  | 
shows "uint q = ua + n' * uint K \<Longrightarrow> p \<le> q - K"  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3235  | 
proof -  | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
3236  | 
have "\<And>w v. uint (w::'a word) \<le> uint v + uint (w - v)"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3237  | 
by (metis (no_types) add_diff_cancel_left' diff_add_cancel uint_add_le)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3238  | 
moreover have "uint K + uint p \<le> uint q"  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3239  | 
using assms by (metis (no_types) add_diff_cancel_left' diff_add_cancel udvd_incr_lem word_less_def)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3240  | 
ultimately show ?thesis  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3241  | 
by (meson add_le_cancel_left order_trans word_less_eq_iff_unsigned)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3242  | 
qed  | 
| 37660 | 3243  | 
|
| 
45816
 
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
 
huffman 
parents: 
45811 
diff
changeset
 | 
3244  | 
lemmas udvd_incr_lem0 = udvd_incr_lem [where ua=0, unfolded add_0_left]  | 
| 
 
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
 
huffman 
parents: 
45811 
diff
changeset
 | 
3245  | 
lemmas udvd_incr0 = udvd_incr' [where ua=0, unfolded add_0_left]  | 
| 
 
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
 
huffman 
parents: 
45811 
diff
changeset
 | 
3246  | 
lemmas udvd_decr0 = udvd_decr' [where ua=0, unfolded add_0_left]  | 
| 37660 | 3247  | 
|
| 65328 | 3248  | 
lemma udvd_minus_le': "xy < k \<Longrightarrow> z udvd xy \<Longrightarrow> z udvd k \<Longrightarrow> xy \<le> k - z"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3249  | 
unfolding udvd_unfold_int  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3250  | 
by (meson udvd_decr0)  | 
| 37660 | 3251  | 
|
| 65268 | 3252  | 
lemma udvd_incr2_K:  | 
| 65328 | 3253  | 
"p < a + s \<Longrightarrow> a \<le> a + s \<Longrightarrow> K udvd s \<Longrightarrow> K udvd p - a \<Longrightarrow> a \<le> p \<Longrightarrow>  | 
3254  | 
0 < K \<Longrightarrow> p \<le> p + K \<and> p + K \<le> a + s"  | 
|
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3255  | 
unfolding udvd_unfold_int  | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
3256  | 
by (smt (verit, best) diff_add_cancel leD udvd_incr_lem uint_plus_if'  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
3257  | 
word_less_eq_iff_unsigned word_sub_le)  | 
| 37660 | 3258  | 
|
3259  | 
||
| 61799 | 3260  | 
subsection \<open>Arithmetic type class instantiations\<close>  | 
| 37660 | 3261  | 
|
3262  | 
lemmas word_le_0_iff [simp] =  | 
|
| 
70749
 
5d06b7bb9d22
More type class generalisations. Note that linorder_antisym_conv1 and linorder_antisym_conv2 no longer exist.
 
paulson <lp15@cam.ac.uk> 
parents: 
70342 
diff
changeset
 | 
3263  | 
word_zero_le [THEN leD, THEN antisym_conv1]  | 
| 37660 | 3264  | 
|
| 65328 | 3265  | 
lemma word_of_int_nat: "0 \<le> x \<Longrightarrow> word_of_int x = of_nat (nat x)"  | 
| 72262 | 3266  | 
by simp  | 
| 37660 | 3267  | 
|
| 67408 | 3268  | 
text \<open>  | 
3269  | 
note that \<open>iszero_def\<close> is only for class \<open>comm_semiring_1_cancel\<close>,  | 
|
3270  | 
which requires word length \<open>\<ge> 1\<close>, ie \<open>'a::len word\<close>  | 
|
3271  | 
\<close>  | 
|
| 46603 | 3272  | 
lemma iszero_word_no [simp]:  | 
| 65268 | 3273  | 
"iszero (numeral bin :: 'a::len word) =  | 
| 72128 | 3274  | 
    iszero (take_bit LENGTH('a) (numeral bin :: int))"
 | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3275  | 
by (metis iszero_def uint_0_iff uint_bintrunc)  | 
| 65268 | 3276  | 
|
| 61799 | 3277  | 
text \<open>Use \<open>iszero\<close> to simplify equalities between word numerals.\<close>  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
3278  | 
|
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
3279  | 
lemmas word_eq_numeral_iff_iszero [simp] =  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
3280  | 
eq_numeral_iff_iszero [where 'a="'a::len word"]  | 
| 46603 | 3281  | 
|
| 79590 | 3282  | 
lemma word_less_eq_imp_half_less_eq:  | 
3283  | 
\<open>v div 2 \<le> w div 2\<close> if \<open>v \<le> w\<close> for v w :: \<open>'a::len word\<close>  | 
|
3284  | 
using that by (simp add: word_le_nat_alt unat_div div_le_mono)  | 
|
3285  | 
||
3286  | 
lemma word_half_less_imp_less_eq:  | 
|
3287  | 
\<open>v \<le> w\<close> if \<open>v div 2 < w div 2\<close> for v w :: \<open>'a::len word\<close>  | 
|
3288  | 
using that linorder_linear word_less_eq_imp_half_less_eq by fastforce  | 
|
3289  | 
||
| 37660 | 3290  | 
|
| 61799 | 3291  | 
subsection \<open>Word and nat\<close>  | 
| 37660 | 3292  | 
|
| 70185 | 3293  | 
lemma word_nchotomy: "\<forall>w :: 'a::len word. \<exists>n. w = of_nat n \<and> n < 2 ^ LENGTH('a)"
 | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3294  | 
by (metis of_nat_unat ucast_id unsigned_less)  | 
| 37660 | 3295  | 
|
| 70185 | 3296  | 
lemma of_nat_eq: "of_nat n = w \<longleftrightarrow> (\<exists>q. n = unat w + q * 2 ^ LENGTH('a))"
 | 
| 65328 | 3297  | 
for w :: "'a::len word"  | 
| 68157 | 3298  | 
  using mod_div_mult_eq [of n "2 ^ LENGTH('a)", symmetric]
 | 
| 74496 | 3299  | 
by (auto simp flip: take_bit_eq_mod simp add: unsigned_of_nat)  | 
| 37660 | 3300  | 
|
| 65328 | 3301  | 
lemma of_nat_eq_size: "of_nat n = w \<longleftrightarrow> (\<exists>q. n = unat w + q * 2 ^ size w)"  | 
| 37660 | 3302  | 
unfolding word_size by (rule of_nat_eq)  | 
3303  | 
||
| 70185 | 3304  | 
lemma of_nat_0: "of_nat m = (0::'a::len word) \<longleftrightarrow> (\<exists>q. m = q * 2 ^ LENGTH('a))"
 | 
| 37660 | 3305  | 
by (simp add: of_nat_eq)  | 
3306  | 
||
| 70185 | 3307  | 
lemma of_nat_2p [simp]: "of_nat (2 ^ LENGTH('a)) = (0::'a::len word)"
 | 
| 45805 | 3308  | 
by (fact mult_1 [symmetric, THEN iffD2 [OF of_nat_0 exI]])  | 
| 37660 | 3309  | 
|
| 65328 | 3310  | 
lemma of_nat_gt_0: "of_nat k \<noteq> 0 \<Longrightarrow> 0 < k"  | 
| 37660 | 3311  | 
by (cases k) auto  | 
3312  | 
||
| 70185 | 3313  | 
lemma of_nat_neq_0: "0 < k \<Longrightarrow> k < 2 ^ LENGTH('a::len) \<Longrightarrow> of_nat k \<noteq> (0 :: 'a word)"
 | 
| 80777 | 3314  | 
by (auto simp : of_nat_0)  | 
| 65328 | 3315  | 
|
3316  | 
lemma Abs_fnat_hom_add: "of_nat a + of_nat b = of_nat (a + b)"  | 
|
| 37660 | 3317  | 
by simp  | 
3318  | 
||
| 65328 | 3319  | 
lemma Abs_fnat_hom_mult: "of_nat a * of_nat b = (of_nat (a * b) :: 'a::len word)"  | 
| 72262 | 3320  | 
by (simp add: wi_hom_mult)  | 
| 37660 | 3321  | 
|
| 65328 | 3322  | 
lemma Abs_fnat_hom_Suc: "word_succ (of_nat a) = of_nat (Suc a)"  | 
| 72262 | 3323  | 
by transfer (simp add: ac_simps)  | 
| 37660 | 3324  | 
|
3325  | 
lemma Abs_fnat_hom_0: "(0::'a::len word) = of_nat 0"  | 
|
| 
45995
 
b16070689726
declare word_of_int_{0,1} [simp], for consistency with word_of_int_bin
 
huffman 
parents: 
45958 
diff
changeset
 | 
3326  | 
by simp  | 
| 37660 | 3327  | 
|
3328  | 
lemma Abs_fnat_hom_1: "(1::'a::len word) = of_nat (Suc 0)"  | 
|
| 
45995
 
b16070689726
declare word_of_int_{0,1} [simp], for consistency with word_of_int_bin
 
huffman 
parents: 
45958 
diff
changeset
 | 
3329  | 
by simp  | 
| 37660 | 3330  | 
|
| 65268 | 3331  | 
lemmas Abs_fnat_homs =  | 
3332  | 
Abs_fnat_hom_add Abs_fnat_hom_mult Abs_fnat_hom_Suc  | 
|
| 37660 | 3333  | 
Abs_fnat_hom_0 Abs_fnat_hom_1  | 
3334  | 
||
| 65328 | 3335  | 
lemma word_arith_nat_add: "a + b = of_nat (unat a + unat b)"  | 
3336  | 
by simp  | 
|
3337  | 
||
3338  | 
lemma word_arith_nat_mult: "a * b = of_nat (unat a * unat b)"  | 
|
| 37660 | 3339  | 
by simp  | 
3340  | 
||
| 65328 | 3341  | 
lemma word_arith_nat_Suc: "word_succ a = of_nat (Suc (unat a))"  | 
| 37660 | 3342  | 
by (subst Abs_fnat_hom_Suc [symmetric]) simp  | 
3343  | 
||
| 65328 | 3344  | 
lemma word_arith_nat_div: "a div b = of_nat (unat a div unat b)"  | 
| 72262 | 3345  | 
by (metis of_int_of_nat_eq of_nat_unat of_nat_div word_div_def)  | 
3346  | 
||
| 65328 | 3347  | 
lemma word_arith_nat_mod: "a mod b = of_nat (unat a mod unat b)"  | 
| 72262 | 3348  | 
by (metis of_int_of_nat_eq of_nat_mod of_nat_unat word_mod_def)  | 
| 37660 | 3349  | 
|
3350  | 
lemmas word_arith_nat_defs =  | 
|
3351  | 
word_arith_nat_add word_arith_nat_mult  | 
|
3352  | 
word_arith_nat_Suc Abs_fnat_hom_0  | 
|
3353  | 
Abs_fnat_hom_1 word_arith_nat_div  | 
|
| 65268 | 3354  | 
word_arith_nat_mod  | 
| 37660 | 3355  | 
|
| 72292 | 3356  | 
lemma unat_of_nat:  | 
3357  | 
  \<open>unat (word_of_nat x :: 'a::len word) = x mod 2 ^ LENGTH('a)\<close>
 | 
|
3358  | 
by transfer (simp flip: take_bit_eq_mod add: nat_take_bit_eq)  | 
|
| 65268 | 3359  | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
3360  | 
lemma unat_cong: "x = y \<Longrightarrow> unat x = unat y"  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
3361  | 
by (fact arg_cong)  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
3362  | 
|
| 37660 | 3363  | 
lemmas unat_word_ariths = word_arith_nat_defs  | 
| 45604 | 3364  | 
[THEN trans [OF unat_cong unat_of_nat]]  | 
| 37660 | 3365  | 
|
3366  | 
lemmas word_sub_less_iff = word_sub_le_iff  | 
|
| 
45816
 
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
 
huffman 
parents: 
45811 
diff
changeset
 | 
3367  | 
[unfolded linorder_not_less [symmetric] Not_eq_iff]  | 
| 37660 | 3368  | 
|
| 65268 | 3369  | 
lemma unat_add_lem:  | 
| 70185 | 3370  | 
  "unat x + unat y < 2 ^ LENGTH('a) \<longleftrightarrow> unat (x + y) = unat x + unat y"
 | 
| 65328 | 3371  | 
for x y :: "'a::len word"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3372  | 
by (metis mod_less unat_word_ariths(1) unsigned_less)  | 
| 37660 | 3373  | 
|
| 65268 | 3374  | 
lemma unat_mult_lem:  | 
| 70185 | 3375  | 
  "unat x * unat y < 2 ^ LENGTH('a) \<longleftrightarrow> unat (x * y) = unat x * unat y"
 | 
| 65363 | 3376  | 
for x y :: "'a::len word"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3377  | 
by (metis mod_less unat_word_ariths(2) unsigned_less)  | 
| 71997 | 3378  | 
|
| 65328 | 3379  | 
lemma le_no_overflow: "x \<le> b \<Longrightarrow> a \<le> a + b \<Longrightarrow> x \<le> a + b"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3380  | 
for a b x :: "'a::len word"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3381  | 
using word_le_plus_either by blast  | 
| 37660 | 3382  | 
|
| 71997 | 3383  | 
lemma uint_div:  | 
3384  | 
\<open>uint (x div y) = uint x div uint y\<close>  | 
|
| 72262 | 3385  | 
by (fact uint_div_distrib)  | 
| 71997 | 3386  | 
|
3387  | 
lemma uint_mod:  | 
|
3388  | 
\<open>uint (x mod y) = uint x mod uint y\<close>  | 
|
| 72262 | 3389  | 
by (fact uint_mod_distrib)  | 
| 71997 | 3390  | 
|
| 65328 | 3391  | 
lemma no_plus_overflow_unat_size: "x \<le> x + y \<longleftrightarrow> unat x + unat y < 2 ^ size x"  | 
3392  | 
for x y :: "'a::len word"  | 
|
| 37660 | 3393  | 
unfolding word_size by unat_arith  | 
3394  | 
||
| 65328 | 3395  | 
lemmas no_olen_add_nat =  | 
3396  | 
no_plus_overflow_unat_size [unfolded word_size]  | 
|
3397  | 
||
3398  | 
lemmas unat_plus_simple =  | 
|
3399  | 
trans [OF no_olen_add_nat unat_add_lem]  | 
|
3400  | 
||
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3401  | 
lemma word_div_mult: "\<lbrakk>0 < y; unat x * unat y < 2 ^ LENGTH('a)\<rbrakk> \<Longrightarrow> x * y div y = x"
 | 
| 65328 | 3402  | 
for x y :: "'a::len word"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3403  | 
by (simp add: unat_eq_zero unat_mult_lem word_arith_nat_div)  | 
| 37660 | 3404  | 
|
| 70185 | 3405  | 
lemma div_lt': "i \<le> k div x \<Longrightarrow> unat i * unat x < 2 ^ LENGTH('a)"
 | 
| 65328 | 3406  | 
for i k x :: "'a::len word"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3407  | 
by unat_arith (meson le_less_trans less_mult_imp_div_less not_le unsigned_less)  | 
| 37660 | 3408  | 
|
3409  | 
lemmas div_lt'' = order_less_imp_le [THEN div_lt']  | 
|
3410  | 
||
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3411  | 
lemma div_lt_mult: "\<lbrakk>i < k div x; 0 < x\<rbrakk> \<Longrightarrow> i * x < k"  | 
| 65328 | 3412  | 
for i k x :: "'a::len word"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3413  | 
by (metis div_le_mono div_lt'' not_le unat_div word_div_mult word_less_iff_unsigned)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3414  | 
|
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3415  | 
lemma div_le_mult: "\<lbrakk>i \<le> k div x; 0 < x\<rbrakk> \<Longrightarrow> i * x \<le> k"  | 
| 65328 | 3416  | 
for i k x :: "'a::len word"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3417  | 
by (metis div_lt' less_mult_imp_div_less not_less unat_arith_simps(2) unat_div unat_mult_lem)  | 
| 37660 | 3418  | 
|
| 70185 | 3419  | 
lemma div_lt_uint': "i \<le> k div x \<Longrightarrow> uint i * uint x < 2 ^ LENGTH('a)"
 | 
| 65328 | 3420  | 
for i k x :: "'a::len word"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3421  | 
unfolding uint_nat  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3422  | 
by (metis div_lt' int_ops(7) of_nat_unat uint_mult_lem unat_mult_lem)  | 
| 37660 | 3423  | 
|
3424  | 
lemmas div_lt_uint'' = order_less_imp_le [THEN div_lt_uint']  | 
|
3425  | 
||
| 70185 | 3426  | 
lemma word_le_exists': "x \<le> y \<Longrightarrow> \<exists>z. y = x + z \<and> uint x + uint z < 2 ^ LENGTH('a)"
 | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3427  | 
for x y z :: "'a::len word"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3428  | 
by (metis add.commute diff_add_cancel no_olen_add)  | 
| 71997 | 3429  | 
|
| 37660 | 3430  | 
lemmas plus_minus_not_NULL = order_less_imp_le [THEN plus_minus_not_NULL_ab]  | 
3431  | 
||
3432  | 
lemmas plus_minus_no_overflow =  | 
|
3433  | 
order_less_imp_le [THEN plus_minus_no_overflow_ab]  | 
|
| 65268 | 3434  | 
|
| 37660 | 3435  | 
lemmas mcs = word_less_minus_cancel word_less_minus_mono_left  | 
3436  | 
word_le_minus_cancel word_le_minus_mono_left  | 
|
3437  | 
||
| 45604 | 3438  | 
lemmas word_l_diffs = mcs [where y = "w + x", unfolded add_diff_cancel] for w x  | 
3439  | 
lemmas word_diff_ls = mcs [where z = "w + x", unfolded add_diff_cancel] for w x  | 
|
3440  | 
lemmas word_plus_mcs = word_diff_ls [where y = "v + x", unfolded add_diff_cancel] for v x  | 
|
| 37660 | 3441  | 
|
| 72292 | 3442  | 
lemma le_unat_uoi:  | 
3443  | 
\<open>y \<le> unat z \<Longrightarrow> unat (word_of_nat y :: 'a word) = y\<close>  | 
|
3444  | 
for z :: \<open>'a::len word\<close>  | 
|
3445  | 
by transfer (simp add: nat_take_bit_eq take_bit_nat_eq_self_iff le_less_trans)  | 
|
| 37660 | 3446  | 
|
| 
66808
 
1907167b6038
elementary definition of division on natural numbers
 
haftmann 
parents: 
66453 
diff
changeset
 | 
3447  | 
lemmas thd = times_div_less_eq_dividend  | 
| 37660 | 3448  | 
|
| 71997 | 3449  | 
lemmas uno_simps [THEN le_unat_uoi] = mod_le_divisor div_le_dividend  | 
| 37660 | 3450  | 
|
| 65328 | 3451  | 
lemma word_mod_div_equality: "(n div b) * b + (n mod b) = n"  | 
3452  | 
for n b :: "'a::len word"  | 
|
| 71997 | 3453  | 
by (fact div_mult_mod_eq)  | 
| 37660 | 3454  | 
|
| 65328 | 3455  | 
lemma word_div_mult_le: "a div b * b \<le> a"  | 
3456  | 
for a b :: "'a::len word"  | 
|
| 71997 | 3457  | 
by (metis div_le_mult mult_not_zero order.not_eq_order_implies_strict order_refl word_zero_le)  | 
| 37660 | 3458  | 
|
| 65328 | 3459  | 
lemma word_mod_less_divisor: "0 < n \<Longrightarrow> m mod n < n"  | 
3460  | 
for m n :: "'a::len word"  | 
|
| 71997 | 3461  | 
by (simp add: unat_arith_simps)  | 
3462  | 
||
| 65328 | 3463  | 
lemma word_of_int_power_hom: "word_of_int a ^ n = (word_of_int (a ^ n) :: 'a::len word)"  | 
| 
45995
 
b16070689726
declare word_of_int_{0,1} [simp], for consistency with word_of_int_bin
 
huffman 
parents: 
45958 
diff
changeset
 | 
3464  | 
by (induct n) (simp_all add: wi_hom_mult [symmetric])  | 
| 37660 | 3465  | 
|
| 65328 | 3466  | 
lemma word_arith_power_alt: "a ^ n = (word_of_int (uint a ^ n) :: 'a::len word)"  | 
| 37660 | 3467  | 
by (simp add : word_of_int_power_hom [symmetric])  | 
3468  | 
||
| 
70183
 
3ea80c950023
incorporated various material from the AFP into the distribution
 
haftmann 
parents: 
70175 
diff
changeset
 | 
3469  | 
lemma unatSuc: "1 + n \<noteq> 0 \<Longrightarrow> unat (1 + n) = Suc (unat n)"  | 
| 
 
3ea80c950023
incorporated various material from the AFP into the distribution
 
haftmann 
parents: 
70175 
diff
changeset
 | 
3470  | 
for n :: "'a::len word"  | 
| 
 
3ea80c950023
incorporated various material from the AFP into the distribution
 
haftmann 
parents: 
70175 
diff
changeset
 | 
3471  | 
by unat_arith  | 
| 
 
3ea80c950023
incorporated various material from the AFP into the distribution
 
haftmann 
parents: 
70175 
diff
changeset
 | 
3472  | 
|
| 37660 | 3473  | 
|
| 61799 | 3474  | 
subsection \<open>Cardinality, finiteness of set of words\<close>  | 
| 37660 | 3475  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3476  | 
lemma inj_on_word_of_int: \<open>inj_on (word_of_int :: int \<Rightarrow> 'a word) {0..<2 ^ LENGTH('a::len)}\<close>
 | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3477  | 
unfolding inj_on_def  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3478  | 
by (metis atLeastLessThan_iff word_of_int_inverse)  | 
| 
71948
 
6ede899d26d3
fundamental construction of word type following existing transfer rules
 
haftmann 
parents: 
71947 
diff
changeset
 | 
3479  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3480  | 
lemma range_uint: \<open>range (uint :: 'a word \<Rightarrow> int) = {0..<2 ^ LENGTH('a::len)}\<close>
 | 
| 72488 | 3481  | 
apply transfer  | 
| 80777 | 3482  | 
apply (auto simp: image_iff)  | 
| 72488 | 3483  | 
apply (metis take_bit_int_eq_self_iff)  | 
3484  | 
done  | 
|
| 
71948
 
6ede899d26d3
fundamental construction of word type following existing transfer rules
 
haftmann 
parents: 
71947 
diff
changeset
 | 
3485  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3486  | 
lemma UNIV_eq: \<open>(UNIV :: 'a word set) = word_of_int ` {0..<2 ^ LENGTH('a::len)}\<close>
 | 
| 80777 | 3487  | 
by (auto simp: image_iff) (metis atLeastLessThan_iff linorder_not_le uint_split)  | 
| 
45809
 
2bee94cbae72
finite class instance for word type; remove unused lemmas
 
huffman 
parents: 
45808 
diff
changeset
 | 
3488  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3489  | 
lemma card_word: "CARD('a word) = 2 ^ LENGTH('a::len)"
 | 
| 
71948
 
6ede899d26d3
fundamental construction of word type following existing transfer rules
 
haftmann 
parents: 
71947 
diff
changeset
 | 
3490  | 
by (simp add: UNIV_eq card_image inj_on_word_of_int)  | 
| 37660 | 3491  | 
|
| 
70183
 
3ea80c950023
incorporated various material from the AFP into the distribution
 
haftmann 
parents: 
70175 
diff
changeset
 | 
3492  | 
lemma card_word_size: "CARD('a word) = 2 ^ size x"
 | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3493  | 
for x :: "'a::len word"  | 
| 65328 | 3494  | 
unfolding word_size by (rule card_word)  | 
| 37660 | 3495  | 
|
| 74097 | 3496  | 
end  | 
3497  | 
||
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3498  | 
instance word :: (len) finite  | 
| 
71948
 
6ede899d26d3
fundamental construction of word type following existing transfer rules
 
haftmann 
parents: 
71947 
diff
changeset
 | 
3499  | 
by standard (simp add: UNIV_eq)  | 
| 
 
6ede899d26d3
fundamental construction of word type following existing transfer rules
 
haftmann 
parents: 
71947 
diff
changeset
 | 
3500  | 
|
| 37660 | 3501  | 
|
| 61799 | 3502  | 
subsection \<open>Bitwise Operations on Words\<close>  | 
| 37660 | 3503  | 
|
| 74097 | 3504  | 
context  | 
3505  | 
includes bit_operations_syntax  | 
|
3506  | 
begin  | 
|
3507  | 
||
| 46011 | 3508  | 
lemma word_wi_log_defs:  | 
| 71149 | 3509  | 
"NOT (word_of_int a) = word_of_int (NOT a)"  | 
| 46011 | 3510  | 
"word_of_int a AND word_of_int b = word_of_int (a AND b)"  | 
3511  | 
"word_of_int a OR word_of_int b = word_of_int (a OR b)"  | 
|
3512  | 
"word_of_int a XOR word_of_int b = word_of_int (a XOR b)"  | 
|
| 
47374
 
9475d524bafb
set up and use lift_definition for word operations
 
huffman 
parents: 
47372 
diff
changeset
 | 
3513  | 
by (transfer, rule refl)+  | 
| 47372 | 3514  | 
|
| 46011 | 3515  | 
lemma word_no_log_defs [simp]:  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
3516  | 
"NOT (numeral a) = word_of_int (NOT (numeral a))"  | 
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54225 
diff
changeset
 | 
3517  | 
"NOT (- numeral a) = word_of_int (NOT (- numeral a))"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
3518  | 
"numeral a AND numeral b = word_of_int (numeral a AND numeral b)"  | 
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54225 
diff
changeset
 | 
3519  | 
"numeral a AND - numeral b = word_of_int (numeral a AND - numeral b)"  | 
| 
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54225 
diff
changeset
 | 
3520  | 
"- numeral a AND numeral b = word_of_int (- numeral a AND numeral b)"  | 
| 
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54225 
diff
changeset
 | 
3521  | 
"- numeral a AND - numeral b = word_of_int (- numeral a AND - numeral b)"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
3522  | 
"numeral a OR numeral b = word_of_int (numeral a OR numeral b)"  | 
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54225 
diff
changeset
 | 
3523  | 
"numeral a OR - numeral b = word_of_int (numeral a OR - numeral b)"  | 
| 
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54225 
diff
changeset
 | 
3524  | 
"- numeral a OR numeral b = word_of_int (- numeral a OR numeral b)"  | 
| 
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54225 
diff
changeset
 | 
3525  | 
"- numeral a OR - numeral b = word_of_int (- numeral a OR - numeral b)"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
3526  | 
"numeral a XOR numeral b = word_of_int (numeral a XOR numeral b)"  | 
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54225 
diff
changeset
 | 
3527  | 
"numeral a XOR - numeral b = word_of_int (numeral a XOR - numeral b)"  | 
| 
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54225 
diff
changeset
 | 
3528  | 
"- numeral a XOR numeral b = word_of_int (- numeral a XOR numeral b)"  | 
| 
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54225 
diff
changeset
 | 
3529  | 
"- numeral a XOR - numeral b = word_of_int (- numeral a XOR - numeral b)"  | 
| 47372 | 3530  | 
by (transfer, rule refl)+  | 
| 37660 | 3531  | 
|
| 61799 | 3532  | 
text \<open>Special cases for when one of the arguments equals 1.\<close>  | 
| 
46064
 
88ef116e0522
add simp rules for bitwise word operations with 1
 
huffman 
parents: 
46057 
diff
changeset
 | 
3533  | 
|
| 
 
88ef116e0522
add simp rules for bitwise word operations with 1
 
huffman 
parents: 
46057 
diff
changeset
 | 
3534  | 
lemma word_bitwise_1_simps [simp]:  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3535  | 
"NOT (1::'a::len word) = -2"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
3536  | 
"1 AND numeral b = word_of_int (1 AND numeral b)"  | 
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54225 
diff
changeset
 | 
3537  | 
"1 AND - numeral b = word_of_int (1 AND - numeral b)"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
3538  | 
"numeral a AND 1 = word_of_int (numeral a AND 1)"  | 
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54225 
diff
changeset
 | 
3539  | 
"- numeral a AND 1 = word_of_int (- numeral a AND 1)"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
3540  | 
"1 OR numeral b = word_of_int (1 OR numeral b)"  | 
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54225 
diff
changeset
 | 
3541  | 
"1 OR - numeral b = word_of_int (1 OR - numeral b)"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
3542  | 
"numeral a OR 1 = word_of_int (numeral a OR 1)"  | 
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54225 
diff
changeset
 | 
3543  | 
"- numeral a OR 1 = word_of_int (- numeral a OR 1)"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
3544  | 
"1 XOR numeral b = word_of_int (1 XOR numeral b)"  | 
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54225 
diff
changeset
 | 
3545  | 
"1 XOR - numeral b = word_of_int (1 XOR - numeral b)"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
3546  | 
"numeral a XOR 1 = word_of_int (numeral a XOR 1)"  | 
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54225 
diff
changeset
 | 
3547  | 
"- numeral a XOR 1 = word_of_int (- numeral a XOR 1)"  | 
| 74496 | 3548  | 
apply (simp_all add: word_uint_eq_iff unsigned_not_eq unsigned_and_eq unsigned_or_eq  | 
3549  | 
unsigned_xor_eq of_nat_take_bit ac_simps unsigned_of_int)  | 
|
| 74163 | 3550  | 
apply (simp_all add: minus_numeral_eq_not_sub_one)  | 
3551  | 
apply (simp_all only: sub_one_eq_not_neg bit.xor_compl_right take_bit_xor bit.double_compl)  | 
|
3552  | 
apply simp_all  | 
|
3553  | 
done  | 
|
| 
46064
 
88ef116e0522
add simp rules for bitwise word operations with 1
 
huffman 
parents: 
46057 
diff
changeset
 | 
3554  | 
|
| 61799 | 3555  | 
text \<open>Special cases for when one of the arguments equals -1.\<close>  | 
| 56979 | 3556  | 
|
3557  | 
lemma word_bitwise_m1_simps [simp]:  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3558  | 
"NOT (-1::'a::len word) = 0"  | 
| 
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3559  | 
"(-1::'a::len word) AND x = x"  | 
| 
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3560  | 
"x AND (-1::'a::len word) = x"  | 
| 
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3561  | 
"(-1::'a::len word) OR x = -1"  | 
| 
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3562  | 
"x OR (-1::'a::len word) = -1"  | 
| 
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3563  | 
" (-1::'a::len word) XOR x = NOT x"  | 
| 
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3564  | 
"x XOR (-1::'a::len word) = NOT x"  | 
| 56979 | 3565  | 
by (transfer, simp)+  | 
3566  | 
||
| 74163 | 3567  | 
lemma word_of_int_not_numeral_eq [simp]:  | 
3568  | 
\<open>(word_of_int (NOT (numeral bin)) :: 'a::len word) = - numeral bin - 1\<close>  | 
|
3569  | 
by transfer (simp add: not_eq_complement)  | 
|
3570  | 
||
| 
71957
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
3571  | 
lemma uint_and:  | 
| 
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
3572  | 
\<open>uint (x AND y) = uint x AND uint y\<close>  | 
| 
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
3573  | 
by transfer simp  | 
| 
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
3574  | 
|
| 
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
3575  | 
lemma uint_or:  | 
| 
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
3576  | 
\<open>uint (x OR y) = uint x OR uint y\<close>  | 
| 
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
3577  | 
by transfer simp  | 
| 
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
3578  | 
|
| 
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
3579  | 
lemma uint_xor:  | 
| 
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
3580  | 
\<open>uint (x XOR y) = uint x XOR uint y\<close>  | 
| 
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
3581  | 
by transfer simp  | 
| 47372 | 3582  | 
|
| 67408 | 3583  | 
\<comment> \<open>get from commutativity, associativity etc of \<open>int_and\<close> etc to same for \<open>word_and etc\<close>\<close>  | 
| 65268 | 3584  | 
lemmas bwsimps =  | 
| 46013 | 3585  | 
wi_hom_add  | 
| 37660 | 3586  | 
word_wi_log_defs  | 
3587  | 
||
3588  | 
lemma word_bw_assocs:  | 
|
3589  | 
"(x AND y) AND z = x AND y AND z"  | 
|
3590  | 
"(x OR y) OR z = x OR y OR z"  | 
|
3591  | 
"(x XOR y) XOR z = x XOR y XOR z"  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3592  | 
for x :: "'a::len word"  | 
| 72508 | 3593  | 
by (fact ac_simps)+  | 
| 65268 | 3594  | 
|
| 37660 | 3595  | 
lemma word_bw_comms:  | 
3596  | 
"x AND y = y AND x"  | 
|
3597  | 
"x OR y = y OR x"  | 
|
3598  | 
"x XOR y = y XOR x"  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3599  | 
for x :: "'a::len word"  | 
| 72508 | 3600  | 
by (fact ac_simps)+  | 
| 65268 | 3601  | 
|
| 37660 | 3602  | 
lemma word_bw_lcs:  | 
3603  | 
"y AND x AND z = x AND y AND z"  | 
|
3604  | 
"y OR x OR z = x OR y OR z"  | 
|
3605  | 
"y XOR x XOR z = x XOR y XOR z"  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3606  | 
for x :: "'a::len word"  | 
| 72508 | 3607  | 
by (fact ac_simps)+  | 
| 37660 | 3608  | 
|
| 
71957
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
3609  | 
lemma word_log_esimps:  | 
| 37660 | 3610  | 
"x AND 0 = 0"  | 
3611  | 
"x AND -1 = x"  | 
|
3612  | 
"x OR 0 = x"  | 
|
3613  | 
"x OR -1 = -1"  | 
|
3614  | 
"x XOR 0 = x"  | 
|
3615  | 
"x XOR -1 = NOT x"  | 
|
3616  | 
"0 AND x = 0"  | 
|
3617  | 
"-1 AND x = x"  | 
|
3618  | 
"0 OR x = x"  | 
|
3619  | 
"-1 OR x = -1"  | 
|
3620  | 
"0 XOR x = x"  | 
|
3621  | 
"-1 XOR x = NOT x"  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3622  | 
for x :: "'a::len word"  | 
| 
71957
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
3623  | 
by simp_all  | 
| 37660 | 3624  | 
|
3625  | 
lemma word_not_dist:  | 
|
3626  | 
"NOT (x OR y) = NOT x AND NOT y"  | 
|
3627  | 
"NOT (x AND y) = NOT x OR NOT y"  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3628  | 
for x :: "'a::len word"  | 
| 
71957
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
3629  | 
by simp_all  | 
| 37660 | 3630  | 
|
3631  | 
lemma word_bw_same:  | 
|
3632  | 
"x AND x = x"  | 
|
3633  | 
"x OR x = x"  | 
|
3634  | 
"x XOR x = 0"  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3635  | 
for x :: "'a::len word"  | 
| 
71957
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
3636  | 
by simp_all  | 
| 37660 | 3637  | 
|
3638  | 
lemma word_ao_absorbs [simp]:  | 
|
3639  | 
"x AND (y OR x) = x"  | 
|
3640  | 
"x OR y AND x = x"  | 
|
3641  | 
"x AND (x OR y) = x"  | 
|
3642  | 
"y AND x OR x = x"  | 
|
3643  | 
"(y OR x) AND x = x"  | 
|
3644  | 
"x OR x AND y = x"  | 
|
3645  | 
"(x OR y) AND x = x"  | 
|
3646  | 
"x AND y OR x = x"  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3647  | 
for x :: "'a::len word"  | 
| 72508 | 3648  | 
by (auto intro: bit_eqI simp add: bit_and_iff bit_or_iff)  | 
| 37660 | 3649  | 
|
| 71149 | 3650  | 
lemma word_not_not [simp]: "NOT (NOT x) = x"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3651  | 
for x :: "'a::len word"  | 
| 72508 | 3652  | 
by (fact bit.double_compl)  | 
| 37660 | 3653  | 
|
| 65328 | 3654  | 
lemma word_ao_dist: "(x OR y) AND z = x AND z OR y AND z"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3655  | 
for x :: "'a::len word"  | 
| 72508 | 3656  | 
by (fact bit.conj_disj_distrib2)  | 
| 37660 | 3657  | 
|
| 65328 | 3658  | 
lemma word_oa_dist: "x AND y OR z = (x OR z) AND (y OR z)"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3659  | 
for x :: "'a::len word"  | 
| 72508 | 3660  | 
by (fact bit.disj_conj_distrib2)  | 
3661  | 
||
| 65328 | 3662  | 
lemma word_add_not [simp]: "x + NOT x = -1"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3663  | 
for x :: "'a::len word"  | 
| 72508 | 3664  | 
by (simp add: not_eq_complement)  | 
3665  | 
||
| 65328 | 3666  | 
lemma word_plus_and_or [simp]: "(x AND y) + (x OR y) = x + y"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3667  | 
for x :: "'a::len word"  | 
| 47372 | 3668  | 
by transfer (simp add: plus_and_or)  | 
| 37660 | 3669  | 
|
| 65328 | 3670  | 
lemma leoa: "w = x OR y \<Longrightarrow> y = w AND y"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3671  | 
for x :: "'a::len word"  | 
| 65328 | 3672  | 
by auto  | 
3673  | 
||
3674  | 
lemma leao: "w' = x' AND y' \<Longrightarrow> x' = x' OR w'"  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3675  | 
for x' :: "'a::len word"  | 
| 65328 | 3676  | 
by auto  | 
3677  | 
||
3678  | 
lemma word_ao_equiv: "w = w OR w' \<longleftrightarrow> w' = w AND w'"  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3679  | 
for w w' :: "'a::len word"  | 
| 48196 | 3680  | 
by (auto intro: leoa leao)  | 
| 37660 | 3681  | 
|
| 65328 | 3682  | 
lemma le_word_or2: "x \<le> x OR y"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
3683  | 
for x y :: "'a::len word"  | 
| 72488 | 3684  | 
by (simp add: or_greater_eq uint_or word_le_def)  | 
| 37660 | 3685  | 
|
| 71997 | 3686  | 
lemmas le_word_or1 = xtrans(3) [OF word_bw_comms (2) le_word_or2]  | 
3687  | 
lemmas word_and_le1 = xtrans(3) [OF word_ao_absorbs (4) [symmetric] le_word_or2]  | 
|
3688  | 
lemmas word_and_le2 = xtrans(3) [OF word_ao_absorbs (8) [symmetric] le_word_or2]  | 
|
| 37660 | 3689  | 
|
| 
72611
 
c7bc3e70a8c7
official collection for bit projection simplifications
 
haftmann 
parents: 
72515 
diff
changeset
 | 
3690  | 
lemma bit_horner_sum_bit_word_iff [bit_simps]:  | 
| 
72027
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3691  | 
\<open>bit (horner_sum of_bool (2 :: 'a::len word) bs) n  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3692  | 
    \<longleftrightarrow> n < min LENGTH('a) (length bs) \<and> bs ! n\<close>
 | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3693  | 
by transfer (simp add: bit_horner_sum_bit_iff)  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3694  | 
|
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3695  | 
definition word_reverse :: \<open>'a::len word \<Rightarrow> 'a word\<close>  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3696  | 
  where \<open>word_reverse w = horner_sum of_bool 2 (rev (map (bit w) [0..<LENGTH('a)]))\<close>
 | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3697  | 
|
| 
72611
 
c7bc3e70a8c7
official collection for bit projection simplifications
 
haftmann 
parents: 
72515 
diff
changeset
 | 
3698  | 
lemma bit_word_reverse_iff [bit_simps]:  | 
| 71990 | 3699  | 
  \<open>bit (word_reverse w) n \<longleftrightarrow> n < LENGTH('a) \<and> bit w (LENGTH('a) - Suc n)\<close>
 | 
3700  | 
for w :: \<open>'a::len word\<close>  | 
|
| 
72027
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3701  | 
  by (cases \<open>n < LENGTH('a)\<close>)
 | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3702  | 
(simp_all add: word_reverse_def bit_horner_sum_bit_word_iff rev_nth)  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3703  | 
|
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3704  | 
lemma word_rev_rev [simp] : "word_reverse (word_reverse w) = w"  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3705  | 
by (rule bit_word_eqI)  | 
| 80777 | 3706  | 
(auto simp: bit_word_reverse_iff bit_imp_le_length Suc_diff_Suc)  | 
| 
72027
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3707  | 
|
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3708  | 
lemma word_rev_gal: "word_reverse w = u \<Longrightarrow> word_reverse u = w"  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3709  | 
by (metis word_rev_rev)  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3710  | 
|
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3711  | 
lemma word_rev_gal': "u = word_reverse w \<Longrightarrow> w = word_reverse u"  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3712  | 
by simp  | 
| 37660 | 3713  | 
|
| 
80401
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
3714  | 
lemma word_eq_reverseI:  | 
| 
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
3715  | 
\<open>v = w\<close> if \<open>word_reverse v = word_reverse w\<close>  | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
3716  | 
by (metis that word_rev_rev)  | 
| 
80401
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
3717  | 
|
| 65328 | 3718  | 
lemma uint_2p: "(0::'a::len word) < 2 ^ n \<Longrightarrow> uint (2 ^ n::'a::len word) = 2 ^ n"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3719  | 
  by (cases \<open>n < LENGTH('a)\<close>; transfer; force)
 | 
| 37660 | 3720  | 
|
| 65268 | 3721  | 
lemma word_of_int_2p: "(word_of_int (2 ^ n) :: 'a::len word) = 2 ^ n"  | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
3722  | 
by simp  | 
| 37660 | 3723  | 
|
3724  | 
||
| 61799 | 3725  | 
subsubsection \<open>shift functions in terms of lists of bools\<close>  | 
| 37660 | 3726  | 
|
| 71997 | 3727  | 
lemma drop_bit_word_numeral [simp]:  | 
3728  | 
\<open>drop_bit (numeral n) (numeral k) =  | 
|
3729  | 
    (word_of_int (drop_bit (numeral n) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\<close>
 | 
|
3730  | 
by transfer simp  | 
|
3731  | 
||
| 74498 | 3732  | 
lemma drop_bit_word_Suc_numeral [simp]:  | 
3733  | 
\<open>drop_bit (Suc n) (numeral k) =  | 
|
3734  | 
    (word_of_int (drop_bit (Suc n) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\<close>
 | 
|
3735  | 
by transfer simp  | 
|
3736  | 
||
3737  | 
lemma drop_bit_word_minus_numeral [simp]:  | 
|
3738  | 
\<open>drop_bit (numeral n) (- numeral k) =  | 
|
3739  | 
    (word_of_int (drop_bit (numeral n) (take_bit LENGTH('a) (- numeral k))) :: 'a::len word)\<close>
 | 
|
3740  | 
by transfer simp  | 
|
3741  | 
||
3742  | 
lemma drop_bit_word_Suc_minus_numeral [simp]:  | 
|
3743  | 
\<open>drop_bit (Suc n) (- numeral k) =  | 
|
3744  | 
    (word_of_int (drop_bit (Suc n) (take_bit LENGTH('a) (- numeral k))) :: 'a::len word)\<close>
 | 
|
3745  | 
by transfer simp  | 
|
3746  | 
||
| 73853 | 3747  | 
lemma signed_drop_bit_word_numeral [simp]:  | 
3748  | 
\<open>signed_drop_bit (numeral n) (numeral k) =  | 
|
3749  | 
    (word_of_int (drop_bit (numeral n) (signed_take_bit (LENGTH('a) - 1) (numeral k))) :: 'a::len word)\<close>
 | 
|
3750  | 
by transfer simp  | 
|
3751  | 
||
| 74498 | 3752  | 
lemma signed_drop_bit_word_Suc_numeral [simp]:  | 
3753  | 
\<open>signed_drop_bit (Suc n) (numeral k) =  | 
|
3754  | 
    (word_of_int (drop_bit (Suc n) (signed_take_bit (LENGTH('a) - 1) (numeral k))) :: 'a::len word)\<close>
 | 
|
3755  | 
by transfer simp  | 
|
3756  | 
||
3757  | 
lemma signed_drop_bit_word_minus_numeral [simp]:  | 
|
3758  | 
\<open>signed_drop_bit (numeral n) (- numeral k) =  | 
|
3759  | 
    (word_of_int (drop_bit (numeral n) (signed_take_bit (LENGTH('a) - 1) (- numeral k))) :: 'a::len word)\<close>
 | 
|
3760  | 
by transfer simp  | 
|
3761  | 
||
3762  | 
lemma signed_drop_bit_word_Suc_minus_numeral [simp]:  | 
|
3763  | 
\<open>signed_drop_bit (Suc n) (- numeral k) =  | 
|
3764  | 
    (word_of_int (drop_bit (Suc n) (signed_take_bit (LENGTH('a) - 1) (- numeral k))) :: 'a::len word)\<close>
 | 
|
3765  | 
by transfer simp  | 
|
3766  | 
||
3767  | 
lemma take_bit_word_numeral [simp]:  | 
|
3768  | 
\<open>take_bit (numeral n) (numeral k) =  | 
|
3769  | 
    (word_of_int (take_bit (min LENGTH('a) (numeral n)) (numeral k)) :: 'a::len word)\<close>
 | 
|
3770  | 
by transfer rule  | 
|
3771  | 
||
3772  | 
lemma take_bit_word_Suc_numeral [simp]:  | 
|
3773  | 
\<open>take_bit (Suc n) (numeral k) =  | 
|
3774  | 
    (word_of_int (take_bit (min LENGTH('a) (Suc n)) (numeral k)) :: 'a::len word)\<close>
 | 
|
3775  | 
by transfer rule  | 
|
3776  | 
||
3777  | 
lemma take_bit_word_minus_numeral [simp]:  | 
|
3778  | 
\<open>take_bit (numeral n) (- numeral k) =  | 
|
3779  | 
    (word_of_int (take_bit (min LENGTH('a) (numeral n)) (- numeral k)) :: 'a::len word)\<close>
 | 
|
3780  | 
by transfer rule  | 
|
3781  | 
||
3782  | 
lemma take_bit_word_Suc_minus_numeral [simp]:  | 
|
3783  | 
\<open>take_bit (Suc n) (- numeral k) =  | 
|
3784  | 
    (word_of_int (take_bit (min LENGTH('a) (Suc n)) (- numeral k)) :: 'a::len word)\<close>
 | 
|
3785  | 
by transfer rule  | 
|
3786  | 
||
3787  | 
lemma signed_take_bit_word_numeral [simp]:  | 
|
3788  | 
\<open>signed_take_bit (numeral n) (numeral k) =  | 
|
3789  | 
    (word_of_int (signed_take_bit (numeral n) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\<close>
 | 
|
3790  | 
by transfer rule  | 
|
3791  | 
||
3792  | 
lemma signed_take_bit_word_Suc_numeral [simp]:  | 
|
3793  | 
\<open>signed_take_bit (Suc n) (numeral k) =  | 
|
3794  | 
    (word_of_int (signed_take_bit (Suc n) (take_bit LENGTH('a) (numeral k))) :: 'a::len word)\<close>
 | 
|
3795  | 
by transfer rule  | 
|
3796  | 
||
3797  | 
lemma signed_take_bit_word_minus_numeral [simp]:  | 
|
3798  | 
\<open>signed_take_bit (numeral n) (- numeral k) =  | 
|
3799  | 
    (word_of_int (signed_take_bit (numeral n) (take_bit LENGTH('a) (- numeral k))) :: 'a::len word)\<close>
 | 
|
3800  | 
by transfer rule  | 
|
3801  | 
||
3802  | 
lemma signed_take_bit_word_Suc_minus_numeral [simp]:  | 
|
3803  | 
\<open>signed_take_bit (Suc n) (- numeral k) =  | 
|
3804  | 
    (word_of_int (signed_take_bit (Suc n) (take_bit LENGTH('a) (- numeral k))) :: 'a::len word)\<close>
 | 
|
3805  | 
by transfer rule  | 
|
3806  | 
||
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3807  | 
lemma False_map2_or: "\<lbrakk>set xs \<subseteq> {False}; length ys = length xs\<rbrakk> \<Longrightarrow> map2 (\<or>) xs ys = ys"
 | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3808  | 
by (induction xs arbitrary: ys) (auto simp: length_Suc_conv)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3809  | 
|
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3810  | 
lemma align_lem_or:  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3811  | 
assumes "length xs = n + m" "length ys = n + m"  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3812  | 
and "drop m xs = replicate n False" "take m ys = replicate m False"  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3813  | 
shows "map2 (\<or>) xs ys = take m xs @ drop m ys"  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3814  | 
using assms  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3815  | 
proof (induction xs arbitrary: ys m)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3816  | 
case (Cons a xs)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3817  | 
then show ?case  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3818  | 
by (cases m) (auto simp: length_Suc_conv False_map2_or)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3819  | 
qed auto  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3820  | 
|
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3821  | 
lemma False_map2_and: "\<lbrakk>set xs \<subseteq> {False}; length ys = length xs\<rbrakk> \<Longrightarrow> map2 (\<and>) xs ys = xs"
 | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3822  | 
by (induction xs arbitrary: ys) (auto simp: length_Suc_conv)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3823  | 
|
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3824  | 
lemma align_lem_and:  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3825  | 
assumes "length xs = n + m" "length ys = n + m"  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3826  | 
and "drop m xs = replicate n False" "take m ys = replicate m False"  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3827  | 
shows "map2 (\<and>) xs ys = replicate (n + m) False"  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3828  | 
using assms  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3829  | 
proof (induction xs arbitrary: ys m)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3830  | 
case (Cons a xs)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3831  | 
then show ?case  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3832  | 
by (cases m) (auto simp: length_Suc_conv set_replicate_conv_if False_map2_and)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3833  | 
qed auto  | 
| 37660 | 3834  | 
|
| 
55816
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55415 
diff
changeset
 | 
3835  | 
|
| 61799 | 3836  | 
subsubsection \<open>Mask\<close>  | 
| 37660 | 3837  | 
|
| 
71957
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
3838  | 
lemma minus_1_eq_mask:  | 
| 72082 | 3839  | 
  \<open>- 1 = (mask LENGTH('a) :: 'a::len word)\<close>
 | 
| 74101 | 3840  | 
by (rule bit_eqI) (simp add: bit_exp_iff bit_mask_iff)  | 
| 72079 | 3841  | 
|
3842  | 
lemma mask_eq_decr_exp:  | 
|
| 72082 | 3843  | 
\<open>mask n = 2 ^ n - (1 :: 'a::len word)\<close>  | 
3844  | 
by (fact mask_eq_exp_minus_1)  | 
|
| 71953 | 3845  | 
|
3846  | 
lemma mask_Suc_rec:  | 
|
| 72082 | 3847  | 
\<open>mask (Suc n) = 2 * mask n + (1 :: 'a::len word)\<close>  | 
3848  | 
by (simp add: mask_eq_exp_minus_1)  | 
|
| 71953 | 3849  | 
|
3850  | 
context  | 
|
3851  | 
begin  | 
|
3852  | 
||
| 
72611
 
c7bc3e70a8c7
official collection for bit projection simplifications
 
haftmann 
parents: 
72515 
diff
changeset
 | 
3853  | 
qualified lemma bit_mask_iff [bit_simps]:  | 
| 71990 | 3854  | 
  \<open>bit (mask m :: 'a::len word) n \<longleftrightarrow> n < min LENGTH('a) m\<close>
 | 
| 74101 | 3855  | 
by (simp add: bit_mask_iff not_le)  | 
| 71953 | 3856  | 
|
3857  | 
end  | 
|
3858  | 
||
| 72128 | 3859  | 
lemma mask_bin: "mask n = word_of_int (take_bit n (- 1))"  | 
| 74592 | 3860  | 
by transfer simp  | 
| 37660 | 3861  | 
|
| 72128 | 3862  | 
lemma and_mask_bintr: "w AND mask n = word_of_int (take_bit n (uint w))"  | 
| 72488 | 3863  | 
by transfer (simp add: ac_simps take_bit_eq_mask)  | 
| 37660 | 3864  | 
|
| 72128 | 3865  | 
lemma and_mask_wi: "word_of_int i AND mask n = word_of_int (take_bit n i)"  | 
| 74496 | 3866  | 
by (simp add: take_bit_eq_mask of_int_and_eq of_int_mask_eq)  | 
| 
46023
 
fad87bb608fc
restate some lemmas to respect int/bin distinction
 
huffman 
parents: 
46022 
diff
changeset
 | 
3867  | 
|
| 65328 | 3868  | 
lemma and_mask_wi':  | 
| 72128 | 3869  | 
  "word_of_int i AND mask n = (word_of_int (take_bit (min LENGTH('a) n) i) :: 'a::len word)"
 | 
| 80777 | 3870  | 
by (auto simp: and_mask_wi min_def wi_bintr)  | 
| 
64593
 
50c715579715
reoriented congruence rules in non-explosive direction
 
haftmann 
parents: 
64243 
diff
changeset
 | 
3871  | 
|
| 72128 | 3872  | 
lemma and_mask_no: "numeral i AND mask n = word_of_int (take_bit n (numeral i))"  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
3873  | 
unfolding word_numeral_alt by (rule and_mask_wi)  | 
| 37660 | 3874  | 
|
| 45811 | 3875  | 
lemma and_mask_mod_2p: "w AND mask n = word_of_int (uint w mod 2 ^ n)"  | 
| 72128 | 3876  | 
by (simp only: and_mask_bintr take_bit_eq_mod)  | 
| 37660 | 3877  | 
|
| 
72130
 
9e5862223442
dedicated symbols for code generation, to pave way for generic conversions from and to word
 
haftmann 
parents: 
72128 
diff
changeset
 | 
3878  | 
lemma uint_mask_eq:  | 
| 
 
9e5862223442
dedicated symbols for code generation, to pave way for generic conversions from and to word
 
haftmann 
parents: 
72128 
diff
changeset
 | 
3879  | 
  \<open>uint (mask n :: 'a::len word) = mask (min LENGTH('a) n)\<close>
 | 
| 
 
9e5862223442
dedicated symbols for code generation, to pave way for generic conversions from and to word
 
haftmann 
parents: 
72128 
diff
changeset
 | 
3880  | 
by transfer simp  | 
| 
 
9e5862223442
dedicated symbols for code generation, to pave way for generic conversions from and to word
 
haftmann 
parents: 
72128 
diff
changeset
 | 
3881  | 
|
| 37660 | 3882  | 
lemma and_mask_lt_2p: "uint (w AND mask n) < 2 ^ n"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3883  | 
by (metis take_bit_eq_mask take_bit_int_less_exp unsigned_take_bit_eq)  | 
| 37660 | 3884  | 
|
| 65363 | 3885  | 
lemma mask_eq_iff: "w AND mask n = w \<longleftrightarrow> uint w < 2 ^ n"  | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
3886  | 
by (metis and_mask_bintr and_mask_lt_2p take_bit_int_eq_self take_bit_nonnegative  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
3887  | 
uint_sint word_of_int_uint)  | 
| 37660 | 3888  | 
|
| 65328 | 3889  | 
lemma and_mask_dvd: "2 ^ n dvd uint w \<longleftrightarrow> w AND mask n = 0"  | 
| 72262 | 3890  | 
by (simp flip: take_bit_eq_mask take_bit_eq_mod unsigned_take_bit_eq add: dvd_eq_mod_eq_0 uint_0_iff)  | 
| 37660 | 3891  | 
|
| 65328 | 3892  | 
lemma and_mask_dvd_nat: "2 ^ n dvd unat w \<longleftrightarrow> w AND mask n = 0"  | 
| 72262 | 3893  | 
by (simp flip: take_bit_eq_mask take_bit_eq_mod unsigned_take_bit_eq add: dvd_eq_mod_eq_0 unat_0_iff uint_0_iff)  | 
| 37660 | 3894  | 
|
| 65328 | 3895  | 
lemma word_2p_lem: "n < size w \<Longrightarrow> w < 2 ^ n = (uint w < 2 ^ n)"  | 
3896  | 
for w :: "'a::len word"  | 
|
| 72262 | 3897  | 
by transfer simp  | 
| 37660 | 3898  | 
|
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3899  | 
lemma less_mask_eq:  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3900  | 
fixes x :: "'a::len word"  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3901  | 
assumes "x < 2 ^ n" shows "x AND mask n = x"  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3902  | 
by (metis (no_types) assms lt2p_lem mask_eq_iff not_less word_2p_lem word_size)  | 
| 37660 | 3903  | 
|
| 45604 | 3904  | 
lemmas mask_eq_iff_w2p = trans [OF mask_eq_iff word_2p_lem [symmetric]]  | 
3905  | 
||
3906  | 
lemmas and_mask_less' = iffD2 [OF word_2p_lem and_mask_lt_2p, simplified word_size]  | 
|
| 37660 | 3907  | 
|
| 72082 | 3908  | 
lemma and_mask_less_size: "n < size x \<Longrightarrow> x AND mask n < 2 ^ n"  | 
3909  | 
for x :: \<open>'a::len word\<close>  | 
|
| 37660 | 3910  | 
unfolding word_size by (erule and_mask_less')  | 
3911  | 
||
| 65328 | 3912  | 
lemma word_mod_2p_is_mask [OF refl]: "c = 2 ^ n \<Longrightarrow> c > 0 \<Longrightarrow> x mod c = x AND mask n"  | 
3913  | 
for c x :: "'a::len word"  | 
|
3914  | 
by (auto simp: word_mod_def uint_2p and_mask_mod_2p)  | 
|
| 37660 | 3915  | 
|
3916  | 
lemma mask_eqs:  | 
|
3917  | 
"(a AND mask n) + b AND mask n = a + b AND mask n"  | 
|
3918  | 
"a + (b AND mask n) AND mask n = a + b AND mask n"  | 
|
3919  | 
"(a AND mask n) - b AND mask n = a - b AND mask n"  | 
|
3920  | 
"a - (b AND mask n) AND mask n = a - b AND mask n"  | 
|
3921  | 
"a * (b AND mask n) AND mask n = a * b AND mask n"  | 
|
3922  | 
"(b AND mask n) * a AND mask n = b * a AND mask n"  | 
|
3923  | 
"(a AND mask n) + (b AND mask n) AND mask n = a + b AND mask n"  | 
|
3924  | 
"(a AND mask n) - (b AND mask n) AND mask n = a - b AND mask n"  | 
|
3925  | 
"(a AND mask n) * (b AND mask n) AND mask n = a * b AND mask n"  | 
|
3926  | 
"- (a AND mask n) AND mask n = - a AND mask n"  | 
|
3927  | 
"word_succ (a AND mask n) AND mask n = word_succ a AND mask n"  | 
|
3928  | 
"word_pred (a AND mask n) AND mask n = word_pred a AND mask n"  | 
|
3929  | 
using word_of_int_Ex [where x=a] word_of_int_Ex [where x=b]  | 
|
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3930  | 
unfolding take_bit_eq_mask [symmetric]  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3931  | 
by (transfer; simp add: take_bit_eq_mod mod_simps)+  | 
| 65328 | 3932  | 
|
3933  | 
lemma mask_power_eq: "(x AND mask n) ^ k AND mask n = x ^ k AND mask n"  | 
|
| 72082 | 3934  | 
for x :: \<open>'a::len word\<close>  | 
| 37660 | 3935  | 
using word_of_int_Ex [where x=x]  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3936  | 
unfolding take_bit_eq_mask [symmetric]  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3937  | 
by (transfer; simp add: take_bit_eq_mod mod_simps)+  | 
| 37660 | 3938  | 
|
| 
70183
 
3ea80c950023
incorporated various material from the AFP into the distribution
 
haftmann 
parents: 
70175 
diff
changeset
 | 
3939  | 
lemma mask_full [simp]: "mask LENGTH('a) = (- 1 :: 'a::len word)"
 | 
| 74592 | 3940  | 
by transfer simp  | 
| 
70183
 
3ea80c950023
incorporated various material from the AFP into the distribution
 
haftmann 
parents: 
70175 
diff
changeset
 | 
3941  | 
|
| 37660 | 3942  | 
|
| 
72027
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3943  | 
subsubsection \<open>Slices\<close>  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3944  | 
|
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3945  | 
definition slice1 :: \<open>nat \<Rightarrow> 'a::len word \<Rightarrow> 'b::len word\<close>  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3946  | 
  where \<open>slice1 n w = (if n < LENGTH('a)
 | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3947  | 
    then ucast (drop_bit (LENGTH('a) - n) w)
 | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3948  | 
    else push_bit (n - LENGTH('a)) (ucast w))\<close>
 | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3949  | 
|
| 
72611
 
c7bc3e70a8c7
official collection for bit projection simplifications
 
haftmann 
parents: 
72515 
diff
changeset
 | 
3950  | 
lemma bit_slice1_iff [bit_simps]:  | 
| 
72027
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3951  | 
  \<open>bit (slice1 m w :: 'b::len word) n \<longleftrightarrow> m - LENGTH('a) \<le> n \<and> n < min LENGTH('b) m
 | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3952  | 
    \<and> bit w (n + (LENGTH('a) - m) - (m - LENGTH('a)))\<close>
 | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3953  | 
for w :: \<open>'a::len word\<close>  | 
| 80777 | 3954  | 
by (auto simp: slice1_def bit_ucast_iff bit_drop_bit_eq bit_push_bit_iff not_less not_le ac_simps  | 
| 
72027
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3955  | 
dest: bit_imp_le_length)  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3956  | 
|
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3957  | 
definition slice :: \<open>nat \<Rightarrow> 'a::len word \<Rightarrow> 'b::len word\<close>  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3958  | 
  where \<open>slice n = slice1 (LENGTH('a) - n)\<close>
 | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3959  | 
|
| 
72611
 
c7bc3e70a8c7
official collection for bit projection simplifications
 
haftmann 
parents: 
72515 
diff
changeset
 | 
3960  | 
lemma bit_slice_iff [bit_simps]:  | 
| 
72027
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3961  | 
  \<open>bit (slice m w :: 'b::len word) n \<longleftrightarrow> n < min LENGTH('b) (LENGTH('a) - m) \<and> bit w (n + LENGTH('a) - (LENGTH('a) - m))\<close>
 | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3962  | 
for w :: \<open>'a::len word\<close>  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3963  | 
by (simp add: slice_def word_size bit_slice1_iff)  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3964  | 
|
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3965  | 
lemma slice1_0 [simp] : "slice1 n 0 = 0"  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3966  | 
unfolding slice1_def by simp  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3967  | 
|
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3968  | 
lemma slice_0 [simp] : "slice n 0 = 0"  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3969  | 
unfolding slice_def by auto  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3970  | 
|
| 
72088
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
3971  | 
lemma ucast_slice1: "ucast w = slice1 (size w) w"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3972  | 
unfolding slice1_def by (simp add: size_word.rep_eq)  | 
| 
72027
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3973  | 
|
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3974  | 
lemma ucast_slice: "ucast w = slice 0 w"  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3975  | 
by (simp add: slice_def slice1_def)  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3976  | 
|
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3977  | 
lemma slice_id: "slice 0 t = t"  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3978  | 
by (simp only: ucast_slice [symmetric] ucast_id)  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3979  | 
|
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3980  | 
lemma rev_slice1:  | 
| 
72088
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
3981  | 
\<open>slice1 n (word_reverse w :: 'b::len word) = word_reverse (slice1 k w :: 'a::len word)\<close>  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
3982  | 
  if \<open>n + k = LENGTH('a) + LENGTH('b)\<close>
 | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
3983  | 
proof (rule bit_word_eqI)  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
3984  | 
fix m  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
3985  | 
  assume *: \<open>m < LENGTH('a)\<close>
 | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
3986  | 
  from that have **: \<open>LENGTH('b) = n + k - LENGTH('a)\<close>
 | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
3987  | 
by simp  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
3988  | 
show \<open>bit (slice1 n (word_reverse w :: 'b word) :: 'a word) m \<longleftrightarrow> bit (word_reverse (slice1 k w :: 'a word)) m\<close>  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3989  | 
unfolding bit_slice1_iff bit_word_reverse_iff  | 
| 
72088
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
3990  | 
using * **  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3991  | 
    by (cases \<open>n \<le> LENGTH('a)\<close>; cases \<open>k \<le> LENGTH('a)\<close>) auto
 | 
| 
72088
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
3992  | 
qed  | 
| 
72027
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3993  | 
|
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3994  | 
lemma rev_slice:  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3995  | 
  "n + k + LENGTH('a::len) = LENGTH('b::len) \<Longrightarrow>
 | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3996  | 
slice n (word_reverse (w::'b word)) = word_reverse (slice k w :: 'a word)"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3997  | 
unfolding slice_def word_size  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
3998  | 
by (simp add: rev_slice1)  | 
| 
72027
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
3999  | 
|
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
4000  | 
|
| 61799 | 4001  | 
subsubsection \<open>Revcast\<close>  | 
| 37660 | 4002  | 
|
| 
72027
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
4003  | 
definition revcast :: \<open>'a::len word \<Rightarrow> 'b::len word\<close>  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
4004  | 
  where \<open>revcast = slice1 LENGTH('b)\<close>
 | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
4005  | 
|
| 
72611
 
c7bc3e70a8c7
official collection for bit projection simplifications
 
haftmann 
parents: 
72515 
diff
changeset
 | 
4006  | 
lemma bit_revcast_iff [bit_simps]:  | 
| 
72027
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
4007  | 
  \<open>bit (revcast w :: 'b::len word) n \<longleftrightarrow> LENGTH('b) - LENGTH('a) \<le> n \<and> n < LENGTH('b)
 | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
4008  | 
    \<and> bit w (n + (LENGTH('a) - LENGTH('b)) - (LENGTH('b) - LENGTH('a)))\<close>
 | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
4009  | 
for w :: \<open>'a::len word\<close>  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
4010  | 
by (simp add: revcast_def bit_slice1_iff)  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
4011  | 
|
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
4012  | 
lemma revcast_slice1 [OF refl]: "rc = revcast w \<Longrightarrow> slice1 (size rc) w = rc"  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
4013  | 
by (simp add: revcast_def word_size)  | 
| 
 
759532ef0885
prefer canonically oriented lists of bits and more direct characterizations in definitions
 
haftmann 
parents: 
72010 
diff
changeset
 | 
4014  | 
|
| 65268 | 4015  | 
lemma revcast_rev_ucast [OF refl refl refl]:  | 
4016  | 
"cs = [rc, uc] \<Longrightarrow> rc = revcast (word_reverse w) \<Longrightarrow> uc = ucast w \<Longrightarrow>  | 
|
| 37660 | 4017  | 
rc = word_reverse uc"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4018  | 
by (metis rev_slice1 revcast_slice1 ucast_slice1 word_size)  | 
| 37660 | 4019  | 
|
| 45811 | 4020  | 
lemma revcast_ucast: "revcast w = word_reverse (ucast (word_reverse w))"  | 
4021  | 
using revcast_rev_ucast [of "word_reverse w"] by simp  | 
|
4022  | 
||
4023  | 
lemma ucast_revcast: "ucast w = word_reverse (revcast (word_reverse w))"  | 
|
4024  | 
by (fact revcast_rev_ucast [THEN word_rev_gal'])  | 
|
4025  | 
||
4026  | 
lemma ucast_rev_revcast: "ucast (word_reverse w) = word_reverse (revcast w)"  | 
|
4027  | 
by (fact revcast_ucast [THEN word_rev_gal'])  | 
|
| 37660 | 4028  | 
|
4029  | 
||
| 65328 | 4030  | 
text "linking revcast and cast via shift"  | 
| 37660 | 4031  | 
|
4032  | 
lemmas wsst_TYs = source_size target_size word_size  | 
|
4033  | 
||
| 65268 | 4034  | 
lemmas sym_notr =  | 
| 37660 | 4035  | 
not_iff [THEN iffD2, THEN not_sym, THEN not_iff [THEN iffD1]]  | 
4036  | 
||
4037  | 
||
| 61799 | 4038  | 
subsection \<open>Split and cat\<close>  | 
| 37660 | 4039  | 
|
| 
40827
 
abbc05c20e24
code preprocessor setup for numerals on word type;
 
haftmann 
parents: 
39910 
diff
changeset
 | 
4040  | 
lemmas word_split_bin' = word_split_def  | 
| 
72088
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4041  | 
lemmas word_cat_bin' = word_cat_eq  | 
| 37660 | 4042  | 
|
| 65268 | 4043  | 
\<comment> \<open>this odd result is analogous to \<open>ucast_id\<close>,  | 
| 61799 | 4044  | 
result to the length given by the result type\<close>  | 
| 37660 | 4045  | 
|
4046  | 
lemma word_cat_id: "word_cat a b = b"  | 
|
| 72488 | 4047  | 
by transfer (simp add: take_bit_concat_bit_eq)  | 
| 65336 | 4048  | 
|
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4049  | 
lemma word_cat_split_alt: "\<lbrakk>size w \<le> size u + size v; word_split w = (u,v)\<rbrakk> \<Longrightarrow> word_cat u v = w"  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4050  | 
unfolding word_split_def  | 
| 80777 | 4051  | 
by (rule bit_word_eqI) (auto simp: bit_word_cat_iff not_less word_size bit_ucast_iff bit_drop_bit_eq)  | 
| 37660 | 4052  | 
|
| 45604 | 4053  | 
lemmas word_cat_split_size = sym [THEN [2] word_cat_split_alt [symmetric]]  | 
| 37660 | 4054  | 
|
4055  | 
||
| 61799 | 4056  | 
subsubsection \<open>Split and slice\<close>  | 
| 37660 | 4057  | 
|
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4058  | 
lemma split_slices:  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4059  | 
assumes "word_split w = (u, v)"  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4060  | 
shows "u = slice (size v) w \<and> v = slice 0 w"  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4061  | 
unfolding word_size  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4062  | 
proof (intro conjI)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4063  | 
  have \<section>: "\<And>n. \<lbrakk>ucast (drop_bit LENGTH('b) w) = u; LENGTH('c) < LENGTH('b)\<rbrakk> \<Longrightarrow> \<not> bit u n"
 | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4064  | 
by (metis bit_take_bit_iff bit_word_of_int_iff diff_is_0_eq' drop_bit_take_bit less_imp_le less_nat_zero_code of_int_uint unsigned_drop_bit_eq)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4065  | 
  show "u = slice LENGTH('b) w"
 | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4066  | 
proof (rule bit_word_eqI)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4067  | 
    show "bit u n = bit ((slice LENGTH('b) w)::'a word) n" if "n < LENGTH('a)" for n
 | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4068  | 
using assms bit_imp_le_length  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4069  | 
unfolding word_split_def bit_slice_iff  | 
| 80777 | 4070  | 
by (fastforce simp: \<section> ac_simps word_size bit_ucast_iff bit_drop_bit_eq)  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4071  | 
qed  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4072  | 
show "v = slice 0 w"  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4073  | 
by (metis Pair_inject assms ucast_slice word_split_bin')  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4074  | 
qed  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4075  | 
|
| 37660 | 4076  | 
|
| 
45816
 
6a04efd99f25
replace more uses of 'lemmas' with explicit 'lemma';
 
huffman 
parents: 
45811 
diff
changeset
 | 
4077  | 
lemma slice_cat1 [OF refl]:  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4078  | 
"\<lbrakk>wc = word_cat a b; size a + size b \<le> size wc\<rbrakk> \<Longrightarrow> slice (size b) wc = a"  | 
| 80777 | 4079  | 
by (rule bit_word_eqI) (auto simp: bit_slice_iff bit_word_cat_iff word_size)  | 
| 37660 | 4080  | 
|
4081  | 
lemmas slice_cat2 = trans [OF slice_id word_cat_id]  | 
|
4082  | 
||
4083  | 
lemma cat_slices:  | 
|
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4084  | 
"\<lbrakk>a = slice n c; b = slice 0 c; n = size b; size c \<le> size a + size b\<rbrakk> \<Longrightarrow> word_cat a b = c"  | 
| 80777 | 4085  | 
by (rule bit_word_eqI) (auto simp: bit_slice_iff bit_word_cat_iff word_size)  | 
| 37660 | 4086  | 
|
4087  | 
lemma word_split_cat_alt:  | 
|
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4088  | 
assumes "w = word_cat u v" and size: "size u + size v \<le> size w"  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4089  | 
shows "word_split w = (u,v)"  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4090  | 
proof -  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4091  | 
  have "ucast ((drop_bit LENGTH('c) (word_cat u v))::'a word) = u" "ucast ((word_cat u v)::'a word) = v"
 | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4092  | 
using assms  | 
| 80777 | 4093  | 
by (auto simp: word_size bit_ucast_iff bit_drop_bit_eq bit_word_cat_iff intro: bit_eqI)  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4094  | 
then show ?thesis  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4095  | 
by (simp add: assms(1) word_split_bin')  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4096  | 
qed  | 
| 37660 | 4097  | 
|
| 
72088
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4098  | 
lemma horner_sum_uint_exp_Cons_eq:  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4099  | 
  \<open>horner_sum uint (2 ^ LENGTH('a)) (w # ws) =
 | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4100  | 
    concat_bit LENGTH('a) (uint w) (horner_sum uint (2 ^ LENGTH('a)) ws)\<close>
 | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4101  | 
for ws :: \<open>'a::len word list\<close>  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4102  | 
by (simp add: bintr_uint concat_bit_eq push_bit_eq_mult)  | 
| 
72088
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4103  | 
|
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4104  | 
lemma bit_horner_sum_uint_exp_iff:  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4105  | 
  \<open>bit (horner_sum uint (2 ^ LENGTH('a)) ws) n \<longleftrightarrow>
 | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4106  | 
    n div LENGTH('a) < length ws \<and> bit (ws ! (n div LENGTH('a))) (n mod LENGTH('a))\<close>
 | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4107  | 
for ws :: \<open>'a::len word list\<close>  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4108  | 
proof (induction ws arbitrary: n)  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4109  | 
case Nil  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4110  | 
then show ?case  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4111  | 
by simp  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4112  | 
next  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4113  | 
case (Cons w ws)  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4114  | 
then show ?case  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4115  | 
    by (cases \<open>n \<ge> LENGTH('a)\<close>)
 | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4116  | 
(simp_all only: horner_sum_uint_exp_Cons_eq, simp_all add: bit_concat_bit_iff le_div_geq le_mod_geq bit_uint_iff Cons)  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4117  | 
qed  | 
| 37660 | 4118  | 
|
4119  | 
||
| 61799 | 4120  | 
subsection \<open>Rotation\<close>  | 
| 37660 | 4121  | 
|
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4122  | 
lemma word_rotr_word_rotr_eq: \<open>word_rotr m (word_rotr n w) = word_rotr (m + n) w\<close>  | 
| 
72088
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4123  | 
by (rule bit_word_eqI) (simp add: bit_word_rotr_iff ac_simps mod_add_right_eq)  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4124  | 
|
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4125  | 
lemma word_rot_lem: "\<lbrakk>l + k = d + k mod l; n < l\<rbrakk> \<Longrightarrow> ((d + n) mod l) = n" for l::nat  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4126  | 
by (metis (no_types, lifting) add.commute add.right_neutral add_diff_cancel_left' mod_if mod_mult_div_eq mod_mult_self2 mod_self)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4127  | 
|
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4128  | 
lemma word_rot_rl [simp]: \<open>word_rotl k (word_rotr k v) = v\<close>  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4129  | 
proof (rule bit_word_eqI)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4130  | 
  show "bit (word_rotl k (word_rotr k v)) n = bit v n" if "n < LENGTH('a)" for n
 | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4131  | 
using that  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4132  | 
by (auto simp: word_rot_lem word_rotl_eq_word_rotr word_rotr_word_rotr_eq bit_word_rotr_iff algebra_simps split: nat_diff_split)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4133  | 
qed  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4134  | 
|
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4135  | 
lemma word_rot_lr [simp]: \<open>word_rotr k (word_rotl k v) = v\<close>  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4136  | 
proof (rule bit_word_eqI)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4137  | 
  show "bit (word_rotr k (word_rotl k v)) n = bit v n" if "n < LENGTH('a)" for n
 | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4138  | 
using that  | 
| 80777 | 4139  | 
by (auto simp: word_rot_lem word_rotl_eq_word_rotr word_rotr_word_rotr_eq bit_word_rotr_iff algebra_simps split: nat_diff_split)  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4140  | 
qed  | 
| 
72088
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4141  | 
|
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4142  | 
lemma word_rot_gal:  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4143  | 
\<open>word_rotr n v = w \<longleftrightarrow> word_rotl n w = v\<close>  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4144  | 
by auto  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4145  | 
|
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4146  | 
lemma word_rot_gal':  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4147  | 
\<open>w = word_rotr n v \<longleftrightarrow> v = word_rotl n w\<close>  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4148  | 
by auto  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4149  | 
|
| 
80401
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
4150  | 
lemma word_reverse_word_rotl:  | 
| 
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
4151  | 
\<open>word_reverse (word_rotl n w) = word_rotr n (word_reverse w)\<close> (is \<open>?lhs = ?rhs\<close>)  | 
| 
72088
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4152  | 
proof (rule bit_word_eqI)  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4153  | 
fix m  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4154  | 
  assume \<open>m < LENGTH('a)\<close>
 | 
| 
80401
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
4155  | 
  then have \<open>int (LENGTH('a) - Suc ((m + n) mod LENGTH('a))) =
 | 
| 
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
4156  | 
    int ((LENGTH('a) + LENGTH('a) - Suc (m + n mod LENGTH('a))) mod LENGTH('a))\<close>
 | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
4157  | 
unfolding of_nat_diff of_nat_mod  | 
| 
80401
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
4158  | 
apply (simp add: Suc_le_eq add_less_le_mono of_nat_mod algebra_simps)  | 
| 
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
4159  | 
    apply (simp only: mod_diff_left_eq [symmetric, of \<open>int LENGTH('a) * 2\<close>] mod_mult_self1_is_0 diff_0 minus_mod_int_eq)
 | 
| 
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
4160  | 
apply (simp add: mod_simps)  | 
| 
72088
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4161  | 
done  | 
| 
80401
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
4162  | 
  then have \<open>LENGTH('a) - Suc ((m + n) mod LENGTH('a)) =
 | 
| 80777 | 4163  | 
            (LENGTH('a) + LENGTH('a) - Suc (m + n mod LENGTH('a))) mod LENGTH('a)\<close>
 | 
| 
72088
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4164  | 
by simp  | 
| 
80401
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
4165  | 
  with \<open>m < LENGTH('a)\<close> show \<open>bit ?lhs m \<longleftrightarrow> bit ?rhs m\<close>
 | 
| 
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
4166  | 
by (simp add: bit_simps)  | 
| 
72088
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4167  | 
qed  | 
| 65268 | 4168  | 
|
| 
80401
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
4169  | 
lemma word_reverse_word_rotr:  | 
| 
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
4170  | 
\<open>word_reverse (word_rotr n w) = word_rotl n (word_reverse w)\<close>  | 
| 
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
4171  | 
by (rule word_eq_reverseI) (simp add: word_reverse_word_rotl)  | 
| 
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
4172  | 
|
| 
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
4173  | 
lemma word_rotl_rev:  | 
| 
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
4174  | 
\<open>word_rotl n w = word_reverse (word_rotr n (word_reverse w))\<close>  | 
| 
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
4175  | 
by (simp add: word_reverse_word_rotr)  | 
| 
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
4176  | 
|
| 
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
4177  | 
lemma word_rotr_rev:  | 
| 
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
4178  | 
\<open>word_rotr n w = word_reverse (word_rotl n (word_reverse w))\<close>  | 
| 
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
4179  | 
by (simp add: word_reverse_word_rotl)  | 
| 
 
31bf95336f16
dropped references to theorems from transitional theory Divides.thy
 
haftmann 
parents: 
79950 
diff
changeset
 | 
4180  | 
|
| 37660 | 4181  | 
lemma word_roti_0 [simp]: "word_roti 0 w = w"  | 
| 72079 | 4182  | 
by transfer simp  | 
| 37660 | 4183  | 
|
| 65336 | 4184  | 
lemma word_roti_add: "word_roti (m + n) w = word_roti m (word_roti n w)"  | 
| 
72088
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4185  | 
by (rule bit_word_eqI)  | 
| 
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4186  | 
(simp add: bit_word_roti_iff nat_less_iff mod_simps ac_simps)  | 
| 65268 | 4187  | 
|
| 67118 | 4188  | 
lemma word_roti_conv_mod':  | 
4189  | 
"word_roti n w = word_roti (n mod int (size w)) w"  | 
|
| 72079 | 4190  | 
by transfer simp  | 
| 37660 | 4191  | 
|
4192  | 
lemmas word_roti_conv_mod = word_roti_conv_mod' [unfolded word_size]  | 
|
4193  | 
||
| 74097 | 4194  | 
end  | 
4195  | 
||
| 37660 | 4196  | 
|
| 61799 | 4197  | 
subsubsection \<open>"Word rotation commutes with bit-wise operations\<close>  | 
| 37660 | 4198  | 
|
| 67408 | 4199  | 
\<comment> \<open>using locale to not pollute lemma namespace\<close>  | 
| 65268 | 4200  | 
locale word_rotate  | 
| 37660 | 4201  | 
begin  | 
4202  | 
||
| 74097 | 4203  | 
context  | 
4204  | 
includes bit_operations_syntax  | 
|
4205  | 
begin  | 
|
4206  | 
||
| 37660 | 4207  | 
lemma word_rot_logs:  | 
| 71149 | 4208  | 
"word_rotl n (NOT v) = NOT (word_rotl n v)"  | 
4209  | 
"word_rotr n (NOT v) = NOT (word_rotr n v)"  | 
|
| 37660 | 4210  | 
"word_rotl n (x AND y) = word_rotl n x AND word_rotl n y"  | 
4211  | 
"word_rotr n (x AND y) = word_rotr n x AND word_rotr n y"  | 
|
4212  | 
"word_rotl n (x OR y) = word_rotl n x OR word_rotl n y"  | 
|
4213  | 
"word_rotr n (x OR y) = word_rotr n x OR word_rotr n y"  | 
|
4214  | 
"word_rotl n (x XOR y) = word_rotl n x XOR word_rotl n y"  | 
|
| 65268 | 4215  | 
"word_rotr n (x XOR y) = word_rotr n x XOR word_rotr n y"  | 
| 80777 | 4216  | 
by (rule bit_word_eqI, auto simp: bit_word_rotl_iff bit_word_rotr_iff bit_and_iff bit_or_iff bit_xor_iff bit_not_iff algebra_simps not_le)+  | 
| 
72088
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4217  | 
|
| 37660 | 4218  | 
end  | 
4219  | 
||
| 74097 | 4220  | 
end  | 
4221  | 
||
| 37660 | 4222  | 
lemmas word_rot_logs = word_rotate.word_rot_logs  | 
4223  | 
||
| 65336 | 4224  | 
lemma word_rotx_0 [simp] : "word_rotr i 0 = 0 \<and> word_rotl i 0 = 0"  | 
| 
72088
 
a36db1c8238e
separation of reversed bit lists from other material
 
haftmann 
parents: 
72083 
diff
changeset
 | 
4225  | 
by transfer simp_all  | 
| 37660 | 4226  | 
|
4227  | 
lemma word_roti_0' [simp] : "word_roti n 0 = 0"  | 
|
| 72079 | 4228  | 
by transfer simp  | 
| 37660 | 4229  | 
|
| 72079 | 4230  | 
declare word_roti_eq_word_rotr_word_rotl [simp]  | 
| 37660 | 4231  | 
|
4232  | 
||
| 61799 | 4233  | 
subsection \<open>Maximum machine word\<close>  | 
| 37660 | 4234  | 
|
| 74097 | 4235  | 
context  | 
4236  | 
includes bit_operations_syntax  | 
|
4237  | 
begin  | 
|
4238  | 
||
| 37660 | 4239  | 
lemma word_int_cases:  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
4240  | 
fixes x :: "'a::len word"  | 
| 70185 | 4241  | 
  obtains n where "x = word_of_int n" and "0 \<le> n" and "n < 2^LENGTH('a)"
 | 
| 72292 | 4242  | 
by (rule that [of \<open>uint x\<close>]) simp_all  | 
| 37660 | 4243  | 
|
4244  | 
lemma word_nat_cases [cases type: word]:  | 
|
| 65336 | 4245  | 
fixes x :: "'a::len word"  | 
| 70185 | 4246  | 
  obtains n where "x = of_nat n" and "n < 2^LENGTH('a)"
 | 
| 72292 | 4247  | 
by (rule that [of \<open>unat x\<close>]) simp_all  | 
| 37660 | 4248  | 
|
| 73788 | 4249  | 
lemma max_word_max [intro!]:  | 
4250  | 
\<open>n \<le> - 1\<close> for n :: \<open>'a::len word\<close>  | 
|
| 
71957
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
4251  | 
by (fact word_order.extremum)  | 
| 65268 | 4252  | 
|
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
4253  | 
lemma word_of_int_2p_len: "word_of_int (2 ^ LENGTH('a)) = (0::'a::len word)"
 | 
| 72292 | 4254  | 
by simp  | 
| 37660 | 4255  | 
|
| 70185 | 4256  | 
lemma word_pow_0: "(2::'a::len word) ^ LENGTH('a) = 0"
 | 
| 
71957
 
3e162c63371a
build bit operations on word on library theory on bit operations
 
haftmann 
parents: 
71955 
diff
changeset
 | 
4257  | 
by (fact word_exp_length_eq_0)  | 
| 37660 | 4258  | 
|
| 73788 | 4259  | 
lemma max_word_wrap:  | 
4260  | 
\<open>x + 1 = 0 \<Longrightarrow> x = - 1\<close> for x :: \<open>'a::len word\<close>  | 
|
| 71946 | 4261  | 
by (simp add: eq_neg_iff_add_eq_0)  | 
4262  | 
||
| 73788 | 4263  | 
lemma word_and_max:  | 
4264  | 
\<open>x AND - 1 = x\<close> for x :: \<open>'a::len word\<close>  | 
|
| 71946 | 4265  | 
by (fact word_log_esimps)  | 
4266  | 
||
| 73788 | 4267  | 
lemma word_or_max:  | 
4268  | 
\<open>x OR - 1 = - 1\<close> for x :: \<open>'a::len word\<close>  | 
|
| 71946 | 4269  | 
by (fact word_log_esimps)  | 
| 37660 | 4270  | 
|
| 65336 | 4271  | 
lemma word_ao_dist2: "x AND (y OR z) = x AND y OR x AND z"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
4272  | 
for x y z :: "'a::len word"  | 
| 72508 | 4273  | 
by (fact bit.conj_disj_distrib)  | 
| 37660 | 4274  | 
|
| 65336 | 4275  | 
lemma word_oa_dist2: "x OR y AND z = (x OR y) AND (x OR z)"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
4276  | 
for x y z :: "'a::len word"  | 
| 72508 | 4277  | 
by (fact bit.disj_conj_distrib)  | 
| 37660 | 4278  | 
|
| 65336 | 4279  | 
lemma word_and_not [simp]: "x AND NOT x = 0"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
4280  | 
for x :: "'a::len word"  | 
| 72508 | 4281  | 
by (fact bit.conj_cancel_right)  | 
| 37660 | 4282  | 
|
| 73788 | 4283  | 
lemma word_or_not [simp]:  | 
4284  | 
\<open>x OR NOT x = - 1\<close> for x :: \<open>'a::len word\<close>  | 
|
| 72508 | 4285  | 
by (fact bit.disj_cancel_right)  | 
| 37660 | 4286  | 
|
| 65336 | 4287  | 
lemma word_xor_and_or: "x XOR y = x AND NOT y OR NOT x AND y"  | 
| 
71954
 
13bb3f5cdc5b
pragmatically ruled out word types of length zero: a bit string with no bits is not bit string at all
 
haftmann 
parents: 
71953 
diff
changeset
 | 
4288  | 
for x y :: "'a::len word"  | 
| 72508 | 4289  | 
by (fact bit.xor_def)  | 
| 37660 | 4290  | 
|
| 65336 | 4291  | 
lemma uint_lt_0 [simp]: "uint x < 0 = False"  | 
| 37660 | 4292  | 
by (simp add: linorder_not_less)  | 
4293  | 
||
| 65336 | 4294  | 
lemma word_less_1 [simp]: "x < 1 \<longleftrightarrow> x = 0"  | 
4295  | 
for x :: "'a::len word"  | 
|
| 37660 | 4296  | 
by (simp add: word_less_nat_alt unat_0_iff)  | 
4297  | 
||
4298  | 
lemma uint_plus_if_size:  | 
|
| 65268 | 4299  | 
"uint (x + y) =  | 
| 65336 | 4300  | 
(if uint x + uint y < 2^size x  | 
4301  | 
then uint x + uint y  | 
|
4302  | 
else uint x + uint y - 2^size x)"  | 
|
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4303  | 
by (simp add: take_bit_eq_mod word_size uint_word_of_int_eq uint_plus_if')  | 
| 37660 | 4304  | 
|
4305  | 
lemma unat_plus_if_size:  | 
|
| 65363 | 4306  | 
"unat (x + y) =  | 
| 65336 | 4307  | 
(if unat x + unat y < 2^size x  | 
4308  | 
then unat x + unat y  | 
|
4309  | 
else unat x + unat y - 2^size x)"  | 
|
| 65363 | 4310  | 
for x y :: "'a::len word"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4311  | 
by (simp add: size_word.rep_eq unat_arith_simps)  | 
| 37660 | 4312  | 
|
| 65336 | 4313  | 
lemma word_neq_0_conv: "w \<noteq> 0 \<longleftrightarrow> 0 < w"  | 
4314  | 
for w :: "'a::len word"  | 
|
| 72262 | 4315  | 
by (fact word_coorder.not_eq_extremum)  | 
| 65336 | 4316  | 
|
4317  | 
lemma max_lt: "unat (max a b div c) = unat (max a b) div unat c"  | 
|
4318  | 
for c :: "'a::len word"  | 
|
| 55818 | 4319  | 
by (fact unat_div)  | 
| 37660 | 4320  | 
|
4321  | 
lemma uint_sub_if_size:  | 
|
| 65268 | 4322  | 
"uint (x - y) =  | 
| 65336 | 4323  | 
(if uint y \<le> uint x  | 
4324  | 
then uint x - uint y  | 
|
4325  | 
else uint x - uint y + 2^size x)"  | 
|
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4326  | 
by (simp add: size_word.rep_eq uint_sub_if')  | 
| 65336 | 4327  | 
|
| 
72130
 
9e5862223442
dedicated symbols for code generation, to pave way for generic conversions from and to word
 
haftmann 
parents: 
72128 
diff
changeset
 | 
4328  | 
lemma unat_sub:  | 
| 
 
9e5862223442
dedicated symbols for code generation, to pave way for generic conversions from and to word
 
haftmann 
parents: 
72128 
diff
changeset
 | 
4329  | 
\<open>unat (a - b) = unat a - unat b\<close>  | 
| 
 
9e5862223442
dedicated symbols for code generation, to pave way for generic conversions from and to word
 
haftmann 
parents: 
72128 
diff
changeset
 | 
4330  | 
if \<open>b \<le> a\<close>  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4331  | 
by (meson that unat_sub_if_size word_le_nat_alt)  | 
| 37660 | 4332  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
4333  | 
lemmas word_less_sub1_numberof [simp] = word_less_sub1 [of "numeral w"] for w  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46962 
diff
changeset
 | 
4334  | 
lemmas word_le_sub1_numberof [simp] = word_le_sub1 [of "numeral w"] for w  | 
| 65268 | 4335  | 
|
| 70185 | 4336  | 
lemma word_of_int_minus: "word_of_int (2^LENGTH('a) - i) = (word_of_int (-i)::'a::len word)"
 | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4337  | 
by simp  | 
| 72292 | 4338  | 
|
4339  | 
lemma word_of_int_inj:  | 
|
4340  | 
\<open>(word_of_int x :: 'a::len word) = word_of_int y \<longleftrightarrow> x = y\<close>  | 
|
4341  | 
  if \<open>0 \<le> x \<and> x < 2 ^ LENGTH('a)\<close> \<open>0 \<le> y \<and> y < 2 ^ LENGTH('a)\<close>
 | 
|
4342  | 
using that by (transfer fixing: x y) (simp add: take_bit_int_eq_self)  | 
|
| 37660 | 4343  | 
|
| 65336 | 4344  | 
lemma word_le_less_eq: "x \<le> y \<longleftrightarrow> x = y \<or> x < y"  | 
4345  | 
for x y :: "'z::len word"  | 
|
| 80777 | 4346  | 
by (auto simp: order_class.le_less)  | 
| 37660 | 4347  | 
|
4348  | 
lemma mod_plus_cong:  | 
|
| 65336 | 4349  | 
fixes b b' :: int  | 
4350  | 
assumes 1: "b = b'"  | 
|
4351  | 
and 2: "x mod b' = x' mod b'"  | 
|
4352  | 
and 3: "y mod b' = y' mod b'"  | 
|
4353  | 
and 4: "x' + y' = z'"  | 
|
| 37660 | 4354  | 
shows "(x + y) mod b = z' mod b'"  | 
4355  | 
proof -  | 
|
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
4356  | 
from 1 2[symmetric] 3[symmetric]  | 
| 
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
4357  | 
have "(x + y) mod b = (x' mod b' + y' mod b') mod b'"  | 
| 
64593
 
50c715579715
reoriented congruence rules in non-explosive direction
 
haftmann 
parents: 
64243 
diff
changeset
 | 
4358  | 
by (simp add: mod_add_eq)  | 
| 37660 | 4359  | 
also have "\<dots> = (x' + y') mod b'"  | 
| 
64593
 
50c715579715
reoriented congruence rules in non-explosive direction
 
haftmann 
parents: 
64243 
diff
changeset
 | 
4360  | 
by (simp add: mod_add_eq)  | 
| 65336 | 4361  | 
finally show ?thesis  | 
4362  | 
by (simp add: 4)  | 
|
| 37660 | 4363  | 
qed  | 
4364  | 
||
4365  | 
lemma mod_minus_cong:  | 
|
| 65336 | 4366  | 
fixes b b' :: int  | 
4367  | 
assumes "b = b'"  | 
|
4368  | 
and "x mod b' = x' mod b'"  | 
|
4369  | 
and "y mod b' = y' mod b'"  | 
|
4370  | 
and "x' - y' = z'"  | 
|
| 37660 | 4371  | 
shows "(x - y) mod b = z' mod b'"  | 
| 65336 | 4372  | 
using assms [symmetric] by (auto intro: mod_diff_cong)  | 
4373  | 
||
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4374  | 
lemma word_induct_less [case_names zero less]:  | 
| 72262 | 4375  | 
\<open>P m\<close> if zero: \<open>P 0\<close> and less: \<open>\<And>n. n < m \<Longrightarrow> P n \<Longrightarrow> P (1 + n)\<close>  | 
4376  | 
for m :: \<open>'a::len word\<close>  | 
|
4377  | 
proof -  | 
|
4378  | 
define q where \<open>q = unat m\<close>  | 
|
4379  | 
with less have \<open>\<And>n. n < word_of_nat q \<Longrightarrow> P n \<Longrightarrow> P (1 + n)\<close>  | 
|
4380  | 
by simp  | 
|
4381  | 
then have \<open>P (word_of_nat q :: 'a word)\<close>  | 
|
4382  | 
proof (induction q)  | 
|
4383  | 
case 0  | 
|
4384  | 
show ?case  | 
|
4385  | 
by (simp add: zero)  | 
|
4386  | 
next  | 
|
4387  | 
case (Suc q)  | 
|
4388  | 
show ?case  | 
|
4389  | 
proof (cases \<open>1 + word_of_nat q = (0 :: 'a word)\<close>)  | 
|
4390  | 
case True  | 
|
4391  | 
then show ?thesis  | 
|
4392  | 
by (simp add: zero)  | 
|
4393  | 
next  | 
|
4394  | 
case False  | 
|
4395  | 
then have *: \<open>word_of_nat q < (word_of_nat (Suc q) :: 'a word)\<close>  | 
|
4396  | 
by (simp add: unatSuc word_less_nat_alt)  | 
|
4397  | 
then have **: \<open>n < (1 + word_of_nat q :: 'a word) \<longleftrightarrow> n \<le> (word_of_nat q :: 'a word)\<close> for n  | 
|
4398  | 
by (metis (no_types, lifting) add.commute inc_le le_less_trans not_less of_nat_Suc)  | 
|
4399  | 
have \<open>P (word_of_nat q)\<close>  | 
|
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4400  | 
by (simp add: "**" Suc.IH Suc.prems)  | 
| 72262 | 4401  | 
with * have \<open>P (1 + word_of_nat q)\<close>  | 
4402  | 
by (rule Suc.prems)  | 
|
4403  | 
then show ?thesis  | 
|
4404  | 
by simp  | 
|
4405  | 
qed  | 
|
4406  | 
qed  | 
|
4407  | 
with \<open>q = unat m\<close> show ?thesis  | 
|
4408  | 
by simp  | 
|
4409  | 
qed  | 
|
| 65268 | 4410  | 
|
| 65363 | 4411  | 
lemma word_induct: "P 0 \<Longrightarrow> (\<And>n. P n \<Longrightarrow> P (1 + n)) \<Longrightarrow> P m"  | 
| 65336 | 4412  | 
for P :: "'a::len word \<Rightarrow> bool"  | 
| 72262 | 4413  | 
by (rule word_induct_less)  | 
| 65336 | 4414  | 
|
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4415  | 
lemma word_induct2 [case_names zero suc, induct type]: "P 0 \<Longrightarrow> (\<And>n. 1 + n \<noteq> 0 \<Longrightarrow> P n \<Longrightarrow> P (1 + n)) \<Longrightarrow> P n"  | 
| 65336 | 4416  | 
for P :: "'b::len word \<Rightarrow> bool"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4417  | 
by (induction rule: word_induct_less; force)  | 
| 37660 | 4418  | 
|
| 
55816
 
e8dd03241e86
cursory polishing: tuned proofs, tuned symbols, tuned headings
 
haftmann 
parents: 
55415 
diff
changeset
 | 
4419  | 
|
| 61799 | 4420  | 
subsection \<open>Recursion combinator for words\<close>  | 
| 46010 | 4421  | 
|
| 54848 | 4422  | 
definition word_rec :: "'a \<Rightarrow> ('b::len word \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b word \<Rightarrow> 'a"
 | 
| 65336 | 4423  | 
where "word_rec forZero forSuc n = rec_nat forZero (forSuc \<circ> of_nat) (unat n)"  | 
| 37660 | 4424  | 
|
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4425  | 
lemma word_rec_0 [simp]: "word_rec z s 0 = z"  | 
| 37660 | 4426  | 
by (simp add: word_rec_def)  | 
4427  | 
||
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4428  | 
lemma word_rec_Suc [simp]: "1 + n \<noteq> 0 \<Longrightarrow> word_rec z s (1 + n) = s n (word_rec z s n)"  | 
| 65363 | 4429  | 
for n :: "'a::len word"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4430  | 
by (simp add: unatSuc word_rec_def)  | 
| 37660 | 4431  | 
|
| 65363 | 4432  | 
lemma word_rec_Pred: "n \<noteq> 0 \<Longrightarrow> word_rec z s n = s (n - 1) (word_rec z s (n - 1))"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4433  | 
by (metis add.commute diff_add_cancel word_rec_Suc)  | 
| 37660 | 4434  | 
|
| 65336 | 4435  | 
lemma word_rec_in: "f (word_rec z (\<lambda>_. f) n) = word_rec (f z) (\<lambda>_. f) n"  | 
| 74101 | 4436  | 
by (induct n) simp_all  | 
| 37660 | 4437  | 
|
| 67399 | 4438  | 
lemma word_rec_in2: "f n (word_rec z f n) = word_rec (f 0 z) (f \<circ> (+) 1) n"  | 
| 74101 | 4439  | 
by (induct n) simp_all  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4440  | 
|
| 65268 | 4441  | 
lemma word_rec_twice:  | 
| 67399 | 4442  | 
"m \<le> n \<Longrightarrow> word_rec z f n = word_rec (word_rec z f (n - m)) (f \<circ> (+) (n - m)) m"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4443  | 
proof (induction n arbitrary: z f)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4444  | 
case zero  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4445  | 
then show ?case  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4446  | 
by (metis diff_0_right word_le_0_iff word_rec_0)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4447  | 
next  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4448  | 
case (suc n z f)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4449  | 
show ?case  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4450  | 
proof (cases "1 + (n - m) = 0")  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4451  | 
case True  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4452  | 
then show ?thesis  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4453  | 
by (simp add: add_diff_eq)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4454  | 
next  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4455  | 
case False  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4456  | 
then have eq: "1 + n - m = 1 + (n - m)"  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4457  | 
by simp  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4458  | 
with False have "m \<le> n"  | 
| 
81763
 
2cf8f8e4c1fd
Simplified a lot of messy proofs
 
paulson <lp15@cam.ac.uk> 
parents: 
81609 
diff
changeset
 | 
4459  | 
using inc_le linorder_not_le suc.prems word_le_minus_mono_left by fastforce  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4460  | 
with False "suc.hyps" show ?thesis  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4461  | 
using suc.IH [of "f 0 z" "f \<circ> (+) 1"]  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4462  | 
by (simp add: word_rec_in2 eq add.assoc o_def)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4463  | 
qed  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4464  | 
qed  | 
| 37660 | 4465  | 
|
4466  | 
lemma word_rec_id: "word_rec z (\<lambda>_. id) n = z"  | 
|
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4467  | 
by (induct n) auto  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4468  | 
|
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4469  | 
lemma word_rec_id_eq: "(\<And>m. m < n \<Longrightarrow> f m = id) \<Longrightarrow> word_rec z f n = z"  | 
| 80777 | 4470  | 
by (induction n) (auto simp: unatSuc unat_arith_simps(2))  | 
| 37660 | 4471  | 
|
| 65268 | 4472  | 
lemma word_rec_max:  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4473  | 
assumes "\<forall>m\<ge>n. m \<noteq> - 1 \<longrightarrow> f m = id"  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4474  | 
shows "word_rec z f (- 1) = word_rec z f n"  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4475  | 
proof -  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4476  | 
have \<section>: "\<And>m. \<lbrakk>m < - 1 - n\<rbrakk> \<Longrightarrow> (f \<circ> (+) n) m = id"  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4477  | 
using assms  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4478  | 
by (metis (mono_tags, lifting) add.commute add_diff_cancel_left' comp_apply less_le olen_add_eqv plus_minus_no_overflow word_n1_ge)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4479  | 
have "word_rec z f (- 1) = word_rec (word_rec z f (- 1 - (- 1 - n))) (f \<circ> (+) (- 1 - (- 1 - n))) (- 1 - n)"  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4480  | 
by (meson word_n1_ge word_rec_twice)  | 
| 80777 | 4481  | 
also have "\<dots> = word_rec z f n"  | 
| 
72735
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4482  | 
by (metis (no_types, lifting) \<section> diff_add_cancel minus_diff_eq uminus_add_conv_diff word_rec_id_eq)  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4483  | 
finally show ?thesis .  | 
| 
 
bbe5d3ef2052
Stepan Holub's stronger version of comm_append_are_replicate, and a de-applied Word.thy
 
paulson <lp15@cam.ac.uk> 
parents: 
72611 
diff
changeset
 | 
4484  | 
qed  | 
| 65336 | 4485  | 
|
| 74097 | 4486  | 
end  | 
4487  | 
||
| 72512 | 4488  | 
|
| 82523 | 4489  | 
subsection \<open>Some more naive computations rules\<close>  | 
4490  | 
||
4491  | 
lemma drop_bit_of_minus_1_eq [simp]:  | 
|
4492  | 
  \<open>drop_bit n (- 1 :: 'a::len word) = mask (LENGTH('a) - n)\<close>
 | 
|
4493  | 
by (rule bit_word_eqI) (auto simp add: bit_simps)  | 
|
4494  | 
||
4495  | 
context  | 
|
4496  | 
includes bit_operations_syntax  | 
|
4497  | 
begin  | 
|
4498  | 
||
4499  | 
lemma word_cat_eq_push_bit_or:  | 
|
4500  | 
  \<open>word_cat v w = (push_bit LENGTH('b) (ucast v) OR ucast w :: 'c::len word)\<close>
 | 
|
4501  | 
for v :: \<open>'a::len word\<close> and w :: \<open>'b::len word\<close>  | 
|
4502  | 
by transfer (simp add: concat_bit_def ac_simps)  | 
|
4503  | 
||
4504  | 
end  | 
|
4505  | 
||
4506  | 
context semiring_bit_operations  | 
|
4507  | 
begin  | 
|
4508  | 
||
4509  | 
lemma of_nat_take_bit_numeral_eq [simp]:  | 
|
4510  | 
\<open>of_nat (take_bit m (numeral n)) = take_bit m (numeral n)\<close>  | 
|
4511  | 
by (simp add: of_nat_take_bit)  | 
|
4512  | 
||
4513  | 
end  | 
|
4514  | 
||
4515  | 
context ring_bit_operations  | 
|
4516  | 
begin  | 
|
4517  | 
||
4518  | 
lemma signed_take_bit_of_int:  | 
|
4519  | 
\<open>signed_take_bit n (of_int k) = of_int (signed_take_bit n k)\<close>  | 
|
4520  | 
by (rule bit_eqI) (simp add: bit_simps)  | 
|
4521  | 
||
4522  | 
lemma of_int_signed_take_bit:  | 
|
4523  | 
\<open>of_int (signed_take_bit n k) = signed_take_bit n (of_int k)\<close>  | 
|
4524  | 
by (simp add: signed_take_bit_of_int)  | 
|
4525  | 
||
4526  | 
lemma of_int_take_bit_minus_numeral_eq [simp]:  | 
|
4527  | 
\<open>of_int (take_bit m (numeral n)) = take_bit m (numeral n)\<close>  | 
|
4528  | 
\<open>of_int (take_bit m (- numeral n)) = take_bit m (- numeral n)\<close>  | 
|
4529  | 
by (simp_all add: of_int_take_bit)  | 
|
4530  | 
||
4531  | 
end  | 
|
4532  | 
||
4533  | 
context  | 
|
4534  | 
includes bit_operations_syntax  | 
|
4535  | 
begin  | 
|
4536  | 
||
4537  | 
lemma concat_bit_numeral_of_one_1 [simp]:  | 
|
4538  | 
\<open>concat_bit (numeral m) 1 l = 1 OR push_bit (numeral m) l\<close>  | 
|
4539  | 
by (rule bit_eqI) (auto simp add: bit_simps)  | 
|
4540  | 
||
4541  | 
lemma concat_bit_of_one_2 [simp]:  | 
|
4542  | 
\<open>concat_bit n k 1 = set_bit n (take_bit n k)\<close>  | 
|
4543  | 
by (rule bit_eqI) (auto simp add: bit_simps)  | 
|
4544  | 
||
4545  | 
lemma concat_bit_numeral_of_minus_one_1 [simp]:  | 
|
4546  | 
\<open>concat_bit (numeral m) (- 1) l = push_bit (numeral m) l OR mask (numeral m)\<close>  | 
|
4547  | 
by (rule bit_eqI) (auto simp add: bit_simps)  | 
|
4548  | 
||
4549  | 
lemma concat_bit_numeral_of_minus_one_2 [simp]:  | 
|
4550  | 
\<open>concat_bit (numeral m) k (- 1) = take_bit (numeral m) k OR NOT (mask (numeral m))\<close>  | 
|
4551  | 
by (rule bit_eqI) (auto simp add: bit_simps)  | 
|
4552  | 
||
4553  | 
lemma concat_bit_numeral [simp]:  | 
|
4554  | 
\<open>concat_bit (numeral m) (numeral n) (numeral q) = take_bit (numeral m) (numeral n) OR push_bit (numeral m) (numeral q)\<close>  | 
|
4555  | 
\<open>concat_bit (numeral m) (- numeral n) (numeral q) = take_bit (numeral m) (- numeral n) OR push_bit (numeral m) (numeral q)\<close>  | 
|
4556  | 
\<open>concat_bit (numeral m) (numeral n) (- numeral q) = take_bit (numeral m) (numeral n) OR push_bit (numeral m) (- numeral q)\<close>  | 
|
4557  | 
\<open>concat_bit (numeral m) (- numeral n) (- numeral q) = take_bit (numeral m) (- numeral n) OR push_bit (numeral m) (- numeral q)\<close>  | 
|
4558  | 
by (fact concat_bit_def)+  | 
|
4559  | 
||
4560  | 
end  | 
|
4561  | 
||
4562  | 
lemma word_cat_0_left [simp]:  | 
|
4563  | 
\<open>word_cat 0 w = ucast w\<close>  | 
|
4564  | 
by (simp add: word_cat_eq)  | 
|
4565  | 
||
4566  | 
||
| 83274 | 4567  | 
subsection \<open>Computation on ranges\<close>  | 
4568  | 
||
4569  | 
context  | 
|
4570  | 
begin  | 
|
4571  | 
||
4572  | 
qualified definition range :: \<open>'a::len word \<Rightarrow> 'a word \<Rightarrow> 'a word list\<close> \<comment> \<open>only for code generation\<close>  | 
|
4573  | 
  where range_eq: \<open>range a b = sorted_list_of_set {a..<b}\<close>
 | 
|
4574  | 
||
4575  | 
qualified lemma set_range_eq [simp]:  | 
|
4576  | 
  \<open>set (range a b) = {a..<b}\<close>
 | 
|
4577  | 
using finite_atLeastLessThan by (simp add: range_eq)  | 
|
4578  | 
||
4579  | 
qualified lemma distinct_range [simp]:  | 
|
4580  | 
\<open>distinct (range a b)\<close>  | 
|
4581  | 
by (simp add: range_eq)  | 
|
4582  | 
||
4583  | 
qualified lemma range_code [code]:  | 
|
4584  | 
\<open>range a b = (if a < b then a # range (a + 1) b else [])\<close>  | 
|
4585  | 
proof (cases \<open>a < b\<close>)  | 
|
4586  | 
case True  | 
|
4587  | 
  then have \<open>{a..<b} = insert a {a + 1..<b}\<close>
 | 
|
4588  | 
apply (auto simp add: not_le)  | 
|
4589  | 
apply (metis inc_le leD word_le_less_eq)  | 
|
4590  | 
apply (metis dual_order.strict_implies_order inc_less_eq_iff word_not_simps(3))  | 
|
4591  | 
done  | 
|
4592  | 
then show ?thesis  | 
|
4593  | 
apply (auto simp add: range_eq finite_atLeastLessThan intro!: insort_is_Cons)  | 
|
4594  | 
apply (subst insort_is_Cons)  | 
|
4595  | 
apply auto  | 
|
4596  | 
apply (metis Diff_insert Diff_insert0 atLeastAtMost_iff  | 
|
4597  | 
atLeastLessThan_eq_atLeastAtMost_diff inc_less_eq_triv_imp  | 
|
4598  | 
word_not_simps(3))  | 
|
4599  | 
done  | 
|
4600  | 
next  | 
|
4601  | 
case False  | 
|
4602  | 
then show ?thesis  | 
|
4603  | 
by (simp add: range_eq)  | 
|
4604  | 
qed  | 
|
4605  | 
||
4606  | 
qualified lemma atLeast_eq_range [code_unfold]:  | 
|
4607  | 
  \<open>{a..} = insert (- 1) (set (range a (- 1)))\<close>
 | 
|
4608  | 
by (auto simp add: not_less word_order.extremum_unique)  | 
|
4609  | 
||
4610  | 
qualified lemma greaterThan_eq_range [code_unfold]:  | 
|
4611  | 
  \<open>{a<..} = (if a = - 1 then {} else insert (- 1) (set (range (a + 1) (- 1))))\<close>
 | 
|
4612  | 
apply (auto simp add: not_less not_le word_order.extremum_unique inc_less_eq_iff)  | 
|
4613  | 
apply (simp add: order_neq_le_trans)  | 
|
4614  | 
done  | 
|
4615  | 
||
4616  | 
qualified lemma lessThan_eq_range [code_unfold]:  | 
|
4617  | 
  \<open>{..<b} = set (range 0 b)\<close>
 | 
|
4618  | 
by auto  | 
|
4619  | 
||
4620  | 
qualified lemma atMost_eq_range [code_unfold]:  | 
|
4621  | 
  \<open>{..b} = insert b (set (range 0 b))\<close> for b :: \<open>'a::len word\<close>
 | 
|
4622  | 
by auto  | 
|
4623  | 
||
4624  | 
qualified lemma atLeastLessThan_eq_range [code_unfold]:  | 
|
4625  | 
  \<open>{a..<b} = set (range a b)\<close>
 | 
|
4626  | 
by simp  | 
|
4627  | 
||
4628  | 
qualified lemma atLeastAtMost_eq_range [code_unfold]:  | 
|
4629  | 
  \<open>{a..b} = (if b = - 1 then {a..} else set (range a (b + 1)))\<close>
 | 
|
4630  | 
apply auto  | 
|
4631  | 
using inc_less_eq_iff linorder_not_le apply blast  | 
|
4632  | 
using inc_le linorder_not_less apply blast  | 
|
4633  | 
done  | 
|
4634  | 
||
4635  | 
qualified lemma greaterThanLessThan_eq_range [code_unfold]:  | 
|
4636  | 
  \<open>{a<..<b} = (if a = - 1 then {} else set (range (a + 1) b))\<close>
 | 
|
4637  | 
by (auto simp add: inc_less_eq_iff)  | 
|
4638  | 
||
4639  | 
qualified lemma greaterThanAtMost_eq_range [code_unfold]:  | 
|
4640  | 
  \<open>{a<..b} = (if b = - 1 then {a<..} else if a = - 1 then {} else set (range (a + 1) (b + 1)))\<close>
 | 
|
4641  | 
apply (auto simp add: inc_less_eq_iff)  | 
|
4642  | 
apply (metis add_diff_cancel add_eq_0_iff2 less_eq_dec_iff)  | 
|
4643  | 
using inc_le linorder_not_less apply blast  | 
|
4644  | 
done  | 
|
4645  | 
||
4646  | 
qualified definition all_range :: \<open>('a::len word \<Rightarrow> bool) \<Rightarrow> 'a word \<Rightarrow> 'a word \<Rightarrow> bool\<close> \<comment> \<open>only for code generation\<close>
 | 
|
4647  | 
  where all_range_iff [code_abbrev, simp]: \<open>all_range P a b \<longleftrightarrow> (\<forall>n\<in>{a..<b}. P n)\<close>
 | 
|
4648  | 
||
4649  | 
qualified lemma all_range_code [code]:  | 
|
4650  | 
\<open>all_range P a b \<longleftrightarrow> (a < b \<longrightarrow> P a \<and> all_range P (a + 1) b)\<close>  | 
|
4651  | 
apply (auto simp add: Ball_def simp flip: less_iff_succ_less_eq)  | 
|
4652  | 
apply (metis inc_less_eq_iff order_le_less word_not_simps(3))  | 
|
4653  | 
apply (metis antisym_conv2 inc_le)  | 
|
4654  | 
done  | 
|
4655  | 
||
4656  | 
qualified lemma forall_atLeast_iff [code_unfold]:  | 
|
4657  | 
  \<open>(\<forall>n\<in>{a..}. P n) \<longleftrightarrow> all_range P a (- 1) \<and> P (- 1)\<close>
 | 
|
4658  | 
apply auto  | 
|
4659  | 
using atLeastLessThan_iff word_order.not_eq_extremum apply blast  | 
|
4660  | 
done  | 
|
4661  | 
||
4662  | 
qualified lemma forall_greater_eq_iff [code_unfold]:  | 
|
4663  | 
  \<open>(\<forall>n\<ge>a. P n) \<longleftrightarrow> (\<forall>n\<in>{a..}. P n)\<close>
 | 
|
4664  | 
by auto  | 
|
4665  | 
||
4666  | 
qualified lemma forall_greaterThan_iff [code_unfold]:  | 
|
4667  | 
  \<open>(\<forall>n\<in>{a<..}. P n) \<longleftrightarrow> a = - 1 \<or> all_range P (a + 1) (- 1) \<and> P (- 1)\<close>
 | 
|
4668  | 
apply (auto simp add: inc_less_eq_iff word_order.not_eq_extremum)  | 
|
4669  | 
apply (metis atLeastLessThan_iff inc_le word_order.not_eq_extremum)  | 
|
4670  | 
done  | 
|
4671  | 
||
4672  | 
qualified lemma forall_greater_iff [code_unfold]:  | 
|
4673  | 
  \<open>(\<forall>n>a. P n) \<longleftrightarrow> (\<forall>n\<in>{a<..}. P n)\<close>
 | 
|
4674  | 
by auto  | 
|
4675  | 
||
4676  | 
qualified lemma forall_lessThan_iff [code_unfold]:  | 
|
4677  | 
  \<open>(\<forall>n\<in>{..<b}. P n) \<longleftrightarrow> all_range P 0 b\<close>
 | 
|
4678  | 
by auto  | 
|
4679  | 
||
4680  | 
qualified lemma forall_less_iff [code_unfold]:  | 
|
4681  | 
  \<open>(\<forall>n<b. P n) \<longleftrightarrow> (\<forall>n\<in>{..<b}. P n)\<close>
 | 
|
4682  | 
by auto  | 
|
4683  | 
||
4684  | 
qualified lemma forall_atMost_iff [code_unfold]:  | 
|
4685  | 
  \<open>(\<forall>n\<in>{..b}. P n) \<longleftrightarrow> P b \<and> all_range P 0 b\<close>
 | 
|
4686  | 
by auto  | 
|
4687  | 
||
4688  | 
qualified lemma forall_less_eq_iff [code_unfold]:  | 
|
4689  | 
  \<open>(\<forall>n\<le>b. P n) \<longleftrightarrow> (\<forall>n\<in>{..b}. P n)\<close>
 | 
|
4690  | 
by auto  | 
|
4691  | 
||
4692  | 
qualified lemma forall_atLeastLessThan_iff [code_unfold]:  | 
|
4693  | 
  \<open>(\<forall>n\<in>{a..<b}. P n) \<longleftrightarrow> all_range P a b\<close>
 | 
|
4694  | 
by simp  | 
|
4695  | 
||
4696  | 
qualified lemma forall_atLeastAtMost_iff [code_unfold]:  | 
|
4697  | 
  \<open>(\<forall>n\<in>{a..b}. P n) \<longleftrightarrow> (if b = - 1 then (\<forall>n\<in>{a..}. P n) else all_range P a (b + 1))\<close>
 | 
|
4698  | 
apply auto  | 
|
4699  | 
apply (metis atLeastAtMost_iff inc_le not_less_iff_gr_or_eq order_le_less)  | 
|
4700  | 
apply (metis (no_types, lifting) antisym_conv2 atLeastLessThan_iff dual_order.trans inc_less_eq_triv_imp  | 
|
4701  | 
nle_le)  | 
|
4702  | 
done  | 
|
4703  | 
||
4704  | 
qualified lemma forall_greaterThanLessThan_iff [code_unfold]:  | 
|
4705  | 
  \<open>(\<forall>n\<in>{a<..<b}. P n) \<longleftrightarrow> a = - 1 \<or> all_range P (a + 1) b\<close>
 | 
|
4706  | 
by (auto simp add: inc_less_eq_iff)  | 
|
4707  | 
||
4708  | 
qualified lemma forall_greaterThanAtMost_iff [code_unfold]:  | 
|
4709  | 
  \<open>(\<forall>n\<in>{a<..b}. P n) \<longleftrightarrow> (if b = - 1 then (\<forall>n\<in>{a<..}. P n) else a = - 1 \<or> all_range P (a + 1) (b + 1))\<close>
 | 
|
4710  | 
apply auto  | 
|
4711  | 
apply (metis greaterThanAtMost_iff inc_less_eq_iff linorder_not_less)  | 
|
4712  | 
apply (meson atLeastLessThan_iff inc_less_eq_iff order_le_less_trans word_le_less_eq)  | 
|
4713  | 
done  | 
|
4714  | 
||
4715  | 
qualified lemma exists_atLeast_iff [code_unfold]:  | 
|
4716  | 
  \<open>(\<exists>n\<in>{a..}. P n) \<longleftrightarrow> \<not> all_range (Not \<circ> P) a (- 1) \<or> P (- 1)\<close>
 | 
|
4717  | 
using forall_atLeast_iff [of a \<open>Not \<circ> P\<close>] by auto  | 
|
4718  | 
||
4719  | 
qualified lemma exists_greater_eq_iff [code_unfold]:  | 
|
4720  | 
  \<open>(\<exists>n\<ge>a. P n) \<longleftrightarrow> (\<exists>n\<in>{a..}. P n)\<close>
 | 
|
4721  | 
by auto  | 
|
4722  | 
||
4723  | 
qualified lemma exists_greaterThan_iff [code_unfold]:  | 
|
4724  | 
  \<open>(\<exists>n\<in>{a<..}. P n) \<longleftrightarrow> a \<noteq> - 1 \<and> (\<not> all_range (Not \<circ> P) (a + 1) (- 1) \<or> P (- 1))\<close>
 | 
|
4725  | 
using forall_greaterThan_iff [of a \<open>Not \<circ> P\<close>] by auto  | 
|
4726  | 
||
4727  | 
qualified lemma exists_greater_iff [code_unfold]:  | 
|
4728  | 
  \<open>(\<exists>n>a. P n) \<longleftrightarrow> (\<exists>n\<in>{a<..}. P n)\<close>
 | 
|
4729  | 
by auto  | 
|
4730  | 
||
4731  | 
qualified lemma exists_lessThan_iff [code_unfold]:  | 
|
4732  | 
  \<open>(\<exists>n\<in>{..<b}. P n) \<longleftrightarrow> \<not> all_range (Not \<circ> P) 0 b\<close>
 | 
|
4733  | 
using forall_lessThan_iff [of b \<open>Not \<circ> P\<close>] by auto  | 
|
4734  | 
||
4735  | 
qualified lemma exists_less_iff [code_unfold]:  | 
|
4736  | 
  \<open>(\<exists>n<b. P n) \<longleftrightarrow> (\<exists>n\<in>{..<b}. P n)\<close>
 | 
|
4737  | 
by auto  | 
|
4738  | 
||
4739  | 
qualified lemma exists_atMost_iff [code_unfold]:  | 
|
4740  | 
  \<open>(\<exists>n\<in>{..b}. P n) \<longleftrightarrow> P b \<or> \<not> all_range (Not \<circ> P) 0 b\<close>
 | 
|
4741  | 
using forall_atMost_iff [of b \<open>Not \<circ> P\<close>] by auto  | 
|
4742  | 
||
4743  | 
qualified lemma exists_less_eq_iff [code_unfold]:  | 
|
4744  | 
  \<open>(\<exists>n\<le>b. P n) \<longleftrightarrow> (\<exists>n\<in>{..b}. P n)\<close>
 | 
|
4745  | 
by auto  | 
|
4746  | 
||
4747  | 
qualified lemma exists_atLeastLessThan_iff [code_unfold]:  | 
|
4748  | 
  \<open>(\<exists>n\<in>{a..<b}. P n) \<longleftrightarrow> \<not> all_range (Not \<circ> P) a b\<close>
 | 
|
4749  | 
using forall_atLeastLessThan_iff [of a b \<open>Not \<circ> P\<close>] by simp  | 
|
4750  | 
||
4751  | 
qualified lemma exists_atLeastAtMost_iff [code_unfold]:  | 
|
4752  | 
  \<open>(\<exists>n\<in>{a..b}. P n) \<longleftrightarrow> (if b = - 1 then \<exists>n\<in>{a..}. P n else \<not> all_range (Not \<circ> P) a (b + 1))\<close>
 | 
|
4753  | 
using forall_atLeastAtMost_iff [of a b \<open>Not \<circ> P\<close>] by auto  | 
|
4754  | 
||
4755  | 
qualified lemma exists_greaterThanLessThan_iff [code_unfold]:  | 
|
4756  | 
  \<open>(\<exists>n\<in>{a<..<b}. P n) \<longleftrightarrow> a \<noteq> - 1 \<and> \<not> all_range (Not \<circ> P) (a + 1) b\<close>
 | 
|
4757  | 
using forall_greaterThanLessThan_iff [of a b \<open>Not \<circ> P\<close>] by auto  | 
|
4758  | 
||
4759  | 
qualified lemma exists_greaterThanAtMost_iff [code_unfold]:  | 
|
4760  | 
  \<open>(\<exists>n\<in>{a<..b}. P n) \<longleftrightarrow> (if b = - 1 then \<exists>n\<in>{a<..}. P n else a \<noteq> - 1 \<and> \<not> all_range (Not \<circ> P) (a + 1) (b + 1))\<close>
 | 
|
4761  | 
using forall_greaterThanAtMost_iff [of a b \<open>Not \<circ> P\<close>] by auto  | 
|
4762  | 
||
4763  | 
end  | 
|
4764  | 
||
| 74592 | 4765  | 
subsection \<open>Tool support\<close>  | 
| 72489 | 4766  | 
|
| 69605 | 4767  | 
ML_file \<open>Tools/smt_word.ML\<close>  | 
| 
36899
 
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
 
boehmes 
parents: 
35049 
diff
changeset
 | 
4768  | 
|
| 
41060
 
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
 
boehmes 
parents: 
40827 
diff
changeset
 | 
4769  | 
end  |