author | boehmes |
Tue, 07 Dec 2010 14:53:44 +0100 | |
changeset 41060 | 4199fdcfa3c0 |
parent 40827 | abbc05c20e24 |
child 41413 | 64cd30d6b0b8 |
permissions | -rw-r--r-- |
29628 | 1 |
(* Title: HOL/Word/Word.thy |
37660 | 2 |
Author: Jeremy Dawson and Gerwin Klein, NICTA |
24333 | 3 |
*) |
4 |
||
37660 | 5 |
header {* A type of finite bit strings *} |
24350 | 6 |
|
29628 | 7 |
theory Word |
37660 | 8 |
imports Type_Length Misc_Typedef Boolean_Algebra Bool_List_Representation |
41060
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
boehmes
parents:
40827
diff
changeset
|
9 |
uses ("~~/src/HOL/Word/Tools/smt_word.ML") |
37660 | 10 |
begin |
11 |
||
12 |
text {* see @{text "Examples/WordExamples.thy"} for examples *} |
|
13 |
||
14 |
subsection {* Type definition *} |
|
15 |
||
16 |
typedef (open word) 'a word = "{(0::int) ..< 2^len_of TYPE('a::len0)}" |
|
17 |
morphisms uint Abs_word by auto |
|
18 |
||
19 |
definition word_of_int :: "int \<Rightarrow> 'a\<Colon>len0 word" where |
|
20 |
-- {* representation of words using unsigned or signed bins, |
|
21 |
only difference in these is the type class *} |
|
22 |
"word_of_int w = Abs_word (bintrunc (len_of TYPE ('a)) w)" |
|
23 |
||
24 |
lemma uint_word_of_int [code]: "uint (word_of_int w \<Colon> 'a\<Colon>len0 word) = w mod 2 ^ len_of TYPE('a)" |
|
25 |
by (auto simp add: word_of_int_def bintrunc_mod2p intro: Abs_word_inverse) |
|
26 |
||
27 |
code_datatype word_of_int |
|
28 |
||
37751 | 29 |
notation fcomp (infixl "\<circ>>" 60) |
30 |
notation scomp (infixl "\<circ>\<rightarrow>" 60) |
|
37660 | 31 |
|
32 |
instantiation word :: ("{len0, typerep}") random |
|
33 |
begin |
|
34 |
||
35 |
definition |
|
37751 | 36 |
"random_word i = Random.range (max i (2 ^ len_of TYPE('a))) \<circ>\<rightarrow> (\<lambda>k. Pair ( |
37660 | 37 |
let j = word_of_int (Code_Numeral.int_of k) :: 'a word |
38 |
in (j, \<lambda>_::unit. Code_Evaluation.term_of j)))" |
|
39 |
||
40 |
instance .. |
|
41 |
||
42 |
end |
|
43 |
||
37751 | 44 |
no_notation fcomp (infixl "\<circ>>" 60) |
45 |
no_notation scomp (infixl "\<circ>\<rightarrow>" 60) |
|
37660 | 46 |
|
47 |
||
48 |
subsection {* Type conversions and casting *} |
|
49 |
||
50 |
definition sint :: "'a :: len word => int" where |
|
51 |
-- {* treats the most-significant-bit as a sign bit *} |
|
52 |
sint_uint: "sint w = sbintrunc (len_of TYPE ('a) - 1) (uint w)" |
|
53 |
||
54 |
definition unat :: "'a :: len0 word => nat" where |
|
55 |
"unat w = nat (uint w)" |
|
56 |
||
57 |
definition uints :: "nat => int set" where |
|
58 |
-- "the sets of integers representing the words" |
|
59 |
"uints n = range (bintrunc n)" |
|
60 |
||
61 |
definition sints :: "nat => int set" where |
|
62 |
"sints n = range (sbintrunc (n - 1))" |
|
63 |
||
64 |
definition unats :: "nat => nat set" where |
|
65 |
"unats n = {i. i < 2 ^ n}" |
|
66 |
||
67 |
definition norm_sint :: "nat => int => int" where |
|
68 |
"norm_sint n w = (w + 2 ^ (n - 1)) mod 2 ^ n - 2 ^ (n - 1)" |
|
69 |
||
70 |
definition scast :: "'a :: len word => 'b :: len word" where |
|
71 |
-- "cast a word to a different length" |
|
72 |
"scast w = word_of_int (sint w)" |
|
73 |
||
74 |
definition ucast :: "'a :: len0 word => 'b :: len0 word" where |
|
75 |
"ucast w = word_of_int (uint w)" |
|
76 |
||
77 |
instantiation word :: (len0) size |
|
78 |
begin |
|
79 |
||
80 |
definition |
|
81 |
word_size: "size (w :: 'a word) = len_of TYPE('a)" |
|
82 |
||
83 |
instance .. |
|
84 |
||
85 |
end |
|
86 |
||
87 |
definition source_size :: "('a :: len0 word => 'b) => nat" where |
|
88 |
-- "whether a cast (or other) function is to a longer or shorter length" |
|
89 |
"source_size c = (let arb = undefined ; x = c arb in size arb)" |
|
90 |
||
91 |
definition target_size :: "('a => 'b :: len0 word) => nat" where |
|
92 |
"target_size c = size (c undefined)" |
|
93 |
||
94 |
definition is_up :: "('a :: len0 word => 'b :: len0 word) => bool" where |
|
95 |
"is_up c \<longleftrightarrow> source_size c <= target_size c" |
|
96 |
||
97 |
definition is_down :: "('a :: len0 word => 'b :: len0 word) => bool" where |
|
98 |
"is_down c \<longleftrightarrow> target_size c <= source_size c" |
|
99 |
||
100 |
definition of_bl :: "bool list => 'a :: len0 word" where |
|
101 |
"of_bl bl = word_of_int (bl_to_bin bl)" |
|
102 |
||
103 |
definition to_bl :: "'a :: len0 word => bool list" where |
|
104 |
"to_bl w = bin_to_bl (len_of TYPE ('a)) (uint w)" |
|
105 |
||
106 |
definition word_reverse :: "'a :: len0 word => 'a word" where |
|
107 |
"word_reverse w = of_bl (rev (to_bl w))" |
|
108 |
||
109 |
definition word_int_case :: "(int => 'b) => ('a :: len0 word) => 'b" where |
|
110 |
"word_int_case f w = f (uint w)" |
|
111 |
||
112 |
syntax |
|
113 |
of_int :: "int => 'a" |
|
114 |
translations |
|
115 |
"case x of CONST of_int y => b" == "CONST word_int_case (%y. b) x" |
|
116 |
||
117 |
||
118 |
subsection "Arithmetic operations" |
|
119 |
||
120 |
instantiation word :: (len0) "{number, uminus, minus, plus, one, zero, times, Divides.div, ord, bit}" |
|
121 |
begin |
|
122 |
||
123 |
definition |
|
124 |
word_0_wi: "0 = word_of_int 0" |
|
125 |
||
126 |
definition |
|
127 |
word_1_wi: "1 = word_of_int 1" |
|
128 |
||
129 |
definition |
|
130 |
word_add_def: "a + b = word_of_int (uint a + uint b)" |
|
131 |
||
132 |
definition |
|
133 |
word_sub_wi: "a - b = word_of_int (uint a - uint b)" |
|
134 |
||
135 |
definition |
|
136 |
word_minus_def: "- a = word_of_int (- uint a)" |
|
137 |
||
138 |
definition |
|
139 |
word_mult_def: "a * b = word_of_int (uint a * uint b)" |
|
140 |
||
141 |
definition |
|
142 |
word_div_def: "a div b = word_of_int (uint a div uint b)" |
|
143 |
||
144 |
definition |
|
145 |
word_mod_def: "a mod b = word_of_int (uint a mod uint b)" |
|
146 |
||
147 |
definition |
|
148 |
word_number_of_def: "number_of w = word_of_int w" |
|
149 |
||
150 |
definition |
|
151 |
word_le_def: "a \<le> b \<longleftrightarrow> uint a \<le> uint b" |
|
152 |
||
153 |
definition |
|
154 |
word_less_def: "x < y \<longleftrightarrow> x \<le> y \<and> x \<noteq> (y \<Colon> 'a word)" |
|
155 |
||
156 |
definition |
|
157 |
word_and_def: |
|
158 |
"(a::'a word) AND b = word_of_int (uint a AND uint b)" |
|
159 |
||
160 |
definition |
|
161 |
word_or_def: |
|
162 |
"(a::'a word) OR b = word_of_int (uint a OR uint b)" |
|
163 |
||
164 |
definition |
|
165 |
word_xor_def: |
|
166 |
"(a::'a word) XOR b = word_of_int (uint a XOR uint b)" |
|
167 |
||
168 |
definition |
|
169 |
word_not_def: |
|
170 |
"NOT (a::'a word) = word_of_int (NOT (uint a))" |
|
171 |
||
172 |
instance .. |
|
173 |
||
174 |
end |
|
175 |
||
176 |
definition |
|
177 |
word_succ :: "'a :: len0 word => 'a word" |
|
178 |
where |
|
179 |
"word_succ a = word_of_int (Int.succ (uint a))" |
|
180 |
||
181 |
definition |
|
182 |
word_pred :: "'a :: len0 word => 'a word" |
|
183 |
where |
|
184 |
"word_pred a = word_of_int (Int.pred (uint a))" |
|
185 |
||
186 |
definition udvd :: "'a::len word => 'a::len word => bool" (infixl "udvd" 50) where |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
187 |
"a udvd b = (EX n>=0. uint b = n * uint a)" |
37660 | 188 |
|
189 |
definition word_sle :: "'a :: len word => 'a word => bool" ("(_/ <=s _)" [50, 51] 50) where |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
190 |
"a <=s b = (sint a <= sint b)" |
37660 | 191 |
|
192 |
definition word_sless :: "'a :: len word => 'a word => bool" ("(_/ <s _)" [50, 51] 50) where |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
193 |
"(x <s y) = (x <=s y & x ~= y)" |
37660 | 194 |
|
195 |
||
196 |
||
197 |
subsection "Bit-wise operations" |
|
198 |
||
199 |
instantiation word :: (len0) bits |
|
200 |
begin |
|
201 |
||
202 |
definition |
|
203 |
word_test_bit_def: "test_bit a = bin_nth (uint a)" |
|
204 |
||
205 |
definition |
|
206 |
word_set_bit_def: "set_bit a n x = |
|
207 |
word_of_int (bin_sc n (If x 1 0) (uint a))" |
|
208 |
||
209 |
definition |
|
210 |
word_set_bits_def: "(BITS n. f n) = of_bl (bl_of_nth (len_of TYPE ('a)) f)" |
|
211 |
||
212 |
definition |
|
213 |
word_lsb_def: "lsb a \<longleftrightarrow> bin_last (uint a) = 1" |
|
214 |
||
215 |
definition shiftl1 :: "'a word \<Rightarrow> 'a word" where |
|
216 |
"shiftl1 w = word_of_int (uint w BIT 0)" |
|
217 |
||
218 |
definition shiftr1 :: "'a word \<Rightarrow> 'a word" where |
|
219 |
-- "shift right as unsigned or as signed, ie logical or arithmetic" |
|
220 |
"shiftr1 w = word_of_int (bin_rest (uint w))" |
|
221 |
||
222 |
definition |
|
223 |
shiftl_def: "w << n = (shiftl1 ^^ n) w" |
|
224 |
||
225 |
definition |
|
226 |
shiftr_def: "w >> n = (shiftr1 ^^ n) w" |
|
227 |
||
228 |
instance .. |
|
229 |
||
230 |
end |
|
231 |
||
232 |
instantiation word :: (len) bitss |
|
233 |
begin |
|
234 |
||
235 |
definition |
|
236 |
word_msb_def: |
|
237 |
"msb a \<longleftrightarrow> bin_sign (sint a) = Int.Min" |
|
238 |
||
239 |
instance .. |
|
240 |
||
241 |
end |
|
242 |
||
37667 | 243 |
lemma [code]: |
244 |
"msb a \<longleftrightarrow> bin_sign (sint a) = (- 1 :: int)" |
|
245 |
by (simp only: word_msb_def Min_def) |
|
246 |
||
37660 | 247 |
definition setBit :: "'a :: len0 word => nat => 'a word" where |
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
248 |
"setBit w n = set_bit w n True" |
37660 | 249 |
|
250 |
definition clearBit :: "'a :: len0 word => nat => 'a word" where |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
251 |
"clearBit w n = set_bit w n False" |
37660 | 252 |
|
253 |
||
254 |
subsection "Shift operations" |
|
255 |
||
256 |
definition sshiftr1 :: "'a :: len word => 'a word" where |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
257 |
"sshiftr1 w = word_of_int (bin_rest (sint w))" |
37660 | 258 |
|
259 |
definition bshiftr1 :: "bool => 'a :: len word => 'a word" where |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
260 |
"bshiftr1 b w = of_bl (b # butlast (to_bl w))" |
37660 | 261 |
|
262 |
definition sshiftr :: "'a :: len word => nat => 'a word" (infixl ">>>" 55) where |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
263 |
"w >>> n = (sshiftr1 ^^ n) w" |
37660 | 264 |
|
265 |
definition mask :: "nat => 'a::len word" where |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
266 |
"mask n = (1 << n) - 1" |
37660 | 267 |
|
268 |
definition revcast :: "'a :: len0 word => 'b :: len0 word" where |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
269 |
"revcast w = of_bl (takefill False (len_of TYPE('b)) (to_bl w))" |
37660 | 270 |
|
271 |
definition slice1 :: "nat => 'a :: len0 word => 'b :: len0 word" where |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
272 |
"slice1 n w = of_bl (takefill False n (to_bl w))" |
37660 | 273 |
|
274 |
definition slice :: "nat => 'a :: len0 word => 'b :: len0 word" where |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
275 |
"slice n w = slice1 (size w - n) w" |
37660 | 276 |
|
277 |
||
278 |
subsection "Rotation" |
|
279 |
||
280 |
definition rotater1 :: "'a list => 'a list" where |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
281 |
"rotater1 ys = |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
282 |
(case ys of [] => [] | x # xs => last ys # butlast ys)" |
37660 | 283 |
|
284 |
definition rotater :: "nat => 'a list => 'a list" where |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
285 |
"rotater n = rotater1 ^^ n" |
37660 | 286 |
|
287 |
definition word_rotr :: "nat => 'a :: len0 word => 'a :: len0 word" where |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
288 |
"word_rotr n w = of_bl (rotater n (to_bl w))" |
37660 | 289 |
|
290 |
definition word_rotl :: "nat => 'a :: len0 word => 'a :: len0 word" where |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
291 |
"word_rotl n w = of_bl (rotate n (to_bl w))" |
37660 | 292 |
|
293 |
definition word_roti :: "int => 'a :: len0 word => 'a :: len0 word" where |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
294 |
"word_roti i w = (if i >= 0 then word_rotr (nat i) w |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
295 |
else word_rotl (nat (- i)) w)" |
37660 | 296 |
|
297 |
||
298 |
subsection "Split and cat operations" |
|
299 |
||
300 |
definition word_cat :: "'a :: len0 word => 'b :: len0 word => 'c :: len0 word" where |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
301 |
"word_cat a b = word_of_int (bin_cat (uint a) (len_of TYPE ('b)) (uint b))" |
37660 | 302 |
|
303 |
definition word_split :: "'a :: len0 word => ('b :: len0 word) * ('c :: len0 word)" where |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
304 |
"word_split a = |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
305 |
(case bin_split (len_of TYPE ('c)) (uint a) of |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
306 |
(u, v) => (word_of_int u, word_of_int v))" |
37660 | 307 |
|
308 |
definition word_rcat :: "'a :: len0 word list => 'b :: len0 word" where |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
309 |
"word_rcat ws = |
37660 | 310 |
word_of_int (bin_rcat (len_of TYPE ('a)) (map uint ws))" |
311 |
||
312 |
definition word_rsplit :: "'a :: len0 word => 'b :: len word list" where |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
313 |
"word_rsplit w = |
37660 | 314 |
map word_of_int (bin_rsplit (len_of TYPE ('b)) (len_of TYPE ('a), uint w))" |
315 |
||
316 |
definition max_word :: "'a::len word" -- "Largest representable machine integer." where |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
317 |
"max_word = word_of_int (2 ^ len_of TYPE('a) - 1)" |
37660 | 318 |
|
319 |
primrec of_bool :: "bool \<Rightarrow> 'a::len word" where |
|
320 |
"of_bool False = 0" |
|
321 |
| "of_bool True = 1" |
|
322 |
||
323 |
||
324 |
lemmas of_nth_def = word_set_bits_def |
|
325 |
||
326 |
lemmas word_size_gt_0 [iff] = |
|
327 |
xtr1 [OF word_size len_gt_0, standard] |
|
328 |
lemmas lens_gt_0 = word_size_gt_0 len_gt_0 |
|
329 |
lemmas lens_not_0 [iff] = lens_gt_0 [THEN gr_implies_not0, standard] |
|
330 |
||
331 |
lemma uints_num: "uints n = {i. 0 \<le> i \<and> i < 2 ^ n}" |
|
332 |
by (simp add: uints_def range_bintrunc) |
|
333 |
||
334 |
lemma sints_num: "sints n = {i. - (2 ^ (n - 1)) \<le> i \<and> i < 2 ^ (n - 1)}" |
|
335 |
by (simp add: sints_def range_sbintrunc) |
|
336 |
||
337 |
lemmas atLeastLessThan_alt = atLeastLessThan_def [unfolded |
|
338 |
atLeast_def lessThan_def Collect_conj_eq [symmetric]] |
|
339 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
340 |
lemma mod_in_reps: "m > 0 \<Longrightarrow> y mod m : {0::int ..< m}" |
37660 | 341 |
unfolding atLeastLessThan_alt by auto |
342 |
||
343 |
lemma |
|
344 |
uint_0:"0 <= uint x" and |
|
345 |
uint_lt: "uint (x::'a::len0 word) < 2 ^ len_of TYPE('a)" |
|
346 |
by (auto simp: uint [simplified]) |
|
347 |
||
348 |
lemma uint_mod_same: |
|
349 |
"uint x mod 2 ^ len_of TYPE('a) = uint (x::'a::len0 word)" |
|
350 |
by (simp add: int_mod_eq uint_lt uint_0) |
|
351 |
||
352 |
lemma td_ext_uint: |
|
353 |
"td_ext (uint :: 'a word => int) word_of_int (uints (len_of TYPE('a::len0))) |
|
354 |
(%w::int. w mod 2 ^ len_of TYPE('a))" |
|
355 |
apply (unfold td_ext_def') |
|
356 |
apply (simp add: uints_num word_of_int_def bintrunc_mod2p) |
|
357 |
apply (simp add: uint_mod_same uint_0 uint_lt |
|
358 |
word.uint_inverse word.Abs_word_inverse int_mod_lem) |
|
359 |
done |
|
360 |
||
361 |
lemmas int_word_uint = td_ext_uint [THEN td_ext.eq_norm, standard] |
|
362 |
||
363 |
interpretation word_uint: |
|
364 |
td_ext "uint::'a::len0 word \<Rightarrow> int" |
|
365 |
word_of_int |
|
366 |
"uints (len_of TYPE('a::len0))" |
|
367 |
"\<lambda>w. w mod 2 ^ len_of TYPE('a::len0)" |
|
368 |
by (rule td_ext_uint) |
|
369 |
||
370 |
lemmas td_uint = word_uint.td_thm |
|
371 |
||
372 |
lemmas td_ext_ubin = td_ext_uint |
|
373 |
[simplified len_gt_0 no_bintr_alt1 [symmetric]] |
|
374 |
||
375 |
interpretation word_ubin: |
|
376 |
td_ext "uint::'a::len0 word \<Rightarrow> int" |
|
377 |
word_of_int |
|
378 |
"uints (len_of TYPE('a::len0))" |
|
379 |
"bintrunc (len_of TYPE('a::len0))" |
|
380 |
by (rule td_ext_ubin) |
|
381 |
||
382 |
lemma sint_sbintrunc': |
|
383 |
"sint (word_of_int bin :: 'a word) = |
|
384 |
(sbintrunc (len_of TYPE ('a :: len) - 1) bin)" |
|
385 |
unfolding sint_uint |
|
386 |
by (auto simp: word_ubin.eq_norm sbintrunc_bintrunc_lt) |
|
387 |
||
388 |
lemma uint_sint: |
|
389 |
"uint w = bintrunc (len_of TYPE('a)) (sint (w :: 'a :: len word))" |
|
390 |
unfolding sint_uint by (auto simp: bintrunc_sbintrunc_le) |
|
391 |
||
392 |
lemma bintr_uint': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
393 |
"n >= size w \<Longrightarrow> bintrunc n (uint w) = uint w" |
37660 | 394 |
apply (unfold word_size) |
395 |
apply (subst word_ubin.norm_Rep [symmetric]) |
|
396 |
apply (simp only: bintrunc_bintrunc_min word_size) |
|
397 |
apply (simp add: min_max.inf_absorb2) |
|
398 |
done |
|
399 |
||
400 |
lemma wi_bintr': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
401 |
"wb = word_of_int bin \<Longrightarrow> n >= size wb \<Longrightarrow> |
37660 | 402 |
word_of_int (bintrunc n bin) = wb" |
403 |
unfolding word_size |
|
404 |
by (clarsimp simp add: word_ubin.norm_eq_iff [symmetric] min_max.inf_absorb1) |
|
405 |
||
406 |
lemmas bintr_uint = bintr_uint' [unfolded word_size] |
|
407 |
lemmas wi_bintr = wi_bintr' [unfolded word_size] |
|
408 |
||
409 |
lemma td_ext_sbin: |
|
410 |
"td_ext (sint :: 'a word => int) word_of_int (sints (len_of TYPE('a::len))) |
|
411 |
(sbintrunc (len_of TYPE('a) - 1))" |
|
412 |
apply (unfold td_ext_def' sint_uint) |
|
413 |
apply (simp add : word_ubin.eq_norm) |
|
414 |
apply (cases "len_of TYPE('a)") |
|
415 |
apply (auto simp add : sints_def) |
|
416 |
apply (rule sym [THEN trans]) |
|
417 |
apply (rule word_ubin.Abs_norm) |
|
418 |
apply (simp only: bintrunc_sbintrunc) |
|
419 |
apply (drule sym) |
|
420 |
apply simp |
|
421 |
done |
|
422 |
||
423 |
lemmas td_ext_sint = td_ext_sbin |
|
424 |
[simplified len_gt_0 no_sbintr_alt2 Suc_pred' [symmetric]] |
|
425 |
||
426 |
(* We do sint before sbin, before sint is the user version |
|
427 |
and interpretations do not produce thm duplicates. I.e. |
|
428 |
we get the name word_sint.Rep_eqD, but not word_sbin.Req_eqD, |
|
429 |
because the latter is the same thm as the former *) |
|
430 |
interpretation word_sint: |
|
431 |
td_ext "sint ::'a::len word => int" |
|
432 |
word_of_int |
|
433 |
"sints (len_of TYPE('a::len))" |
|
434 |
"%w. (w + 2^(len_of TYPE('a::len) - 1)) mod 2^len_of TYPE('a::len) - |
|
435 |
2 ^ (len_of TYPE('a::len) - 1)" |
|
436 |
by (rule td_ext_sint) |
|
437 |
||
438 |
interpretation word_sbin: |
|
439 |
td_ext "sint ::'a::len word => int" |
|
440 |
word_of_int |
|
441 |
"sints (len_of TYPE('a::len))" |
|
442 |
"sbintrunc (len_of TYPE('a::len) - 1)" |
|
443 |
by (rule td_ext_sbin) |
|
444 |
||
445 |
lemmas int_word_sint = td_ext_sint [THEN td_ext.eq_norm, standard] |
|
446 |
||
447 |
lemmas td_sint = word_sint.td |
|
448 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
449 |
lemma word_number_of_alt [code_unfold_post]: |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
450 |
"number_of b = word_of_int (number_of b)" |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
451 |
by (simp add: number_of_eq word_number_of_def) |
37660 | 452 |
|
453 |
lemma word_no_wi: "number_of = word_of_int" |
|
454 |
by (auto simp: word_number_of_def intro: ext) |
|
455 |
||
456 |
lemma to_bl_def': |
|
457 |
"(to_bl :: 'a :: len0 word => bool list) = |
|
458 |
bin_to_bl (len_of TYPE('a)) o uint" |
|
459 |
by (auto simp: to_bl_def intro: ext) |
|
460 |
||
461 |
lemmas word_reverse_no_def [simp] = word_reverse_def [of "number_of w", standard] |
|
462 |
||
463 |
lemmas uints_mod = uints_def [unfolded no_bintr_alt1] |
|
464 |
||
465 |
lemma uint_bintrunc: "uint (number_of bin :: 'a word) = |
|
466 |
number_of (bintrunc (len_of TYPE ('a :: len0)) bin)" |
|
467 |
unfolding word_number_of_def number_of_eq |
|
468 |
by (auto intro: word_ubin.eq_norm) |
|
469 |
||
470 |
lemma sint_sbintrunc: "sint (number_of bin :: 'a word) = |
|
471 |
number_of (sbintrunc (len_of TYPE ('a :: len) - 1) bin)" |
|
472 |
unfolding word_number_of_def number_of_eq |
|
473 |
by (subst word_sbin.eq_norm) simp |
|
474 |
||
475 |
lemma unat_bintrunc: |
|
476 |
"unat (number_of bin :: 'a :: len0 word) = |
|
477 |
number_of (bintrunc (len_of TYPE('a)) bin)" |
|
478 |
unfolding unat_def nat_number_of_def |
|
479 |
by (simp only: uint_bintrunc) |
|
480 |
||
481 |
(* WARNING - these may not always be helpful *) |
|
482 |
declare |
|
483 |
uint_bintrunc [simp] |
|
484 |
sint_sbintrunc [simp] |
|
485 |
unat_bintrunc [simp] |
|
486 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
487 |
lemma size_0_eq: "size (w :: 'a :: len0 word) = 0 \<Longrightarrow> v = w" |
37660 | 488 |
apply (unfold word_size) |
489 |
apply (rule word_uint.Rep_eqD) |
|
490 |
apply (rule box_equals) |
|
491 |
defer |
|
492 |
apply (rule word_ubin.norm_Rep)+ |
|
493 |
apply simp |
|
494 |
done |
|
495 |
||
496 |
lemmas uint_lem = word_uint.Rep [unfolded uints_num mem_Collect_eq] |
|
497 |
lemmas sint_lem = word_sint.Rep [unfolded sints_num mem_Collect_eq] |
|
498 |
lemmas uint_ge_0 [iff] = uint_lem [THEN conjunct1, standard] |
|
499 |
lemmas uint_lt2p [iff] = uint_lem [THEN conjunct2, standard] |
|
500 |
lemmas sint_ge = sint_lem [THEN conjunct1, standard] |
|
501 |
lemmas sint_lt = sint_lem [THEN conjunct2, standard] |
|
502 |
||
503 |
lemma sign_uint_Pls [simp]: |
|
504 |
"bin_sign (uint x) = Int.Pls" |
|
505 |
by (simp add: sign_Pls_ge_0 number_of_eq) |
|
506 |
||
507 |
lemmas uint_m2p_neg = iffD2 [OF diff_less_0_iff_less uint_lt2p, standard] |
|
508 |
lemmas uint_m2p_not_non_neg = |
|
509 |
iffD2 [OF linorder_not_le uint_m2p_neg, standard] |
|
510 |
||
511 |
lemma lt2p_lem: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
512 |
"len_of TYPE('a) <= n \<Longrightarrow> uint (w :: 'a :: len0 word) < 2 ^ n" |
37660 | 513 |
by (rule xtr8 [OF _ uint_lt2p]) simp |
514 |
||
515 |
lemmas uint_le_0_iff [simp] = |
|
516 |
uint_ge_0 [THEN leD, THEN linorder_antisym_conv1, standard] |
|
517 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
518 |
lemma uint_nat: "uint w = int (unat w)" |
37660 | 519 |
unfolding unat_def by auto |
520 |
||
521 |
lemma uint_number_of: |
|
522 |
"uint (number_of b :: 'a :: len0 word) = number_of b mod 2 ^ len_of TYPE('a)" |
|
523 |
unfolding word_number_of_alt |
|
524 |
by (simp only: int_word_uint) |
|
525 |
||
526 |
lemma unat_number_of: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
527 |
"bin_sign b = Int.Pls \<Longrightarrow> |
37660 | 528 |
unat (number_of b::'a::len0 word) = number_of b mod 2 ^ len_of TYPE ('a)" |
529 |
apply (unfold unat_def) |
|
530 |
apply (clarsimp simp only: uint_number_of) |
|
531 |
apply (rule nat_mod_distrib [THEN trans]) |
|
532 |
apply (erule sign_Pls_ge_0 [THEN iffD1]) |
|
533 |
apply (simp_all add: nat_power_eq) |
|
534 |
done |
|
535 |
||
536 |
lemma sint_number_of: "sint (number_of b :: 'a :: len word) = (number_of b + |
|
537 |
2 ^ (len_of TYPE('a) - 1)) mod 2 ^ len_of TYPE('a) - |
|
538 |
2 ^ (len_of TYPE('a) - 1)" |
|
539 |
unfolding word_number_of_alt by (rule int_word_sint) |
|
540 |
||
541 |
lemma word_of_int_bin [simp] : |
|
542 |
"(word_of_int (number_of bin) :: 'a :: len0 word) = (number_of bin)" |
|
543 |
unfolding word_number_of_alt by auto |
|
544 |
||
545 |
lemma word_int_case_wi: |
|
546 |
"word_int_case f (word_of_int i :: 'b word) = |
|
547 |
f (i mod 2 ^ len_of TYPE('b::len0))" |
|
548 |
unfolding word_int_case_def by (simp add: word_uint.eq_norm) |
|
549 |
||
550 |
lemma word_int_split: |
|
551 |
"P (word_int_case f x) = |
|
552 |
(ALL i. x = (word_of_int i :: 'b :: len0 word) & |
|
553 |
0 <= i & i < 2 ^ len_of TYPE('b) --> P (f i))" |
|
554 |
unfolding word_int_case_def |
|
555 |
by (auto simp: word_uint.eq_norm int_mod_eq') |
|
556 |
||
557 |
lemma word_int_split_asm: |
|
558 |
"P (word_int_case f x) = |
|
559 |
(~ (EX n. x = (word_of_int n :: 'b::len0 word) & |
|
560 |
0 <= n & n < 2 ^ len_of TYPE('b::len0) & ~ P (f n)))" |
|
561 |
unfolding word_int_case_def |
|
562 |
by (auto simp: word_uint.eq_norm int_mod_eq') |
|
563 |
||
564 |
lemmas uint_range' = |
|
565 |
word_uint.Rep [unfolded uints_num mem_Collect_eq, standard] |
|
566 |
lemmas sint_range' = word_sint.Rep [unfolded One_nat_def |
|
567 |
sints_num mem_Collect_eq, standard] |
|
568 |
||
569 |
lemma uint_range_size: "0 <= uint w & uint w < 2 ^ size w" |
|
570 |
unfolding word_size by (rule uint_range') |
|
571 |
||
572 |
lemma sint_range_size: |
|
573 |
"- (2 ^ (size w - Suc 0)) <= sint w & sint w < 2 ^ (size w - Suc 0)" |
|
574 |
unfolding word_size by (rule sint_range') |
|
575 |
||
576 |
lemmas sint_above_size = sint_range_size |
|
577 |
[THEN conjunct2, THEN [2] xtr8, folded One_nat_def, standard] |
|
578 |
||
579 |
lemmas sint_below_size = sint_range_size |
|
580 |
[THEN conjunct1, THEN [2] order_trans, folded One_nat_def, standard] |
|
581 |
||
582 |
lemma test_bit_eq_iff: "(test_bit (u::'a::len0 word) = test_bit v) = (u = v)" |
|
583 |
unfolding word_test_bit_def by (simp add: bin_nth_eq_iff) |
|
584 |
||
585 |
lemma test_bit_size [rule_format] : "(w::'a::len0 word) !! n --> n < size w" |
|
586 |
apply (unfold word_test_bit_def) |
|
587 |
apply (subst word_ubin.norm_Rep [symmetric]) |
|
588 |
apply (simp only: nth_bintr word_size) |
|
589 |
apply fast |
|
590 |
done |
|
591 |
||
592 |
lemma word_eqI [rule_format] : |
|
593 |
fixes u :: "'a::len0 word" |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
594 |
shows "(ALL n. n < size u --> u !! n = v !! n) \<Longrightarrow> u = v" |
37660 | 595 |
apply (rule test_bit_eq_iff [THEN iffD1]) |
596 |
apply (rule ext) |
|
597 |
apply (erule allE) |
|
598 |
apply (erule impCE) |
|
599 |
prefer 2 |
|
600 |
apply assumption |
|
601 |
apply (auto dest!: test_bit_size simp add: word_size) |
|
602 |
done |
|
603 |
||
604 |
lemmas word_eqD = test_bit_eq_iff [THEN iffD2, THEN fun_cong, standard] |
|
605 |
||
606 |
lemma test_bit_bin': "w !! n = (n < size w & bin_nth (uint w) n)" |
|
607 |
unfolding word_test_bit_def word_size |
|
608 |
by (simp add: nth_bintr [symmetric]) |
|
609 |
||
610 |
lemmas test_bit_bin = test_bit_bin' [unfolded word_size] |
|
611 |
||
612 |
lemma bin_nth_uint_imp': "bin_nth (uint w) n --> n < size w" |
|
613 |
apply (unfold word_size) |
|
614 |
apply (rule impI) |
|
615 |
apply (rule nth_bintr [THEN iffD1, THEN conjunct1]) |
|
616 |
apply (subst word_ubin.norm_Rep) |
|
617 |
apply assumption |
|
618 |
done |
|
619 |
||
620 |
lemma bin_nth_sint': |
|
621 |
"n >= size w --> bin_nth (sint w) n = bin_nth (sint w) (size w - 1)" |
|
622 |
apply (rule impI) |
|
623 |
apply (subst word_sbin.norm_Rep [symmetric]) |
|
624 |
apply (simp add : nth_sbintr word_size) |
|
625 |
apply auto |
|
626 |
done |
|
627 |
||
628 |
lemmas bin_nth_uint_imp = bin_nth_uint_imp' [rule_format, unfolded word_size] |
|
629 |
lemmas bin_nth_sint = bin_nth_sint' [rule_format, unfolded word_size] |
|
630 |
||
631 |
(* type definitions theorem for in terms of equivalent bool list *) |
|
632 |
lemma td_bl: |
|
633 |
"type_definition (to_bl :: 'a::len0 word => bool list) |
|
634 |
of_bl |
|
635 |
{bl. length bl = len_of TYPE('a)}" |
|
636 |
apply (unfold type_definition_def of_bl_def to_bl_def) |
|
637 |
apply (simp add: word_ubin.eq_norm) |
|
638 |
apply safe |
|
639 |
apply (drule sym) |
|
640 |
apply simp |
|
641 |
done |
|
642 |
||
643 |
interpretation word_bl: |
|
644 |
type_definition "to_bl :: 'a::len0 word => bool list" |
|
645 |
of_bl |
|
646 |
"{bl. length bl = len_of TYPE('a::len0)}" |
|
647 |
by (rule td_bl) |
|
648 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
649 |
lemma word_size_bl: "size w = size (to_bl w)" |
37660 | 650 |
unfolding word_size by auto |
651 |
||
652 |
lemma to_bl_use_of_bl: |
|
653 |
"(to_bl w = bl) = (w = of_bl bl \<and> length bl = length (to_bl w))" |
|
654 |
by (fastsimp elim!: word_bl.Abs_inverse [simplified]) |
|
655 |
||
656 |
lemma to_bl_word_rev: "to_bl (word_reverse w) = rev (to_bl w)" |
|
657 |
unfolding word_reverse_def by (simp add: word_bl.Abs_inverse) |
|
658 |
||
659 |
lemma word_rev_rev [simp] : "word_reverse (word_reverse w) = w" |
|
660 |
unfolding word_reverse_def by (simp add : word_bl.Abs_inverse) |
|
661 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
662 |
lemma word_rev_gal: "word_reverse w = u \<Longrightarrow> word_reverse u = w" |
37660 | 663 |
by auto |
664 |
||
665 |
lemmas word_rev_gal' = sym [THEN word_rev_gal, symmetric, standard] |
|
666 |
||
667 |
lemmas length_bl_gt_0 [iff] = xtr1 [OF word_bl.Rep' len_gt_0, standard] |
|
668 |
lemmas bl_not_Nil [iff] = |
|
669 |
length_bl_gt_0 [THEN length_greater_0_conv [THEN iffD1], standard] |
|
670 |
lemmas length_bl_neq_0 [iff] = length_bl_gt_0 [THEN gr_implies_not0] |
|
671 |
||
672 |
lemma hd_bl_sign_sint: "hd (to_bl w) = (bin_sign (sint w) = Int.Min)" |
|
673 |
apply (unfold to_bl_def sint_uint) |
|
674 |
apply (rule trans [OF _ bl_sbin_sign]) |
|
675 |
apply simp |
|
676 |
done |
|
677 |
||
678 |
lemma of_bl_drop': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
679 |
"lend = length bl - len_of TYPE ('a :: len0) \<Longrightarrow> |
37660 | 680 |
of_bl (drop lend bl) = (of_bl bl :: 'a word)" |
681 |
apply (unfold of_bl_def) |
|
682 |
apply (clarsimp simp add : trunc_bl2bin [symmetric]) |
|
683 |
done |
|
684 |
||
685 |
lemmas of_bl_no = of_bl_def [folded word_number_of_def] |
|
686 |
||
687 |
lemma test_bit_of_bl: |
|
688 |
"(of_bl bl::'a::len0 word) !! n = (rev bl ! n \<and> n < len_of TYPE('a) \<and> n < length bl)" |
|
689 |
apply (unfold of_bl_def word_test_bit_def) |
|
690 |
apply (auto simp add: word_size word_ubin.eq_norm nth_bintr bin_nth_of_bl) |
|
691 |
done |
|
692 |
||
693 |
lemma no_of_bl: |
|
694 |
"(number_of bin ::'a::len0 word) = of_bl (bin_to_bl (len_of TYPE ('a)) bin)" |
|
695 |
unfolding word_size of_bl_no by (simp add : word_number_of_def) |
|
696 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
697 |
lemma uint_bl: "to_bl w = bin_to_bl (size w) (uint w)" |
37660 | 698 |
unfolding word_size to_bl_def by auto |
699 |
||
700 |
lemma to_bl_bin: "bl_to_bin (to_bl w) = uint w" |
|
701 |
unfolding uint_bl by (simp add : word_size) |
|
702 |
||
703 |
lemma to_bl_of_bin: |
|
704 |
"to_bl (word_of_int bin::'a::len0 word) = bin_to_bl (len_of TYPE('a)) bin" |
|
705 |
unfolding uint_bl by (clarsimp simp add: word_ubin.eq_norm word_size) |
|
706 |
||
707 |
lemmas to_bl_no_bin [simp] = to_bl_of_bin [folded word_number_of_def] |
|
708 |
||
709 |
lemma to_bl_to_bin [simp] : "bl_to_bin (to_bl w) = uint w" |
|
710 |
unfolding uint_bl by (simp add : word_size) |
|
711 |
||
712 |
lemmas uint_bl_bin [simp] = trans [OF bin_bl_bin word_ubin.norm_Rep, standard] |
|
713 |
||
714 |
lemmas num_AB_u [simp] = word_uint.Rep_inverse |
|
715 |
[unfolded o_def word_number_of_def [symmetric], standard] |
|
716 |
lemmas num_AB_s [simp] = word_sint.Rep_inverse |
|
717 |
[unfolded o_def word_number_of_def [symmetric], standard] |
|
718 |
||
719 |
(* naturals *) |
|
720 |
lemma uints_unats: "uints n = int ` unats n" |
|
721 |
apply (unfold unats_def uints_num) |
|
722 |
apply safe |
|
723 |
apply (rule_tac image_eqI) |
|
724 |
apply (erule_tac nat_0_le [symmetric]) |
|
725 |
apply auto |
|
726 |
apply (erule_tac nat_less_iff [THEN iffD2]) |
|
727 |
apply (rule_tac [2] zless_nat_eq_int_zless [THEN iffD1]) |
|
728 |
apply (auto simp add : nat_power_eq int_power) |
|
729 |
done |
|
730 |
||
731 |
lemma unats_uints: "unats n = nat ` uints n" |
|
732 |
by (auto simp add : uints_unats image_iff) |
|
733 |
||
734 |
lemmas bintr_num = word_ubin.norm_eq_iff |
|
735 |
[symmetric, folded word_number_of_def, standard] |
|
736 |
lemmas sbintr_num = word_sbin.norm_eq_iff |
|
737 |
[symmetric, folded word_number_of_def, standard] |
|
738 |
||
739 |
lemmas num_of_bintr = word_ubin.Abs_norm [folded word_number_of_def, standard] |
|
740 |
lemmas num_of_sbintr = word_sbin.Abs_norm [folded word_number_of_def, standard]; |
|
741 |
||
742 |
(* don't add these to simpset, since may want bintrunc n w to be simplified; |
|
743 |
may want these in reverse, but loop as simp rules, so use following *) |
|
744 |
||
745 |
lemma num_of_bintr': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
746 |
"bintrunc (len_of TYPE('a :: len0)) a = b \<Longrightarrow> |
37660 | 747 |
number_of a = (number_of b :: 'a word)" |
748 |
apply safe |
|
749 |
apply (rule_tac num_of_bintr [symmetric]) |
|
750 |
done |
|
751 |
||
752 |
lemma num_of_sbintr': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
753 |
"sbintrunc (len_of TYPE('a :: len) - 1) a = b \<Longrightarrow> |
37660 | 754 |
number_of a = (number_of b :: 'a word)" |
755 |
apply safe |
|
756 |
apply (rule_tac num_of_sbintr [symmetric]) |
|
757 |
done |
|
758 |
||
759 |
lemmas num_abs_bintr = sym [THEN trans, |
|
760 |
OF num_of_bintr word_number_of_def, standard] |
|
761 |
lemmas num_abs_sbintr = sym [THEN trans, |
|
762 |
OF num_of_sbintr word_number_of_def, standard] |
|
763 |
||
764 |
(** cast - note, no arg for new length, as it's determined by type of result, |
|
765 |
thus in "cast w = w, the type means cast to length of w! **) |
|
766 |
||
767 |
lemma ucast_id: "ucast w = w" |
|
768 |
unfolding ucast_def by auto |
|
769 |
||
770 |
lemma scast_id: "scast w = w" |
|
771 |
unfolding scast_def by auto |
|
772 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
773 |
lemma ucast_bl: "ucast w = of_bl (to_bl w)" |
37660 | 774 |
unfolding ucast_def of_bl_def uint_bl |
775 |
by (auto simp add : word_size) |
|
776 |
||
777 |
lemma nth_ucast: |
|
778 |
"(ucast w::'a::len0 word) !! n = (w !! n & n < len_of TYPE('a))" |
|
779 |
apply (unfold ucast_def test_bit_bin) |
|
780 |
apply (simp add: word_ubin.eq_norm nth_bintr word_size) |
|
781 |
apply (fast elim!: bin_nth_uint_imp) |
|
782 |
done |
|
783 |
||
784 |
(* for literal u(s)cast *) |
|
785 |
||
786 |
lemma ucast_bintr [simp]: |
|
787 |
"ucast (number_of w ::'a::len0 word) = |
|
788 |
number_of (bintrunc (len_of TYPE('a)) w)" |
|
789 |
unfolding ucast_def by simp |
|
790 |
||
791 |
lemma scast_sbintr [simp]: |
|
792 |
"scast (number_of w ::'a::len word) = |
|
793 |
number_of (sbintrunc (len_of TYPE('a) - Suc 0) w)" |
|
794 |
unfolding scast_def by simp |
|
795 |
||
796 |
lemmas source_size = source_size_def [unfolded Let_def word_size] |
|
797 |
lemmas target_size = target_size_def [unfolded Let_def word_size] |
|
798 |
lemmas is_down = is_down_def [unfolded source_size target_size] |
|
799 |
lemmas is_up = is_up_def [unfolded source_size target_size] |
|
800 |
||
801 |
lemmas is_up_down = trans [OF is_up is_down [symmetric], standard] |
|
802 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
803 |
lemma down_cast_same': "uc = ucast \<Longrightarrow> is_down uc \<Longrightarrow> uc = scast" |
37660 | 804 |
apply (unfold is_down) |
805 |
apply safe |
|
806 |
apply (rule ext) |
|
807 |
apply (unfold ucast_def scast_def uint_sint) |
|
808 |
apply (rule word_ubin.norm_eq_iff [THEN iffD1]) |
|
809 |
apply simp |
|
810 |
done |
|
811 |
||
812 |
lemma word_rev_tf': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
813 |
"r = to_bl (of_bl bl) \<Longrightarrow> r = rev (takefill False (length r) (rev bl))" |
37660 | 814 |
unfolding of_bl_def uint_bl |
815 |
by (clarsimp simp add: bl_bin_bl_rtf word_ubin.eq_norm word_size) |
|
816 |
||
817 |
lemmas word_rev_tf = refl [THEN word_rev_tf', unfolded word_bl.Rep', standard] |
|
818 |
||
819 |
lemmas word_rep_drop = word_rev_tf [simplified takefill_alt, |
|
820 |
simplified, simplified rev_take, simplified] |
|
821 |
||
822 |
lemma to_bl_ucast: |
|
823 |
"to_bl (ucast (w::'b::len0 word) ::'a::len0 word) = |
|
824 |
replicate (len_of TYPE('a) - len_of TYPE('b)) False @ |
|
825 |
drop (len_of TYPE('b) - len_of TYPE('a)) (to_bl w)" |
|
826 |
apply (unfold ucast_bl) |
|
827 |
apply (rule trans) |
|
828 |
apply (rule word_rep_drop) |
|
829 |
apply simp |
|
830 |
done |
|
831 |
||
832 |
lemma ucast_up_app': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
833 |
"uc = ucast \<Longrightarrow> source_size uc + n = target_size uc \<Longrightarrow> |
37660 | 834 |
to_bl (uc w) = replicate n False @ (to_bl w)" |
835 |
by (auto simp add : source_size target_size to_bl_ucast) |
|
836 |
||
837 |
lemma ucast_down_drop': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
838 |
"uc = ucast \<Longrightarrow> source_size uc = target_size uc + n \<Longrightarrow> |
37660 | 839 |
to_bl (uc w) = drop n (to_bl w)" |
840 |
by (auto simp add : source_size target_size to_bl_ucast) |
|
841 |
||
842 |
lemma scast_down_drop': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
843 |
"sc = scast \<Longrightarrow> source_size sc = target_size sc + n \<Longrightarrow> |
37660 | 844 |
to_bl (sc w) = drop n (to_bl w)" |
845 |
apply (subgoal_tac "sc = ucast") |
|
846 |
apply safe |
|
847 |
apply simp |
|
848 |
apply (erule refl [THEN ucast_down_drop']) |
|
849 |
apply (rule refl [THEN down_cast_same', symmetric]) |
|
850 |
apply (simp add : source_size target_size is_down) |
|
851 |
done |
|
852 |
||
853 |
lemma sint_up_scast': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
854 |
"sc = scast \<Longrightarrow> is_up sc \<Longrightarrow> sint (sc w) = sint w" |
37660 | 855 |
apply (unfold is_up) |
856 |
apply safe |
|
857 |
apply (simp add: scast_def word_sbin.eq_norm) |
|
858 |
apply (rule box_equals) |
|
859 |
prefer 3 |
|
860 |
apply (rule word_sbin.norm_Rep) |
|
861 |
apply (rule sbintrunc_sbintrunc_l) |
|
862 |
defer |
|
863 |
apply (subst word_sbin.norm_Rep) |
|
864 |
apply (rule refl) |
|
865 |
apply simp |
|
866 |
done |
|
867 |
||
868 |
lemma uint_up_ucast': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
869 |
"uc = ucast \<Longrightarrow> is_up uc \<Longrightarrow> uint (uc w) = uint w" |
37660 | 870 |
apply (unfold is_up) |
871 |
apply safe |
|
872 |
apply (rule bin_eqI) |
|
873 |
apply (fold word_test_bit_def) |
|
874 |
apply (auto simp add: nth_ucast) |
|
875 |
apply (auto simp add: test_bit_bin) |
|
876 |
done |
|
877 |
||
878 |
lemmas down_cast_same = refl [THEN down_cast_same'] |
|
879 |
lemmas ucast_up_app = refl [THEN ucast_up_app'] |
|
880 |
lemmas ucast_down_drop = refl [THEN ucast_down_drop'] |
|
881 |
lemmas scast_down_drop = refl [THEN scast_down_drop'] |
|
882 |
lemmas uint_up_ucast = refl [THEN uint_up_ucast'] |
|
883 |
lemmas sint_up_scast = refl [THEN sint_up_scast'] |
|
884 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
885 |
lemma ucast_up_ucast': "uc = ucast \<Longrightarrow> is_up uc \<Longrightarrow> ucast (uc w) = ucast w" |
37660 | 886 |
apply (simp (no_asm) add: ucast_def) |
887 |
apply (clarsimp simp add: uint_up_ucast) |
|
888 |
done |
|
889 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
890 |
lemma scast_up_scast': "sc = scast \<Longrightarrow> is_up sc \<Longrightarrow> scast (sc w) = scast w" |
37660 | 891 |
apply (simp (no_asm) add: scast_def) |
892 |
apply (clarsimp simp add: sint_up_scast) |
|
893 |
done |
|
894 |
||
895 |
lemma ucast_of_bl_up': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
896 |
"w = of_bl bl \<Longrightarrow> size bl <= size w \<Longrightarrow> ucast w = of_bl bl" |
37660 | 897 |
by (auto simp add : nth_ucast word_size test_bit_of_bl intro!: word_eqI) |
898 |
||
899 |
lemmas ucast_up_ucast = refl [THEN ucast_up_ucast'] |
|
900 |
lemmas scast_up_scast = refl [THEN scast_up_scast'] |
|
901 |
lemmas ucast_of_bl_up = refl [THEN ucast_of_bl_up'] |
|
902 |
||
903 |
lemmas ucast_up_ucast_id = trans [OF ucast_up_ucast ucast_id] |
|
904 |
lemmas scast_up_scast_id = trans [OF scast_up_scast scast_id] |
|
905 |
||
906 |
lemmas isduu = is_up_down [where c = "ucast", THEN iffD2] |
|
907 |
lemmas isdus = is_up_down [where c = "scast", THEN iffD2] |
|
908 |
lemmas ucast_down_ucast_id = isduu [THEN ucast_up_ucast_id] |
|
909 |
lemmas scast_down_scast_id = isdus [THEN ucast_up_ucast_id] |
|
910 |
||
911 |
lemma up_ucast_surj: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
912 |
"is_up (ucast :: 'b::len0 word => 'a::len0 word) \<Longrightarrow> |
37660 | 913 |
surj (ucast :: 'a word => 'b word)" |
914 |
by (rule surjI, erule ucast_up_ucast_id) |
|
915 |
||
916 |
lemma up_scast_surj: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
917 |
"is_up (scast :: 'b::len word => 'a::len word) \<Longrightarrow> |
37660 | 918 |
surj (scast :: 'a word => 'b word)" |
919 |
by (rule surjI, erule scast_up_scast_id) |
|
920 |
||
921 |
lemma down_scast_inj: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
922 |
"is_down (scast :: 'b::len word => 'a::len word) \<Longrightarrow> |
37660 | 923 |
inj_on (ucast :: 'a word => 'b word) A" |
924 |
by (rule inj_on_inverseI, erule scast_down_scast_id) |
|
925 |
||
926 |
lemma down_ucast_inj: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
927 |
"is_down (ucast :: 'b::len0 word => 'a::len0 word) \<Longrightarrow> |
37660 | 928 |
inj_on (ucast :: 'a word => 'b word) A" |
929 |
by (rule inj_on_inverseI, erule ucast_down_ucast_id) |
|
930 |
||
931 |
lemma of_bl_append_same: "of_bl (X @ to_bl w) = w" |
|
932 |
by (rule word_bl.Rep_eqD) (simp add: word_rep_drop) |
|
933 |
||
934 |
lemma ucast_down_no': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
935 |
"uc = ucast \<Longrightarrow> is_down uc \<Longrightarrow> uc (number_of bin) = number_of bin" |
37660 | 936 |
apply (unfold word_number_of_def is_down) |
937 |
apply (clarsimp simp add: ucast_def word_ubin.eq_norm) |
|
938 |
apply (rule word_ubin.norm_eq_iff [THEN iffD1]) |
|
939 |
apply (erule bintrunc_bintrunc_ge) |
|
940 |
done |
|
941 |
||
942 |
lemmas ucast_down_no = ucast_down_no' [OF refl] |
|
943 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
944 |
lemma ucast_down_bl': "uc = ucast \<Longrightarrow> is_down uc \<Longrightarrow> uc (of_bl bl) = of_bl bl" |
37660 | 945 |
unfolding of_bl_no by clarify (erule ucast_down_no) |
946 |
||
947 |
lemmas ucast_down_bl = ucast_down_bl' [OF refl] |
|
948 |
||
949 |
lemmas slice_def' = slice_def [unfolded word_size] |
|
950 |
lemmas test_bit_def' = word_test_bit_def [THEN fun_cong] |
|
951 |
||
952 |
lemmas word_log_defs = word_and_def word_or_def word_xor_def word_not_def |
|
953 |
lemmas word_log_bin_defs = word_log_defs |
|
954 |
||
955 |
text {* Executable equality *} |
|
956 |
||
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38527
diff
changeset
|
957 |
instantiation word :: (len0) equal |
24333 | 958 |
begin |
959 |
||
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38527
diff
changeset
|
960 |
definition equal_word :: "'a word \<Rightarrow> 'a word \<Rightarrow> bool" where |
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38527
diff
changeset
|
961 |
"equal_word k l \<longleftrightarrow> HOL.equal (uint k) (uint l)" |
37660 | 962 |
|
963 |
instance proof |
|
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38527
diff
changeset
|
964 |
qed (simp add: equal equal_word_def) |
37660 | 965 |
|
966 |
end |
|
967 |
||
968 |
||
969 |
subsection {* Word Arithmetic *} |
|
970 |
||
971 |
lemma word_less_alt: "(a < b) = (uint a < uint b)" |
|
972 |
unfolding word_less_def word_le_def |
|
973 |
by (auto simp del: word_uint.Rep_inject |
|
974 |
simp: word_uint.Rep_inject [symmetric]) |
|
975 |
||
976 |
lemma signed_linorder: "class.linorder word_sle word_sless" |
|
977 |
proof |
|
978 |
qed (unfold word_sle_def word_sless_def, auto) |
|
979 |
||
980 |
interpretation signed: linorder "word_sle" "word_sless" |
|
981 |
by (rule signed_linorder) |
|
982 |
||
983 |
lemmas word_arith_wis = |
|
984 |
word_add_def word_mult_def word_minus_def |
|
985 |
word_succ_def word_pred_def word_0_wi word_1_wi |
|
986 |
||
987 |
lemma udvdI: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
988 |
"0 \<le> n \<Longrightarrow> uint b = n * uint a \<Longrightarrow> a udvd b" |
37660 | 989 |
by (auto simp: udvd_def) |
990 |
||
991 |
lemmas word_div_no [simp] = |
|
992 |
word_div_def [of "number_of a" "number_of b", standard] |
|
993 |
||
994 |
lemmas word_mod_no [simp] = |
|
995 |
word_mod_def [of "number_of a" "number_of b", standard] |
|
996 |
||
997 |
lemmas word_less_no [simp] = |
|
998 |
word_less_def [of "number_of a" "number_of b", standard] |
|
999 |
||
1000 |
lemmas word_le_no [simp] = |
|
1001 |
word_le_def [of "number_of a" "number_of b", standard] |
|
1002 |
||
1003 |
lemmas word_sless_no [simp] = |
|
1004 |
word_sless_def [of "number_of a" "number_of b", standard] |
|
1005 |
||
1006 |
lemmas word_sle_no [simp] = |
|
1007 |
word_sle_def [of "number_of a" "number_of b", standard] |
|
1008 |
||
1009 |
(* following two are available in class number_ring, |
|
1010 |
but convenient to have them here here; |
|
1011 |
note - the number_ring versions, numeral_0_eq_0 and numeral_1_eq_1 |
|
1012 |
are in the default simpset, so to use the automatic simplifications for |
|
1013 |
(eg) sint (number_of bin) on sint 1, must do |
|
1014 |
(simp add: word_1_no del: numeral_1_eq_1) |
|
1015 |
*) |
|
1016 |
lemmas word_0_wi_Pls = word_0_wi [folded Pls_def] |
|
1017 |
lemmas word_0_no = word_0_wi_Pls [folded word_no_wi] |
|
1018 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1019 |
lemma int_one_bin: "(1 :: int) = (Int.Pls BIT 1)" |
37660 | 1020 |
unfolding Pls_def Bit_def by auto |
1021 |
||
1022 |
lemma word_1_no: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1023 |
"(1 :: 'a :: len0 word) = number_of (Int.Pls BIT 1)" |
37660 | 1024 |
unfolding word_1_wi word_number_of_def int_one_bin by auto |
1025 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1026 |
lemma word_m1_wi: "-1 = word_of_int -1" |
37660 | 1027 |
by (rule word_number_of_alt) |
1028 |
||
1029 |
lemma word_m1_wi_Min: "-1 = word_of_int Int.Min" |
|
1030 |
by (simp add: word_m1_wi number_of_eq) |
|
1031 |
||
1032 |
lemma word_0_bl: "of_bl [] = 0" |
|
1033 |
unfolding word_0_wi of_bl_def by (simp add : Pls_def) |
|
1034 |
||
1035 |
lemma word_1_bl: "of_bl [True] = 1" |
|
1036 |
unfolding word_1_wi of_bl_def |
|
1037 |
by (simp add : bl_to_bin_def Bit_def Pls_def) |
|
1038 |
||
1039 |
lemma uint_eq_0 [simp] : "(uint 0 = 0)" |
|
1040 |
unfolding word_0_wi |
|
1041 |
by (simp add: word_ubin.eq_norm Pls_def [symmetric]) |
|
1042 |
||
1043 |
lemma of_bl_0 [simp] : "of_bl (replicate n False) = 0" |
|
1044 |
by (simp add : word_0_wi of_bl_def bl_to_bin_rep_False Pls_def) |
|
1045 |
||
1046 |
lemma to_bl_0: |
|
1047 |
"to_bl (0::'a::len0 word) = replicate (len_of TYPE('a)) False" |
|
1048 |
unfolding uint_bl |
|
1049 |
by (simp add : word_size bin_to_bl_Pls Pls_def [symmetric]) |
|
1050 |
||
1051 |
lemma uint_0_iff: "(uint x = 0) = (x = 0)" |
|
1052 |
by (auto intro!: word_uint.Rep_eqD) |
|
1053 |
||
1054 |
lemma unat_0_iff: "(unat x = 0) = (x = 0)" |
|
1055 |
unfolding unat_def by (auto simp add : nat_eq_iff uint_0_iff) |
|
1056 |
||
1057 |
lemma unat_0 [simp]: "unat 0 = 0" |
|
1058 |
unfolding unat_def by auto |
|
1059 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1060 |
lemma size_0_same': "size w = 0 \<Longrightarrow> w = (v :: 'a :: len0 word)" |
37660 | 1061 |
apply (unfold word_size) |
1062 |
apply (rule box_equals) |
|
1063 |
defer |
|
1064 |
apply (rule word_uint.Rep_inverse)+ |
|
1065 |
apply (rule word_ubin.norm_eq_iff [THEN iffD1]) |
|
1066 |
apply simp |
|
1067 |
done |
|
1068 |
||
1069 |
lemmas size_0_same = size_0_same' [folded word_size] |
|
1070 |
||
1071 |
lemmas unat_eq_0 = unat_0_iff |
|
1072 |
lemmas unat_eq_zero = unat_0_iff |
|
1073 |
||
1074 |
lemma unat_gt_0: "(0 < unat x) = (x ~= 0)" |
|
1075 |
by (auto simp: unat_0_iff [symmetric]) |
|
1076 |
||
1077 |
lemma ucast_0 [simp] : "ucast 0 = 0" |
|
1078 |
unfolding ucast_def |
|
1079 |
by simp (simp add: word_0_wi) |
|
1080 |
||
1081 |
lemma sint_0 [simp] : "sint 0 = 0" |
|
1082 |
unfolding sint_uint |
|
1083 |
by (simp add: Pls_def [symmetric]) |
|
1084 |
||
1085 |
lemma scast_0 [simp] : "scast 0 = 0" |
|
1086 |
apply (unfold scast_def) |
|
1087 |
apply simp |
|
1088 |
apply (simp add: word_0_wi) |
|
1089 |
done |
|
1090 |
||
1091 |
lemma sint_n1 [simp] : "sint -1 = -1" |
|
1092 |
apply (unfold word_m1_wi_Min) |
|
1093 |
apply (simp add: word_sbin.eq_norm) |
|
1094 |
apply (unfold Min_def number_of_eq) |
|
1095 |
apply simp |
|
1096 |
done |
|
1097 |
||
1098 |
lemma scast_n1 [simp] : "scast -1 = -1" |
|
1099 |
apply (unfold scast_def sint_n1) |
|
1100 |
apply (unfold word_number_of_alt) |
|
1101 |
apply (rule refl) |
|
1102 |
done |
|
1103 |
||
1104 |
lemma uint_1 [simp] : "uint (1 :: 'a :: len word) = 1" |
|
1105 |
unfolding word_1_wi |
|
1106 |
by (simp add: word_ubin.eq_norm int_one_bin bintrunc_minus_simps) |
|
1107 |
||
1108 |
lemma unat_1 [simp] : "unat (1 :: 'a :: len word) = 1" |
|
1109 |
by (unfold unat_def uint_1) auto |
|
1110 |
||
1111 |
lemma ucast_1 [simp] : "ucast (1 :: 'a :: len word) = 1" |
|
1112 |
unfolding ucast_def word_1_wi |
|
1113 |
by (simp add: word_ubin.eq_norm int_one_bin bintrunc_minus_simps) |
|
1114 |
||
1115 |
(* abstraction preserves the operations |
|
1116 |
(the definitions tell this for bins in range uint) *) |
|
1117 |
||
1118 |
lemmas arths = |
|
1119 |
bintr_ariths [THEN word_ubin.norm_eq_iff [THEN iffD1], |
|
1120 |
folded word_ubin.eq_norm, standard] |
|
1121 |
||
1122 |
lemma wi_homs: |
|
1123 |
shows |
|
1124 |
wi_hom_add: "word_of_int a + word_of_int b = word_of_int (a + b)" and |
|
1125 |
wi_hom_mult: "word_of_int a * word_of_int b = word_of_int (a * b)" and |
|
1126 |
wi_hom_neg: "- word_of_int a = word_of_int (- a)" and |
|
1127 |
wi_hom_succ: "word_succ (word_of_int a) = word_of_int (Int.succ a)" and |
|
1128 |
wi_hom_pred: "word_pred (word_of_int a) = word_of_int (Int.pred a)" |
|
1129 |
by (auto simp: word_arith_wis arths) |
|
1130 |
||
1131 |
lemmas wi_hom_syms = wi_homs [symmetric] |
|
1132 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1133 |
lemma word_sub_def: "a - b = a + - (b :: 'a :: len0 word)" |
37887 | 1134 |
unfolding word_sub_wi diff_minus |
37660 | 1135 |
by (simp only : word_uint.Rep_inverse wi_hom_syms) |
1136 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1137 |
lemmas word_diff_minus = word_sub_def [standard] |
37660 | 1138 |
|
1139 |
lemma word_of_int_sub_hom: |
|
1140 |
"(word_of_int a) - word_of_int b = word_of_int (a - b)" |
|
37887 | 1141 |
unfolding word_sub_def diff_minus by (simp only : wi_homs) |
37660 | 1142 |
|
1143 |
lemmas new_word_of_int_homs = |
|
1144 |
word_of_int_sub_hom wi_homs word_0_wi word_1_wi |
|
1145 |
||
1146 |
lemmas new_word_of_int_hom_syms = new_word_of_int_homs [symmetric, standard] |
|
1147 |
||
1148 |
lemmas word_of_int_hom_syms = |
|
1149 |
new_word_of_int_hom_syms [unfolded succ_def pred_def] |
|
1150 |
||
1151 |
lemmas word_of_int_homs = |
|
1152 |
new_word_of_int_homs [unfolded succ_def pred_def] |
|
1153 |
||
1154 |
lemmas word_of_int_add_hom = word_of_int_homs (2) |
|
1155 |
lemmas word_of_int_mult_hom = word_of_int_homs (3) |
|
1156 |
lemmas word_of_int_minus_hom = word_of_int_homs (4) |
|
1157 |
lemmas word_of_int_succ_hom = word_of_int_homs (5) |
|
1158 |
lemmas word_of_int_pred_hom = word_of_int_homs (6) |
|
1159 |
lemmas word_of_int_0_hom = word_of_int_homs (7) |
|
1160 |
lemmas word_of_int_1_hom = word_of_int_homs (8) |
|
1161 |
||
1162 |
(* now, to get the weaker results analogous to word_div/mod_def *) |
|
1163 |
||
1164 |
lemmas word_arith_alts = |
|
1165 |
word_sub_wi [unfolded succ_def pred_def, standard] |
|
1166 |
word_arith_wis [unfolded succ_def pred_def, standard] |
|
1167 |
||
1168 |
lemmas word_sub_alt = word_arith_alts (1) |
|
1169 |
lemmas word_add_alt = word_arith_alts (2) |
|
1170 |
lemmas word_mult_alt = word_arith_alts (3) |
|
1171 |
lemmas word_minus_alt = word_arith_alts (4) |
|
1172 |
lemmas word_succ_alt = word_arith_alts (5) |
|
1173 |
lemmas word_pred_alt = word_arith_alts (6) |
|
1174 |
lemmas word_0_alt = word_arith_alts (7) |
|
1175 |
lemmas word_1_alt = word_arith_alts (8) |
|
1176 |
||
1177 |
subsection "Transferring goals from words to ints" |
|
1178 |
||
1179 |
lemma word_ths: |
|
1180 |
shows |
|
1181 |
word_succ_p1: "word_succ a = a + 1" and |
|
1182 |
word_pred_m1: "word_pred a = a - 1" and |
|
1183 |
word_pred_succ: "word_pred (word_succ a) = a" and |
|
1184 |
word_succ_pred: "word_succ (word_pred a) = a" and |
|
1185 |
word_mult_succ: "word_succ a * b = b + a * b" |
|
1186 |
by (rule word_uint.Abs_cases [of b], |
|
1187 |
rule word_uint.Abs_cases [of a], |
|
1188 |
simp add: pred_def succ_def add_commute mult_commute |
|
1189 |
ring_distribs new_word_of_int_homs)+ |
|
1190 |
||
1191 |
lemmas uint_cong = arg_cong [where f = uint] |
|
1192 |
||
1193 |
lemmas uint_word_ariths = |
|
1194 |
word_arith_alts [THEN trans [OF uint_cong int_word_uint], standard] |
|
1195 |
||
1196 |
lemmas uint_word_arith_bintrs = uint_word_ariths [folded bintrunc_mod2p] |
|
1197 |
||
1198 |
(* similar expressions for sint (arith operations) *) |
|
1199 |
lemmas sint_word_ariths = uint_word_arith_bintrs |
|
1200 |
[THEN uint_sint [symmetric, THEN trans], |
|
1201 |
unfolded uint_sint bintr_arith1s bintr_ariths |
|
1202 |
len_gt_0 [THEN bin_sbin_eq_iff'] word_sbin.norm_Rep, standard] |
|
1203 |
||
1204 |
lemmas uint_div_alt = word_div_def |
|
1205 |
[THEN trans [OF uint_cong int_word_uint], standard] |
|
1206 |
lemmas uint_mod_alt = word_mod_def |
|
1207 |
[THEN trans [OF uint_cong int_word_uint], standard] |
|
1208 |
||
1209 |
lemma word_pred_0_n1: "word_pred 0 = word_of_int -1" |
|
1210 |
unfolding word_pred_def number_of_eq |
|
1211 |
by (simp add : pred_def word_no_wi) |
|
1212 |
||
1213 |
lemma word_pred_0_Min: "word_pred 0 = word_of_int Int.Min" |
|
1214 |
by (simp add: word_pred_0_n1 number_of_eq) |
|
1215 |
||
1216 |
lemma word_m1_Min: "- 1 = word_of_int Int.Min" |
|
1217 |
unfolding Min_def by (simp only: word_of_int_hom_syms) |
|
1218 |
||
1219 |
lemma succ_pred_no [simp]: |
|
1220 |
"word_succ (number_of bin) = number_of (Int.succ bin) & |
|
1221 |
word_pred (number_of bin) = number_of (Int.pred bin)" |
|
1222 |
unfolding word_number_of_def by (simp add : new_word_of_int_homs) |
|
1223 |
||
1224 |
lemma word_sp_01 [simp] : |
|
1225 |
"word_succ -1 = 0 & word_succ 0 = 1 & word_pred 0 = -1 & word_pred 1 = 0" |
|
1226 |
by (unfold word_0_no word_1_no) auto |
|
1227 |
||
1228 |
(* alternative approach to lifting arithmetic equalities *) |
|
1229 |
lemma word_of_int_Ex: |
|
1230 |
"\<exists>y. x = word_of_int y" |
|
1231 |
by (rule_tac x="uint x" in exI) simp |
|
1232 |
||
1233 |
lemma word_arith_eqs: |
|
1234 |
fixes a :: "'a::len0 word" |
|
1235 |
fixes b :: "'a::len0 word" |
|
1236 |
shows |
|
1237 |
word_add_0: "0 + a = a" and |
|
1238 |
word_add_0_right: "a + 0 = a" and |
|
1239 |
word_mult_1: "1 * a = a" and |
|
1240 |
word_mult_1_right: "a * 1 = a" and |
|
1241 |
word_add_commute: "a + b = b + a" and |
|
1242 |
word_add_assoc: "a + b + c = a + (b + c)" and |
|
1243 |
word_add_left_commute: "a + (b + c) = b + (a + c)" and |
|
1244 |
word_mult_commute: "a * b = b * a" and |
|
1245 |
word_mult_assoc: "a * b * c = a * (b * c)" and |
|
1246 |
word_mult_left_commute: "a * (b * c) = b * (a * c)" and |
|
1247 |
word_left_distrib: "(a + b) * c = a * c + b * c" and |
|
1248 |
word_right_distrib: "a * (b + c) = a * b + a * c" and |
|
1249 |
word_left_minus: "- a + a = 0" and |
|
1250 |
word_diff_0_right: "a - 0 = a" and |
|
1251 |
word_diff_self: "a - a = 0" |
|
1252 |
using word_of_int_Ex [of a] |
|
1253 |
word_of_int_Ex [of b] |
|
1254 |
word_of_int_Ex [of c] |
|
1255 |
by (auto simp: word_of_int_hom_syms [symmetric] |
|
1256 |
zadd_0_right add_commute add_assoc add_left_commute |
|
1257 |
mult_commute mult_assoc mult_left_commute |
|
1258 |
left_distrib right_distrib) |
|
1259 |
||
1260 |
lemmas word_add_ac = word_add_commute word_add_assoc word_add_left_commute |
|
1261 |
lemmas word_mult_ac = word_mult_commute word_mult_assoc word_mult_left_commute |
|
1262 |
||
1263 |
lemmas word_plus_ac0 = word_add_0 word_add_0_right word_add_ac |
|
1264 |
lemmas word_times_ac1 = word_mult_1 word_mult_1_right word_mult_ac |
|
1265 |
||
1266 |
||
1267 |
subsection "Order on fixed-length words" |
|
1268 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1269 |
lemma word_order_trans: "x <= y \<Longrightarrow> y <= z \<Longrightarrow> x <= (z :: 'a :: len0 word)" |
37660 | 1270 |
unfolding word_le_def by auto |
1271 |
||
1272 |
lemma word_order_refl: "z <= (z :: 'a :: len0 word)" |
|
1273 |
unfolding word_le_def by auto |
|
1274 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1275 |
lemma word_order_antisym: "x <= y \<Longrightarrow> y <= x \<Longrightarrow> x = (y :: 'a :: len0 word)" |
37660 | 1276 |
unfolding word_le_def by (auto intro!: word_uint.Rep_eqD) |
1277 |
||
1278 |
lemma word_order_linear: |
|
1279 |
"y <= x | x <= (y :: 'a :: len0 word)" |
|
1280 |
unfolding word_le_def by auto |
|
1281 |
||
1282 |
lemma word_zero_le [simp] : |
|
1283 |
"0 <= (y :: 'a :: len0 word)" |
|
1284 |
unfolding word_le_def by auto |
|
1285 |
||
1286 |
instance word :: (len0) semigroup_add |
|
1287 |
by intro_classes (simp add: word_add_assoc) |
|
1288 |
||
1289 |
instance word :: (len0) linorder |
|
1290 |
by intro_classes (auto simp: word_less_def word_le_def) |
|
1291 |
||
1292 |
instance word :: (len0) ring |
|
1293 |
by intro_classes |
|
1294 |
(auto simp: word_arith_eqs word_diff_minus |
|
1295 |
word_diff_self [unfolded word_diff_minus]) |
|
1296 |
||
1297 |
lemma word_m1_ge [simp] : "word_pred 0 >= y" |
|
1298 |
unfolding word_le_def |
|
1299 |
by (simp only : word_pred_0_n1 word_uint.eq_norm m1mod2k) auto |
|
1300 |
||
1301 |
lemmas word_n1_ge [simp] = word_m1_ge [simplified word_sp_01] |
|
1302 |
||
1303 |
lemmas word_not_simps [simp] = |
|
1304 |
word_zero_le [THEN leD] word_m1_ge [THEN leD] word_n1_ge [THEN leD] |
|
1305 |
||
1306 |
lemma word_gt_0: "0 < y = (0 ~= (y :: 'a :: len0 word))" |
|
1307 |
unfolding word_less_def by auto |
|
1308 |
||
1309 |
lemmas word_gt_0_no [simp] = word_gt_0 [of "number_of y", standard] |
|
1310 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1311 |
lemma word_sless_alt: "(a <s b) = (sint a < sint b)" |
37660 | 1312 |
unfolding word_sle_def word_sless_def |
1313 |
by (auto simp add: less_le) |
|
1314 |
||
1315 |
lemma word_le_nat_alt: "(a <= b) = (unat a <= unat b)" |
|
1316 |
unfolding unat_def word_le_def |
|
1317 |
by (rule nat_le_eq_zle [symmetric]) simp |
|
1318 |
||
1319 |
lemma word_less_nat_alt: "(a < b) = (unat a < unat b)" |
|
1320 |
unfolding unat_def word_less_alt |
|
1321 |
by (rule nat_less_eq_zless [symmetric]) simp |
|
1322 |
||
1323 |
lemma wi_less: |
|
1324 |
"(word_of_int n < (word_of_int m :: 'a :: len0 word)) = |
|
1325 |
(n mod 2 ^ len_of TYPE('a) < m mod 2 ^ len_of TYPE('a))" |
|
1326 |
unfolding word_less_alt by (simp add: word_uint.eq_norm) |
|
1327 |
||
1328 |
lemma wi_le: |
|
1329 |
"(word_of_int n <= (word_of_int m :: 'a :: len0 word)) = |
|
1330 |
(n mod 2 ^ len_of TYPE('a) <= m mod 2 ^ len_of TYPE('a))" |
|
1331 |
unfolding word_le_def by (simp add: word_uint.eq_norm) |
|
1332 |
||
1333 |
lemma udvd_nat_alt: "a udvd b = (EX n>=0. unat b = n * unat a)" |
|
1334 |
apply (unfold udvd_def) |
|
1335 |
apply safe |
|
1336 |
apply (simp add: unat_def nat_mult_distrib) |
|
1337 |
apply (simp add: uint_nat int_mult) |
|
1338 |
apply (rule exI) |
|
1339 |
apply safe |
|
1340 |
prefer 2 |
|
1341 |
apply (erule notE) |
|
1342 |
apply (rule refl) |
|
1343 |
apply force |
|
1344 |
done |
|
1345 |
||
1346 |
lemma udvd_iff_dvd: "x udvd y <-> unat x dvd unat y" |
|
1347 |
unfolding dvd_def udvd_nat_alt by force |
|
1348 |
||
1349 |
lemmas unat_mono = word_less_nat_alt [THEN iffD1, standard] |
|
1350 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1351 |
lemma word_zero_neq_one: "0 < len_of TYPE ('a :: len0) \<Longrightarrow> (0 :: 'a word) ~= 1"; |
37660 | 1352 |
unfolding word_arith_wis |
1353 |
by (auto simp add: word_ubin.norm_eq_iff [symmetric] gr0_conv_Suc) |
|
1354 |
||
1355 |
lemmas lenw1_zero_neq_one = len_gt_0 [THEN word_zero_neq_one] |
|
1356 |
||
1357 |
lemma no_no [simp] : "number_of (number_of b) = number_of b" |
|
1358 |
by (simp add: number_of_eq) |
|
1359 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1360 |
lemma unat_minus_one: "x ~= 0 \<Longrightarrow> unat (x - 1) = unat x - 1" |
37660 | 1361 |
apply (unfold unat_def) |
1362 |
apply (simp only: int_word_uint word_arith_alts rdmods) |
|
1363 |
apply (subgoal_tac "uint x >= 1") |
|
1364 |
prefer 2 |
|
1365 |
apply (drule contrapos_nn) |
|
1366 |
apply (erule word_uint.Rep_inverse' [symmetric]) |
|
1367 |
apply (insert uint_ge_0 [of x])[1] |
|
1368 |
apply arith |
|
1369 |
apply (rule box_equals) |
|
1370 |
apply (rule nat_diff_distrib) |
|
1371 |
prefer 2 |
|
1372 |
apply assumption |
|
1373 |
apply simp |
|
1374 |
apply (subst mod_pos_pos_trivial) |
|
1375 |
apply arith |
|
1376 |
apply (insert uint_lt2p [of x])[1] |
|
1377 |
apply arith |
|
1378 |
apply (rule refl) |
|
1379 |
apply simp |
|
1380 |
done |
|
1381 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1382 |
lemma measure_unat: "p ~= 0 \<Longrightarrow> unat (p - 1) < unat p" |
37660 | 1383 |
by (simp add: unat_minus_one) (simp add: unat_0_iff [symmetric]) |
1384 |
||
1385 |
lemmas uint_add_ge0 [simp] = |
|
1386 |
add_nonneg_nonneg [OF uint_ge_0 uint_ge_0, standard] |
|
1387 |
lemmas uint_mult_ge0 [simp] = |
|
1388 |
mult_nonneg_nonneg [OF uint_ge_0 uint_ge_0, standard] |
|
1389 |
||
1390 |
lemma uint_sub_lt2p [simp]: |
|
1391 |
"uint (x :: 'a :: len0 word) - uint (y :: 'b :: len0 word) < |
|
1392 |
2 ^ len_of TYPE('a)" |
|
1393 |
using uint_ge_0 [of y] uint_lt2p [of x] by arith |
|
1394 |
||
1395 |
||
1396 |
subsection "Conditions for the addition (etc) of two words to overflow" |
|
1397 |
||
1398 |
lemma uint_add_lem: |
|
1399 |
"(uint x + uint y < 2 ^ len_of TYPE('a)) = |
|
1400 |
(uint (x + y :: 'a :: len0 word) = uint x + uint y)" |
|
1401 |
by (unfold uint_word_ariths) (auto intro!: trans [OF _ int_mod_lem]) |
|
1402 |
||
1403 |
lemma uint_mult_lem: |
|
1404 |
"(uint x * uint y < 2 ^ len_of TYPE('a)) = |
|
1405 |
(uint (x * y :: 'a :: len0 word) = uint x * uint y)" |
|
1406 |
by (unfold uint_word_ariths) (auto intro!: trans [OF _ int_mod_lem]) |
|
1407 |
||
1408 |
lemma uint_sub_lem: |
|
1409 |
"(uint x >= uint y) = (uint (x - y) = uint x - uint y)" |
|
1410 |
by (unfold uint_word_ariths) (auto intro!: trans [OF _ int_mod_lem]) |
|
1411 |
||
1412 |
lemma uint_add_le: "uint (x + y) <= uint x + uint y" |
|
1413 |
unfolding uint_word_ariths by (auto simp: mod_add_if_z) |
|
1414 |
||
1415 |
lemma uint_sub_ge: "uint (x - y) >= uint x - uint y" |
|
1416 |
unfolding uint_word_ariths by (auto simp: mod_sub_if_z) |
|
1417 |
||
1418 |
lemmas uint_sub_if' = |
|
1419 |
trans [OF uint_word_ariths(1) mod_sub_if_z, simplified, standard] |
|
1420 |
lemmas uint_plus_if' = |
|
1421 |
trans [OF uint_word_ariths(2) mod_add_if_z, simplified, standard] |
|
1422 |
||
1423 |
||
1424 |
subsection {* Definition of uint\_arith *} |
|
1425 |
||
1426 |
lemma word_of_int_inverse: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1427 |
"word_of_int r = a \<Longrightarrow> 0 <= r \<Longrightarrow> r < 2 ^ len_of TYPE('a) \<Longrightarrow> |
37660 | 1428 |
uint (a::'a::len0 word) = r" |
1429 |
apply (erule word_uint.Abs_inverse' [rotated]) |
|
1430 |
apply (simp add: uints_num) |
|
1431 |
done |
|
1432 |
||
1433 |
lemma uint_split: |
|
1434 |
fixes x::"'a::len0 word" |
|
1435 |
shows "P (uint x) = |
|
1436 |
(ALL i. word_of_int i = x & 0 <= i & i < 2^len_of TYPE('a) --> P i)" |
|
1437 |
apply (fold word_int_case_def) |
|
1438 |
apply (auto dest!: word_of_int_inverse simp: int_word_uint int_mod_eq' |
|
1439 |
split: word_int_split) |
|
1440 |
done |
|
1441 |
||
1442 |
lemma uint_split_asm: |
|
1443 |
fixes x::"'a::len0 word" |
|
1444 |
shows "P (uint x) = |
|
1445 |
(~(EX i. word_of_int i = x & 0 <= i & i < 2^len_of TYPE('a) & ~ P i))" |
|
1446 |
by (auto dest!: word_of_int_inverse |
|
1447 |
simp: int_word_uint int_mod_eq' |
|
1448 |
split: uint_split) |
|
1449 |
||
1450 |
lemmas uint_splits = uint_split uint_split_asm |
|
1451 |
||
1452 |
lemmas uint_arith_simps = |
|
1453 |
word_le_def word_less_alt |
|
1454 |
word_uint.Rep_inject [symmetric] |
|
1455 |
uint_sub_if' uint_plus_if' |
|
1456 |
||
1457 |
(* use this to stop, eg, 2 ^ len_of TYPE (32) being simplified *) |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1458 |
lemma power_False_cong: "False \<Longrightarrow> a ^ b = c ^ d" |
37660 | 1459 |
by auto |
1460 |
||
1461 |
(* uint_arith_tac: reduce to arithmetic on int, try to solve by arith *) |
|
1462 |
ML {* |
|
1463 |
fun uint_arith_ss_of ss = |
|
1464 |
ss addsimps @{thms uint_arith_simps} |
|
1465 |
delsimps @{thms word_uint.Rep_inject} |
|
1466 |
addsplits @{thms split_if_asm} |
|
1467 |
addcongs @{thms power_False_cong} |
|
1468 |
||
1469 |
fun uint_arith_tacs ctxt = |
|
1470 |
let |
|
1471 |
fun arith_tac' n t = |
|
1472 |
Arith_Data.verbose_arith_tac ctxt n t |
|
1473 |
handle Cooper.COOPER _ => Seq.empty; |
|
1474 |
val cs = claset_of ctxt; |
|
1475 |
val ss = simpset_of ctxt; |
|
1476 |
in |
|
1477 |
[ clarify_tac cs 1, |
|
1478 |
full_simp_tac (uint_arith_ss_of ss) 1, |
|
1479 |
ALLGOALS (full_simp_tac (HOL_ss addsplits @{thms uint_splits} |
|
1480 |
addcongs @{thms power_False_cong})), |
|
1481 |
rewrite_goals_tac @{thms word_size}, |
|
1482 |
ALLGOALS (fn n => REPEAT (resolve_tac [allI, impI] n) THEN |
|
1483 |
REPEAT (etac conjE n) THEN |
|
1484 |
REPEAT (dtac @{thm word_of_int_inverse} n |
|
1485 |
THEN atac n |
|
1486 |
THEN atac n)), |
|
1487 |
TRYALL arith_tac' ] |
|
1488 |
end |
|
1489 |
||
1490 |
fun uint_arith_tac ctxt = SELECT_GOAL (EVERY (uint_arith_tacs ctxt)) |
|
1491 |
*} |
|
1492 |
||
1493 |
method_setup uint_arith = |
|
1494 |
{* Scan.succeed (SIMPLE_METHOD' o uint_arith_tac) *} |
|
1495 |
"solving word arithmetic via integers and arith" |
|
1496 |
||
1497 |
||
1498 |
subsection "More on overflows and monotonicity" |
|
1499 |
||
1500 |
lemma no_plus_overflow_uint_size: |
|
1501 |
"((x :: 'a :: len0 word) <= x + y) = (uint x + uint y < 2 ^ size x)" |
|
1502 |
unfolding word_size by uint_arith |
|
1503 |
||
1504 |
lemmas no_olen_add = no_plus_overflow_uint_size [unfolded word_size] |
|
1505 |
||
1506 |
lemma no_ulen_sub: "((x :: 'a :: len0 word) >= x - y) = (uint y <= uint x)" |
|
1507 |
by uint_arith |
|
1508 |
||
1509 |
lemma no_olen_add': |
|
1510 |
fixes x :: "'a::len0 word" |
|
1511 |
shows "(x \<le> y + x) = (uint y + uint x < 2 ^ len_of TYPE('a))" |
|
1512 |
by (simp add: word_add_ac add_ac no_olen_add) |
|
1513 |
||
1514 |
lemmas olen_add_eqv = trans [OF no_olen_add no_olen_add' [symmetric], standard] |
|
1515 |
||
1516 |
lemmas uint_plus_simple_iff = trans [OF no_olen_add uint_add_lem, standard] |
|
1517 |
lemmas uint_plus_simple = uint_plus_simple_iff [THEN iffD1, standard] |
|
1518 |
lemmas uint_minus_simple_iff = trans [OF no_ulen_sub uint_sub_lem, standard] |
|
1519 |
lemmas uint_minus_simple_alt = uint_sub_lem [folded word_le_def] |
|
1520 |
lemmas word_sub_le_iff = no_ulen_sub [folded word_le_def] |
|
1521 |
lemmas word_sub_le = word_sub_le_iff [THEN iffD2, standard] |
|
1522 |
||
1523 |
lemma word_less_sub1: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1524 |
"(x :: 'a :: len word) ~= 0 \<Longrightarrow> (1 < x) = (0 < x - 1)" |
37660 | 1525 |
by uint_arith |
1526 |
||
1527 |
lemma word_le_sub1: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1528 |
"(x :: 'a :: len word) ~= 0 \<Longrightarrow> (1 <= x) = (0 <= x - 1)" |
37660 | 1529 |
by uint_arith |
1530 |
||
1531 |
lemma sub_wrap_lt: |
|
1532 |
"((x :: 'a :: len0 word) < x - z) = (x < z)" |
|
1533 |
by uint_arith |
|
1534 |
||
1535 |
lemma sub_wrap: |
|
1536 |
"((x :: 'a :: len0 word) <= x - z) = (z = 0 | x < z)" |
|
1537 |
by uint_arith |
|
1538 |
||
1539 |
lemma plus_minus_not_NULL_ab: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1540 |
"(x :: 'a :: len0 word) <= ab - c \<Longrightarrow> c <= ab \<Longrightarrow> c ~= 0 \<Longrightarrow> x + c ~= 0" |
37660 | 1541 |
by uint_arith |
1542 |
||
1543 |
lemma plus_minus_no_overflow_ab: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1544 |
"(x :: 'a :: len0 word) <= ab - c \<Longrightarrow> c <= ab \<Longrightarrow> x <= x + c" |
37660 | 1545 |
by uint_arith |
1546 |
||
1547 |
lemma le_minus': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1548 |
"(a :: 'a :: len0 word) + c <= b \<Longrightarrow> a <= a + c \<Longrightarrow> c <= b - a" |
37660 | 1549 |
by uint_arith |
1550 |
||
1551 |
lemma le_plus': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1552 |
"(a :: 'a :: len0 word) <= b \<Longrightarrow> c <= b - a \<Longrightarrow> a + c <= b" |
37660 | 1553 |
by uint_arith |
1554 |
||
1555 |
lemmas le_plus = le_plus' [rotated] |
|
1556 |
||
1557 |
lemmas le_minus = leD [THEN thin_rl, THEN le_minus', standard] |
|
1558 |
||
1559 |
lemma word_plus_mono_right: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1560 |
"(y :: 'a :: len0 word) <= z \<Longrightarrow> x <= x + z \<Longrightarrow> x + y <= x + z" |
37660 | 1561 |
by uint_arith |
1562 |
||
1563 |
lemma word_less_minus_cancel: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1564 |
"y - x < z - x \<Longrightarrow> x <= z \<Longrightarrow> (y :: 'a :: len0 word) < z" |
37660 | 1565 |
by uint_arith |
1566 |
||
1567 |
lemma word_less_minus_mono_left: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1568 |
"(y :: 'a :: len0 word) < z \<Longrightarrow> x <= y \<Longrightarrow> y - x < z - x" |
37660 | 1569 |
by uint_arith |
1570 |
||
1571 |
lemma word_less_minus_mono: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1572 |
"a < c \<Longrightarrow> d < b \<Longrightarrow> a - b < a \<Longrightarrow> c - d < c |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1573 |
\<Longrightarrow> a - b < c - (d::'a::len word)" |
37660 | 1574 |
by uint_arith |
1575 |
||
1576 |
lemma word_le_minus_cancel: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1577 |
"y - x <= z - x \<Longrightarrow> x <= z \<Longrightarrow> (y :: 'a :: len0 word) <= z" |
37660 | 1578 |
by uint_arith |
1579 |
||
1580 |
lemma word_le_minus_mono_left: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1581 |
"(y :: 'a :: len0 word) <= z \<Longrightarrow> x <= y \<Longrightarrow> y - x <= z - x" |
37660 | 1582 |
by uint_arith |
1583 |
||
1584 |
lemma word_le_minus_mono: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1585 |
"a <= c \<Longrightarrow> d <= b \<Longrightarrow> a - b <= a \<Longrightarrow> c - d <= c |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1586 |
\<Longrightarrow> a - b <= c - (d::'a::len word)" |
37660 | 1587 |
by uint_arith |
1588 |
||
1589 |
lemma plus_le_left_cancel_wrap: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1590 |
"(x :: 'a :: len0 word) + y' < x \<Longrightarrow> x + y < x \<Longrightarrow> (x + y' < x + y) = (y' < y)" |
37660 | 1591 |
by uint_arith |
1592 |
||
1593 |
lemma plus_le_left_cancel_nowrap: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1594 |
"(x :: 'a :: len0 word) <= x + y' \<Longrightarrow> x <= x + y \<Longrightarrow> |
37660 | 1595 |
(x + y' < x + y) = (y' < y)" |
1596 |
by uint_arith |
|
1597 |
||
1598 |
lemma word_plus_mono_right2: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1599 |
"(a :: 'a :: len0 word) <= a + b \<Longrightarrow> c <= b \<Longrightarrow> a <= a + c" |
37660 | 1600 |
by uint_arith |
1601 |
||
1602 |
lemma word_less_add_right: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1603 |
"(x :: 'a :: len0 word) < y - z \<Longrightarrow> z <= y \<Longrightarrow> x + z < y" |
37660 | 1604 |
by uint_arith |
1605 |
||
1606 |
lemma word_less_sub_right: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1607 |
"(x :: 'a :: len0 word) < y + z \<Longrightarrow> y <= x \<Longrightarrow> x - y < z" |
37660 | 1608 |
by uint_arith |
1609 |
||
1610 |
lemma word_le_plus_either: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1611 |
"(x :: 'a :: len0 word) <= y | x <= z \<Longrightarrow> y <= y + z \<Longrightarrow> x <= y + z" |
37660 | 1612 |
by uint_arith |
1613 |
||
1614 |
lemma word_less_nowrapI: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1615 |
"(x :: 'a :: len0 word) < z - k \<Longrightarrow> k <= z \<Longrightarrow> 0 < k \<Longrightarrow> x < x + k" |
37660 | 1616 |
by uint_arith |
1617 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1618 |
lemma inc_le: "(i :: 'a :: len word) < m \<Longrightarrow> i + 1 <= m" |
37660 | 1619 |
by uint_arith |
1620 |
||
1621 |
lemma inc_i: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1622 |
"(1 :: 'a :: len word) <= i \<Longrightarrow> i < m \<Longrightarrow> 1 <= (i + 1) & i + 1 <= m" |
37660 | 1623 |
by uint_arith |
1624 |
||
1625 |
lemma udvd_incr_lem: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1626 |
"up < uq \<Longrightarrow> up = ua + n * uint K \<Longrightarrow> |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1627 |
uq = ua + n' * uint K \<Longrightarrow> up + uint K <= uq" |
37660 | 1628 |
apply clarsimp |
1629 |
apply (drule less_le_mult) |
|
1630 |
apply safe |
|
1631 |
done |
|
1632 |
||
1633 |
lemma udvd_incr': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1634 |
"p < q \<Longrightarrow> uint p = ua + n * uint K \<Longrightarrow> |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1635 |
uint q = ua + n' * uint K \<Longrightarrow> p + K <= q" |
37660 | 1636 |
apply (unfold word_less_alt word_le_def) |
1637 |
apply (drule (2) udvd_incr_lem) |
|
1638 |
apply (erule uint_add_le [THEN order_trans]) |
|
1639 |
done |
|
1640 |
||
1641 |
lemma udvd_decr': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1642 |
"p < q \<Longrightarrow> uint p = ua + n * uint K \<Longrightarrow> |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1643 |
uint q = ua + n' * uint K \<Longrightarrow> p <= q - K" |
37660 | 1644 |
apply (unfold word_less_alt word_le_def) |
1645 |
apply (drule (2) udvd_incr_lem) |
|
1646 |
apply (drule le_diff_eq [THEN iffD2]) |
|
1647 |
apply (erule order_trans) |
|
1648 |
apply (rule uint_sub_ge) |
|
1649 |
done |
|
1650 |
||
1651 |
lemmas udvd_incr_lem0 = udvd_incr_lem [where ua=0, simplified] |
|
1652 |
lemmas udvd_incr0 = udvd_incr' [where ua=0, simplified] |
|
1653 |
lemmas udvd_decr0 = udvd_decr' [where ua=0, simplified] |
|
1654 |
||
1655 |
lemma udvd_minus_le': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1656 |
"xy < k \<Longrightarrow> z udvd xy \<Longrightarrow> z udvd k \<Longrightarrow> xy <= k - z" |
37660 | 1657 |
apply (unfold udvd_def) |
1658 |
apply clarify |
|
1659 |
apply (erule (2) udvd_decr0) |
|
1660 |
done |
|
1661 |
||
1662 |
ML {* Delsimprocs Numeral_Simprocs.cancel_factors *} |
|
1663 |
||
1664 |
lemma udvd_incr2_K: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1665 |
"p < a + s \<Longrightarrow> a <= a + s \<Longrightarrow> K udvd s \<Longrightarrow> K udvd p - a \<Longrightarrow> a <= p \<Longrightarrow> |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1666 |
0 < K \<Longrightarrow> p <= p + K & p + K <= a + s" |
37660 | 1667 |
apply (unfold udvd_def) |
1668 |
apply clarify |
|
1669 |
apply (simp add: uint_arith_simps split: split_if_asm) |
|
1670 |
prefer 2 |
|
1671 |
apply (insert uint_range' [of s])[1] |
|
1672 |
apply arith |
|
1673 |
apply (drule add_commute [THEN xtr1]) |
|
1674 |
apply (simp add: diff_less_eq [symmetric]) |
|
1675 |
apply (drule less_le_mult) |
|
1676 |
apply arith |
|
1677 |
apply simp |
|
1678 |
done |
|
1679 |
||
1680 |
ML {* Addsimprocs Numeral_Simprocs.cancel_factors *} |
|
1681 |
||
1682 |
(* links with rbl operations *) |
|
1683 |
lemma word_succ_rbl: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1684 |
"to_bl w = bl \<Longrightarrow> to_bl (word_succ w) = (rev (rbl_succ (rev bl)))" |
37660 | 1685 |
apply (unfold word_succ_def) |
1686 |
apply clarify |
|
1687 |
apply (simp add: to_bl_of_bin) |
|
1688 |
apply (simp add: to_bl_def rbl_succ) |
|
1689 |
done |
|
1690 |
||
1691 |
lemma word_pred_rbl: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1692 |
"to_bl w = bl \<Longrightarrow> to_bl (word_pred w) = (rev (rbl_pred (rev bl)))" |
37660 | 1693 |
apply (unfold word_pred_def) |
1694 |
apply clarify |
|
1695 |
apply (simp add: to_bl_of_bin) |
|
1696 |
apply (simp add: to_bl_def rbl_pred) |
|
1697 |
done |
|
1698 |
||
1699 |
lemma word_add_rbl: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1700 |
"to_bl v = vbl \<Longrightarrow> to_bl w = wbl \<Longrightarrow> |
37660 | 1701 |
to_bl (v + w) = (rev (rbl_add (rev vbl) (rev wbl)))" |
1702 |
apply (unfold word_add_def) |
|
1703 |
apply clarify |
|
1704 |
apply (simp add: to_bl_of_bin) |
|
1705 |
apply (simp add: to_bl_def rbl_add) |
|
1706 |
done |
|
1707 |
||
1708 |
lemma word_mult_rbl: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1709 |
"to_bl v = vbl \<Longrightarrow> to_bl w = wbl \<Longrightarrow> |
37660 | 1710 |
to_bl (v * w) = (rev (rbl_mult (rev vbl) (rev wbl)))" |
1711 |
apply (unfold word_mult_def) |
|
1712 |
apply clarify |
|
1713 |
apply (simp add: to_bl_of_bin) |
|
1714 |
apply (simp add: to_bl_def rbl_mult) |
|
1715 |
done |
|
1716 |
||
1717 |
lemma rtb_rbl_ariths: |
|
1718 |
"rev (to_bl w) = ys \<Longrightarrow> rev (to_bl (word_succ w)) = rbl_succ ys" |
|
1719 |
"rev (to_bl w) = ys \<Longrightarrow> rev (to_bl (word_pred w)) = rbl_pred ys" |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1720 |
"rev (to_bl v) = ys \<Longrightarrow> rev (to_bl w) = xs \<Longrightarrow> rev (to_bl (v * w)) = rbl_mult ys xs" |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1721 |
"rev (to_bl v) = ys \<Longrightarrow> rev (to_bl w) = xs \<Longrightarrow> rev (to_bl (v + w)) = rbl_add ys xs" |
37660 | 1722 |
by (auto simp: rev_swap [symmetric] word_succ_rbl |
1723 |
word_pred_rbl word_mult_rbl word_add_rbl) |
|
1724 |
||
1725 |
||
1726 |
subsection "Arithmetic type class instantiations" |
|
1727 |
||
1728 |
instance word :: (len0) comm_monoid_add .. |
|
1729 |
||
1730 |
instance word :: (len0) comm_monoid_mult |
|
1731 |
apply (intro_classes) |
|
1732 |
apply (simp add: word_mult_commute) |
|
1733 |
apply (simp add: word_mult_1) |
|
1734 |
done |
|
1735 |
||
1736 |
instance word :: (len0) comm_semiring |
|
1737 |
by (intro_classes) (simp add : word_left_distrib) |
|
1738 |
||
1739 |
instance word :: (len0) ab_group_add .. |
|
1740 |
||
1741 |
instance word :: (len0) comm_ring .. |
|
1742 |
||
1743 |
instance word :: (len) comm_semiring_1 |
|
1744 |
by (intro_classes) (simp add: lenw1_zero_neq_one) |
|
1745 |
||
1746 |
instance word :: (len) comm_ring_1 .. |
|
1747 |
||
1748 |
instance word :: (len0) comm_semiring_0 .. |
|
1749 |
||
1750 |
instance word :: (len0) order .. |
|
1751 |
||
1752 |
(* note that iszero_def is only for class comm_semiring_1_cancel, |
|
1753 |
which requires word length >= 1, ie 'a :: len word *) |
|
1754 |
lemma zero_bintrunc: |
|
1755 |
"iszero (number_of x :: 'a :: len word) = |
|
1756 |
(bintrunc (len_of TYPE('a)) x = Int.Pls)" |
|
1757 |
apply (unfold iszero_def word_0_wi word_no_wi) |
|
1758 |
apply (rule word_ubin.norm_eq_iff [symmetric, THEN trans]) |
|
1759 |
apply (simp add : Pls_def [symmetric]) |
|
1760 |
done |
|
1761 |
||
1762 |
lemmas word_le_0_iff [simp] = |
|
1763 |
word_zero_le [THEN leD, THEN linorder_antisym_conv1] |
|
1764 |
||
1765 |
lemma word_of_nat: "of_nat n = word_of_int (int n)" |
|
1766 |
by (induct n) (auto simp add : word_of_int_hom_syms) |
|
1767 |
||
1768 |
lemma word_of_int: "of_int = word_of_int" |
|
1769 |
apply (rule ext) |
|
1770 |
apply (unfold of_int_def) |
|
39910 | 1771 |
apply (rule the_elemI) |
37660 | 1772 |
apply safe |
1773 |
apply (simp_all add: word_of_nat word_of_int_homs) |
|
1774 |
defer |
|
1775 |
apply (rule Rep_Integ_ne [THEN nonemptyE]) |
|
1776 |
apply (rule bexI) |
|
1777 |
prefer 2 |
|
1778 |
apply assumption |
|
1779 |
apply (auto simp add: RI_eq_diff) |
|
1780 |
done |
|
1781 |
||
1782 |
lemma word_of_int_nat: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1783 |
"0 <= x \<Longrightarrow> word_of_int x = of_nat (nat x)" |
37660 | 1784 |
by (simp add: of_nat_nat word_of_int) |
1785 |
||
1786 |
lemma word_number_of_eq: |
|
1787 |
"number_of w = (of_int w :: 'a :: len word)" |
|
1788 |
unfolding word_number_of_def word_of_int by auto |
|
1789 |
||
1790 |
instance word :: (len) number_ring |
|
1791 |
by (intro_classes) (simp add : word_number_of_eq) |
|
1792 |
||
1793 |
lemma iszero_word_no [simp] : |
|
1794 |
"iszero (number_of bin :: 'a :: len word) = |
|
1795 |
iszero (number_of (bintrunc (len_of TYPE('a)) bin) :: int)" |
|
1796 |
apply (simp add: zero_bintrunc number_of_is_id) |
|
1797 |
apply (unfold iszero_def Pls_def) |
|
1798 |
apply (rule refl) |
|
1799 |
done |
|
1800 |
||
1801 |
||
1802 |
subsection "Word and nat" |
|
1803 |
||
1804 |
lemma td_ext_unat': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1805 |
"n = len_of TYPE ('a :: len) \<Longrightarrow> |
37660 | 1806 |
td_ext (unat :: 'a word => nat) of_nat |
1807 |
(unats n) (%i. i mod 2 ^ n)" |
|
1808 |
apply (unfold td_ext_def' unat_def word_of_nat unats_uints) |
|
1809 |
apply (auto intro!: imageI simp add : word_of_int_hom_syms) |
|
1810 |
apply (erule word_uint.Abs_inverse [THEN arg_cong]) |
|
1811 |
apply (simp add: int_word_uint nat_mod_distrib nat_power_eq) |
|
1812 |
done |
|
1813 |
||
1814 |
lemmas td_ext_unat = refl [THEN td_ext_unat'] |
|
1815 |
lemmas unat_of_nat = td_ext_unat [THEN td_ext.eq_norm, standard] |
|
1816 |
||
1817 |
interpretation word_unat: |
|
1818 |
td_ext "unat::'a::len word => nat" |
|
1819 |
of_nat |
|
1820 |
"unats (len_of TYPE('a::len))" |
|
1821 |
"%i. i mod 2 ^ len_of TYPE('a::len)" |
|
1822 |
by (rule td_ext_unat) |
|
1823 |
||
1824 |
lemmas td_unat = word_unat.td_thm |
|
1825 |
||
1826 |
lemmas unat_lt2p [iff] = word_unat.Rep [unfolded unats_def mem_Collect_eq] |
|
1827 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1828 |
lemma unat_le: "y <= unat (z :: 'a :: len word) \<Longrightarrow> y : unats (len_of TYPE ('a))" |
37660 | 1829 |
apply (unfold unats_def) |
1830 |
apply clarsimp |
|
1831 |
apply (rule xtrans, rule unat_lt2p, assumption) |
|
1832 |
done |
|
1833 |
||
1834 |
lemma word_nchotomy: |
|
1835 |
"ALL w. EX n. (w :: 'a :: len word) = of_nat n & n < 2 ^ len_of TYPE ('a)" |
|
1836 |
apply (rule allI) |
|
1837 |
apply (rule word_unat.Abs_cases) |
|
1838 |
apply (unfold unats_def) |
|
1839 |
apply auto |
|
1840 |
done |
|
1841 |
||
1842 |
lemma of_nat_eq: |
|
1843 |
fixes w :: "'a::len word" |
|
1844 |
shows "(of_nat n = w) = (\<exists>q. n = unat w + q * 2 ^ len_of TYPE('a))" |
|
1845 |
apply (rule trans) |
|
1846 |
apply (rule word_unat.inverse_norm) |
|
1847 |
apply (rule iffI) |
|
1848 |
apply (rule mod_eqD) |
|
1849 |
apply simp |
|
1850 |
apply clarsimp |
|
1851 |
done |
|
1852 |
||
1853 |
lemma of_nat_eq_size: |
|
1854 |
"(of_nat n = w) = (EX q. n = unat w + q * 2 ^ size w)" |
|
1855 |
unfolding word_size by (rule of_nat_eq) |
|
1856 |
||
1857 |
lemma of_nat_0: |
|
1858 |
"(of_nat m = (0::'a::len word)) = (\<exists>q. m = q * 2 ^ len_of TYPE('a))" |
|
1859 |
by (simp add: of_nat_eq) |
|
1860 |
||
1861 |
lemmas of_nat_2p = mult_1 [symmetric, THEN iffD2 [OF of_nat_0 exI]] |
|
1862 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1863 |
lemma of_nat_gt_0: "of_nat k ~= 0 \<Longrightarrow> 0 < k" |
37660 | 1864 |
by (cases k) auto |
1865 |
||
1866 |
lemma of_nat_neq_0: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1867 |
"0 < k \<Longrightarrow> k < 2 ^ len_of TYPE ('a :: len) \<Longrightarrow> of_nat k ~= (0 :: 'a word)" |
37660 | 1868 |
by (clarsimp simp add : of_nat_0) |
1869 |
||
1870 |
lemma Abs_fnat_hom_add: |
|
1871 |
"of_nat a + of_nat b = of_nat (a + b)" |
|
1872 |
by simp |
|
1873 |
||
1874 |
lemma Abs_fnat_hom_mult: |
|
1875 |
"of_nat a * of_nat b = (of_nat (a * b) :: 'a :: len word)" |
|
1876 |
by (simp add: word_of_nat word_of_int_mult_hom zmult_int) |
|
1877 |
||
1878 |
lemma Abs_fnat_hom_Suc: |
|
1879 |
"word_succ (of_nat a) = of_nat (Suc a)" |
|
1880 |
by (simp add: word_of_nat word_of_int_succ_hom add_ac) |
|
1881 |
||
1882 |
lemma Abs_fnat_hom_0: "(0::'a::len word) = of_nat 0" |
|
1883 |
by (simp add: word_of_nat word_0_wi) |
|
1884 |
||
1885 |
lemma Abs_fnat_hom_1: "(1::'a::len word) = of_nat (Suc 0)" |
|
1886 |
by (simp add: word_of_nat word_1_wi) |
|
1887 |
||
1888 |
lemmas Abs_fnat_homs = |
|
1889 |
Abs_fnat_hom_add Abs_fnat_hom_mult Abs_fnat_hom_Suc |
|
1890 |
Abs_fnat_hom_0 Abs_fnat_hom_1 |
|
1891 |
||
1892 |
lemma word_arith_nat_add: |
|
1893 |
"a + b = of_nat (unat a + unat b)" |
|
1894 |
by simp |
|
1895 |
||
1896 |
lemma word_arith_nat_mult: |
|
1897 |
"a * b = of_nat (unat a * unat b)" |
|
1898 |
by (simp add: Abs_fnat_hom_mult [symmetric]) |
|
1899 |
||
1900 |
lemma word_arith_nat_Suc: |
|
1901 |
"word_succ a = of_nat (Suc (unat a))" |
|
1902 |
by (subst Abs_fnat_hom_Suc [symmetric]) simp |
|
1903 |
||
1904 |
lemma word_arith_nat_div: |
|
1905 |
"a div b = of_nat (unat a div unat b)" |
|
1906 |
by (simp add: word_div_def word_of_nat zdiv_int uint_nat) |
|
1907 |
||
1908 |
lemma word_arith_nat_mod: |
|
1909 |
"a mod b = of_nat (unat a mod unat b)" |
|
1910 |
by (simp add: word_mod_def word_of_nat zmod_int uint_nat) |
|
1911 |
||
1912 |
lemmas word_arith_nat_defs = |
|
1913 |
word_arith_nat_add word_arith_nat_mult |
|
1914 |
word_arith_nat_Suc Abs_fnat_hom_0 |
|
1915 |
Abs_fnat_hom_1 word_arith_nat_div |
|
1916 |
word_arith_nat_mod |
|
1917 |
||
1918 |
lemmas unat_cong = arg_cong [where f = "unat"] |
|
1919 |
||
1920 |
lemmas unat_word_ariths = word_arith_nat_defs |
|
1921 |
[THEN trans [OF unat_cong unat_of_nat], standard] |
|
1922 |
||
1923 |
lemmas word_sub_less_iff = word_sub_le_iff |
|
1924 |
[simplified linorder_not_less [symmetric], simplified] |
|
1925 |
||
1926 |
lemma unat_add_lem: |
|
1927 |
"(unat x + unat y < 2 ^ len_of TYPE('a)) = |
|
1928 |
(unat (x + y :: 'a :: len word) = unat x + unat y)" |
|
1929 |
unfolding unat_word_ariths |
|
1930 |
by (auto intro!: trans [OF _ nat_mod_lem]) |
|
1931 |
||
1932 |
lemma unat_mult_lem: |
|
1933 |
"(unat x * unat y < 2 ^ len_of TYPE('a)) = |
|
1934 |
(unat (x * y :: 'a :: len word) = unat x * unat y)" |
|
1935 |
unfolding unat_word_ariths |
|
1936 |
by (auto intro!: trans [OF _ nat_mod_lem]) |
|
1937 |
||
1938 |
lemmas unat_plus_if' = |
|
1939 |
trans [OF unat_word_ariths(1) mod_nat_add, simplified, standard] |
|
1940 |
||
1941 |
lemma le_no_overflow: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
1942 |
"x <= b \<Longrightarrow> a <= a + b \<Longrightarrow> x <= a + (b :: 'a :: len0 word)" |
37660 | 1943 |
apply (erule order_trans) |
1944 |
apply (erule olen_add_eqv [THEN iffD1]) |
|
1945 |
done |
|
1946 |
||
1947 |
lemmas un_ui_le = trans |
|
1948 |
[OF word_le_nat_alt [symmetric] |
|
1949 |
word_le_def, |
|
1950 |
standard] |
|
1951 |
||
1952 |
lemma unat_sub_if_size: |
|
1953 |
"unat (x - y) = (if unat y <= unat x |
|
1954 |
then unat x - unat y |
|
1955 |
else unat x + 2 ^ size x - unat y)" |
|
1956 |
apply (unfold word_size) |
|
1957 |
apply (simp add: un_ui_le) |
|
1958 |
apply (auto simp add: unat_def uint_sub_if') |
|
1959 |
apply (rule nat_diff_distrib) |
|
1960 |
prefer 3 |
|
1961 |
apply (simp add: algebra_simps) |
|
1962 |
apply (rule nat_diff_distrib [THEN trans]) |
|
1963 |
prefer 3 |
|
1964 |
apply (subst nat_add_distrib) |
|
1965 |
prefer 3 |
|
1966 |
apply (simp add: nat_power_eq) |
|
1967 |
apply auto |
|
1968 |
apply uint_arith |
|
1969 |
done |
|
1970 |
||
1971 |
lemmas unat_sub_if' = unat_sub_if_size [unfolded word_size] |
|
1972 |
||
1973 |
lemma unat_div: "unat ((x :: 'a :: len word) div y) = unat x div unat y" |
|
1974 |
apply (simp add : unat_word_ariths) |
|
1975 |
apply (rule unat_lt2p [THEN xtr7, THEN nat_mod_eq']) |
|
1976 |
apply (rule div_le_dividend) |
|
1977 |
done |
|
1978 |
||
1979 |
lemma unat_mod: "unat ((x :: 'a :: len word) mod y) = unat x mod unat y" |
|
1980 |
apply (clarsimp simp add : unat_word_ariths) |
|
1981 |
apply (cases "unat y") |
|
1982 |
prefer 2 |
|
1983 |
apply (rule unat_lt2p [THEN xtr7, THEN nat_mod_eq']) |
|
1984 |
apply (rule mod_le_divisor) |
|
1985 |
apply auto |
|
1986 |
done |
|
1987 |
||
1988 |
lemma uint_div: "uint ((x :: 'a :: len word) div y) = uint x div uint y" |
|
1989 |
unfolding uint_nat by (simp add : unat_div zdiv_int) |
|
1990 |
||
1991 |
lemma uint_mod: "uint ((x :: 'a :: len word) mod y) = uint x mod uint y" |
|
1992 |
unfolding uint_nat by (simp add : unat_mod zmod_int) |
|
1993 |
||
1994 |
||
1995 |
subsection {* Definition of unat\_arith tactic *} |
|
1996 |
||
1997 |
lemma unat_split: |
|
1998 |
fixes x::"'a::len word" |
|
1999 |
shows "P (unat x) = |
|
2000 |
(ALL n. of_nat n = x & n < 2^len_of TYPE('a) --> P n)" |
|
2001 |
by (auto simp: unat_of_nat) |
|
2002 |
||
2003 |
lemma unat_split_asm: |
|
2004 |
fixes x::"'a::len word" |
|
2005 |
shows "P (unat x) = |
|
2006 |
(~(EX n. of_nat n = x & n < 2^len_of TYPE('a) & ~ P n))" |
|
2007 |
by (auto simp: unat_of_nat) |
|
2008 |
||
2009 |
lemmas of_nat_inverse = |
|
2010 |
word_unat.Abs_inverse' [rotated, unfolded unats_def, simplified] |
|
2011 |
||
2012 |
lemmas unat_splits = unat_split unat_split_asm |
|
2013 |
||
2014 |
lemmas unat_arith_simps = |
|
2015 |
word_le_nat_alt word_less_nat_alt |
|
2016 |
word_unat.Rep_inject [symmetric] |
|
2017 |
unat_sub_if' unat_plus_if' unat_div unat_mod |
|
2018 |
||
2019 |
(* unat_arith_tac: tactic to reduce word arithmetic to nat, |
|
2020 |
try to solve via arith *) |
|
2021 |
ML {* |
|
2022 |
fun unat_arith_ss_of ss = |
|
2023 |
ss addsimps @{thms unat_arith_simps} |
|
2024 |
delsimps @{thms word_unat.Rep_inject} |
|
2025 |
addsplits @{thms split_if_asm} |
|
2026 |
addcongs @{thms power_False_cong} |
|
2027 |
||
2028 |
fun unat_arith_tacs ctxt = |
|
2029 |
let |
|
2030 |
fun arith_tac' n t = |
|
2031 |
Arith_Data.verbose_arith_tac ctxt n t |
|
2032 |
handle Cooper.COOPER _ => Seq.empty; |
|
2033 |
val cs = claset_of ctxt; |
|
2034 |
val ss = simpset_of ctxt; |
|
2035 |
in |
|
2036 |
[ clarify_tac cs 1, |
|
2037 |
full_simp_tac (unat_arith_ss_of ss) 1, |
|
2038 |
ALLGOALS (full_simp_tac (HOL_ss addsplits @{thms unat_splits} |
|
2039 |
addcongs @{thms power_False_cong})), |
|
2040 |
rewrite_goals_tac @{thms word_size}, |
|
2041 |
ALLGOALS (fn n => REPEAT (resolve_tac [allI, impI] n) THEN |
|
2042 |
REPEAT (etac conjE n) THEN |
|
2043 |
REPEAT (dtac @{thm of_nat_inverse} n THEN atac n)), |
|
2044 |
TRYALL arith_tac' ] |
|
2045 |
end |
|
2046 |
||
2047 |
fun unat_arith_tac ctxt = SELECT_GOAL (EVERY (unat_arith_tacs ctxt)) |
|
2048 |
*} |
|
2049 |
||
2050 |
method_setup unat_arith = |
|
2051 |
{* Scan.succeed (SIMPLE_METHOD' o unat_arith_tac) *} |
|
2052 |
"solving word arithmetic via natural numbers and arith" |
|
2053 |
||
2054 |
lemma no_plus_overflow_unat_size: |
|
2055 |
"((x :: 'a :: len word) <= x + y) = (unat x + unat y < 2 ^ size x)" |
|
2056 |
unfolding word_size by unat_arith |
|
2057 |
||
2058 |
lemmas no_olen_add_nat = no_plus_overflow_unat_size [unfolded word_size] |
|
2059 |
||
2060 |
lemmas unat_plus_simple = trans [OF no_olen_add_nat unat_add_lem, standard] |
|
2061 |
||
2062 |
lemma word_div_mult: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2063 |
"(0 :: 'a :: len word) < y \<Longrightarrow> unat x * unat y < 2 ^ len_of TYPE('a) \<Longrightarrow> |
37660 | 2064 |
x * y div y = x" |
2065 |
apply unat_arith |
|
2066 |
apply clarsimp |
|
2067 |
apply (subst unat_mult_lem [THEN iffD1]) |
|
2068 |
apply auto |
|
2069 |
done |
|
2070 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2071 |
lemma div_lt': "(i :: 'a :: len word) <= k div x \<Longrightarrow> |
37660 | 2072 |
unat i * unat x < 2 ^ len_of TYPE('a)" |
2073 |
apply unat_arith |
|
2074 |
apply clarsimp |
|
2075 |
apply (drule mult_le_mono1) |
|
2076 |
apply (erule order_le_less_trans) |
|
2077 |
apply (rule xtr7 [OF unat_lt2p div_mult_le]) |
|
2078 |
done |
|
2079 |
||
2080 |
lemmas div_lt'' = order_less_imp_le [THEN div_lt'] |
|
2081 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2082 |
lemma div_lt_mult: "(i :: 'a :: len word) < k div x \<Longrightarrow> 0 < x \<Longrightarrow> i * x < k" |
37660 | 2083 |
apply (frule div_lt'' [THEN unat_mult_lem [THEN iffD1]]) |
2084 |
apply (simp add: unat_arith_simps) |
|
2085 |
apply (drule (1) mult_less_mono1) |
|
2086 |
apply (erule order_less_le_trans) |
|
2087 |
apply (rule div_mult_le) |
|
2088 |
done |
|
2089 |
||
2090 |
lemma div_le_mult: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2091 |
"(i :: 'a :: len word) <= k div x \<Longrightarrow> 0 < x \<Longrightarrow> i * x <= k" |
37660 | 2092 |
apply (frule div_lt' [THEN unat_mult_lem [THEN iffD1]]) |
2093 |
apply (simp add: unat_arith_simps) |
|
2094 |
apply (drule mult_le_mono1) |
|
2095 |
apply (erule order_trans) |
|
2096 |
apply (rule div_mult_le) |
|
2097 |
done |
|
2098 |
||
2099 |
lemma div_lt_uint': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2100 |
"(i :: 'a :: len word) <= k div x \<Longrightarrow> uint i * uint x < 2 ^ len_of TYPE('a)" |
37660 | 2101 |
apply (unfold uint_nat) |
2102 |
apply (drule div_lt') |
|
2103 |
apply (simp add: zmult_int zless_nat_eq_int_zless [symmetric] |
|
2104 |
nat_power_eq) |
|
2105 |
done |
|
2106 |
||
2107 |
lemmas div_lt_uint'' = order_less_imp_le [THEN div_lt_uint'] |
|
2108 |
||
2109 |
lemma word_le_exists': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2110 |
"(x :: 'a :: len0 word) <= y \<Longrightarrow> |
37660 | 2111 |
(EX z. y = x + z & uint x + uint z < 2 ^ len_of TYPE('a))" |
2112 |
apply (rule exI) |
|
2113 |
apply (rule conjI) |
|
2114 |
apply (rule zadd_diff_inverse) |
|
2115 |
apply uint_arith |
|
2116 |
done |
|
2117 |
||
2118 |
lemmas plus_minus_not_NULL = order_less_imp_le [THEN plus_minus_not_NULL_ab] |
|
2119 |
||
2120 |
lemmas plus_minus_no_overflow = |
|
2121 |
order_less_imp_le [THEN plus_minus_no_overflow_ab] |
|
2122 |
||
2123 |
lemmas mcs = word_less_minus_cancel word_less_minus_mono_left |
|
2124 |
word_le_minus_cancel word_le_minus_mono_left |
|
2125 |
||
2126 |
lemmas word_l_diffs = mcs [where y = "w + x", unfolded add_diff_cancel, standard] |
|
2127 |
lemmas word_diff_ls = mcs [where z = "w + x", unfolded add_diff_cancel, standard] |
|
2128 |
lemmas word_plus_mcs = word_diff_ls |
|
2129 |
[where y = "v + x", unfolded add_diff_cancel, standard] |
|
2130 |
||
2131 |
lemmas le_unat_uoi = unat_le [THEN word_unat.Abs_inverse] |
|
2132 |
||
2133 |
lemmas thd = refl [THEN [2] split_div_lemma [THEN iffD2], THEN conjunct1] |
|
2134 |
||
2135 |
lemma thd1: |
|
2136 |
"a div b * b \<le> (a::nat)" |
|
2137 |
using gt_or_eq_0 [of b] |
|
2138 |
apply (rule disjE) |
|
2139 |
apply (erule xtr4 [OF thd mult_commute]) |
|
2140 |
apply clarsimp |
|
2141 |
done |
|
2142 |
||
2143 |
lemmas uno_simps [THEN le_unat_uoi, standard] = |
|
2144 |
mod_le_divisor div_le_dividend thd1 |
|
2145 |
||
2146 |
lemma word_mod_div_equality: |
|
2147 |
"(n div b) * b + (n mod b) = (n :: 'a :: len word)" |
|
2148 |
apply (unfold word_less_nat_alt word_arith_nat_defs) |
|
2149 |
apply (cut_tac y="unat b" in gt_or_eq_0) |
|
2150 |
apply (erule disjE) |
|
2151 |
apply (simp add: mod_div_equality uno_simps) |
|
2152 |
apply simp |
|
2153 |
done |
|
2154 |
||
2155 |
lemma word_div_mult_le: "a div b * b <= (a::'a::len word)" |
|
2156 |
apply (unfold word_le_nat_alt word_arith_nat_defs) |
|
2157 |
apply (cut_tac y="unat b" in gt_or_eq_0) |
|
2158 |
apply (erule disjE) |
|
2159 |
apply (simp add: div_mult_le uno_simps) |
|
2160 |
apply simp |
|
2161 |
done |
|
2162 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2163 |
lemma word_mod_less_divisor: "0 < n \<Longrightarrow> m mod n < (n :: 'a :: len word)" |
37660 | 2164 |
apply (simp only: word_less_nat_alt word_arith_nat_defs) |
2165 |
apply (clarsimp simp add : uno_simps) |
|
2166 |
done |
|
2167 |
||
2168 |
lemma word_of_int_power_hom: |
|
2169 |
"word_of_int a ^ n = (word_of_int (a ^ n) :: 'a :: len word)" |
|
2170 |
by (induct n) (simp_all add : word_of_int_hom_syms power_Suc) |
|
2171 |
||
2172 |
lemma word_arith_power_alt: |
|
2173 |
"a ^ n = (word_of_int (uint a ^ n) :: 'a :: len word)" |
|
2174 |
by (simp add : word_of_int_power_hom [symmetric]) |
|
2175 |
||
2176 |
lemma of_bl_length_less: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2177 |
"length x = k \<Longrightarrow> k < len_of TYPE('a) \<Longrightarrow> (of_bl x :: 'a :: len word) < 2 ^ k" |
37660 | 2178 |
apply (unfold of_bl_no [unfolded word_number_of_def] |
2179 |
word_less_alt word_number_of_alt) |
|
2180 |
apply safe |
|
2181 |
apply (simp (no_asm) add: word_of_int_power_hom word_uint.eq_norm |
|
2182 |
del: word_of_int_bin) |
|
2183 |
apply (simp add: mod_pos_pos_trivial) |
|
2184 |
apply (subst mod_pos_pos_trivial) |
|
2185 |
apply (rule bl_to_bin_ge0) |
|
2186 |
apply (rule order_less_trans) |
|
2187 |
apply (rule bl_to_bin_lt2p) |
|
2188 |
apply simp |
|
2189 |
apply (rule bl_to_bin_lt2p) |
|
2190 |
done |
|
2191 |
||
2192 |
||
2193 |
subsection "Cardinality, finiteness of set of words" |
|
2194 |
||
2195 |
lemmas card_lessThan' = card_lessThan [unfolded lessThan_def] |
|
2196 |
||
2197 |
lemmas card_eq = word_unat.Abs_inj_on [THEN card_image, |
|
2198 |
unfolded word_unat.image, unfolded unats_def, standard] |
|
2199 |
||
2200 |
lemmas card_word = trans [OF card_eq card_lessThan', standard] |
|
2201 |
||
2202 |
lemma finite_word_UNIV: "finite (UNIV :: 'a :: len word set)" |
|
2203 |
apply (rule contrapos_np) |
|
2204 |
prefer 2 |
|
2205 |
apply (erule card_infinite) |
|
2206 |
apply (simp add: card_word) |
|
2207 |
done |
|
2208 |
||
2209 |
lemma card_word_size: |
|
2210 |
"card (UNIV :: 'a :: len word set) = (2 ^ size (x :: 'a word))" |
|
2211 |
unfolding word_size by (rule card_word) |
|
2212 |
||
2213 |
||
2214 |
subsection {* Bitwise Operations on Words *} |
|
2215 |
||
2216 |
lemmas bin_log_bintrs = bin_trunc_not bin_trunc_xor bin_trunc_and bin_trunc_or |
|
2217 |
||
2218 |
(* following definitions require both arithmetic and bit-wise word operations *) |
|
2219 |
||
2220 |
(* to get word_no_log_defs from word_log_defs, using bin_log_bintrs *) |
|
2221 |
lemmas wils1 = bin_log_bintrs [THEN word_ubin.norm_eq_iff [THEN iffD1], |
|
2222 |
folded word_ubin.eq_norm, THEN eq_reflection, standard] |
|
2223 |
||
2224 |
(* the binary operations only *) |
|
2225 |
lemmas word_log_binary_defs = |
|
2226 |
word_and_def word_or_def word_xor_def |
|
2227 |
||
2228 |
lemmas word_no_log_defs [simp] = |
|
2229 |
word_not_def [where a="number_of a", |
|
2230 |
unfolded word_no_wi wils1, folded word_no_wi, standard] |
|
2231 |
word_log_binary_defs [where a="number_of a" and b="number_of b", |
|
2232 |
unfolded word_no_wi wils1, folded word_no_wi, standard] |
|
2233 |
||
2234 |
lemmas word_wi_log_defs = word_no_log_defs [unfolded word_no_wi] |
|
2235 |
||
2236 |
lemma uint_or: "uint (x OR y) = (uint x) OR (uint y)" |
|
2237 |
by (simp add: word_or_def word_no_wi [symmetric] number_of_is_id |
|
2238 |
bin_trunc_ao(2) [symmetric]) |
|
2239 |
||
2240 |
lemma uint_and: "uint (x AND y) = (uint x) AND (uint y)" |
|
2241 |
by (simp add: word_and_def number_of_is_id word_no_wi [symmetric] |
|
2242 |
bin_trunc_ao(1) [symmetric]) |
|
2243 |
||
2244 |
lemma word_ops_nth_size: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2245 |
"n < size (x::'a::len0 word) \<Longrightarrow> |
37660 | 2246 |
(x OR y) !! n = (x !! n | y !! n) & |
2247 |
(x AND y) !! n = (x !! n & y !! n) & |
|
2248 |
(x XOR y) !! n = (x !! n ~= y !! n) & |
|
2249 |
(NOT x) !! n = (~ x !! n)" |
|
2250 |
unfolding word_size word_no_wi word_test_bit_def word_log_defs |
|
2251 |
by (clarsimp simp add : word_ubin.eq_norm nth_bintr bin_nth_ops) |
|
2252 |
||
2253 |
lemma word_ao_nth: |
|
2254 |
fixes x :: "'a::len0 word" |
|
2255 |
shows "(x OR y) !! n = (x !! n | y !! n) & |
|
2256 |
(x AND y) !! n = (x !! n & y !! n)" |
|
2257 |
apply (cases "n < size x") |
|
2258 |
apply (drule_tac y = "y" in word_ops_nth_size) |
|
2259 |
apply simp |
|
2260 |
apply (simp add : test_bit_bin word_size) |
|
2261 |
done |
|
2262 |
||
2263 |
(* get from commutativity, associativity etc of int_and etc |
|
2264 |
to same for word_and etc *) |
|
2265 |
||
2266 |
lemmas bwsimps = |
|
2267 |
word_of_int_homs(2) |
|
2268 |
word_0_wi_Pls |
|
2269 |
word_m1_wi_Min |
|
2270 |
word_wi_log_defs |
|
2271 |
||
2272 |
lemma word_bw_assocs: |
|
2273 |
fixes x :: "'a::len0 word" |
|
2274 |
shows |
|
2275 |
"(x AND y) AND z = x AND y AND z" |
|
2276 |
"(x OR y) OR z = x OR y OR z" |
|
2277 |
"(x XOR y) XOR z = x XOR y XOR z" |
|
2278 |
using word_of_int_Ex [where x=x] |
|
2279 |
word_of_int_Ex [where x=y] |
|
2280 |
word_of_int_Ex [where x=z] |
|
2281 |
by (auto simp: bwsimps bbw_assocs) |
|
2282 |
||
2283 |
lemma word_bw_comms: |
|
2284 |
fixes x :: "'a::len0 word" |
|
2285 |
shows |
|
2286 |
"x AND y = y AND x" |
|
2287 |
"x OR y = y OR x" |
|
2288 |
"x XOR y = y XOR x" |
|
2289 |
using word_of_int_Ex [where x=x] |
|
2290 |
word_of_int_Ex [where x=y] |
|
2291 |
by (auto simp: bwsimps bin_ops_comm) |
|
2292 |
||
2293 |
lemma word_bw_lcs: |
|
2294 |
fixes x :: "'a::len0 word" |
|
2295 |
shows |
|
2296 |
"y AND x AND z = x AND y AND z" |
|
2297 |
"y OR x OR z = x OR y OR z" |
|
2298 |
"y XOR x XOR z = x XOR y XOR z" |
|
2299 |
using word_of_int_Ex [where x=x] |
|
2300 |
word_of_int_Ex [where x=y] |
|
2301 |
word_of_int_Ex [where x=z] |
|
2302 |
by (auto simp: bwsimps) |
|
2303 |
||
2304 |
lemma word_log_esimps [simp]: |
|
2305 |
fixes x :: "'a::len0 word" |
|
2306 |
shows |
|
2307 |
"x AND 0 = 0" |
|
2308 |
"x AND -1 = x" |
|
2309 |
"x OR 0 = x" |
|
2310 |
"x OR -1 = -1" |
|
2311 |
"x XOR 0 = x" |
|
2312 |
"x XOR -1 = NOT x" |
|
2313 |
"0 AND x = 0" |
|
2314 |
"-1 AND x = x" |
|
2315 |
"0 OR x = x" |
|
2316 |
"-1 OR x = -1" |
|
2317 |
"0 XOR x = x" |
|
2318 |
"-1 XOR x = NOT x" |
|
2319 |
using word_of_int_Ex [where x=x] |
|
2320 |
by (auto simp: bwsimps) |
|
2321 |
||
2322 |
lemma word_not_dist: |
|
2323 |
fixes x :: "'a::len0 word" |
|
2324 |
shows |
|
2325 |
"NOT (x OR y) = NOT x AND NOT y" |
|
2326 |
"NOT (x AND y) = NOT x OR NOT y" |
|
2327 |
using word_of_int_Ex [where x=x] |
|
2328 |
word_of_int_Ex [where x=y] |
|
2329 |
by (auto simp: bwsimps bbw_not_dist) |
|
2330 |
||
2331 |
lemma word_bw_same: |
|
2332 |
fixes x :: "'a::len0 word" |
|
2333 |
shows |
|
2334 |
"x AND x = x" |
|
2335 |
"x OR x = x" |
|
2336 |
"x XOR x = 0" |
|
2337 |
using word_of_int_Ex [where x=x] |
|
2338 |
by (auto simp: bwsimps) |
|
2339 |
||
2340 |
lemma word_ao_absorbs [simp]: |
|
2341 |
fixes x :: "'a::len0 word" |
|
2342 |
shows |
|
2343 |
"x AND (y OR x) = x" |
|
2344 |
"x OR y AND x = x" |
|
2345 |
"x AND (x OR y) = x" |
|
2346 |
"y AND x OR x = x" |
|
2347 |
"(y OR x) AND x = x" |
|
2348 |
"x OR x AND y = x" |
|
2349 |
"(x OR y) AND x = x" |
|
2350 |
"x AND y OR x = x" |
|
2351 |
using word_of_int_Ex [where x=x] |
|
2352 |
word_of_int_Ex [where x=y] |
|
2353 |
by (auto simp: bwsimps) |
|
2354 |
||
2355 |
lemma word_not_not [simp]: |
|
2356 |
"NOT NOT (x::'a::len0 word) = x" |
|
2357 |
using word_of_int_Ex [where x=x] |
|
2358 |
by (auto simp: bwsimps) |
|
2359 |
||
2360 |
lemma word_ao_dist: |
|
2361 |
fixes x :: "'a::len0 word" |
|
2362 |
shows "(x OR y) AND z = x AND z OR y AND z" |
|
2363 |
using word_of_int_Ex [where x=x] |
|
2364 |
word_of_int_Ex [where x=y] |
|
2365 |
word_of_int_Ex [where x=z] |
|
2366 |
by (auto simp: bwsimps bbw_ao_dist simp del: bin_ops_comm) |
|
2367 |
||
2368 |
lemma word_oa_dist: |
|
2369 |
fixes x :: "'a::len0 word" |
|
2370 |
shows "x AND y OR z = (x OR z) AND (y OR z)" |
|
2371 |
using word_of_int_Ex [where x=x] |
|
2372 |
word_of_int_Ex [where x=y] |
|
2373 |
word_of_int_Ex [where x=z] |
|
2374 |
by (auto simp: bwsimps bbw_oa_dist simp del: bin_ops_comm) |
|
2375 |
||
2376 |
lemma word_add_not [simp]: |
|
2377 |
fixes x :: "'a::len0 word" |
|
2378 |
shows "x + NOT x = -1" |
|
2379 |
using word_of_int_Ex [where x=x] |
|
2380 |
by (auto simp: bwsimps bin_add_not) |
|
2381 |
||
2382 |
lemma word_plus_and_or [simp]: |
|
2383 |
fixes x :: "'a::len0 word" |
|
2384 |
shows "(x AND y) + (x OR y) = x + y" |
|
2385 |
using word_of_int_Ex [where x=x] |
|
2386 |
word_of_int_Ex [where x=y] |
|
2387 |
by (auto simp: bwsimps plus_and_or) |
|
2388 |
||
2389 |
lemma leoa: |
|
2390 |
fixes x :: "'a::len0 word" |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2391 |
shows "(w = (x OR y)) \<Longrightarrow> (y = (w AND y))" by auto |
37660 | 2392 |
lemma leao: |
2393 |
fixes x' :: "'a::len0 word" |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2394 |
shows "(w' = (x' AND y')) \<Longrightarrow> (x' = (x' OR w'))" by auto |
37660 | 2395 |
|
2396 |
lemmas word_ao_equiv = leao [COMP leoa [COMP iffI]] |
|
2397 |
||
2398 |
lemma le_word_or2: "x <= x OR (y::'a::len0 word)" |
|
2399 |
unfolding word_le_def uint_or |
|
2400 |
by (auto intro: le_int_or) |
|
2401 |
||
2402 |
lemmas le_word_or1 = xtr3 [OF word_bw_comms (2) le_word_or2, standard] |
|
2403 |
lemmas word_and_le1 = |
|
2404 |
xtr3 [OF word_ao_absorbs (4) [symmetric] le_word_or2, standard] |
|
2405 |
lemmas word_and_le2 = |
|
2406 |
xtr3 [OF word_ao_absorbs (8) [symmetric] le_word_or2, standard] |
|
2407 |
||
2408 |
lemma bl_word_not: "to_bl (NOT w) = map Not (to_bl w)" |
|
2409 |
unfolding to_bl_def word_log_defs |
|
2410 |
by (simp add: bl_not_bin number_of_is_id word_no_wi [symmetric] bin_to_bl_def[symmetric]) |
|
2411 |
||
2412 |
lemma bl_word_xor: "to_bl (v XOR w) = map2 op ~= (to_bl v) (to_bl w)" |
|
2413 |
unfolding to_bl_def word_log_defs bl_xor_bin |
|
2414 |
by (simp add: number_of_is_id word_no_wi [symmetric]) |
|
2415 |
||
2416 |
lemma bl_word_or: "to_bl (v OR w) = map2 op | (to_bl v) (to_bl w)" |
|
2417 |
unfolding to_bl_def word_log_defs |
|
2418 |
by (simp add: bl_or_bin number_of_is_id word_no_wi [symmetric]) |
|
2419 |
||
2420 |
lemma bl_word_and: "to_bl (v AND w) = map2 op & (to_bl v) (to_bl w)" |
|
2421 |
unfolding to_bl_def word_log_defs |
|
2422 |
by (simp add: bl_and_bin number_of_is_id word_no_wi [symmetric]) |
|
2423 |
||
2424 |
lemma word_lsb_alt: "lsb (w::'a::len0 word) = test_bit w 0" |
|
2425 |
by (auto simp: word_test_bit_def word_lsb_def) |
|
2426 |
||
2427 |
lemma word_lsb_1_0: "lsb (1::'a::len word) & ~ lsb (0::'b::len0 word)" |
|
2428 |
unfolding word_lsb_def word_1_no word_0_no by auto |
|
2429 |
||
2430 |
lemma word_lsb_last: "lsb (w::'a::len word) = last (to_bl w)" |
|
2431 |
apply (unfold word_lsb_def uint_bl bin_to_bl_def) |
|
2432 |
apply (rule_tac bin="uint w" in bin_exhaust) |
|
2433 |
apply (cases "size w") |
|
2434 |
apply auto |
|
2435 |
apply (auto simp add: bin_to_bl_aux_alt) |
|
2436 |
done |
|
2437 |
||
2438 |
lemma word_lsb_int: "lsb w = (uint w mod 2 = 1)" |
|
2439 |
unfolding word_lsb_def bin_last_mod by auto |
|
2440 |
||
2441 |
lemma word_msb_sint: "msb w = (sint w < 0)" |
|
2442 |
unfolding word_msb_def |
|
2443 |
by (simp add : sign_Min_lt_0 number_of_is_id) |
|
2444 |
||
2445 |
lemma word_msb_no': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2446 |
"w = number_of bin \<Longrightarrow> msb (w::'a::len word) = bin_nth bin (size w - 1)" |
37660 | 2447 |
unfolding word_msb_def word_number_of_def |
2448 |
by (clarsimp simp add: word_sbin.eq_norm word_size bin_sign_lem) |
|
2449 |
||
2450 |
lemmas word_msb_no = refl [THEN word_msb_no', unfolded word_size] |
|
2451 |
||
2452 |
lemma word_msb_nth': "msb (w::'a::len word) = bin_nth (uint w) (size w - 1)" |
|
2453 |
apply (unfold word_size) |
|
2454 |
apply (rule trans [OF _ word_msb_no]) |
|
2455 |
apply (simp add : word_number_of_def) |
|
2456 |
done |
|
2457 |
||
2458 |
lemmas word_msb_nth = word_msb_nth' [unfolded word_size] |
|
2459 |
||
2460 |
lemma word_msb_alt: "msb (w::'a::len word) = hd (to_bl w)" |
|
2461 |
apply (unfold word_msb_nth uint_bl) |
|
2462 |
apply (subst hd_conv_nth) |
|
2463 |
apply (rule length_greater_0_conv [THEN iffD1]) |
|
2464 |
apply simp |
|
2465 |
apply (simp add : nth_bin_to_bl word_size) |
|
2466 |
done |
|
2467 |
||
2468 |
lemma word_set_nth: |
|
2469 |
"set_bit w n (test_bit w n) = (w::'a::len0 word)" |
|
2470 |
unfolding word_test_bit_def word_set_bit_def by auto |
|
2471 |
||
2472 |
lemma bin_nth_uint': |
|
2473 |
"bin_nth (uint w) n = (rev (bin_to_bl (size w) (uint w)) ! n & n < size w)" |
|
2474 |
apply (unfold word_size) |
|
2475 |
apply (safe elim!: bin_nth_uint_imp) |
|
2476 |
apply (frule bin_nth_uint_imp) |
|
2477 |
apply (fast dest!: bin_nth_bl)+ |
|
2478 |
done |
|
2479 |
||
2480 |
lemmas bin_nth_uint = bin_nth_uint' [unfolded word_size] |
|
2481 |
||
2482 |
lemma test_bit_bl: "w !! n = (rev (to_bl w) ! n & n < size w)" |
|
2483 |
unfolding to_bl_def word_test_bit_def word_size |
|
2484 |
by (rule bin_nth_uint) |
|
2485 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2486 |
lemma to_bl_nth: "n < size w \<Longrightarrow> to_bl w ! n = w !! (size w - Suc n)" |
37660 | 2487 |
apply (unfold test_bit_bl) |
2488 |
apply clarsimp |
|
2489 |
apply (rule trans) |
|
2490 |
apply (rule nth_rev_alt) |
|
2491 |
apply (auto simp add: word_size) |
|
2492 |
done |
|
2493 |
||
2494 |
lemma test_bit_set: |
|
2495 |
fixes w :: "'a::len0 word" |
|
2496 |
shows "(set_bit w n x) !! n = (n < size w & x)" |
|
2497 |
unfolding word_size word_test_bit_def word_set_bit_def |
|
2498 |
by (clarsimp simp add : word_ubin.eq_norm nth_bintr) |
|
2499 |
||
2500 |
lemma test_bit_set_gen: |
|
2501 |
fixes w :: "'a::len0 word" |
|
2502 |
shows "test_bit (set_bit w n x) m = |
|
2503 |
(if m = n then n < size w & x else test_bit w m)" |
|
2504 |
apply (unfold word_size word_test_bit_def word_set_bit_def) |
|
2505 |
apply (clarsimp simp add: word_ubin.eq_norm nth_bintr bin_nth_sc_gen) |
|
2506 |
apply (auto elim!: test_bit_size [unfolded word_size] |
|
2507 |
simp add: word_test_bit_def [symmetric]) |
|
2508 |
done |
|
2509 |
||
2510 |
lemma of_bl_rep_False: "of_bl (replicate n False @ bs) = of_bl bs" |
|
2511 |
unfolding of_bl_def bl_to_bin_rep_F by auto |
|
2512 |
||
2513 |
lemma msb_nth': |
|
2514 |
fixes w :: "'a::len word" |
|
2515 |
shows "msb w = w !! (size w - 1)" |
|
2516 |
unfolding word_msb_nth' word_test_bit_def by simp |
|
2517 |
||
2518 |
lemmas msb_nth = msb_nth' [unfolded word_size] |
|
2519 |
||
2520 |
lemmas msb0 = len_gt_0 [THEN diff_Suc_less, THEN |
|
2521 |
word_ops_nth_size [unfolded word_size], standard] |
|
2522 |
lemmas msb1 = msb0 [where i = 0] |
|
2523 |
lemmas word_ops_msb = msb1 [unfolded msb_nth [symmetric, unfolded One_nat_def]] |
|
2524 |
||
2525 |
lemmas lsb0 = len_gt_0 [THEN word_ops_nth_size [unfolded word_size], standard] |
|
2526 |
lemmas word_ops_lsb = lsb0 [unfolded word_lsb_alt] |
|
2527 |
||
2528 |
lemma td_ext_nth': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2529 |
"n = size (w::'a::len0 word) \<Longrightarrow> ofn = set_bits \<Longrightarrow> [w, ofn g] = l \<Longrightarrow> |
37660 | 2530 |
td_ext test_bit ofn {f. ALL i. f i --> i < n} (%h i. h i & i < n)" |
2531 |
apply (unfold word_size td_ext_def') |
|
2532 |
apply (safe del: subset_antisym) |
|
2533 |
apply (rule_tac [3] ext) |
|
2534 |
apply (rule_tac [4] ext) |
|
2535 |
apply (unfold word_size of_nth_def test_bit_bl) |
|
2536 |
apply safe |
|
2537 |
defer |
|
2538 |
apply (clarsimp simp: word_bl.Abs_inverse)+ |
|
2539 |
apply (rule word_bl.Rep_inverse') |
|
2540 |
apply (rule sym [THEN trans]) |
|
2541 |
apply (rule bl_of_nth_nth) |
|
2542 |
apply simp |
|
2543 |
apply (rule bl_of_nth_inj) |
|
2544 |
apply (clarsimp simp add : test_bit_bl word_size) |
|
2545 |
done |
|
2546 |
||
2547 |
lemmas td_ext_nth = td_ext_nth' [OF refl refl refl, unfolded word_size] |
|
2548 |
||
2549 |
interpretation test_bit: |
|
2550 |
td_ext "op !! :: 'a::len0 word => nat => bool" |
|
2551 |
set_bits |
|
2552 |
"{f. \<forall>i. f i \<longrightarrow> i < len_of TYPE('a::len0)}" |
|
2553 |
"(\<lambda>h i. h i \<and> i < len_of TYPE('a::len0))" |
|
2554 |
by (rule td_ext_nth) |
|
2555 |
||
2556 |
declare test_bit.Rep' [simp del] |
|
2557 |
declare test_bit.Rep' [rule del] |
|
2558 |
||
2559 |
lemmas td_nth = test_bit.td_thm |
|
2560 |
||
2561 |
lemma word_set_set_same: |
|
2562 |
fixes w :: "'a::len0 word" |
|
2563 |
shows "set_bit (set_bit w n x) n y = set_bit w n y" |
|
2564 |
by (rule word_eqI) (simp add : test_bit_set_gen word_size) |
|
2565 |
||
2566 |
lemma word_set_set_diff: |
|
2567 |
fixes w :: "'a::len0 word" |
|
2568 |
assumes "m ~= n" |
|
2569 |
shows "set_bit (set_bit w m x) n y = set_bit (set_bit w n y) m x" |
|
2570 |
by (rule word_eqI) (clarsimp simp add : test_bit_set_gen word_size prems) |
|
2571 |
||
2572 |
lemma test_bit_no': |
|
2573 |
fixes w :: "'a::len0 word" |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2574 |
shows "w = number_of bin \<Longrightarrow> test_bit w n = (n < size w & bin_nth bin n)" |
37660 | 2575 |
unfolding word_test_bit_def word_number_of_def word_size |
2576 |
by (simp add : nth_bintr [symmetric] word_ubin.eq_norm) |
|
2577 |
||
2578 |
lemmas test_bit_no = |
|
2579 |
refl [THEN test_bit_no', unfolded word_size, THEN eq_reflection, standard] |
|
2580 |
||
2581 |
lemma nth_0: "~ (0::'a::len0 word) !! n" |
|
2582 |
unfolding test_bit_no word_0_no by auto |
|
2583 |
||
2584 |
lemma nth_sint: |
|
2585 |
fixes w :: "'a::len word" |
|
2586 |
defines "l \<equiv> len_of TYPE ('a)" |
|
2587 |
shows "bin_nth (sint w) n = (if n < l - 1 then w !! n else w !! (l - 1))" |
|
2588 |
unfolding sint_uint l_def |
|
2589 |
by (clarsimp simp add: nth_sbintr word_test_bit_def [symmetric]) |
|
2590 |
||
2591 |
lemma word_lsb_no: |
|
2592 |
"lsb (number_of bin :: 'a :: len word) = (bin_last bin = 1)" |
|
2593 |
unfolding word_lsb_alt test_bit_no by auto |
|
2594 |
||
2595 |
lemma word_set_no: |
|
2596 |
"set_bit (number_of bin::'a::len0 word) n b = |
|
2597 |
number_of (bin_sc n (if b then 1 else 0) bin)" |
|
2598 |
apply (unfold word_set_bit_def word_number_of_def [symmetric]) |
|
2599 |
apply (rule word_eqI) |
|
2600 |
apply (clarsimp simp: word_size bin_nth_sc_gen number_of_is_id |
|
2601 |
test_bit_no nth_bintr) |
|
2602 |
done |
|
2603 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2604 |
lemma setBit_no: |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2605 |
"setBit (number_of bin) n = number_of (bin_sc n 1 bin) " |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2606 |
by (simp add: setBit_def word_set_no) |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2607 |
|
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2608 |
lemma clearBit_no: |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2609 |
"clearBit (number_of bin) n = number_of (bin_sc n 0 bin)" |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2610 |
by (simp add: clearBit_def word_set_no) |
37660 | 2611 |
|
2612 |
lemma to_bl_n1: |
|
2613 |
"to_bl (-1::'a::len0 word) = replicate (len_of TYPE ('a)) True" |
|
2614 |
apply (rule word_bl.Abs_inverse') |
|
2615 |
apply simp |
|
2616 |
apply (rule word_eqI) |
|
2617 |
apply (clarsimp simp add: word_size test_bit_no) |
|
2618 |
apply (auto simp add: word_bl.Abs_inverse test_bit_bl word_size) |
|
2619 |
done |
|
2620 |
||
2621 |
lemma word_msb_n1: "msb (-1::'a::len word)" |
|
2622 |
unfolding word_msb_alt word_msb_alt to_bl_n1 by simp |
|
2623 |
||
2624 |
declare word_set_set_same [simp] word_set_nth [simp] |
|
2625 |
test_bit_no [simp] word_set_no [simp] nth_0 [simp] |
|
2626 |
setBit_no [simp] clearBit_no [simp] |
|
2627 |
word_lsb_no [simp] word_msb_no [simp] word_msb_n1 [simp] word_lsb_1_0 [simp] |
|
2628 |
||
2629 |
lemma word_set_nth_iff: |
|
2630 |
"(set_bit w n b = w) = (w !! n = b | n >= size (w::'a::len0 word))" |
|
2631 |
apply (rule iffI) |
|
2632 |
apply (rule disjCI) |
|
2633 |
apply (drule word_eqD) |
|
2634 |
apply (erule sym [THEN trans]) |
|
2635 |
apply (simp add: test_bit_set) |
|
2636 |
apply (erule disjE) |
|
2637 |
apply clarsimp |
|
2638 |
apply (rule word_eqI) |
|
2639 |
apply (clarsimp simp add : test_bit_set_gen) |
|
2640 |
apply (drule test_bit_size) |
|
2641 |
apply force |
|
2642 |
done |
|
2643 |
||
2644 |
lemma test_bit_2p': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2645 |
"w = word_of_int (2 ^ n) \<Longrightarrow> |
37660 | 2646 |
w !! m = (m = n & m < size (w :: 'a :: len word))" |
2647 |
unfolding word_test_bit_def word_size |
|
2648 |
by (auto simp add: word_ubin.eq_norm nth_bintr nth_2p_bin) |
|
2649 |
||
2650 |
lemmas test_bit_2p = refl [THEN test_bit_2p', unfolded word_size] |
|
2651 |
||
2652 |
lemma nth_w2p: |
|
2653 |
"((2\<Colon>'a\<Colon>len word) ^ n) !! m \<longleftrightarrow> m = n \<and> m < len_of TYPE('a\<Colon>len)" |
|
2654 |
unfolding test_bit_2p [symmetric] word_of_int [symmetric] |
|
2655 |
by (simp add: of_int_power) |
|
2656 |
||
2657 |
lemma uint_2p: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2658 |
"(0::'a::len word) < 2 ^ n \<Longrightarrow> uint (2 ^ n::'a::len word) = 2 ^ n" |
37660 | 2659 |
apply (unfold word_arith_power_alt) |
2660 |
apply (case_tac "len_of TYPE ('a)") |
|
2661 |
apply clarsimp |
|
2662 |
apply (case_tac "nat") |
|
2663 |
apply clarsimp |
|
2664 |
apply (case_tac "n") |
|
2665 |
apply (clarsimp simp add : word_1_wi [symmetric]) |
|
2666 |
apply (clarsimp simp add : word_0_wi [symmetric]) |
|
2667 |
apply (drule word_gt_0 [THEN iffD1]) |
|
2668 |
apply (safe intro!: word_eqI bin_nth_lem ext) |
|
2669 |
apply (auto simp add: test_bit_2p nth_2p_bin word_test_bit_def [symmetric]) |
|
2670 |
done |
|
2671 |
||
2672 |
lemma word_of_int_2p: "(word_of_int (2 ^ n) :: 'a :: len word) = 2 ^ n" |
|
2673 |
apply (unfold word_arith_power_alt) |
|
2674 |
apply (case_tac "len_of TYPE ('a)") |
|
2675 |
apply clarsimp |
|
2676 |
apply (case_tac "nat") |
|
2677 |
apply (rule word_ubin.norm_eq_iff [THEN iffD1]) |
|
2678 |
apply (rule box_equals) |
|
2679 |
apply (rule_tac [2] bintr_ariths (1))+ |
|
2680 |
apply (clarsimp simp add : number_of_is_id) |
|
2681 |
apply simp |
|
2682 |
done |
|
2683 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2684 |
lemma bang_is_le: "x !! m \<Longrightarrow> 2 ^ m <= (x :: 'a :: len word)" |
37660 | 2685 |
apply (rule xtr3) |
2686 |
apply (rule_tac [2] y = "x" in le_word_or2) |
|
2687 |
apply (rule word_eqI) |
|
2688 |
apply (auto simp add: word_ao_nth nth_w2p word_size) |
|
2689 |
done |
|
2690 |
||
2691 |
lemma word_clr_le: |
|
2692 |
fixes w :: "'a::len0 word" |
|
2693 |
shows "w >= set_bit w n False" |
|
2694 |
apply (unfold word_set_bit_def word_le_def word_ubin.eq_norm) |
|
2695 |
apply simp |
|
2696 |
apply (rule order_trans) |
|
2697 |
apply (rule bintr_bin_clr_le) |
|
2698 |
apply simp |
|
2699 |
done |
|
2700 |
||
2701 |
lemma word_set_ge: |
|
2702 |
fixes w :: "'a::len word" |
|
2703 |
shows "w <= set_bit w n True" |
|
2704 |
apply (unfold word_set_bit_def word_le_def word_ubin.eq_norm) |
|
2705 |
apply simp |
|
2706 |
apply (rule order_trans [OF _ bintr_bin_set_ge]) |
|
2707 |
apply simp |
|
2708 |
done |
|
2709 |
||
2710 |
||
2711 |
subsection {* Shifting, Rotating, and Splitting Words *} |
|
2712 |
||
2713 |
lemma shiftl1_number [simp] : |
|
2714 |
"shiftl1 (number_of w) = number_of (w BIT 0)" |
|
2715 |
apply (unfold shiftl1_def word_number_of_def) |
|
2716 |
apply (simp add: word_ubin.norm_eq_iff [symmetric] word_ubin.eq_norm |
|
2717 |
del: BIT_simps) |
|
2718 |
apply (subst refl [THEN bintrunc_BIT_I, symmetric]) |
|
2719 |
apply (subst bintrunc_bintrunc_min) |
|
2720 |
apply simp |
|
2721 |
done |
|
2722 |
||
2723 |
lemma shiftl1_0 [simp] : "shiftl1 0 = 0" |
|
2724 |
unfolding word_0_no shiftl1_number by auto |
|
2725 |
||
2726 |
lemmas shiftl1_def_u = shiftl1_def [folded word_number_of_def] |
|
2727 |
||
2728 |
lemma shiftl1_def_s: "shiftl1 w = number_of (sint w BIT 0)" |
|
2729 |
by (rule trans [OF _ shiftl1_number]) simp |
|
2730 |
||
2731 |
lemma shiftr1_0 [simp] : "shiftr1 0 = 0" |
|
2732 |
unfolding shiftr1_def |
|
2733 |
by simp (simp add: word_0_wi) |
|
2734 |
||
2735 |
lemma sshiftr1_0 [simp] : "sshiftr1 0 = 0" |
|
2736 |
apply (unfold sshiftr1_def) |
|
2737 |
apply simp |
|
2738 |
apply (simp add : word_0_wi) |
|
2739 |
done |
|
2740 |
||
2741 |
lemma sshiftr1_n1 [simp] : "sshiftr1 -1 = -1" |
|
2742 |
unfolding sshiftr1_def by auto |
|
2743 |
||
2744 |
lemma shiftl_0 [simp] : "(0::'a::len0 word) << n = 0" |
|
2745 |
unfolding shiftl_def by (induct n) auto |
|
2746 |
||
2747 |
lemma shiftr_0 [simp] : "(0::'a::len0 word) >> n = 0" |
|
2748 |
unfolding shiftr_def by (induct n) auto |
|
2749 |
||
2750 |
lemma sshiftr_0 [simp] : "0 >>> n = 0" |
|
2751 |
unfolding sshiftr_def by (induct n) auto |
|
2752 |
||
2753 |
lemma sshiftr_n1 [simp] : "-1 >>> n = -1" |
|
2754 |
unfolding sshiftr_def by (induct n) auto |
|
2755 |
||
2756 |
lemma nth_shiftl1: "shiftl1 w !! n = (n < size w & n > 0 & w !! (n - 1))" |
|
2757 |
apply (unfold shiftl1_def word_test_bit_def) |
|
2758 |
apply (simp add: nth_bintr word_ubin.eq_norm word_size) |
|
2759 |
apply (cases n) |
|
2760 |
apply auto |
|
2761 |
done |
|
2762 |
||
2763 |
lemma nth_shiftl' [rule_format]: |
|
2764 |
"ALL n. ((w::'a::len0 word) << m) !! n = (n < size w & n >= m & w !! (n - m))" |
|
2765 |
apply (unfold shiftl_def) |
|
2766 |
apply (induct "m") |
|
2767 |
apply (force elim!: test_bit_size) |
|
2768 |
apply (clarsimp simp add : nth_shiftl1 word_size) |
|
2769 |
apply arith |
|
2770 |
done |
|
2771 |
||
2772 |
lemmas nth_shiftl = nth_shiftl' [unfolded word_size] |
|
2773 |
||
2774 |
lemma nth_shiftr1: "shiftr1 w !! n = w !! Suc n" |
|
2775 |
apply (unfold shiftr1_def word_test_bit_def) |
|
2776 |
apply (simp add: nth_bintr word_ubin.eq_norm) |
|
2777 |
apply safe |
|
2778 |
apply (drule bin_nth.Suc [THEN iffD2, THEN bin_nth_uint_imp]) |
|
2779 |
apply simp |
|
2780 |
done |
|
2781 |
||
2782 |
lemma nth_shiftr: |
|
2783 |
"\<And>n. ((w::'a::len0 word) >> m) !! n = w !! (n + m)" |
|
2784 |
apply (unfold shiftr_def) |
|
2785 |
apply (induct "m") |
|
2786 |
apply (auto simp add : nth_shiftr1) |
|
2787 |
done |
|
2788 |
||
2789 |
(* see paper page 10, (1), (2), shiftr1_def is of the form of (1), |
|
2790 |
where f (ie bin_rest) takes normal arguments to normal results, |
|
2791 |
thus we get (2) from (1) *) |
|
2792 |
||
2793 |
lemma uint_shiftr1: "uint (shiftr1 w) = bin_rest (uint w)" |
|
2794 |
apply (unfold shiftr1_def word_ubin.eq_norm bin_rest_trunc_i) |
|
2795 |
apply (subst bintr_uint [symmetric, OF order_refl]) |
|
2796 |
apply (simp only : bintrunc_bintrunc_l) |
|
2797 |
apply simp |
|
2798 |
done |
|
2799 |
||
2800 |
lemma nth_sshiftr1: |
|
2801 |
"sshiftr1 w !! n = (if n = size w - 1 then w !! n else w !! Suc n)" |
|
2802 |
apply (unfold sshiftr1_def word_test_bit_def) |
|
2803 |
apply (simp add: nth_bintr word_ubin.eq_norm |
|
2804 |
bin_nth.Suc [symmetric] word_size |
|
2805 |
del: bin_nth.simps) |
|
2806 |
apply (simp add: nth_bintr uint_sint del : bin_nth.simps) |
|
2807 |
apply (auto simp add: bin_nth_sint) |
|
2808 |
done |
|
2809 |
||
2810 |
lemma nth_sshiftr [rule_format] : |
|
2811 |
"ALL n. sshiftr w m !! n = (n < size w & |
|
2812 |
(if n + m >= size w then w !! (size w - 1) else w !! (n + m)))" |
|
2813 |
apply (unfold sshiftr_def) |
|
2814 |
apply (induct_tac "m") |
|
2815 |
apply (simp add: test_bit_bl) |
|
2816 |
apply (clarsimp simp add: nth_sshiftr1 word_size) |
|
2817 |
apply safe |
|
2818 |
apply arith |
|
2819 |
apply arith |
|
2820 |
apply (erule thin_rl) |
|
2821 |
apply (case_tac n) |
|
2822 |
apply safe |
|
2823 |
apply simp |
|
2824 |
apply simp |
|
2825 |
apply (erule thin_rl) |
|
2826 |
apply (case_tac n) |
|
2827 |
apply safe |
|
2828 |
apply simp |
|
2829 |
apply simp |
|
2830 |
apply arith+ |
|
2831 |
done |
|
2832 |
||
2833 |
lemma shiftr1_div_2: "uint (shiftr1 w) = uint w div 2" |
|
2834 |
apply (unfold shiftr1_def bin_rest_div) |
|
2835 |
apply (rule word_uint.Abs_inverse) |
|
2836 |
apply (simp add: uints_num pos_imp_zdiv_nonneg_iff) |
|
2837 |
apply (rule xtr7) |
|
2838 |
prefer 2 |
|
2839 |
apply (rule zdiv_le_dividend) |
|
2840 |
apply auto |
|
2841 |
done |
|
2842 |
||
2843 |
lemma sshiftr1_div_2: "sint (sshiftr1 w) = sint w div 2" |
|
2844 |
apply (unfold sshiftr1_def bin_rest_div [symmetric]) |
|
2845 |
apply (simp add: word_sbin.eq_norm) |
|
2846 |
apply (rule trans) |
|
2847 |
defer |
|
2848 |
apply (subst word_sbin.norm_Rep [symmetric]) |
|
2849 |
apply (rule refl) |
|
2850 |
apply (subst word_sbin.norm_Rep [symmetric]) |
|
2851 |
apply (unfold One_nat_def) |
|
2852 |
apply (rule sbintrunc_rest) |
|
2853 |
done |
|
2854 |
||
2855 |
lemma shiftr_div_2n: "uint (shiftr w n) = uint w div 2 ^ n" |
|
2856 |
apply (unfold shiftr_def) |
|
2857 |
apply (induct "n") |
|
2858 |
apply simp |
|
2859 |
apply (simp add: shiftr1_div_2 mult_commute |
|
2860 |
zdiv_zmult2_eq [symmetric]) |
|
2861 |
done |
|
2862 |
||
2863 |
lemma sshiftr_div_2n: "sint (sshiftr w n) = sint w div 2 ^ n" |
|
2864 |
apply (unfold sshiftr_def) |
|
2865 |
apply (induct "n") |
|
2866 |
apply simp |
|
2867 |
apply (simp add: sshiftr1_div_2 mult_commute |
|
2868 |
zdiv_zmult2_eq [symmetric]) |
|
2869 |
done |
|
2870 |
||
2871 |
subsubsection "shift functions in terms of lists of bools" |
|
2872 |
||
2873 |
lemmas bshiftr1_no_bin [simp] = |
|
2874 |
bshiftr1_def [where w="number_of w", unfolded to_bl_no_bin, standard] |
|
2875 |
||
2876 |
lemma bshiftr1_bl: "to_bl (bshiftr1 b w) = b # butlast (to_bl w)" |
|
2877 |
unfolding bshiftr1_def by (rule word_bl.Abs_inverse) simp |
|
2878 |
||
2879 |
lemma shiftl1_of_bl: "shiftl1 (of_bl bl) = of_bl (bl @ [False])" |
|
2880 |
unfolding uint_bl of_bl_no |
|
2881 |
by (simp add: bl_to_bin_aux_append bl_to_bin_def) |
|
2882 |
||
2883 |
lemma shiftl1_bl: "shiftl1 (w::'a::len0 word) = of_bl (to_bl w @ [False])" |
|
2884 |
proof - |
|
2885 |
have "shiftl1 w = shiftl1 (of_bl (to_bl w))" by simp |
|
2886 |
also have "\<dots> = of_bl (to_bl w @ [False])" by (rule shiftl1_of_bl) |
|
2887 |
finally show ?thesis . |
|
2888 |
qed |
|
2889 |
||
2890 |
lemma bl_shiftl1: |
|
2891 |
"to_bl (shiftl1 (w :: 'a :: len word)) = tl (to_bl w) @ [False]" |
|
2892 |
apply (simp add: shiftl1_bl word_rep_drop drop_Suc drop_Cons') |
|
2893 |
apply (fast intro!: Suc_leI) |
|
2894 |
done |
|
2895 |
||
2896 |
lemma shiftr1_bl: "shiftr1 w = of_bl (butlast (to_bl w))" |
|
2897 |
apply (unfold shiftr1_def uint_bl of_bl_def) |
|
2898 |
apply (simp add: butlast_rest_bin word_size) |
|
2899 |
apply (simp add: bin_rest_trunc [symmetric, unfolded One_nat_def]) |
|
2900 |
done |
|
2901 |
||
2902 |
lemma bl_shiftr1: |
|
2903 |
"to_bl (shiftr1 (w :: 'a :: len word)) = False # butlast (to_bl w)" |
|
2904 |
unfolding shiftr1_bl |
|
2905 |
by (simp add : word_rep_drop len_gt_0 [THEN Suc_leI]) |
|
2906 |
||
2907 |
||
2908 |
(* relate the two above : TODO - remove the :: len restriction on |
|
2909 |
this theorem and others depending on it *) |
|
2910 |
lemma shiftl1_rev: |
|
2911 |
"shiftl1 (w :: 'a :: len word) = word_reverse (shiftr1 (word_reverse w))" |
|
2912 |
apply (unfold word_reverse_def) |
|
2913 |
apply (rule word_bl.Rep_inverse' [symmetric]) |
|
2914 |
apply (simp add: bl_shiftl1 bl_shiftr1 word_bl.Abs_inverse) |
|
2915 |
apply (cases "to_bl w") |
|
2916 |
apply auto |
|
2917 |
done |
|
2918 |
||
2919 |
lemma shiftl_rev: |
|
2920 |
"shiftl (w :: 'a :: len word) n = word_reverse (shiftr (word_reverse w) n)" |
|
2921 |
apply (unfold shiftl_def shiftr_def) |
|
2922 |
apply (induct "n") |
|
2923 |
apply (auto simp add : shiftl1_rev) |
|
2924 |
done |
|
2925 |
||
2926 |
lemmas rev_shiftl = |
|
2927 |
shiftl_rev [where w = "word_reverse w", simplified, standard] |
|
2928 |
||
2929 |
lemmas shiftr_rev = rev_shiftl [THEN word_rev_gal', standard] |
|
2930 |
lemmas rev_shiftr = shiftl_rev [THEN word_rev_gal', standard] |
|
2931 |
||
2932 |
lemma bl_sshiftr1: |
|
2933 |
"to_bl (sshiftr1 (w :: 'a :: len word)) = hd (to_bl w) # butlast (to_bl w)" |
|
2934 |
apply (unfold sshiftr1_def uint_bl word_size) |
|
2935 |
apply (simp add: butlast_rest_bin word_ubin.eq_norm) |
|
2936 |
apply (simp add: sint_uint) |
|
2937 |
apply (rule nth_equalityI) |
|
2938 |
apply clarsimp |
|
2939 |
apply clarsimp |
|
2940 |
apply (case_tac i) |
|
2941 |
apply (simp_all add: hd_conv_nth length_0_conv [symmetric] |
|
2942 |
nth_bin_to_bl bin_nth.Suc [symmetric] |
|
2943 |
nth_sbintr |
|
2944 |
del: bin_nth.Suc) |
|
2945 |
apply force |
|
2946 |
apply (rule impI) |
|
2947 |
apply (rule_tac f = "bin_nth (uint w)" in arg_cong) |
|
2948 |
apply simp |
|
2949 |
done |
|
2950 |
||
2951 |
lemma drop_shiftr: |
|
2952 |
"drop n (to_bl ((w :: 'a :: len word) >> n)) = take (size w - n) (to_bl w)" |
|
2953 |
apply (unfold shiftr_def) |
|
2954 |
apply (induct n) |
|
2955 |
prefer 2 |
|
2956 |
apply (simp add: drop_Suc bl_shiftr1 butlast_drop [symmetric]) |
|
2957 |
apply (rule butlast_take [THEN trans]) |
|
2958 |
apply (auto simp: word_size) |
|
2959 |
done |
|
2960 |
||
2961 |
lemma drop_sshiftr: |
|
2962 |
"drop n (to_bl ((w :: 'a :: len word) >>> n)) = take (size w - n) (to_bl w)" |
|
2963 |
apply (unfold sshiftr_def) |
|
2964 |
apply (induct n) |
|
2965 |
prefer 2 |
|
2966 |
apply (simp add: drop_Suc bl_sshiftr1 butlast_drop [symmetric]) |
|
2967 |
apply (rule butlast_take [THEN trans]) |
|
2968 |
apply (auto simp: word_size) |
|
2969 |
done |
|
2970 |
||
2971 |
lemma take_shiftr [rule_format] : |
|
2972 |
"n <= size (w :: 'a :: len word) --> take n (to_bl (w >> n)) = |
|
2973 |
replicate n False" |
|
2974 |
apply (unfold shiftr_def) |
|
2975 |
apply (induct n) |
|
2976 |
prefer 2 |
|
2977 |
apply (simp add: bl_shiftr1) |
|
2978 |
apply (rule impI) |
|
2979 |
apply (rule take_butlast [THEN trans]) |
|
2980 |
apply (auto simp: word_size) |
|
2981 |
done |
|
2982 |
||
2983 |
lemma take_sshiftr' [rule_format] : |
|
2984 |
"n <= size (w :: 'a :: len word) --> hd (to_bl (w >>> n)) = hd (to_bl w) & |
|
2985 |
take n (to_bl (w >>> n)) = replicate n (hd (to_bl w))" |
|
2986 |
apply (unfold sshiftr_def) |
|
2987 |
apply (induct n) |
|
2988 |
prefer 2 |
|
2989 |
apply (simp add: bl_sshiftr1) |
|
2990 |
apply (rule impI) |
|
2991 |
apply (rule take_butlast [THEN trans]) |
|
2992 |
apply (auto simp: word_size) |
|
2993 |
done |
|
2994 |
||
2995 |
lemmas hd_sshiftr = take_sshiftr' [THEN conjunct1, standard] |
|
2996 |
lemmas take_sshiftr = take_sshiftr' [THEN conjunct2, standard] |
|
2997 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
2998 |
lemma atd_lem: "take n xs = t \<Longrightarrow> drop n xs = d \<Longrightarrow> xs = t @ d" |
37660 | 2999 |
by (auto intro: append_take_drop_id [symmetric]) |
3000 |
||
3001 |
lemmas bl_shiftr = atd_lem [OF take_shiftr drop_shiftr] |
|
3002 |
lemmas bl_sshiftr = atd_lem [OF take_sshiftr drop_sshiftr] |
|
3003 |
||
3004 |
lemma shiftl_of_bl: "of_bl bl << n = of_bl (bl @ replicate n False)" |
|
3005 |
unfolding shiftl_def |
|
3006 |
by (induct n) (auto simp: shiftl1_of_bl replicate_app_Cons_same) |
|
3007 |
||
3008 |
lemma shiftl_bl: |
|
3009 |
"(w::'a::len0 word) << (n::nat) = of_bl (to_bl w @ replicate n False)" |
|
3010 |
proof - |
|
3011 |
have "w << n = of_bl (to_bl w) << n" by simp |
|
3012 |
also have "\<dots> = of_bl (to_bl w @ replicate n False)" by (rule shiftl_of_bl) |
|
3013 |
finally show ?thesis . |
|
3014 |
qed |
|
3015 |
||
3016 |
lemmas shiftl_number [simp] = shiftl_def [where w="number_of w", standard] |
|
3017 |
||
3018 |
lemma bl_shiftl: |
|
3019 |
"to_bl (w << n) = drop n (to_bl w) @ replicate (min (size w) n) False" |
|
3020 |
by (simp add: shiftl_bl word_rep_drop word_size) |
|
3021 |
||
3022 |
lemma shiftl_zero_size: |
|
3023 |
fixes x :: "'a::len0 word" |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3024 |
shows "size x <= n \<Longrightarrow> x << n = 0" |
37660 | 3025 |
apply (unfold word_size) |
3026 |
apply (rule word_eqI) |
|
3027 |
apply (clarsimp simp add: shiftl_bl word_size test_bit_of_bl nth_append) |
|
3028 |
done |
|
3029 |
||
3030 |
(* note - the following results use 'a :: len word < number_ring *) |
|
3031 |
||
3032 |
lemma shiftl1_2t: "shiftl1 (w :: 'a :: len word) = 2 * w" |
|
3033 |
apply (simp add: shiftl1_def_u) |
|
3034 |
apply (simp only: double_number_of_Bit0 [symmetric]) |
|
3035 |
apply simp |
|
3036 |
done |
|
3037 |
||
3038 |
lemma shiftl1_p: "shiftl1 (w :: 'a :: len word) = w + w" |
|
3039 |
apply (simp add: shiftl1_def_u) |
|
3040 |
apply (simp only: double_number_of_Bit0 [symmetric]) |
|
3041 |
apply simp |
|
3042 |
done |
|
3043 |
||
3044 |
lemma shiftl_t2n: "shiftl (w :: 'a :: len word) n = 2 ^ n * w" |
|
3045 |
unfolding shiftl_def |
|
3046 |
by (induct n) (auto simp: shiftl1_2t power_Suc) |
|
3047 |
||
3048 |
lemma shiftr1_bintr [simp]: |
|
3049 |
"(shiftr1 (number_of w) :: 'a :: len0 word) = |
|
3050 |
number_of (bin_rest (bintrunc (len_of TYPE ('a)) w))" |
|
3051 |
unfolding shiftr1_def word_number_of_def |
|
3052 |
by (simp add : word_ubin.eq_norm) |
|
3053 |
||
3054 |
lemma sshiftr1_sbintr [simp] : |
|
3055 |
"(sshiftr1 (number_of w) :: 'a :: len word) = |
|
3056 |
number_of (bin_rest (sbintrunc (len_of TYPE ('a) - 1) w))" |
|
3057 |
unfolding sshiftr1_def word_number_of_def |
|
3058 |
by (simp add : word_sbin.eq_norm) |
|
3059 |
||
3060 |
lemma shiftr_no': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3061 |
"w = number_of bin \<Longrightarrow> |
37660 | 3062 |
(w::'a::len0 word) >> n = number_of ((bin_rest ^^ n) (bintrunc (size w) bin))" |
3063 |
apply clarsimp |
|
3064 |
apply (rule word_eqI) |
|
3065 |
apply (auto simp: nth_shiftr nth_rest_power_bin nth_bintr word_size) |
|
3066 |
done |
|
3067 |
||
3068 |
lemma sshiftr_no': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3069 |
"w = number_of bin \<Longrightarrow> w >>> n = number_of ((bin_rest ^^ n) |
37660 | 3070 |
(sbintrunc (size w - 1) bin))" |
3071 |
apply clarsimp |
|
3072 |
apply (rule word_eqI) |
|
3073 |
apply (auto simp: nth_sshiftr nth_rest_power_bin nth_sbintr word_size) |
|
3074 |
apply (subgoal_tac "na + n = len_of TYPE('a) - Suc 0", simp, simp)+ |
|
3075 |
done |
|
3076 |
||
3077 |
lemmas sshiftr_no [simp] = |
|
3078 |
sshiftr_no' [where w = "number_of w", OF refl, unfolded word_size, standard] |
|
3079 |
||
3080 |
lemmas shiftr_no [simp] = |
|
3081 |
shiftr_no' [where w = "number_of w", OF refl, unfolded word_size, standard] |
|
3082 |
||
3083 |
lemma shiftr1_bl_of': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3084 |
"us = shiftr1 (of_bl bl) \<Longrightarrow> length bl <= size us \<Longrightarrow> |
37660 | 3085 |
us = of_bl (butlast bl)" |
3086 |
by (clarsimp simp: shiftr1_def of_bl_def word_size butlast_rest_bl2bin |
|
3087 |
word_ubin.eq_norm trunc_bl2bin) |
|
3088 |
||
3089 |
lemmas shiftr1_bl_of = refl [THEN shiftr1_bl_of', unfolded word_size] |
|
3090 |
||
3091 |
lemma shiftr_bl_of' [rule_format]: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3092 |
"us = of_bl bl >> n \<Longrightarrow> length bl <= size us --> |
37660 | 3093 |
us = of_bl (take (length bl - n) bl)" |
3094 |
apply (unfold shiftr_def) |
|
3095 |
apply hypsubst |
|
3096 |
apply (unfold word_size) |
|
3097 |
apply (induct n) |
|
3098 |
apply clarsimp |
|
3099 |
apply clarsimp |
|
3100 |
apply (subst shiftr1_bl_of) |
|
3101 |
apply simp |
|
3102 |
apply (simp add: butlast_take) |
|
3103 |
done |
|
3104 |
||
3105 |
lemmas shiftr_bl_of = refl [THEN shiftr_bl_of', unfolded word_size] |
|
3106 |
||
3107 |
lemmas shiftr_bl = word_bl.Rep' [THEN eq_imp_le, THEN shiftr_bl_of, |
|
3108 |
simplified word_size, simplified, THEN eq_reflection, standard] |
|
3109 |
||
3110 |
lemma msb_shift': "msb (w::'a::len word) <-> (w >> (size w - 1)) ~= 0" |
|
3111 |
apply (unfold shiftr_bl word_msb_alt) |
|
3112 |
apply (simp add: word_size Suc_le_eq take_Suc) |
|
3113 |
apply (cases "hd (to_bl w)") |
|
3114 |
apply (auto simp: word_1_bl word_0_bl |
|
3115 |
of_bl_rep_False [where n=1 and bs="[]", simplified]) |
|
3116 |
done |
|
3117 |
||
3118 |
lemmas msb_shift = msb_shift' [unfolded word_size] |
|
3119 |
||
3120 |
lemma align_lem_or [rule_format] : |
|
3121 |
"ALL x m. length x = n + m --> length y = n + m --> |
|
3122 |
drop m x = replicate n False --> take m y = replicate m False --> |
|
3123 |
map2 op | x y = take m x @ drop m y" |
|
3124 |
apply (induct_tac y) |
|
3125 |
apply force |
|
3126 |
apply clarsimp |
|
3127 |
apply (case_tac x, force) |
|
3128 |
apply (case_tac m, auto) |
|
3129 |
apply (drule sym) |
|
3130 |
apply auto |
|
3131 |
apply (induct_tac list, auto) |
|
3132 |
done |
|
3133 |
||
3134 |
lemma align_lem_and [rule_format] : |
|
3135 |
"ALL x m. length x = n + m --> length y = n + m --> |
|
3136 |
drop m x = replicate n False --> take m y = replicate m False --> |
|
3137 |
map2 op & x y = replicate (n + m) False" |
|
3138 |
apply (induct_tac y) |
|
3139 |
apply force |
|
3140 |
apply clarsimp |
|
3141 |
apply (case_tac x, force) |
|
3142 |
apply (case_tac m, auto) |
|
3143 |
apply (drule sym) |
|
3144 |
apply auto |
|
3145 |
apply (induct_tac list, auto) |
|
3146 |
done |
|
3147 |
||
3148 |
lemma aligned_bl_add_size': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3149 |
"size x - n = m \<Longrightarrow> n <= size x \<Longrightarrow> drop m (to_bl x) = replicate n False \<Longrightarrow> |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3150 |
take m (to_bl y) = replicate m False \<Longrightarrow> |
37660 | 3151 |
to_bl (x + y) = take m (to_bl x) @ drop m (to_bl y)" |
3152 |
apply (subgoal_tac "x AND y = 0") |
|
3153 |
prefer 2 |
|
3154 |
apply (rule word_bl.Rep_eqD) |
|
3155 |
apply (simp add: bl_word_and to_bl_0) |
|
3156 |
apply (rule align_lem_and [THEN trans]) |
|
3157 |
apply (simp_all add: word_size)[5] |
|
3158 |
apply simp |
|
3159 |
apply (subst word_plus_and_or [symmetric]) |
|
3160 |
apply (simp add : bl_word_or) |
|
3161 |
apply (rule align_lem_or) |
|
3162 |
apply (simp_all add: word_size) |
|
3163 |
done |
|
3164 |
||
3165 |
lemmas aligned_bl_add_size = refl [THEN aligned_bl_add_size'] |
|
3166 |
||
3167 |
subsubsection "Mask" |
|
3168 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3169 |
lemma nth_mask': "m = mask n \<Longrightarrow> test_bit m i = (i < n & i < size m)" |
37660 | 3170 |
apply (unfold mask_def test_bit_bl) |
3171 |
apply (simp only: word_1_bl [symmetric] shiftl_of_bl) |
|
3172 |
apply (clarsimp simp add: word_size) |
|
3173 |
apply (simp only: of_bl_no mask_lem number_of_succ add_diff_cancel2) |
|
3174 |
apply (fold of_bl_no) |
|
3175 |
apply (simp add: word_1_bl) |
|
3176 |
apply (rule test_bit_of_bl [THEN trans, unfolded test_bit_bl word_size]) |
|
3177 |
apply auto |
|
3178 |
done |
|
3179 |
||
3180 |
lemmas nth_mask [simp] = refl [THEN nth_mask'] |
|
3181 |
||
3182 |
lemma mask_bl: "mask n = of_bl (replicate n True)" |
|
3183 |
by (auto simp add : test_bit_of_bl word_size intro: word_eqI) |
|
3184 |
||
3185 |
lemma mask_bin: "mask n = number_of (bintrunc n Int.Min)" |
|
3186 |
by (auto simp add: nth_bintr word_size intro: word_eqI) |
|
3187 |
||
3188 |
lemma and_mask_bintr: "w AND mask n = number_of (bintrunc n (uint w))" |
|
3189 |
apply (rule word_eqI) |
|
3190 |
apply (simp add: nth_bintr word_size word_ops_nth_size) |
|
3191 |
apply (auto simp add: test_bit_bin) |
|
3192 |
done |
|
3193 |
||
3194 |
lemma and_mask_no: "number_of i AND mask n = number_of (bintrunc n i)" |
|
3195 |
by (auto simp add : nth_bintr word_size word_ops_nth_size intro: word_eqI) |
|
3196 |
||
3197 |
lemmas and_mask_wi = and_mask_no [unfolded word_number_of_def] |
|
3198 |
||
3199 |
lemma bl_and_mask': |
|
3200 |
"to_bl (w AND mask n :: 'a :: len word) = |
|
3201 |
replicate (len_of TYPE('a) - n) False @ |
|
3202 |
drop (len_of TYPE('a) - n) (to_bl w)" |
|
3203 |
apply (rule nth_equalityI) |
|
3204 |
apply simp |
|
3205 |
apply (clarsimp simp add: to_bl_nth word_size) |
|
3206 |
apply (simp add: word_size word_ops_nth_size) |
|
3207 |
apply (auto simp add: word_size test_bit_bl nth_append nth_rev) |
|
3208 |
done |
|
3209 |
||
3210 |
lemmas and_mask_mod_2p = |
|
3211 |
and_mask_bintr [unfolded word_number_of_alt no_bintr_alt] |
|
3212 |
||
3213 |
lemma and_mask_lt_2p: "uint (w AND mask n) < 2 ^ n" |
|
3214 |
apply (simp add : and_mask_bintr no_bintr_alt) |
|
3215 |
apply (rule xtr8) |
|
3216 |
prefer 2 |
|
3217 |
apply (rule pos_mod_bound) |
|
3218 |
apply auto |
|
3219 |
done |
|
3220 |
||
3221 |
lemmas eq_mod_iff = trans [symmetric, OF int_mod_lem eq_sym_conv] |
|
3222 |
||
3223 |
lemma mask_eq_iff: "(w AND mask n) = w <-> uint w < 2 ^ n" |
|
3224 |
apply (simp add: and_mask_bintr word_number_of_def) |
|
3225 |
apply (simp add: word_ubin.inverse_norm) |
|
3226 |
apply (simp add: eq_mod_iff bintrunc_mod2p min_def) |
|
3227 |
apply (fast intro!: lt2p_lem) |
|
3228 |
done |
|
3229 |
||
3230 |
lemma and_mask_dvd: "2 ^ n dvd uint w = (w AND mask n = 0)" |
|
3231 |
apply (simp add: dvd_eq_mod_eq_0 and_mask_mod_2p) |
|
3232 |
apply (simp add: word_uint.norm_eq_iff [symmetric] word_of_int_homs) |
|
3233 |
apply (subst word_uint.norm_Rep [symmetric]) |
|
3234 |
apply (simp only: bintrunc_bintrunc_min bintrunc_mod2p [symmetric] min_def) |
|
3235 |
apply auto |
|
3236 |
done |
|
3237 |
||
3238 |
lemma and_mask_dvd_nat: "2 ^ n dvd unat w = (w AND mask n = 0)" |
|
3239 |
apply (unfold unat_def) |
|
3240 |
apply (rule trans [OF _ and_mask_dvd]) |
|
3241 |
apply (unfold dvd_def) |
|
3242 |
apply auto |
|
3243 |
apply (drule uint_ge_0 [THEN nat_int.Abs_inverse' [simplified], symmetric]) |
|
3244 |
apply (simp add : int_mult int_power) |
|
3245 |
apply (simp add : nat_mult_distrib nat_power_eq) |
|
3246 |
done |
|
3247 |
||
3248 |
lemma word_2p_lem: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3249 |
"n < size w \<Longrightarrow> w < 2 ^ n = (uint (w :: 'a :: len word) < 2 ^ n)" |
37660 | 3250 |
apply (unfold word_size word_less_alt word_number_of_alt) |
3251 |
apply (clarsimp simp add: word_of_int_power_hom word_uint.eq_norm |
|
3252 |
int_mod_eq' |
|
3253 |
simp del: word_of_int_bin) |
|
3254 |
done |
|
3255 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3256 |
lemma less_mask_eq: "x < 2 ^ n \<Longrightarrow> x AND mask n = (x :: 'a :: len word)" |
37660 | 3257 |
apply (unfold word_less_alt word_number_of_alt) |
3258 |
apply (clarsimp simp add: and_mask_mod_2p word_of_int_power_hom |
|
3259 |
word_uint.eq_norm |
|
3260 |
simp del: word_of_int_bin) |
|
3261 |
apply (drule xtr8 [rotated]) |
|
3262 |
apply (rule int_mod_le) |
|
3263 |
apply (auto simp add : mod_pos_pos_trivial) |
|
3264 |
done |
|
3265 |
||
3266 |
lemmas mask_eq_iff_w2p = |
|
3267 |
trans [OF mask_eq_iff word_2p_lem [symmetric], standard] |
|
3268 |
||
3269 |
lemmas and_mask_less' = |
|
3270 |
iffD2 [OF word_2p_lem and_mask_lt_2p, simplified word_size, standard] |
|
3271 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3272 |
lemma and_mask_less_size: "n < size x \<Longrightarrow> x AND mask n < 2^n" |
37660 | 3273 |
unfolding word_size by (erule and_mask_less') |
3274 |
||
3275 |
lemma word_mod_2p_is_mask': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3276 |
"c = 2 ^ n \<Longrightarrow> c > 0 \<Longrightarrow> x mod c = (x :: 'a :: len word) AND mask n" |
37660 | 3277 |
by (clarsimp simp add: word_mod_def uint_2p and_mask_mod_2p) |
3278 |
||
3279 |
lemmas word_mod_2p_is_mask = refl [THEN word_mod_2p_is_mask'] |
|
3280 |
||
3281 |
lemma mask_eqs: |
|
3282 |
"(a AND mask n) + b AND mask n = a + b AND mask n" |
|
3283 |
"a + (b AND mask n) AND mask n = a + b AND mask n" |
|
3284 |
"(a AND mask n) - b AND mask n = a - b AND mask n" |
|
3285 |
"a - (b AND mask n) AND mask n = a - b AND mask n" |
|
3286 |
"a * (b AND mask n) AND mask n = a * b AND mask n" |
|
3287 |
"(b AND mask n) * a AND mask n = b * a AND mask n" |
|
3288 |
"(a AND mask n) + (b AND mask n) AND mask n = a + b AND mask n" |
|
3289 |
"(a AND mask n) - (b AND mask n) AND mask n = a - b AND mask n" |
|
3290 |
"(a AND mask n) * (b AND mask n) AND mask n = a * b AND mask n" |
|
3291 |
"- (a AND mask n) AND mask n = - a AND mask n" |
|
3292 |
"word_succ (a AND mask n) AND mask n = word_succ a AND mask n" |
|
3293 |
"word_pred (a AND mask n) AND mask n = word_pred a AND mask n" |
|
3294 |
using word_of_int_Ex [where x=a] word_of_int_Ex [where x=b] |
|
3295 |
by (auto simp: and_mask_wi bintr_ariths bintr_arith1s new_word_of_int_homs) |
|
3296 |
||
3297 |
lemma mask_power_eq: |
|
3298 |
"(x AND mask n) ^ k AND mask n = x ^ k AND mask n" |
|
3299 |
using word_of_int_Ex [where x=x] |
|
3300 |
by (clarsimp simp: and_mask_wi word_of_int_power_hom bintr_ariths) |
|
3301 |
||
3302 |
||
3303 |
subsubsection "Revcast" |
|
3304 |
||
3305 |
lemmas revcast_def' = revcast_def [simplified] |
|
3306 |
lemmas revcast_def'' = revcast_def' [simplified word_size] |
|
3307 |
lemmas revcast_no_def [simp] = |
|
3308 |
revcast_def' [where w="number_of w", unfolded word_size, standard] |
|
3309 |
||
3310 |
lemma to_bl_revcast: |
|
3311 |
"to_bl (revcast w :: 'a :: len0 word) = |
|
3312 |
takefill False (len_of TYPE ('a)) (to_bl w)" |
|
3313 |
apply (unfold revcast_def' word_size) |
|
3314 |
apply (rule word_bl.Abs_inverse) |
|
3315 |
apply simp |
|
3316 |
done |
|
3317 |
||
3318 |
lemma revcast_rev_ucast': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3319 |
"cs = [rc, uc] \<Longrightarrow> rc = revcast (word_reverse w) \<Longrightarrow> uc = ucast w \<Longrightarrow> |
37660 | 3320 |
rc = word_reverse uc" |
3321 |
apply (unfold ucast_def revcast_def' Let_def word_reverse_def) |
|
3322 |
apply (clarsimp simp add : to_bl_of_bin takefill_bintrunc) |
|
3323 |
apply (simp add : word_bl.Abs_inverse word_size) |
|
3324 |
done |
|
3325 |
||
3326 |
lemmas revcast_rev_ucast = revcast_rev_ucast' [OF refl refl refl] |
|
3327 |
||
3328 |
lemmas revcast_ucast = revcast_rev_ucast |
|
3329 |
[where w = "word_reverse w", simplified word_rev_rev, standard] |
|
3330 |
||
3331 |
lemmas ucast_revcast = revcast_rev_ucast [THEN word_rev_gal', standard] |
|
3332 |
lemmas ucast_rev_revcast = revcast_ucast [THEN word_rev_gal', standard] |
|
3333 |
||
3334 |
||
3335 |
-- "linking revcast and cast via shift" |
|
3336 |
||
3337 |
lemmas wsst_TYs = source_size target_size word_size |
|
3338 |
||
3339 |
lemma revcast_down_uu': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3340 |
"rc = revcast \<Longrightarrow> source_size rc = target_size rc + n \<Longrightarrow> |
37660 | 3341 |
rc (w :: 'a :: len word) = ucast (w >> n)" |
3342 |
apply (simp add: revcast_def') |
|
3343 |
apply (rule word_bl.Rep_inverse') |
|
3344 |
apply (rule trans, rule ucast_down_drop) |
|
3345 |
prefer 2 |
|
3346 |
apply (rule trans, rule drop_shiftr) |
|
3347 |
apply (auto simp: takefill_alt wsst_TYs) |
|
3348 |
done |
|
3349 |
||
3350 |
lemma revcast_down_us': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3351 |
"rc = revcast \<Longrightarrow> source_size rc = target_size rc + n \<Longrightarrow> |
37660 | 3352 |
rc (w :: 'a :: len word) = ucast (w >>> n)" |
3353 |
apply (simp add: revcast_def') |
|
3354 |
apply (rule word_bl.Rep_inverse') |
|
3355 |
apply (rule trans, rule ucast_down_drop) |
|
3356 |
prefer 2 |
|
3357 |
apply (rule trans, rule drop_sshiftr) |
|
3358 |
apply (auto simp: takefill_alt wsst_TYs) |
|
3359 |
done |
|
3360 |
||
3361 |
lemma revcast_down_su': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3362 |
"rc = revcast \<Longrightarrow> source_size rc = target_size rc + n \<Longrightarrow> |
37660 | 3363 |
rc (w :: 'a :: len word) = scast (w >> n)" |
3364 |
apply (simp add: revcast_def') |
|
3365 |
apply (rule word_bl.Rep_inverse') |
|
3366 |
apply (rule trans, rule scast_down_drop) |
|
3367 |
prefer 2 |
|
3368 |
apply (rule trans, rule drop_shiftr) |
|
3369 |
apply (auto simp: takefill_alt wsst_TYs) |
|
3370 |
done |
|
3371 |
||
3372 |
lemma revcast_down_ss': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3373 |
"rc = revcast \<Longrightarrow> source_size rc = target_size rc + n \<Longrightarrow> |
37660 | 3374 |
rc (w :: 'a :: len word) = scast (w >>> n)" |
3375 |
apply (simp add: revcast_def') |
|
3376 |
apply (rule word_bl.Rep_inverse') |
|
3377 |
apply (rule trans, rule scast_down_drop) |
|
3378 |
prefer 2 |
|
3379 |
apply (rule trans, rule drop_sshiftr) |
|
3380 |
apply (auto simp: takefill_alt wsst_TYs) |
|
3381 |
done |
|
3382 |
||
3383 |
lemmas revcast_down_uu = refl [THEN revcast_down_uu'] |
|
3384 |
lemmas revcast_down_us = refl [THEN revcast_down_us'] |
|
3385 |
lemmas revcast_down_su = refl [THEN revcast_down_su'] |
|
3386 |
lemmas revcast_down_ss = refl [THEN revcast_down_ss'] |
|
3387 |
||
3388 |
lemma cast_down_rev: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3389 |
"uc = ucast \<Longrightarrow> source_size uc = target_size uc + n \<Longrightarrow> |
37660 | 3390 |
uc w = revcast ((w :: 'a :: len word) << n)" |
3391 |
apply (unfold shiftl_rev) |
|
3392 |
apply clarify |
|
3393 |
apply (simp add: revcast_rev_ucast) |
|
3394 |
apply (rule word_rev_gal') |
|
3395 |
apply (rule trans [OF _ revcast_rev_ucast]) |
|
3396 |
apply (rule revcast_down_uu [symmetric]) |
|
3397 |
apply (auto simp add: wsst_TYs) |
|
3398 |
done |
|
3399 |
||
3400 |
lemma revcast_up': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3401 |
"rc = revcast \<Longrightarrow> source_size rc + n = target_size rc \<Longrightarrow> |
37660 | 3402 |
rc w = (ucast w :: 'a :: len word) << n" |
3403 |
apply (simp add: revcast_def') |
|
3404 |
apply (rule word_bl.Rep_inverse') |
|
3405 |
apply (simp add: takefill_alt) |
|
3406 |
apply (rule bl_shiftl [THEN trans]) |
|
3407 |
apply (subst ucast_up_app) |
|
3408 |
apply (auto simp add: wsst_TYs) |
|
3409 |
done |
|
3410 |
||
3411 |
lemmas revcast_up = refl [THEN revcast_up'] |
|
3412 |
||
3413 |
lemmas rc1 = revcast_up [THEN |
|
3414 |
revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] |
|
3415 |
lemmas rc2 = revcast_down_uu [THEN |
|
3416 |
revcast_rev_ucast [symmetric, THEN trans, THEN word_rev_gal, symmetric]] |
|
3417 |
||
3418 |
lemmas ucast_up = |
|
3419 |
rc1 [simplified rev_shiftr [symmetric] revcast_ucast [symmetric]] |
|
3420 |
lemmas ucast_down = |
|
3421 |
rc2 [simplified rev_shiftr revcast_ucast [symmetric]] |
|
3422 |
||
3423 |
||
3424 |
subsubsection "Slices" |
|
3425 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3426 |
lemma slice1_no_bin [simp]: |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3427 |
"slice1 n (number_of w :: 'b word) = of_bl (takefill False n (bin_to_bl (len_of TYPE('b :: len0)) w))" |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3428 |
by (simp add: slice1_def) |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3429 |
|
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3430 |
lemma slice_no_bin [simp]: |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3431 |
"slice n (number_of w :: 'b word) = of_bl (takefill False (len_of TYPE('b :: len0) - n) |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3432 |
(bin_to_bl (len_of TYPE('b :: len0)) w))" |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3433 |
by (simp add: slice_def word_size) |
37660 | 3434 |
|
3435 |
lemma slice1_0 [simp] : "slice1 n 0 = 0" |
|
3436 |
unfolding slice1_def by (simp add : to_bl_0) |
|
3437 |
||
3438 |
lemma slice_0 [simp] : "slice n 0 = 0" |
|
3439 |
unfolding slice_def by auto |
|
3440 |
||
3441 |
lemma slice_take': "slice n w = of_bl (take (size w - n) (to_bl w))" |
|
3442 |
unfolding slice_def' slice1_def |
|
3443 |
by (simp add : takefill_alt word_size) |
|
3444 |
||
3445 |
lemmas slice_take = slice_take' [unfolded word_size] |
|
3446 |
||
3447 |
-- "shiftr to a word of the same size is just slice, |
|
3448 |
slice is just shiftr then ucast" |
|
3449 |
lemmas shiftr_slice = trans |
|
3450 |
[OF shiftr_bl [THEN meta_eq_to_obj_eq] slice_take [symmetric], standard] |
|
3451 |
||
3452 |
lemma slice_shiftr: "slice n w = ucast (w >> n)" |
|
3453 |
apply (unfold slice_take shiftr_bl) |
|
3454 |
apply (rule ucast_of_bl_up [symmetric]) |
|
3455 |
apply (simp add: word_size) |
|
3456 |
done |
|
3457 |
||
3458 |
lemma nth_slice: |
|
3459 |
"(slice n w :: 'a :: len0 word) !! m = |
|
3460 |
(w !! (m + n) & m < len_of TYPE ('a))" |
|
3461 |
unfolding slice_shiftr |
|
3462 |
by (simp add : nth_ucast nth_shiftr) |
|
3463 |
||
3464 |
lemma slice1_down_alt': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3465 |
"sl = slice1 n w \<Longrightarrow> fs = size sl \<Longrightarrow> fs + k = n \<Longrightarrow> |
37660 | 3466 |
to_bl sl = takefill False fs (drop k (to_bl w))" |
3467 |
unfolding slice1_def word_size of_bl_def uint_bl |
|
3468 |
by (clarsimp simp: word_ubin.eq_norm bl_bin_bl_rep_drop drop_takefill) |
|
3469 |
||
3470 |
lemma slice1_up_alt': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3471 |
"sl = slice1 n w \<Longrightarrow> fs = size sl \<Longrightarrow> fs = n + k \<Longrightarrow> |
37660 | 3472 |
to_bl sl = takefill False fs (replicate k False @ (to_bl w))" |
3473 |
apply (unfold slice1_def word_size of_bl_def uint_bl) |
|
3474 |
apply (clarsimp simp: word_ubin.eq_norm bl_bin_bl_rep_drop |
|
3475 |
takefill_append [symmetric]) |
|
3476 |
apply (rule_tac f = "%k. takefill False (len_of TYPE('a)) |
|
3477 |
(replicate k False @ bin_to_bl (len_of TYPE('b)) (uint w))" in arg_cong) |
|
3478 |
apply arith |
|
3479 |
done |
|
3480 |
||
3481 |
lemmas sd1 = slice1_down_alt' [OF refl refl, unfolded word_size] |
|
3482 |
lemmas su1 = slice1_up_alt' [OF refl refl, unfolded word_size] |
|
3483 |
lemmas slice1_down_alt = le_add_diff_inverse [THEN sd1] |
|
3484 |
lemmas slice1_up_alts = |
|
3485 |
le_add_diff_inverse [symmetric, THEN su1] |
|
3486 |
le_add_diff_inverse2 [symmetric, THEN su1] |
|
3487 |
||
3488 |
lemma ucast_slice1: "ucast w = slice1 (size w) w" |
|
3489 |
unfolding slice1_def ucast_bl |
|
3490 |
by (simp add : takefill_same' word_size) |
|
3491 |
||
3492 |
lemma ucast_slice: "ucast w = slice 0 w" |
|
3493 |
unfolding slice_def by (simp add : ucast_slice1) |
|
3494 |
||
3495 |
lemmas slice_id = trans [OF ucast_slice [symmetric] ucast_id] |
|
3496 |
||
3497 |
lemma revcast_slice1': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3498 |
"rc = revcast w \<Longrightarrow> slice1 (size rc) w = rc" |
37660 | 3499 |
unfolding slice1_def revcast_def' by (simp add : word_size) |
3500 |
||
3501 |
lemmas revcast_slice1 = refl [THEN revcast_slice1'] |
|
3502 |
||
3503 |
lemma slice1_tf_tf': |
|
3504 |
"to_bl (slice1 n w :: 'a :: len0 word) = |
|
3505 |
rev (takefill False (len_of TYPE('a)) (rev (takefill False n (to_bl w))))" |
|
3506 |
unfolding slice1_def by (rule word_rev_tf) |
|
3507 |
||
3508 |
lemmas slice1_tf_tf = slice1_tf_tf' |
|
3509 |
[THEN word_bl.Rep_inverse', symmetric, standard] |
|
3510 |
||
3511 |
lemma rev_slice1: |
|
3512 |
"n + k = len_of TYPE('a) + len_of TYPE('b) \<Longrightarrow> |
|
3513 |
slice1 n (word_reverse w :: 'b :: len0 word) = |
|
3514 |
word_reverse (slice1 k w :: 'a :: len0 word)" |
|
3515 |
apply (unfold word_reverse_def slice1_tf_tf) |
|
3516 |
apply (rule word_bl.Rep_inverse') |
|
3517 |
apply (rule rev_swap [THEN iffD1]) |
|
3518 |
apply (rule trans [symmetric]) |
|
3519 |
apply (rule tf_rev) |
|
3520 |
apply (simp add: word_bl.Abs_inverse) |
|
3521 |
apply (simp add: word_bl.Abs_inverse) |
|
3522 |
done |
|
3523 |
||
3524 |
lemma rev_slice': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3525 |
"res = slice n (word_reverse w) \<Longrightarrow> n + k + size res = size w \<Longrightarrow> |
37660 | 3526 |
res = word_reverse (slice k w)" |
3527 |
apply (unfold slice_def word_size) |
|
3528 |
apply clarify |
|
3529 |
apply (rule rev_slice1) |
|
3530 |
apply arith |
|
3531 |
done |
|
3532 |
||
3533 |
lemmas rev_slice = refl [THEN rev_slice', unfolded word_size] |
|
3534 |
||
3535 |
lemmas sym_notr = |
|
3536 |
not_iff [THEN iffD2, THEN not_sym, THEN not_iff [THEN iffD1]] |
|
3537 |
||
3538 |
-- {* problem posed by TPHOLs referee: |
|
3539 |
criterion for overflow of addition of signed integers *} |
|
3540 |
||
3541 |
lemma sofl_test: |
|
3542 |
"(sint (x :: 'a :: len word) + sint y = sint (x + y)) = |
|
3543 |
((((x+y) XOR x) AND ((x+y) XOR y)) >> (size x - 1) = 0)" |
|
3544 |
apply (unfold word_size) |
|
3545 |
apply (cases "len_of TYPE('a)", simp) |
|
3546 |
apply (subst msb_shift [THEN sym_notr]) |
|
3547 |
apply (simp add: word_ops_msb) |
|
3548 |
apply (simp add: word_msb_sint) |
|
3549 |
apply safe |
|
3550 |
apply simp_all |
|
3551 |
apply (unfold sint_word_ariths) |
|
3552 |
apply (unfold word_sbin.set_iff_norm [symmetric] sints_num) |
|
3553 |
apply safe |
|
3554 |
apply (insert sint_range' [where x=x]) |
|
3555 |
apply (insert sint_range' [where x=y]) |
|
3556 |
defer |
|
3557 |
apply (simp (no_asm), arith) |
|
3558 |
apply (simp (no_asm), arith) |
|
3559 |
defer |
|
3560 |
defer |
|
3561 |
apply (simp (no_asm), arith) |
|
3562 |
apply (simp (no_asm), arith) |
|
3563 |
apply (rule notI [THEN notnotD], |
|
3564 |
drule leI not_leE, |
|
3565 |
drule sbintrunc_inc sbintrunc_dec, |
|
3566 |
simp)+ |
|
3567 |
done |
|
3568 |
||
3569 |
||
3570 |
subsection "Split and cat" |
|
3571 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3572 |
lemmas word_split_bin' = word_split_def |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3573 |
lemmas word_cat_bin' = word_cat_def |
37660 | 3574 |
|
3575 |
lemma word_rsplit_no: |
|
3576 |
"(word_rsplit (number_of bin :: 'b :: len0 word) :: 'a word list) = |
|
3577 |
map number_of (bin_rsplit (len_of TYPE('a :: len)) |
|
3578 |
(len_of TYPE('b), bintrunc (len_of TYPE('b)) bin))" |
|
3579 |
apply (unfold word_rsplit_def word_no_wi) |
|
3580 |
apply (simp add: word_ubin.eq_norm) |
|
3581 |
done |
|
3582 |
||
3583 |
lemmas word_rsplit_no_cl [simp] = word_rsplit_no |
|
3584 |
[unfolded bin_rsplitl_def bin_rsplit_l [symmetric]] |
|
3585 |
||
3586 |
lemma test_bit_cat: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3587 |
"wc = word_cat a b \<Longrightarrow> wc !! n = (n < size wc & |
37660 | 3588 |
(if n < size b then b !! n else a !! (n - size b)))" |
3589 |
apply (unfold word_cat_bin' test_bit_bin) |
|
3590 |
apply (auto simp add : word_ubin.eq_norm nth_bintr bin_nth_cat word_size) |
|
3591 |
apply (erule bin_nth_uint_imp) |
|
3592 |
done |
|
3593 |
||
3594 |
lemma word_cat_bl: "word_cat a b = of_bl (to_bl a @ to_bl b)" |
|
3595 |
apply (unfold of_bl_def to_bl_def word_cat_bin') |
|
3596 |
apply (simp add: bl_to_bin_app_cat) |
|
3597 |
done |
|
3598 |
||
3599 |
lemma of_bl_append: |
|
3600 |
"(of_bl (xs @ ys) :: 'a :: len word) = of_bl xs * 2^(length ys) + of_bl ys" |
|
3601 |
apply (unfold of_bl_def) |
|
3602 |
apply (simp add: bl_to_bin_app_cat bin_cat_num) |
|
3603 |
apply (simp add: word_of_int_power_hom [symmetric] new_word_of_int_hom_syms) |
|
3604 |
done |
|
3605 |
||
3606 |
lemma of_bl_False [simp]: |
|
3607 |
"of_bl (False#xs) = of_bl xs" |
|
3608 |
by (rule word_eqI) |
|
3609 |
(auto simp add: test_bit_of_bl nth_append) |
|
3610 |
||
3611 |
lemma of_bl_True: |
|
3612 |
"(of_bl (True#xs)::'a::len word) = 2^length xs + of_bl xs" |
|
3613 |
by (subst of_bl_append [where xs="[True]", simplified]) |
|
3614 |
(simp add: word_1_bl) |
|
3615 |
||
3616 |
lemma of_bl_Cons: |
|
3617 |
"of_bl (x#xs) = of_bool x * 2^length xs + of_bl xs" |
|
3618 |
by (cases x) (simp_all add: of_bl_True) |
|
3619 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3620 |
lemma split_uint_lem: "bin_split n (uint (w :: 'a :: len0 word)) = (a, b) \<Longrightarrow> |
37660 | 3621 |
a = bintrunc (len_of TYPE('a) - n) a & b = bintrunc (len_of TYPE('a)) b" |
3622 |
apply (frule word_ubin.norm_Rep [THEN ssubst]) |
|
3623 |
apply (drule bin_split_trunc1) |
|
3624 |
apply (drule sym [THEN trans]) |
|
3625 |
apply assumption |
|
3626 |
apply safe |
|
3627 |
done |
|
3628 |
||
3629 |
lemma word_split_bl': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3630 |
"std = size c - size b \<Longrightarrow> (word_split c = (a, b)) \<Longrightarrow> |
37660 | 3631 |
(a = of_bl (take std (to_bl c)) & b = of_bl (drop std (to_bl c)))" |
3632 |
apply (unfold word_split_bin') |
|
3633 |
apply safe |
|
3634 |
defer |
|
3635 |
apply (clarsimp split: prod.splits) |
|
3636 |
apply (drule word_ubin.norm_Rep [THEN ssubst]) |
|
3637 |
apply (drule split_bintrunc) |
|
3638 |
apply (simp add : of_bl_def bl2bin_drop word_size |
|
3639 |
word_ubin.norm_eq_iff [symmetric] min_def del : word_ubin.norm_Rep) |
|
3640 |
apply (clarsimp split: prod.splits) |
|
3641 |
apply (frule split_uint_lem [THEN conjunct1]) |
|
3642 |
apply (unfold word_size) |
|
3643 |
apply (cases "len_of TYPE('a) >= len_of TYPE('b)") |
|
3644 |
defer |
|
3645 |
apply (simp add: word_0_bl word_0_wi_Pls) |
|
3646 |
apply (simp add : of_bl_def to_bl_def) |
|
3647 |
apply (subst bin_split_take1 [symmetric]) |
|
3648 |
prefer 2 |
|
3649 |
apply assumption |
|
3650 |
apply simp |
|
3651 |
apply (erule thin_rl) |
|
3652 |
apply (erule arg_cong [THEN trans]) |
|
3653 |
apply (simp add : word_ubin.norm_eq_iff [symmetric]) |
|
3654 |
done |
|
3655 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3656 |
lemma word_split_bl: "std = size c - size b \<Longrightarrow> |
37660 | 3657 |
(a = of_bl (take std (to_bl c)) & b = of_bl (drop std (to_bl c))) <-> |
3658 |
word_split c = (a, b)" |
|
3659 |
apply (rule iffI) |
|
3660 |
defer |
|
3661 |
apply (erule (1) word_split_bl') |
|
3662 |
apply (case_tac "word_split c") |
|
3663 |
apply (auto simp add : word_size) |
|
3664 |
apply (frule word_split_bl' [rotated]) |
|
3665 |
apply (auto simp add : word_size) |
|
3666 |
done |
|
3667 |
||
3668 |
lemma word_split_bl_eq: |
|
3669 |
"(word_split (c::'a::len word) :: ('c :: len0 word * 'd :: len0 word)) = |
|
3670 |
(of_bl (take (len_of TYPE('a::len) - len_of TYPE('d::len0)) (to_bl c)), |
|
3671 |
of_bl (drop (len_of TYPE('a) - len_of TYPE('d)) (to_bl c)))" |
|
3672 |
apply (rule word_split_bl [THEN iffD1]) |
|
3673 |
apply (unfold word_size) |
|
3674 |
apply (rule refl conjI)+ |
|
3675 |
done |
|
3676 |
||
3677 |
-- "keep quantifiers for use in simplification" |
|
3678 |
lemma test_bit_split': |
|
3679 |
"word_split c = (a, b) --> (ALL n m. b !! n = (n < size b & c !! n) & |
|
3680 |
a !! m = (m < size a & c !! (m + size b)))" |
|
3681 |
apply (unfold word_split_bin' test_bit_bin) |
|
3682 |
apply (clarify) |
|
3683 |
apply (clarsimp simp: word_ubin.eq_norm nth_bintr word_size split: prod.splits) |
|
3684 |
apply (drule bin_nth_split) |
|
3685 |
apply safe |
|
3686 |
apply (simp_all add: add_commute) |
|
3687 |
apply (erule bin_nth_uint_imp)+ |
|
3688 |
done |
|
3689 |
||
3690 |
lemma test_bit_split: |
|
3691 |
"word_split c = (a, b) \<Longrightarrow> |
|
3692 |
(\<forall>n\<Colon>nat. b !! n \<longleftrightarrow> n < size b \<and> c !! n) \<and> (\<forall>m\<Colon>nat. a !! m \<longleftrightarrow> m < size a \<and> c !! (m + size b))" |
|
3693 |
by (simp add: test_bit_split') |
|
3694 |
||
3695 |
lemma test_bit_split_eq: "word_split c = (a, b) <-> |
|
3696 |
((ALL n::nat. b !! n = (n < size b & c !! n)) & |
|
3697 |
(ALL m::nat. a !! m = (m < size a & c !! (m + size b))))" |
|
3698 |
apply (rule_tac iffI) |
|
3699 |
apply (rule_tac conjI) |
|
3700 |
apply (erule test_bit_split [THEN conjunct1]) |
|
3701 |
apply (erule test_bit_split [THEN conjunct2]) |
|
3702 |
apply (case_tac "word_split c") |
|
3703 |
apply (frule test_bit_split) |
|
3704 |
apply (erule trans) |
|
3705 |
apply (fastsimp intro ! : word_eqI simp add : word_size) |
|
3706 |
done |
|
3707 |
||
3708 |
-- {* this odd result is analogous to @{text "ucast_id"}, |
|
3709 |
result to the length given by the result type *} |
|
3710 |
||
3711 |
lemma word_cat_id: "word_cat a b = b" |
|
3712 |
unfolding word_cat_bin' by (simp add: word_ubin.inverse_norm) |
|
3713 |
||
3714 |
-- "limited hom result" |
|
3715 |
lemma word_cat_hom: |
|
3716 |
"len_of TYPE('a::len0) <= len_of TYPE('b::len0) + len_of TYPE ('c::len0) |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3717 |
\<Longrightarrow> |
37660 | 3718 |
(word_cat (word_of_int w :: 'b word) (b :: 'c word) :: 'a word) = |
3719 |
word_of_int (bin_cat w (size b) (uint b))" |
|
3720 |
apply (unfold word_cat_def word_size) |
|
3721 |
apply (clarsimp simp add: word_ubin.norm_eq_iff [symmetric] |
|
3722 |
word_ubin.eq_norm bintr_cat min_max.inf_absorb1) |
|
3723 |
done |
|
3724 |
||
3725 |
lemma word_cat_split_alt: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3726 |
"size w <= size u + size v \<Longrightarrow> word_split w = (u, v) \<Longrightarrow> word_cat u v = w" |
37660 | 3727 |
apply (rule word_eqI) |
3728 |
apply (drule test_bit_split) |
|
3729 |
apply (clarsimp simp add : test_bit_cat word_size) |
|
3730 |
apply safe |
|
3731 |
apply arith |
|
3732 |
done |
|
3733 |
||
3734 |
lemmas word_cat_split_size = |
|
3735 |
sym [THEN [2] word_cat_split_alt [symmetric], standard] |
|
3736 |
||
3737 |
||
3738 |
subsubsection "Split and slice" |
|
3739 |
||
3740 |
lemma split_slices: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3741 |
"word_split w = (u, v) \<Longrightarrow> u = slice (size v) w & v = slice 0 w" |
37660 | 3742 |
apply (drule test_bit_split) |
3743 |
apply (rule conjI) |
|
3744 |
apply (rule word_eqI, clarsimp simp: nth_slice word_size)+ |
|
3745 |
done |
|
3746 |
||
3747 |
lemma slice_cat1': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3748 |
"wc = word_cat a b \<Longrightarrow> size wc >= size a + size b \<Longrightarrow> slice (size b) wc = a" |
37660 | 3749 |
apply safe |
3750 |
apply (rule word_eqI) |
|
3751 |
apply (simp add: nth_slice test_bit_cat word_size) |
|
3752 |
done |
|
3753 |
||
3754 |
lemmas slice_cat1 = refl [THEN slice_cat1'] |
|
3755 |
lemmas slice_cat2 = trans [OF slice_id word_cat_id] |
|
3756 |
||
3757 |
lemma cat_slices: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3758 |
"a = slice n c \<Longrightarrow> b = slice 0 c \<Longrightarrow> n = size b \<Longrightarrow> |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3759 |
size a + size b >= size c \<Longrightarrow> word_cat a b = c" |
37660 | 3760 |
apply safe |
3761 |
apply (rule word_eqI) |
|
3762 |
apply (simp add: nth_slice test_bit_cat word_size) |
|
3763 |
apply safe |
|
3764 |
apply arith |
|
3765 |
done |
|
3766 |
||
3767 |
lemma word_split_cat_alt: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3768 |
"w = word_cat u v \<Longrightarrow> size u + size v <= size w \<Longrightarrow> word_split w = (u, v)" |
37660 | 3769 |
apply (case_tac "word_split ?w") |
3770 |
apply (rule trans, assumption) |
|
3771 |
apply (drule test_bit_split) |
|
3772 |
apply safe |
|
3773 |
apply (rule word_eqI, clarsimp simp: test_bit_cat word_size)+ |
|
3774 |
done |
|
3775 |
||
3776 |
lemmas word_cat_bl_no_bin [simp] = |
|
3777 |
word_cat_bl [where a="number_of a" |
|
3778 |
and b="number_of b", |
|
3779 |
unfolded to_bl_no_bin, standard] |
|
3780 |
||
3781 |
lemmas word_split_bl_no_bin [simp] = |
|
3782 |
word_split_bl_eq [where c="number_of c", unfolded to_bl_no_bin, standard] |
|
3783 |
||
3784 |
-- {* this odd result arises from the fact that the statement of the |
|
3785 |
result implies that the decoded words are of the same type, |
|
3786 |
and therefore of the same length, as the original word *} |
|
3787 |
||
3788 |
lemma word_rsplit_same: "word_rsplit w = [w]" |
|
3789 |
unfolding word_rsplit_def by (simp add : bin_rsplit_all) |
|
3790 |
||
3791 |
lemma word_rsplit_empty_iff_size: |
|
3792 |
"(word_rsplit w = []) = (size w = 0)" |
|
3793 |
unfolding word_rsplit_def bin_rsplit_def word_size |
|
3794 |
by (simp add: bin_rsplit_aux_simp_alt Let_def split: Product_Type.split_split) |
|
3795 |
||
3796 |
lemma test_bit_rsplit: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3797 |
"sw = word_rsplit w \<Longrightarrow> m < size (hd sw :: 'a :: len word) \<Longrightarrow> |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3798 |
k < length sw \<Longrightarrow> (rev sw ! k) !! m = (w !! (k * size (hd sw) + m))" |
37660 | 3799 |
apply (unfold word_rsplit_def word_test_bit_def) |
3800 |
apply (rule trans) |
|
3801 |
apply (rule_tac f = "%x. bin_nth x m" in arg_cong) |
|
3802 |
apply (rule nth_map [symmetric]) |
|
3803 |
apply simp |
|
3804 |
apply (rule bin_nth_rsplit) |
|
3805 |
apply simp_all |
|
3806 |
apply (simp add : word_size rev_map) |
|
3807 |
apply (rule trans) |
|
3808 |
defer |
|
3809 |
apply (rule map_ident [THEN fun_cong]) |
|
3810 |
apply (rule refl [THEN map_cong]) |
|
3811 |
apply (simp add : word_ubin.eq_norm) |
|
3812 |
apply (erule bin_rsplit_size_sign [OF len_gt_0 refl]) |
|
3813 |
done |
|
3814 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3815 |
lemma word_rcat_bl: "word_rcat wl = of_bl (concat (map to_bl wl))" |
37660 | 3816 |
unfolding word_rcat_def to_bl_def' of_bl_def |
3817 |
by (clarsimp simp add : bin_rcat_bl) |
|
3818 |
||
3819 |
lemma size_rcat_lem': |
|
3820 |
"size (concat (map to_bl wl)) = length wl * size (hd wl)" |
|
3821 |
unfolding word_size by (induct wl) auto |
|
3822 |
||
3823 |
lemmas size_rcat_lem = size_rcat_lem' [unfolded word_size] |
|
3824 |
||
3825 |
lemmas td_gal_lt_len = len_gt_0 [THEN td_gal_lt, standard] |
|
3826 |
||
3827 |
lemma nth_rcat_lem' [rule_format] : |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3828 |
"sw = size (hd wl :: 'a :: len word) \<Longrightarrow> (ALL n. n < size wl * sw --> |
37660 | 3829 |
rev (concat (map to_bl wl)) ! n = |
3830 |
rev (to_bl (rev wl ! (n div sw))) ! (n mod sw))" |
|
3831 |
apply (unfold word_size) |
|
3832 |
apply (induct "wl") |
|
3833 |
apply clarsimp |
|
3834 |
apply (clarsimp simp add : nth_append size_rcat_lem) |
|
3835 |
apply (simp (no_asm_use) only: mult_Suc [symmetric] |
|
3836 |
td_gal_lt_len less_Suc_eq_le mod_div_equality') |
|
3837 |
apply clarsimp |
|
3838 |
done |
|
3839 |
||
3840 |
lemmas nth_rcat_lem = refl [THEN nth_rcat_lem', unfolded word_size] |
|
3841 |
||
3842 |
lemma test_bit_rcat: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3843 |
"sw = size (hd wl :: 'a :: len word) \<Longrightarrow> rc = word_rcat wl \<Longrightarrow> rc !! n = |
37660 | 3844 |
(n < size rc & n div sw < size wl & (rev wl) ! (n div sw) !! (n mod sw))" |
3845 |
apply (unfold word_rcat_bl word_size) |
|
3846 |
apply (clarsimp simp add : |
|
3847 |
test_bit_of_bl size_rcat_lem word_size td_gal_lt_len) |
|
3848 |
apply safe |
|
3849 |
apply (auto simp add : |
|
3850 |
test_bit_bl word_size td_gal_lt_len [THEN iffD2, THEN nth_rcat_lem]) |
|
3851 |
done |
|
3852 |
||
3853 |
lemma foldl_eq_foldr [rule_format] : |
|
3854 |
"ALL x. foldl op + x xs = foldr op + (x # xs) (0 :: 'a :: comm_monoid_add)" |
|
3855 |
by (induct xs) (auto simp add : add_assoc) |
|
3856 |
||
3857 |
lemmas test_bit_cong = arg_cong [where f = "test_bit", THEN fun_cong] |
|
3858 |
||
3859 |
lemmas test_bit_rsplit_alt = |
|
3860 |
trans [OF nth_rev_alt [THEN test_bit_cong] |
|
3861 |
test_bit_rsplit [OF refl asm_rl diff_Suc_less]] |
|
3862 |
||
3863 |
-- "lazy way of expressing that u and v, and su and sv, have same types" |
|
3864 |
lemma word_rsplit_len_indep': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3865 |
"[u,v] = p \<Longrightarrow> [su,sv] = q \<Longrightarrow> word_rsplit u = su \<Longrightarrow> |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3866 |
word_rsplit v = sv \<Longrightarrow> length su = length sv" |
37660 | 3867 |
apply (unfold word_rsplit_def) |
3868 |
apply (auto simp add : bin_rsplit_len_indep) |
|
3869 |
done |
|
3870 |
||
3871 |
lemmas word_rsplit_len_indep = word_rsplit_len_indep' [OF refl refl refl refl] |
|
3872 |
||
3873 |
lemma length_word_rsplit_size: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3874 |
"n = len_of TYPE ('a :: len) \<Longrightarrow> |
37660 | 3875 |
(length (word_rsplit w :: 'a word list) <= m) = (size w <= m * n)" |
3876 |
apply (unfold word_rsplit_def word_size) |
|
3877 |
apply (clarsimp simp add : bin_rsplit_len_le) |
|
3878 |
done |
|
3879 |
||
3880 |
lemmas length_word_rsplit_lt_size = |
|
3881 |
length_word_rsplit_size [unfolded Not_eq_iff linorder_not_less [symmetric]] |
|
3882 |
||
3883 |
lemma length_word_rsplit_exp_size: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3884 |
"n = len_of TYPE ('a :: len) \<Longrightarrow> |
37660 | 3885 |
length (word_rsplit w :: 'a word list) = (size w + n - 1) div n" |
3886 |
unfolding word_rsplit_def by (clarsimp simp add : word_size bin_rsplit_len) |
|
3887 |
||
3888 |
lemma length_word_rsplit_even_size: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3889 |
"n = len_of TYPE ('a :: len) \<Longrightarrow> size w = m * n \<Longrightarrow> |
37660 | 3890 |
length (word_rsplit w :: 'a word list) = m" |
3891 |
by (clarsimp simp add : length_word_rsplit_exp_size given_quot_alt) |
|
3892 |
||
3893 |
lemmas length_word_rsplit_exp_size' = refl [THEN length_word_rsplit_exp_size] |
|
3894 |
||
3895 |
(* alternative proof of word_rcat_rsplit *) |
|
3896 |
lemmas tdle = iffD2 [OF split_div_lemma refl, THEN conjunct1] |
|
3897 |
lemmas dtle = xtr4 [OF tdle mult_commute] |
|
3898 |
||
3899 |
lemma word_rcat_rsplit: "word_rcat (word_rsplit w) = w" |
|
3900 |
apply (rule word_eqI) |
|
3901 |
apply (clarsimp simp add : test_bit_rcat word_size) |
|
3902 |
apply (subst refl [THEN test_bit_rsplit]) |
|
3903 |
apply (simp_all add: word_size |
|
3904 |
refl [THEN length_word_rsplit_size [simplified not_less [symmetric], simplified]]) |
|
3905 |
apply safe |
|
3906 |
apply (erule xtr7, rule len_gt_0 [THEN dtle])+ |
|
3907 |
done |
|
3908 |
||
3909 |
lemma size_word_rsplit_rcat_size': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3910 |
"word_rcat (ws :: 'a :: len word list) = frcw \<Longrightarrow> |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3911 |
size frcw = length ws * len_of TYPE ('a) \<Longrightarrow> |
37660 | 3912 |
size (hd [word_rsplit frcw, ws]) = size ws" |
3913 |
apply (clarsimp simp add : word_size length_word_rsplit_exp_size') |
|
3914 |
apply (fast intro: given_quot_alt) |
|
3915 |
done |
|
3916 |
||
3917 |
lemmas size_word_rsplit_rcat_size = |
|
3918 |
size_word_rsplit_rcat_size' [simplified] |
|
3919 |
||
3920 |
lemma msrevs: |
|
3921 |
fixes n::nat |
|
3922 |
shows "0 < n \<Longrightarrow> (k * n + m) div n = m div n + k" |
|
3923 |
and "(k * n + m) mod n = m mod n" |
|
3924 |
by (auto simp: add_commute) |
|
3925 |
||
3926 |
lemma word_rsplit_rcat_size': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3927 |
"word_rcat (ws :: 'a :: len word list) = frcw \<Longrightarrow> |
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3928 |
size frcw = length ws * len_of TYPE ('a) \<Longrightarrow> word_rsplit frcw = ws" |
37660 | 3929 |
apply (frule size_word_rsplit_rcat_size, assumption) |
3930 |
apply (clarsimp simp add : word_size) |
|
3931 |
apply (rule nth_equalityI, assumption) |
|
3932 |
apply clarsimp |
|
3933 |
apply (rule word_eqI) |
|
3934 |
apply (rule trans) |
|
3935 |
apply (rule test_bit_rsplit_alt) |
|
3936 |
apply (clarsimp simp: word_size)+ |
|
3937 |
apply (rule trans) |
|
3938 |
apply (rule test_bit_rcat [OF refl refl]) |
|
3939 |
apply (simp add : word_size msrevs) |
|
3940 |
apply (subst nth_rev) |
|
3941 |
apply arith |
|
3942 |
apply (simp add : le0 [THEN [2] xtr7, THEN diff_Suc_less]) |
|
3943 |
apply safe |
|
3944 |
apply (simp add : diff_mult_distrib) |
|
3945 |
apply (rule mpl_lem) |
|
3946 |
apply (cases "size ws") |
|
3947 |
apply simp_all |
|
3948 |
done |
|
3949 |
||
3950 |
lemmas word_rsplit_rcat_size = refl [THEN word_rsplit_rcat_size'] |
|
3951 |
||
3952 |
||
3953 |
subsection "Rotation" |
|
3954 |
||
3955 |
lemmas rotater_0' [simp] = rotater_def [where n = "0", simplified] |
|
3956 |
||
3957 |
lemmas word_rot_defs = word_roti_def word_rotr_def word_rotl_def |
|
3958 |
||
3959 |
lemma rotate_eq_mod: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
3960 |
"m mod length xs = n mod length xs \<Longrightarrow> rotate m xs = rotate n xs" |
37660 | 3961 |
apply (rule box_equals) |
3962 |
defer |
|
3963 |
apply (rule rotate_conv_mod [symmetric])+ |
|
3964 |
apply simp |
|
3965 |
done |
|
3966 |
||
3967 |
lemmas rotate_eqs [standard] = |
|
3968 |
trans [OF rotate0 [THEN fun_cong] id_apply] |
|
3969 |
rotate_rotate [symmetric] |
|
3970 |
rotate_id |
|
3971 |
rotate_conv_mod |
|
3972 |
rotate_eq_mod |
|
3973 |
||
3974 |
||
3975 |
subsubsection "Rotation of list to right" |
|
3976 |
||
3977 |
lemma rotate1_rl': "rotater1 (l @ [a]) = a # l" |
|
3978 |
unfolding rotater1_def by (cases l) auto |
|
3979 |
||
3980 |
lemma rotate1_rl [simp] : "rotater1 (rotate1 l) = l" |
|
3981 |
apply (unfold rotater1_def) |
|
3982 |
apply (cases "l") |
|
3983 |
apply (case_tac [2] "list") |
|
3984 |
apply auto |
|
3985 |
done |
|
3986 |
||
3987 |
lemma rotate1_lr [simp] : "rotate1 (rotater1 l) = l" |
|
3988 |
unfolding rotater1_def by (cases l) auto |
|
3989 |
||
3990 |
lemma rotater1_rev': "rotater1 (rev xs) = rev (rotate1 xs)" |
|
3991 |
apply (cases "xs") |
|
3992 |
apply (simp add : rotater1_def) |
|
3993 |
apply (simp add : rotate1_rl') |
|
3994 |
done |
|
3995 |
||
3996 |
lemma rotater_rev': "rotater n (rev xs) = rev (rotate n xs)" |
|
3997 |
unfolding rotater_def by (induct n) (auto intro: rotater1_rev') |
|
3998 |
||
3999 |
lemmas rotater_rev = rotater_rev' [where xs = "rev ys", simplified, standard] |
|
4000 |
||
4001 |
lemma rotater_drop_take: |
|
4002 |
"rotater n xs = |
|
4003 |
drop (length xs - n mod length xs) xs @ |
|
4004 |
take (length xs - n mod length xs) xs" |
|
4005 |
by (clarsimp simp add : rotater_rev rotate_drop_take rev_take rev_drop) |
|
4006 |
||
4007 |
lemma rotater_Suc [simp] : |
|
4008 |
"rotater (Suc n) xs = rotater1 (rotater n xs)" |
|
4009 |
unfolding rotater_def by auto |
|
4010 |
||
4011 |
lemma rotate_inv_plus [rule_format] : |
|
4012 |
"ALL k. k = m + n --> rotater k (rotate n xs) = rotater m xs & |
|
4013 |
rotate k (rotater n xs) = rotate m xs & |
|
4014 |
rotater n (rotate k xs) = rotate m xs & |
|
4015 |
rotate n (rotater k xs) = rotater m xs" |
|
4016 |
unfolding rotater_def rotate_def |
|
4017 |
by (induct n) (auto intro: funpow_swap1 [THEN trans]) |
|
4018 |
||
4019 |
lemmas rotate_inv_rel = le_add_diff_inverse2 [symmetric, THEN rotate_inv_plus] |
|
4020 |
||
4021 |
lemmas rotate_inv_eq = order_refl [THEN rotate_inv_rel, simplified] |
|
4022 |
||
4023 |
lemmas rotate_lr [simp] = rotate_inv_eq [THEN conjunct1, standard] |
|
4024 |
lemmas rotate_rl [simp] = |
|
4025 |
rotate_inv_eq [THEN conjunct2, THEN conjunct1, standard] |
|
4026 |
||
4027 |
lemma rotate_gal: "(rotater n xs = ys) = (rotate n ys = xs)" |
|
4028 |
by auto |
|
4029 |
||
4030 |
lemma rotate_gal': "(ys = rotater n xs) = (xs = rotate n ys)" |
|
4031 |
by auto |
|
4032 |
||
4033 |
lemma length_rotater [simp]: |
|
4034 |
"length (rotater n xs) = length xs" |
|
4035 |
by (simp add : rotater_rev) |
|
4036 |
||
38527 | 4037 |
lemma restrict_to_left: |
4038 |
assumes "x = y" |
|
4039 |
shows "(x = z) = (y = z)" |
|
4040 |
using assms by simp |
|
4041 |
||
37660 | 4042 |
lemmas rrs0 = rotate_eqs [THEN restrict_to_left, |
4043 |
simplified rotate_gal [symmetric] rotate_gal' [symmetric], standard] |
|
4044 |
lemmas rrs1 = rrs0 [THEN refl [THEN rev_iffD1]] |
|
4045 |
lemmas rotater_eqs = rrs1 [simplified length_rotater, standard] |
|
4046 |
lemmas rotater_0 = rotater_eqs (1) |
|
4047 |
lemmas rotater_add = rotater_eqs (2) |
|
4048 |
||
4049 |
||
4050 |
subsubsection "map, map2, commuting with rotate(r)" |
|
4051 |
||
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
4052 |
lemma last_map: "xs ~= [] \<Longrightarrow> last (map f xs) = f (last xs)" |
37660 | 4053 |
by (induct xs) auto |
4054 |
||
4055 |
lemma butlast_map: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
4056 |
"xs ~= [] \<Longrightarrow> butlast (map f xs) = map f (butlast xs)" |
37660 | 4057 |
by (induct xs) auto |
4058 |
||
4059 |
lemma rotater1_map: "rotater1 (map f xs) = map f (rotater1 xs)" |
|
4060 |
unfolding rotater1_def |
|
4061 |
by (cases xs) (auto simp add: last_map butlast_map) |
|
4062 |
||
4063 |
lemma rotater_map: |
|
4064 |
"rotater n (map f xs) = map f (rotater n xs)" |
|
4065 |
unfolding rotater_def |
|
4066 |
by (induct n) (auto simp add : rotater1_map) |
|
4067 |
||
4068 |
lemma but_last_zip [rule_format] : |
|
4069 |
"ALL ys. length xs = length ys --> xs ~= [] --> |
|
4070 |
last (zip xs ys) = (last xs, last ys) & |
|
4071 |
butlast (zip xs ys) = zip (butlast xs) (butlast ys)" |
|
4072 |
apply (induct "xs") |
|
4073 |
apply auto |
|
4074 |
apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+ |
|
4075 |
done |
|
4076 |
||
4077 |
lemma but_last_map2 [rule_format] : |
|
4078 |
"ALL ys. length xs = length ys --> xs ~= [] --> |
|
4079 |
last (map2 f xs ys) = f (last xs) (last ys) & |
|
4080 |
butlast (map2 f xs ys) = map2 f (butlast xs) (butlast ys)" |
|
4081 |
apply (induct "xs") |
|
4082 |
apply auto |
|
4083 |
apply (unfold map2_def) |
|
4084 |
apply ((case_tac ys, auto simp: neq_Nil_conv)[1])+ |
|
4085 |
done |
|
4086 |
||
4087 |
lemma rotater1_zip: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
4088 |
"length xs = length ys \<Longrightarrow> |
37660 | 4089 |
rotater1 (zip xs ys) = zip (rotater1 xs) (rotater1 ys)" |
4090 |
apply (unfold rotater1_def) |
|
4091 |
apply (cases "xs") |
|
4092 |
apply auto |
|
4093 |
apply ((case_tac ys, auto simp: neq_Nil_conv but_last_zip)[1])+ |
|
4094 |
done |
|
4095 |
||
4096 |
lemma rotater1_map2: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
4097 |
"length xs = length ys \<Longrightarrow> |
37660 | 4098 |
rotater1 (map2 f xs ys) = map2 f (rotater1 xs) (rotater1 ys)" |
4099 |
unfolding map2_def by (simp add: rotater1_map rotater1_zip) |
|
4100 |
||
4101 |
lemmas lrth = |
|
4102 |
box_equals [OF asm_rl length_rotater [symmetric] |
|
4103 |
length_rotater [symmetric], |
|
4104 |
THEN rotater1_map2] |
|
4105 |
||
4106 |
lemma rotater_map2: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
4107 |
"length xs = length ys \<Longrightarrow> |
37660 | 4108 |
rotater n (map2 f xs ys) = map2 f (rotater n xs) (rotater n ys)" |
4109 |
by (induct n) (auto intro!: lrth) |
|
4110 |
||
4111 |
lemma rotate1_map2: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
4112 |
"length xs = length ys \<Longrightarrow> |
37660 | 4113 |
rotate1 (map2 f xs ys) = map2 f (rotate1 xs) (rotate1 ys)" |
4114 |
apply (unfold map2_def) |
|
4115 |
apply (cases xs) |
|
4116 |
apply (cases ys, auto simp add : rotate1_def)+ |
|
4117 |
done |
|
4118 |
||
4119 |
lemmas lth = box_equals [OF asm_rl length_rotate [symmetric] |
|
4120 |
length_rotate [symmetric], THEN rotate1_map2] |
|
4121 |
||
4122 |
lemma rotate_map2: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
4123 |
"length xs = length ys \<Longrightarrow> |
37660 | 4124 |
rotate n (map2 f xs ys) = map2 f (rotate n xs) (rotate n ys)" |
4125 |
by (induct n) (auto intro!: lth) |
|
4126 |
||
4127 |
||
4128 |
-- "corresponding equalities for word rotation" |
|
4129 |
||
4130 |
lemma to_bl_rotl: |
|
4131 |
"to_bl (word_rotl n w) = rotate n (to_bl w)" |
|
4132 |
by (simp add: word_bl.Abs_inverse' word_rotl_def) |
|
4133 |
||
4134 |
lemmas blrs0 = rotate_eqs [THEN to_bl_rotl [THEN trans]] |
|
4135 |
||
4136 |
lemmas word_rotl_eqs = |
|
4137 |
blrs0 [simplified word_bl.Rep' word_bl.Rep_inject to_bl_rotl [symmetric]] |
|
4138 |
||
4139 |
lemma to_bl_rotr: |
|
4140 |
"to_bl (word_rotr n w) = rotater n (to_bl w)" |
|
4141 |
by (simp add: word_bl.Abs_inverse' word_rotr_def) |
|
4142 |
||
4143 |
lemmas brrs0 = rotater_eqs [THEN to_bl_rotr [THEN trans]] |
|
4144 |
||
4145 |
lemmas word_rotr_eqs = |
|
4146 |
brrs0 [simplified word_bl.Rep' word_bl.Rep_inject to_bl_rotr [symmetric]] |
|
4147 |
||
4148 |
declare word_rotr_eqs (1) [simp] |
|
4149 |
declare word_rotl_eqs (1) [simp] |
|
4150 |
||
4151 |
lemma |
|
4152 |
word_rot_rl [simp]: |
|
4153 |
"word_rotl k (word_rotr k v) = v" and |
|
4154 |
word_rot_lr [simp]: |
|
4155 |
"word_rotr k (word_rotl k v) = v" |
|
4156 |
by (auto simp add: to_bl_rotr to_bl_rotl word_bl.Rep_inject [symmetric]) |
|
4157 |
||
4158 |
lemma |
|
4159 |
word_rot_gal: |
|
4160 |
"(word_rotr n v = w) = (word_rotl n w = v)" and |
|
4161 |
word_rot_gal': |
|
4162 |
"(w = word_rotr n v) = (v = word_rotl n w)" |
|
4163 |
by (auto simp: to_bl_rotr to_bl_rotl word_bl.Rep_inject [symmetric] |
|
4164 |
dest: sym) |
|
4165 |
||
4166 |
lemma word_rotr_rev: |
|
4167 |
"word_rotr n w = word_reverse (word_rotl n (word_reverse w))" |
|
4168 |
by (simp add: word_bl.Rep_inject [symmetric] to_bl_word_rev |
|
4169 |
to_bl_rotr to_bl_rotl rotater_rev) |
|
4170 |
||
4171 |
lemma word_roti_0 [simp]: "word_roti 0 w = w" |
|
4172 |
by (unfold word_rot_defs) auto |
|
4173 |
||
4174 |
lemmas abl_cong = arg_cong [where f = "of_bl"] |
|
4175 |
||
4176 |
lemma word_roti_add: |
|
4177 |
"word_roti (m + n) w = word_roti m (word_roti n w)" |
|
4178 |
proof - |
|
4179 |
have rotater_eq_lem: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
4180 |
"\<And>m n xs. m = n \<Longrightarrow> rotater m xs = rotater n xs" |
37660 | 4181 |
by auto |
4182 |
||
4183 |
have rotate_eq_lem: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
4184 |
"\<And>m n xs. m = n \<Longrightarrow> rotate m xs = rotate n xs" |
37660 | 4185 |
by auto |
4186 |
||
4187 |
note rpts [symmetric, standard] = |
|
4188 |
rotate_inv_plus [THEN conjunct1] |
|
4189 |
rotate_inv_plus [THEN conjunct2, THEN conjunct1] |
|
4190 |
rotate_inv_plus [THEN conjunct2, THEN conjunct2, THEN conjunct1] |
|
4191 |
rotate_inv_plus [THEN conjunct2, THEN conjunct2, THEN conjunct2] |
|
4192 |
||
4193 |
note rrp = trans [symmetric, OF rotate_rotate rotate_eq_lem] |
|
4194 |
note rrrp = trans [symmetric, OF rotater_add [symmetric] rotater_eq_lem] |
|
4195 |
||
4196 |
show ?thesis |
|
4197 |
apply (unfold word_rot_defs) |
|
4198 |
apply (simp only: split: split_if) |
|
4199 |
apply (safe intro!: abl_cong) |
|
4200 |
apply (simp_all only: to_bl_rotl [THEN word_bl.Rep_inverse'] |
|
4201 |
to_bl_rotl |
|
4202 |
to_bl_rotr [THEN word_bl.Rep_inverse'] |
|
4203 |
to_bl_rotr) |
|
4204 |
apply (rule rrp rrrp rpts, |
|
4205 |
simp add: nat_add_distrib [symmetric] |
|
4206 |
nat_diff_distrib [symmetric])+ |
|
4207 |
done |
|
4208 |
qed |
|
4209 |
||
4210 |
lemma word_roti_conv_mod': "word_roti n w = word_roti (n mod int (size w)) w" |
|
4211 |
apply (unfold word_rot_defs) |
|
4212 |
apply (cut_tac y="size w" in gt_or_eq_0) |
|
4213 |
apply (erule disjE) |
|
4214 |
apply simp_all |
|
4215 |
apply (safe intro!: abl_cong) |
|
4216 |
apply (rule rotater_eqs) |
|
4217 |
apply (simp add: word_size nat_mod_distrib) |
|
4218 |
apply (simp add: rotater_add [symmetric] rotate_gal [symmetric]) |
|
4219 |
apply (rule rotater_eqs) |
|
4220 |
apply (simp add: word_size nat_mod_distrib) |
|
4221 |
apply (rule int_eq_0_conv [THEN iffD1]) |
|
4222 |
apply (simp only: zmod_int zadd_int [symmetric]) |
|
4223 |
apply (simp add: rdmods) |
|
4224 |
done |
|
4225 |
||
4226 |
lemmas word_roti_conv_mod = word_roti_conv_mod' [unfolded word_size] |
|
4227 |
||
4228 |
||
4229 |
subsubsection "Word rotation commutes with bit-wise operations" |
|
4230 |
||
4231 |
(* using locale to not pollute lemma namespace *) |
|
4232 |
locale word_rotate |
|
4233 |
||
4234 |
context word_rotate |
|
4235 |
begin |
|
4236 |
||
4237 |
lemmas word_rot_defs' = to_bl_rotl to_bl_rotr |
|
4238 |
||
4239 |
lemmas blwl_syms [symmetric] = bl_word_not bl_word_and bl_word_or bl_word_xor |
|
4240 |
||
4241 |
lemmas lbl_lbl = trans [OF word_bl.Rep' word_bl.Rep' [symmetric]] |
|
4242 |
||
4243 |
lemmas ths_map2 [OF lbl_lbl] = rotate_map2 rotater_map2 |
|
4244 |
||
4245 |
lemmas ths_map [where xs = "to_bl v", standard] = rotate_map rotater_map |
|
4246 |
||
4247 |
lemmas th1s [simplified word_rot_defs' [symmetric]] = ths_map2 ths_map |
|
4248 |
||
4249 |
lemma word_rot_logs: |
|
4250 |
"word_rotl n (NOT v) = NOT word_rotl n v" |
|
4251 |
"word_rotr n (NOT v) = NOT word_rotr n v" |
|
4252 |
"word_rotl n (x AND y) = word_rotl n x AND word_rotl n y" |
|
4253 |
"word_rotr n (x AND y) = word_rotr n x AND word_rotr n y" |
|
4254 |
"word_rotl n (x OR y) = word_rotl n x OR word_rotl n y" |
|
4255 |
"word_rotr n (x OR y) = word_rotr n x OR word_rotr n y" |
|
4256 |
"word_rotl n (x XOR y) = word_rotl n x XOR word_rotl n y" |
|
4257 |
"word_rotr n (x XOR y) = word_rotr n x XOR word_rotr n y" |
|
4258 |
by (rule word_bl.Rep_eqD, |
|
4259 |
rule word_rot_defs' [THEN trans], |
|
4260 |
simp only: blwl_syms [symmetric], |
|
4261 |
rule th1s [THEN trans], |
|
4262 |
rule refl)+ |
|
4263 |
end |
|
4264 |
||
4265 |
lemmas word_rot_logs = word_rotate.word_rot_logs |
|
4266 |
||
4267 |
lemmas bl_word_rotl_dt = trans [OF to_bl_rotl rotate_drop_take, |
|
4268 |
simplified word_bl.Rep', standard] |
|
4269 |
||
4270 |
lemmas bl_word_rotr_dt = trans [OF to_bl_rotr rotater_drop_take, |
|
4271 |
simplified word_bl.Rep', standard] |
|
4272 |
||
4273 |
lemma bl_word_roti_dt': |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
4274 |
"n = nat ((- i) mod int (size (w :: 'a :: len word))) \<Longrightarrow> |
37660 | 4275 |
to_bl (word_roti i w) = drop n (to_bl w) @ take n (to_bl w)" |
4276 |
apply (unfold word_roti_def) |
|
4277 |
apply (simp add: bl_word_rotl_dt bl_word_rotr_dt word_size) |
|
4278 |
apply safe |
|
4279 |
apply (simp add: zmod_zminus1_eq_if) |
|
4280 |
apply safe |
|
4281 |
apply (simp add: nat_mult_distrib) |
|
4282 |
apply (simp add: nat_diff_distrib [OF pos_mod_sign pos_mod_conj |
|
4283 |
[THEN conjunct2, THEN order_less_imp_le]] |
|
4284 |
nat_mod_distrib) |
|
4285 |
apply (simp add: nat_mod_distrib) |
|
4286 |
done |
|
4287 |
||
4288 |
lemmas bl_word_roti_dt = bl_word_roti_dt' [unfolded word_size] |
|
4289 |
||
4290 |
lemmas word_rotl_dt = bl_word_rotl_dt |
|
4291 |
[THEN word_bl.Rep_inverse' [symmetric], standard] |
|
4292 |
lemmas word_rotr_dt = bl_word_rotr_dt |
|
4293 |
[THEN word_bl.Rep_inverse' [symmetric], standard] |
|
4294 |
lemmas word_roti_dt = bl_word_roti_dt |
|
4295 |
[THEN word_bl.Rep_inverse' [symmetric], standard] |
|
4296 |
||
4297 |
lemma word_rotx_0 [simp] : "word_rotr i 0 = 0 & word_rotl i 0 = 0" |
|
4298 |
by (simp add : word_rotr_dt word_rotl_dt to_bl_0 replicate_add [symmetric]) |
|
4299 |
||
4300 |
lemma word_roti_0' [simp] : "word_roti n 0 = 0" |
|
4301 |
unfolding word_roti_def by auto |
|
4302 |
||
4303 |
lemmas word_rotr_dt_no_bin' [simp] = |
|
4304 |
word_rotr_dt [where w="number_of w", unfolded to_bl_no_bin, standard] |
|
4305 |
||
4306 |
lemmas word_rotl_dt_no_bin' [simp] = |
|
4307 |
word_rotl_dt [where w="number_of w", unfolded to_bl_no_bin, standard] |
|
4308 |
||
4309 |
declare word_roti_def [simp] |
|
4310 |
||
4311 |
||
4312 |
subsection {* Miscellaneous *} |
|
4313 |
||
4314 |
declare of_nat_2p [simp] |
|
4315 |
||
4316 |
lemma word_int_cases: |
|
4317 |
"\<lbrakk>\<And>n. \<lbrakk>(x ::'a::len0 word) = word_of_int n; 0 \<le> n; n < 2^len_of TYPE('a)\<rbrakk> \<Longrightarrow> P\<rbrakk> |
|
4318 |
\<Longrightarrow> P" |
|
4319 |
by (cases x rule: word_uint.Abs_cases) (simp add: uints_num) |
|
4320 |
||
4321 |
lemma word_nat_cases [cases type: word]: |
|
4322 |
"\<lbrakk>\<And>n. \<lbrakk>(x ::'a::len word) = of_nat n; n < 2^len_of TYPE('a)\<rbrakk> \<Longrightarrow> P\<rbrakk> |
|
4323 |
\<Longrightarrow> P" |
|
4324 |
by (cases x rule: word_unat.Abs_cases) (simp add: unats_def) |
|
4325 |
||
4326 |
lemma max_word_eq: |
|
4327 |
"(max_word::'a::len word) = 2^len_of TYPE('a) - 1" |
|
4328 |
by (simp add: max_word_def word_of_int_hom_syms word_of_int_2p) |
|
4329 |
||
4330 |
lemma max_word_max [simp,intro!]: |
|
4331 |
"n \<le> max_word" |
|
4332 |
by (cases n rule: word_int_cases) |
|
4333 |
(simp add: max_word_def word_le_def int_word_uint int_mod_eq') |
|
4334 |
||
4335 |
lemma word_of_int_2p_len: |
|
4336 |
"word_of_int (2 ^ len_of TYPE('a)) = (0::'a::len0 word)" |
|
4337 |
by (subst word_uint.Abs_norm [symmetric]) |
|
4338 |
(simp add: word_of_int_hom_syms) |
|
4339 |
||
4340 |
lemma word_pow_0: |
|
4341 |
"(2::'a::len word) ^ len_of TYPE('a) = 0" |
|
4342 |
proof - |
|
4343 |
have "word_of_int (2 ^ len_of TYPE('a)) = (0::'a word)" |
|
4344 |
by (rule word_of_int_2p_len) |
|
4345 |
thus ?thesis by (simp add: word_of_int_2p) |
|
4346 |
qed |
|
4347 |
||
4348 |
lemma max_word_wrap: "x + 1 = 0 \<Longrightarrow> x = max_word" |
|
4349 |
apply (simp add: max_word_eq) |
|
4350 |
apply uint_arith |
|
4351 |
apply auto |
|
4352 |
apply (simp add: word_pow_0) |
|
4353 |
done |
|
4354 |
||
4355 |
lemma max_word_minus: |
|
4356 |
"max_word = (-1::'a::len word)" |
|
4357 |
proof - |
|
4358 |
have "-1 + 1 = (0::'a word)" by simp |
|
4359 |
thus ?thesis by (rule max_word_wrap [symmetric]) |
|
4360 |
qed |
|
4361 |
||
4362 |
lemma max_word_bl [simp]: |
|
4363 |
"to_bl (max_word::'a::len word) = replicate (len_of TYPE('a)) True" |
|
4364 |
by (subst max_word_minus to_bl_n1)+ simp |
|
4365 |
||
4366 |
lemma max_test_bit [simp]: |
|
4367 |
"(max_word::'a::len word) !! n = (n < len_of TYPE('a))" |
|
4368 |
by (auto simp add: test_bit_bl word_size) |
|
4369 |
||
4370 |
lemma word_and_max [simp]: |
|
4371 |
"x AND max_word = x" |
|
4372 |
by (rule word_eqI) (simp add: word_ops_nth_size word_size) |
|
4373 |
||
4374 |
lemma word_or_max [simp]: |
|
4375 |
"x OR max_word = max_word" |
|
4376 |
by (rule word_eqI) (simp add: word_ops_nth_size word_size) |
|
4377 |
||
4378 |
lemma word_ao_dist2: |
|
4379 |
"x AND (y OR z) = x AND y OR x AND (z::'a::len0 word)" |
|
4380 |
by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) |
|
4381 |
||
4382 |
lemma word_oa_dist2: |
|
4383 |
"x OR y AND z = (x OR y) AND (x OR (z::'a::len0 word))" |
|
4384 |
by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) |
|
4385 |
||
4386 |
lemma word_and_not [simp]: |
|
4387 |
"x AND NOT x = (0::'a::len0 word)" |
|
4388 |
by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) |
|
4389 |
||
4390 |
lemma word_or_not [simp]: |
|
4391 |
"x OR NOT x = max_word" |
|
4392 |
by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) |
|
4393 |
||
4394 |
lemma word_boolean: |
|
4395 |
"boolean (op AND) (op OR) bitNOT 0 max_word" |
|
4396 |
apply (rule boolean.intro) |
|
4397 |
apply (rule word_bw_assocs) |
|
4398 |
apply (rule word_bw_assocs) |
|
4399 |
apply (rule word_bw_comms) |
|
4400 |
apply (rule word_bw_comms) |
|
4401 |
apply (rule word_ao_dist2) |
|
4402 |
apply (rule word_oa_dist2) |
|
4403 |
apply (rule word_and_max) |
|
4404 |
apply (rule word_log_esimps) |
|
4405 |
apply (rule word_and_not) |
|
4406 |
apply (rule word_or_not) |
|
4407 |
done |
|
4408 |
||
4409 |
interpretation word_bool_alg: |
|
4410 |
boolean "op AND" "op OR" bitNOT 0 max_word |
|
4411 |
by (rule word_boolean) |
|
4412 |
||
4413 |
lemma word_xor_and_or: |
|
4414 |
"x XOR y = x AND NOT y OR NOT x AND (y::'a::len0 word)" |
|
4415 |
by (rule word_eqI) (auto simp add: word_ops_nth_size word_size) |
|
4416 |
||
4417 |
interpretation word_bool_alg: |
|
4418 |
boolean_xor "op AND" "op OR" bitNOT 0 max_word "op XOR" |
|
4419 |
apply (rule boolean_xor.intro) |
|
4420 |
apply (rule word_boolean) |
|
4421 |
apply (rule boolean_xor_axioms.intro) |
|
4422 |
apply (rule word_xor_and_or) |
|
4423 |
done |
|
4424 |
||
4425 |
lemma shiftr_x_0 [iff]: |
|
4426 |
"(x::'a::len0 word) >> 0 = x" |
|
4427 |
by (simp add: shiftr_bl) |
|
4428 |
||
4429 |
lemma shiftl_x_0 [simp]: |
|
4430 |
"(x :: 'a :: len word) << 0 = x" |
|
4431 |
by (simp add: shiftl_t2n) |
|
4432 |
||
4433 |
lemma shiftl_1 [simp]: |
|
4434 |
"(1::'a::len word) << n = 2^n" |
|
4435 |
by (simp add: shiftl_t2n) |
|
4436 |
||
4437 |
lemma uint_lt_0 [simp]: |
|
4438 |
"uint x < 0 = False" |
|
4439 |
by (simp add: linorder_not_less) |
|
4440 |
||
4441 |
lemma shiftr1_1 [simp]: |
|
4442 |
"shiftr1 (1::'a::len word) = 0" |
|
4443 |
by (simp add: shiftr1_def word_0_alt) |
|
4444 |
||
4445 |
lemma shiftr_1[simp]: |
|
4446 |
"(1::'a::len word) >> n = (if n = 0 then 1 else 0)" |
|
4447 |
by (induct n) (auto simp: shiftr_def) |
|
4448 |
||
4449 |
lemma word_less_1 [simp]: |
|
4450 |
"((x::'a::len word) < 1) = (x = 0)" |
|
4451 |
by (simp add: word_less_nat_alt unat_0_iff) |
|
4452 |
||
4453 |
lemma to_bl_mask: |
|
4454 |
"to_bl (mask n :: 'a::len word) = |
|
4455 |
replicate (len_of TYPE('a) - n) False @ |
|
4456 |
replicate (min (len_of TYPE('a)) n) True" |
|
4457 |
by (simp add: mask_bl word_rep_drop min_def) |
|
4458 |
||
4459 |
lemma map_replicate_True: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
4460 |
"n = length xs \<Longrightarrow> |
37660 | 4461 |
map (\<lambda>(x,y). x & y) (zip xs (replicate n True)) = xs" |
4462 |
by (induct xs arbitrary: n) auto |
|
4463 |
||
4464 |
lemma map_replicate_False: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
4465 |
"n = length xs \<Longrightarrow> map (\<lambda>(x,y). x & y) |
37660 | 4466 |
(zip xs (replicate n False)) = replicate n False" |
4467 |
by (induct xs arbitrary: n) auto |
|
4468 |
||
4469 |
lemma bl_and_mask: |
|
4470 |
fixes w :: "'a::len word" |
|
4471 |
fixes n |
|
4472 |
defines "n' \<equiv> len_of TYPE('a) - n" |
|
4473 |
shows "to_bl (w AND mask n) = replicate n' False @ drop n' (to_bl w)" |
|
4474 |
proof - |
|
4475 |
note [simp] = map_replicate_True map_replicate_False |
|
4476 |
have "to_bl (w AND mask n) = |
|
4477 |
map2 op & (to_bl w) (to_bl (mask n::'a::len word))" |
|
4478 |
by (simp add: bl_word_and) |
|
4479 |
also |
|
4480 |
have "to_bl w = take n' (to_bl w) @ drop n' (to_bl w)" by simp |
|
4481 |
also |
|
4482 |
have "map2 op & \<dots> (to_bl (mask n::'a::len word)) = |
|
4483 |
replicate n' False @ drop n' (to_bl w)" |
|
4484 |
unfolding to_bl_mask n'_def map2_def |
|
4485 |
by (subst zip_append) auto |
|
4486 |
finally |
|
4487 |
show ?thesis . |
|
4488 |
qed |
|
4489 |
||
4490 |
lemma drop_rev_takefill: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
4491 |
"length xs \<le> n \<Longrightarrow> |
37660 | 4492 |
drop (n - length xs) (rev (takefill False n (rev xs))) = xs" |
4493 |
by (simp add: takefill_alt rev_take) |
|
4494 |
||
4495 |
lemma map_nth_0 [simp]: |
|
4496 |
"map (op !! (0::'a::len0 word)) xs = replicate (length xs) False" |
|
4497 |
by (induct xs) auto |
|
4498 |
||
4499 |
lemma uint_plus_if_size: |
|
4500 |
"uint (x + y) = |
|
4501 |
(if uint x + uint y < 2^size x then |
|
4502 |
uint x + uint y |
|
4503 |
else |
|
4504 |
uint x + uint y - 2^size x)" |
|
4505 |
by (simp add: word_arith_alts int_word_uint mod_add_if_z |
|
4506 |
word_size) |
|
4507 |
||
4508 |
lemma unat_plus_if_size: |
|
4509 |
"unat (x + (y::'a::len word)) = |
|
4510 |
(if unat x + unat y < 2^size x then |
|
4511 |
unat x + unat y |
|
4512 |
else |
|
4513 |
unat x + unat y - 2^size x)" |
|
4514 |
apply (subst word_arith_nat_defs) |
|
4515 |
apply (subst unat_of_nat) |
|
4516 |
apply (simp add: mod_nat_add word_size) |
|
4517 |
done |
|
4518 |
||
4519 |
lemma word_neq_0_conv [simp]: |
|
4520 |
fixes w :: "'a :: len word" |
|
4521 |
shows "(w \<noteq> 0) = (0 < w)" |
|
4522 |
proof - |
|
4523 |
have "0 \<le> w" by (rule word_zero_le) |
|
4524 |
thus ?thesis by (auto simp add: word_less_def) |
|
4525 |
qed |
|
4526 |
||
4527 |
lemma max_lt: |
|
4528 |
"unat (max a b div c) = unat (max a b) div unat (c:: 'a :: len word)" |
|
4529 |
apply (subst word_arith_nat_defs) |
|
4530 |
apply (subst word_unat.eq_norm) |
|
4531 |
apply (subst mod_if) |
|
4532 |
apply clarsimp |
|
4533 |
apply (erule notE) |
|
4534 |
apply (insert div_le_dividend [of "unat (max a b)" "unat c"]) |
|
4535 |
apply (erule order_le_less_trans) |
|
4536 |
apply (insert unat_lt2p [of "max a b"]) |
|
4537 |
apply simp |
|
4538 |
done |
|
4539 |
||
4540 |
lemma uint_sub_if_size: |
|
4541 |
"uint (x - y) = |
|
4542 |
(if uint y \<le> uint x then |
|
4543 |
uint x - uint y |
|
4544 |
else |
|
4545 |
uint x - uint y + 2^size x)" |
|
4546 |
by (simp add: word_arith_alts int_word_uint mod_sub_if_z |
|
4547 |
word_size) |
|
4548 |
||
4549 |
lemma unat_sub: |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
4550 |
"b <= a \<Longrightarrow> unat (a - b) = unat a - unat b" |
37660 | 4551 |
by (simp add: unat_def uint_sub_if_size word_le_def nat_diff_distrib) |
4552 |
||
4553 |
lemmas word_less_sub1_numberof [simp] = |
|
4554 |
word_less_sub1 [of "number_of w", standard] |
|
4555 |
lemmas word_le_sub1_numberof [simp] = |
|
4556 |
word_le_sub1 [of "number_of w", standard] |
|
4557 |
||
4558 |
lemma word_of_int_minus: |
|
4559 |
"word_of_int (2^len_of TYPE('a) - i) = (word_of_int (-i)::'a::len word)" |
|
4560 |
proof - |
|
4561 |
have x: "2^len_of TYPE('a) - i = -i + 2^len_of TYPE('a)" by simp |
|
4562 |
show ?thesis |
|
4563 |
apply (subst x) |
|
4564 |
apply (subst word_uint.Abs_norm [symmetric], subst mod_add_self2) |
|
4565 |
apply simp |
|
4566 |
done |
|
4567 |
qed |
|
4568 |
||
4569 |
lemmas word_of_int_inj = |
|
4570 |
word_uint.Abs_inject [unfolded uints_num, simplified] |
|
4571 |
||
4572 |
lemma word_le_less_eq: |
|
4573 |
"(x ::'z::len word) \<le> y = (x = y \<or> x < y)" |
|
4574 |
by (auto simp add: word_less_def) |
|
4575 |
||
4576 |
lemma mod_plus_cong: |
|
4577 |
assumes 1: "(b::int) = b'" |
|
4578 |
and 2: "x mod b' = x' mod b'" |
|
4579 |
and 3: "y mod b' = y' mod b'" |
|
4580 |
and 4: "x' + y' = z'" |
|
4581 |
shows "(x + y) mod b = z' mod b'" |
|
4582 |
proof - |
|
4583 |
from 1 2[symmetric] 3[symmetric] have "(x + y) mod b = (x' mod b' + y' mod b') mod b'" |
|
4584 |
by (simp add: mod_add_eq[symmetric]) |
|
4585 |
also have "\<dots> = (x' + y') mod b'" |
|
4586 |
by (simp add: mod_add_eq[symmetric]) |
|
4587 |
finally show ?thesis by (simp add: 4) |
|
4588 |
qed |
|
4589 |
||
4590 |
lemma mod_minus_cong: |
|
4591 |
assumes 1: "(b::int) = b'" |
|
4592 |
and 2: "x mod b' = x' mod b'" |
|
4593 |
and 3: "y mod b' = y' mod b'" |
|
4594 |
and 4: "x' - y' = z'" |
|
4595 |
shows "(x - y) mod b = z' mod b'" |
|
4596 |
using assms |
|
4597 |
apply (subst zmod_zsub_left_eq) |
|
4598 |
apply (subst zmod_zsub_right_eq) |
|
4599 |
apply (simp add: zmod_zsub_left_eq [symmetric] zmod_zsub_right_eq [symmetric]) |
|
4600 |
done |
|
4601 |
||
4602 |
lemma word_induct_less: |
|
4603 |
"\<lbrakk>P (0::'a::len word); \<And>n. \<lbrakk>n < m; P n\<rbrakk> \<Longrightarrow> P (1 + n)\<rbrakk> \<Longrightarrow> P m" |
|
4604 |
apply (cases m) |
|
4605 |
apply atomize |
|
4606 |
apply (erule rev_mp)+ |
|
4607 |
apply (rule_tac x=m in spec) |
|
4608 |
apply (induct_tac n) |
|
4609 |
apply simp |
|
4610 |
apply clarsimp |
|
4611 |
apply (erule impE) |
|
4612 |
apply clarsimp |
|
4613 |
apply (erule_tac x=n in allE) |
|
4614 |
apply (erule impE) |
|
4615 |
apply (simp add: unat_arith_simps) |
|
4616 |
apply (clarsimp simp: unat_of_nat) |
|
4617 |
apply simp |
|
4618 |
apply (erule_tac x="of_nat na" in allE) |
|
4619 |
apply (erule impE) |
|
4620 |
apply (simp add: unat_arith_simps) |
|
4621 |
apply (clarsimp simp: unat_of_nat) |
|
4622 |
apply simp |
|
4623 |
done |
|
4624 |
||
4625 |
lemma word_induct: |
|
4626 |
"\<lbrakk>P (0::'a::len word); \<And>n. P n \<Longrightarrow> P (1 + n)\<rbrakk> \<Longrightarrow> P m" |
|
4627 |
by (erule word_induct_less, simp) |
|
4628 |
||
4629 |
lemma word_induct2 [induct type]: |
|
4630 |
"\<lbrakk>P 0; \<And>n. \<lbrakk>1 + n \<noteq> 0; P n\<rbrakk> \<Longrightarrow> P (1 + n)\<rbrakk> \<Longrightarrow> P (n::'b::len word)" |
|
4631 |
apply (rule word_induct, simp) |
|
4632 |
apply (case_tac "1+n = 0", auto) |
|
4633 |
done |
|
4634 |
||
4635 |
definition word_rec :: "'a \<Rightarrow> ('b::len word \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b word \<Rightarrow> 'a" where |
|
40827
abbc05c20e24
code preprocessor setup for numerals on word type;
haftmann
parents:
39910
diff
changeset
|
4636 |
"word_rec forZero forSuc n = nat_rec forZero (forSuc \<circ> of_nat) (unat n)" |
37660 | 4637 |
|
4638 |
lemma word_rec_0: "word_rec z s 0 = z" |
|
4639 |
by (simp add: word_rec_def) |
|
4640 |
||
4641 |
lemma word_rec_Suc: |
|
4642 |
"1 + n \<noteq> (0::'a::len word) \<Longrightarrow> word_rec z s (1 + n) = s n (word_rec z s n)" |
|
4643 |
apply (simp add: word_rec_def unat_word_ariths) |
|
4644 |
apply (subst nat_mod_eq') |
|
4645 |
apply (cut_tac x=n in unat_lt2p) |
|
4646 |
apply (drule Suc_mono) |
|
4647 |
apply (simp add: less_Suc_eq_le) |
|
4648 |
apply (simp only: order_less_le, simp) |
|
4649 |
apply (erule contrapos_pn, simp) |
|
4650 |
apply (drule arg_cong[where f=of_nat]) |
|
4651 |
apply simp |
|
4652 |
apply (subst (asm) word_unat.Rep_inverse[of n]) |
|
4653 |
apply simp |
|
4654 |
apply simp |
|
4655 |
done |
|
4656 |
||
4657 |
lemma word_rec_Pred: |
|
4658 |
"n \<noteq> 0 \<Longrightarrow> word_rec z s n = s (n - 1) (word_rec z s (n - 1))" |
|
4659 |
apply (rule subst[where t="n" and s="1 + (n - 1)"]) |
|
4660 |
apply simp |
|
4661 |
apply (subst word_rec_Suc) |
|
4662 |
apply simp |
|
4663 |
apply simp |
|
4664 |
done |
|
4665 |
||
4666 |
lemma word_rec_in: |
|
4667 |
"f (word_rec z (\<lambda>_. f) n) = word_rec (f z) (\<lambda>_. f) n" |
|
4668 |
by (induct n) (simp_all add: word_rec_0 word_rec_Suc) |
|
4669 |
||
4670 |
lemma word_rec_in2: |
|
4671 |
"f n (word_rec z f n) = word_rec (f 0 z) (f \<circ> op + 1) n" |
|
4672 |
by (induct n) (simp_all add: word_rec_0 word_rec_Suc) |
|
4673 |
||
4674 |
lemma word_rec_twice: |
|
4675 |
"m \<le> n \<Longrightarrow> word_rec z f n = word_rec (word_rec z f (n - m)) (f \<circ> op + (n - m)) m" |
|
4676 |
apply (erule rev_mp) |
|
4677 |
apply (rule_tac x=z in spec) |
|
4678 |
apply (rule_tac x=f in spec) |
|
4679 |
apply (induct n) |
|
4680 |
apply (simp add: word_rec_0) |
|
4681 |
apply clarsimp |
|
4682 |
apply (rule_tac t="1 + n - m" and s="1 + (n - m)" in subst) |
|
4683 |
apply simp |
|
4684 |
apply (case_tac "1 + (n - m) = 0") |
|
4685 |
apply (simp add: word_rec_0) |
|
4686 |
apply (rule_tac f = "word_rec ?a ?b" in arg_cong) |
|
4687 |
apply (rule_tac t="m" and s="m + (1 + (n - m))" in subst) |
|
4688 |
apply simp |
|
4689 |
apply (simp (no_asm_use)) |
|
4690 |
apply (simp add: word_rec_Suc word_rec_in2) |
|
4691 |
apply (erule impE) |
|
4692 |
apply uint_arith |
|
4693 |
apply (drule_tac x="x \<circ> op + 1" in spec) |
|
4694 |
apply (drule_tac x="x 0 xa" in spec) |
|
4695 |
apply simp |
|
4696 |
apply (rule_tac t="\<lambda>a. x (1 + (n - m + a))" and s="\<lambda>a. x (1 + (n - m) + a)" |
|
4697 |
in subst) |
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
4698 |
apply (clarsimp simp add: fun_eq_iff) |
37660 | 4699 |
apply (rule_tac t="(1 + (n - m + xb))" and s="1 + (n - m) + xb" in subst) |
4700 |
apply simp |
|
4701 |
apply (rule refl) |
|
4702 |
apply (rule refl) |
|
4703 |
done |
|
4704 |
||
4705 |
lemma word_rec_id: "word_rec z (\<lambda>_. id) n = z" |
|
4706 |
by (induct n) (auto simp add: word_rec_0 word_rec_Suc) |
|
4707 |
||
4708 |
lemma word_rec_id_eq: "\<forall>m < n. f m = id \<Longrightarrow> word_rec z f n = z" |
|
4709 |
apply (erule rev_mp) |
|
4710 |
apply (induct n) |
|
4711 |
apply (auto simp add: word_rec_0 word_rec_Suc) |
|
4712 |
apply (drule spec, erule mp) |
|
4713 |
apply uint_arith |
|
4714 |
apply (drule_tac x=n in spec, erule impE) |
|
4715 |
apply uint_arith |
|
4716 |
apply simp |
|
4717 |
done |
|
4718 |
||
4719 |
lemma word_rec_max: |
|
4720 |
"\<forall>m\<ge>n. m \<noteq> -1 \<longrightarrow> f m = id \<Longrightarrow> word_rec z f -1 = word_rec z f n" |
|
4721 |
apply (subst word_rec_twice[where n="-1" and m="-1 - n"]) |
|
4722 |
apply simp |
|
4723 |
apply simp |
|
4724 |
apply (rule word_rec_id_eq) |
|
4725 |
apply clarsimp |
|
4726 |
apply (drule spec, rule mp, erule mp) |
|
4727 |
apply (rule word_plus_mono_right2[OF _ order_less_imp_le]) |
|
4728 |
prefer 2 |
|
4729 |
apply assumption |
|
4730 |
apply simp |
|
4731 |
apply (erule contrapos_pn) |
|
4732 |
apply simp |
|
4733 |
apply (drule arg_cong[where f="\<lambda>x. x - n"]) |
|
4734 |
apply simp |
|
4735 |
done |
|
4736 |
||
4737 |
lemma unatSuc: |
|
4738 |
"1 + n \<noteq> (0::'a::len word) \<Longrightarrow> unat (1 + n) = Suc (unat n)" |
|
4739 |
by unat_arith |
|
4740 |
||
4741 |
||
4742 |
lemmas word_no_1 [simp] = word_1_no [symmetric, unfolded BIT_simps] |
|
4743 |
lemmas word_no_0 [simp] = word_0_no [symmetric] |
|
4744 |
||
4745 |
declare word_0_bl [simp] |
|
4746 |
declare bin_to_bl_def [simp] |
|
4747 |
declare to_bl_0 [simp] |
|
4748 |
declare of_bl_True [simp] |
|
4749 |
||
4750 |
||
41060
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
boehmes
parents:
40827
diff
changeset
|
4751 |
use "~~/src/HOL/Word/Tools/smt_word.ML" |
37660 | 4752 |
|
36899
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
35049
diff
changeset
|
4753 |
setup {* SMT_Word.setup *} |
bcd6fce5bf06
layered SMT setup, adapted SMT clients, added further tests, made Z3 proof abstraction configurable
boehmes
parents:
35049
diff
changeset
|
4754 |
|
41060
4199fdcfa3c0
moved smt_word.ML into the directory of the Word library
boehmes
parents:
40827
diff
changeset
|
4755 |
end |