| author | urbanc | 
| Thu, 12 Jun 2008 09:56:28 +0200 | |
| changeset 27162 | 8d747de5c73e | 
| parent 27139 | a1f3c7b5ce9c | 
| child 28524 | 644b62cf678f | 
| permissions | -rw-r--r-- | 
| 24333 | 1  | 
(*  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Jeremy Dawson and Gerwin Klein, NICTA  | 
|
4  | 
||
5  | 
Basic definition of word type and basic theorems following from  | 
|
6  | 
the definition of the word type  | 
|
7  | 
*)  | 
|
8  | 
||
| 24350 | 9  | 
header {* Definition of Word Type *}
 | 
| 24333 | 10  | 
|
| 26559 | 11  | 
theory WordDefinition  | 
| 
27139
 
a1f3c7b5ce9c
back to original import order -- thanks to proper deletion of nat cases/induct rules from type_definition;
 
wenzelm 
parents: 
27134 
diff
changeset
 | 
12  | 
imports Size BinBoolList TdThs  | 
| 26559 | 13  | 
begin  | 
| 24333 | 14  | 
|
15  | 
typedef (open word) 'a word  | 
|
| 24465 | 16  | 
  = "{(0::int) ..< 2^len_of TYPE('a::len0)}" by auto
 | 
| 24333 | 17  | 
|
| 25762 | 18  | 
instantiation word :: (len0) size  | 
19  | 
begin  | 
|
20  | 
||
21  | 
definition  | 
|
22  | 
  word_size: "size (w :: 'a word) = len_of TYPE('a)"
 | 
|
23  | 
||
24  | 
instance ..  | 
|
25  | 
||
26  | 
end  | 
|
27  | 
||
28  | 
definition  | 
|
29  | 
  -- {* representation of words using unsigned or signed bins, 
 | 
|
30  | 
only difference in these is the type class *}  | 
|
31  | 
word_of_int :: "int \<Rightarrow> 'a\<Colon>len0 word"  | 
|
32  | 
where  | 
|
| 26514 | 33  | 
  [code func del]: "word_of_int w = Abs_word (bintrunc (len_of TYPE ('a)) w)" 
 | 
34  | 
||
35  | 
code_datatype word_of_int  | 
|
| 24333 | 36  | 
|
37  | 
||
| 24350 | 38  | 
subsection "Type conversions and casting"  | 
| 24333 | 39  | 
|
40  | 
constdefs  | 
|
41  | 
  -- {* uint and sint cast a word to an integer,
 | 
|
42  | 
uint treats the word as unsigned,  | 
|
43  | 
sint treats the most-significant-bit as a sign bit *}  | 
|
| 24465 | 44  | 
uint :: "'a :: len0 word => int"  | 
| 24333 | 45  | 
"uint w == Rep_word w"  | 
| 24465 | 46  | 
sint :: "'a :: len word => int"  | 
47  | 
  sint_uint: "sint w == sbintrunc (len_of TYPE ('a) - 1) (uint w)"
 | 
|
48  | 
unat :: "'a :: len0 word => nat"  | 
|
| 24333 | 49  | 
"unat w == nat (uint w)"  | 
50  | 
||
51  | 
-- "the sets of integers representing the words"  | 
|
52  | 
uints :: "nat => int set"  | 
|
53  | 
"uints n == range (bintrunc n)"  | 
|
54  | 
sints :: "nat => int set"  | 
|
55  | 
"sints n == range (sbintrunc (n - 1))"  | 
|
56  | 
unats :: "nat => nat set"  | 
|
57  | 
  "unats n == {i. i < 2 ^ n}"
 | 
|
58  | 
norm_sint :: "nat => int => int"  | 
|
59  | 
"norm_sint n w == (w + 2 ^ (n - 1)) mod 2 ^ n - 2 ^ (n - 1)"  | 
|
60  | 
||
| 24465 | 61  | 
-- "cast a word to a different length"  | 
62  | 
scast :: "'a :: len word => 'b :: len word"  | 
|
63  | 
"scast w == word_of_int (sint w)"  | 
|
64  | 
ucast :: "'a :: len0 word => 'b :: len0 word"  | 
|
65  | 
"ucast w == word_of_int (uint w)"  | 
|
66  | 
||
67  | 
-- "whether a cast (or other) function is to a longer or shorter length"  | 
|
68  | 
  source_size :: "('a :: len0 word => 'b) => nat"
 | 
|
69  | 
"source_size c == let arb = arbitrary ; x = c arb in size arb"  | 
|
70  | 
  target_size :: "('a => 'b :: len0 word) => nat"
 | 
|
71  | 
"target_size c == size (c arbitrary)"  | 
|
72  | 
  is_up :: "('a :: len0 word => 'b :: len0 word) => bool"
 | 
|
73  | 
"is_up c == source_size c <= target_size c"  | 
|
74  | 
  is_down :: "('a :: len0 word => 'b :: len0 word) => bool"
 | 
|
75  | 
"is_down c == target_size c <= source_size c"  | 
|
76  | 
||
77  | 
constdefs  | 
|
78  | 
of_bl :: "bool list => 'a :: len0 word"  | 
|
79  | 
"of_bl bl == word_of_int (bl_to_bin bl)"  | 
|
80  | 
to_bl :: "'a :: len0 word => bool list"  | 
|
81  | 
"to_bl w ==  | 
|
82  | 
  bin_to_bl (len_of TYPE ('a)) (uint w)"
 | 
|
83  | 
||
84  | 
word_reverse :: "'a :: len0 word => 'a word"  | 
|
85  | 
"word_reverse w == of_bl (rev (to_bl w))"  | 
|
86  | 
||
| 24333 | 87  | 
constdefs  | 
| 24465 | 88  | 
  word_int_case :: "(int => 'b) => ('a :: len0 word) => 'b"
 | 
| 24333 | 89  | 
"word_int_case f w == f (uint w)"  | 
90  | 
||
91  | 
syntax  | 
|
92  | 
of_int :: "int => 'a"  | 
|
93  | 
translations  | 
|
94  | 
"case x of of_int y => b" == "word_int_case (%y. b) x"  | 
|
95  | 
||
96  | 
||
| 24350 | 97  | 
subsection "Arithmetic operations"  | 
| 24333 | 98  | 
|
| 26514 | 99  | 
declare uint_def [code func del]  | 
100  | 
||
101  | 
lemma [code func]: "uint (word_of_int w \<Colon> 'a\<Colon>len0 word) = bintrunc (len_of TYPE('a)) w"
 | 
|
| 26559 | 102  | 
by (auto simp add: uint_def word_of_int_def intro!: Abs_word_inverse)  | 
103  | 
(insert range_bintrunc, auto)  | 
|
| 26514 | 104  | 
|
| 25762 | 105  | 
instantiation word :: (len0) "{number, uminus, minus, plus, one, zero, times, Divides.div, power, ord, bit}"
 | 
106  | 
begin  | 
|
107  | 
||
108  | 
definition  | 
|
109  | 
word_0_wi: "0 = word_of_int 0"  | 
|
110  | 
||
111  | 
definition  | 
|
112  | 
word_1_wi: "1 = word_of_int 1"  | 
|
113  | 
||
114  | 
definition  | 
|
115  | 
word_add_def: "a + b = word_of_int (uint a + uint b)"  | 
|
116  | 
||
117  | 
definition  | 
|
118  | 
word_sub_wi: "a - b = word_of_int (uint a - uint b)"  | 
|
119  | 
||
120  | 
definition  | 
|
121  | 
word_minus_def: "- a = word_of_int (- uint a)"  | 
|
122  | 
||
123  | 
definition  | 
|
124  | 
word_mult_def: "a * b = word_of_int (uint a * uint b)"  | 
|
125  | 
||
126  | 
definition  | 
|
127  | 
word_div_def: "a div b = word_of_int (uint a div uint b)"  | 
|
128  | 
||
129  | 
definition  | 
|
130  | 
word_mod_def: "a mod b = word_of_int (uint a mod uint b)"  | 
|
131  | 
||
132  | 
primrec power_word where  | 
|
133  | 
"(a\<Colon>'a word) ^ 0 = 1"  | 
|
134  | 
| "(a\<Colon>'a word) ^ Suc n = a * a ^ n"  | 
|
135  | 
||
136  | 
definition  | 
|
137  | 
word_number_of_def: "number_of w = word_of_int w"  | 
|
138  | 
||
139  | 
definition  | 
|
140  | 
word_le_def: "a \<le> b \<longleftrightarrow> uint a \<le> uint b"  | 
|
| 24415 | 141  | 
|
| 25762 | 142  | 
definition  | 
143  | 
word_less_def: "x < y \<longleftrightarrow> x \<le> y \<and> x \<noteq> (y \<Colon> 'a word)"  | 
|
144  | 
||
145  | 
definition  | 
|
146  | 
word_and_def:  | 
|
147  | 
"(a::'a word) AND b = word_of_int (uint a AND uint b)"  | 
|
148  | 
||
149  | 
definition  | 
|
150  | 
word_or_def:  | 
|
151  | 
"(a::'a word) OR b = word_of_int (uint a OR uint b)"  | 
|
152  | 
||
153  | 
definition  | 
|
154  | 
word_xor_def:  | 
|
155  | 
"(a::'a word) XOR b = word_of_int (uint a XOR uint b)"  | 
|
156  | 
||
157  | 
definition  | 
|
158  | 
word_not_def:  | 
|
159  | 
"NOT (a::'a word) = word_of_int (NOT (uint a))"  | 
|
160  | 
||
161  | 
instance ..  | 
|
162  | 
||
163  | 
end  | 
|
164  | 
||
165  | 
definition  | 
|
166  | 
word_succ :: "'a :: len0 word => 'a word"  | 
|
167  | 
where  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents: 
25762 
diff
changeset
 | 
168  | 
"word_succ a = word_of_int (Int.succ (uint a))"  | 
| 25762 | 169  | 
|
170  | 
definition  | 
|
171  | 
word_pred :: "'a :: len0 word => 'a word"  | 
|
172  | 
where  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents: 
25762 
diff
changeset
 | 
173  | 
"word_pred a = word_of_int (Int.pred (uint a))"  | 
| 24333 | 174  | 
|
175  | 
constdefs  | 
|
| 24465 | 176  | 
udvd :: "'a::len word => 'a::len word => bool" (infixl "udvd" 50)  | 
177  | 
"a udvd b == EX n>=0. uint b = n * uint a"  | 
|
178  | 
||
179  | 
  word_sle :: "'a :: len word => 'a word => bool" ("(_/ <=s _)" [50, 51] 50)
 | 
|
180  | 
"a <=s b == sint a <= sint b"  | 
|
181  | 
||
182  | 
  word_sless :: "'a :: len word => 'a word => bool" ("(_/ <s _)" [50, 51] 50)
 | 
|
183  | 
"(x <s y) == (x <=s y & x ~= y)"  | 
|
184  | 
||
| 24333 | 185  | 
|
| 26559 | 186  | 
|
| 24350 | 187  | 
subsection "Bit-wise operations"  | 
| 24333 | 188  | 
|
| 26559 | 189  | 
|
190  | 
||
191  | 
instantiation word :: (len0) bits  | 
|
192  | 
begin  | 
|
| 24333 | 193  | 
|
| 26559 | 194  | 
definition  | 
195  | 
word_test_bit_def: "test_bit a = bin_nth (uint a)"  | 
|
196  | 
||
197  | 
definition  | 
|
198  | 
word_set_bit_def: "set_bit a n x =  | 
|
| 24333 | 199  | 
word_of_int (bin_sc n (If x bit.B1 bit.B0) (uint a))"  | 
200  | 
||
| 26559 | 201  | 
definition  | 
202  | 
  word_set_bits_def: "(BITS n. f n) = of_bl (bl_of_nth (len_of TYPE ('a)) f)"
 | 
|
203  | 
||
204  | 
definition  | 
|
205  | 
word_lsb_def: "lsb a \<longleftrightarrow> bin_last (uint a) = bit.B1"  | 
|
206  | 
||
207  | 
definition shiftl1 :: "'a word \<Rightarrow> 'a word" where  | 
|
208  | 
"shiftl1 w = word_of_int (uint w BIT bit.B0)"  | 
|
209  | 
||
210  | 
definition shiftr1 :: "'a word \<Rightarrow> 'a word" where  | 
|
211  | 
-- "shift right as unsigned or as signed, ie logical or arithmetic"  | 
|
212  | 
"shiftr1 w = word_of_int (bin_rest (uint w))"  | 
|
213  | 
||
214  | 
definition  | 
|
215  | 
shiftl_def: "w << n = (shiftl1 ^ n) w"  | 
|
| 24465 | 216  | 
|
| 26559 | 217  | 
definition  | 
218  | 
shiftr_def: "w >> n = (shiftr1 ^ n) w"  | 
|
219  | 
||
220  | 
instance ..  | 
|
221  | 
||
222  | 
end  | 
|
| 24333 | 223  | 
|
| 26559 | 224  | 
instantiation word :: (len) bitss  | 
225  | 
begin  | 
|
226  | 
||
227  | 
definition  | 
|
| 24333 | 228  | 
word_msb_def:  | 
| 26559 | 229  | 
"msb a \<longleftrightarrow> bin_sign (sint a) = Int.Min"  | 
| 24333 | 230  | 
|
| 26559 | 231  | 
instance ..  | 
232  | 
||
233  | 
end  | 
|
| 24333 | 234  | 
|
235  | 
constdefs  | 
|
| 24465 | 236  | 
setBit :: "'a :: len0 word => nat => 'a word"  | 
| 24333 | 237  | 
"setBit w n == set_bit w n True"  | 
238  | 
||
| 24465 | 239  | 
clearBit :: "'a :: len0 word => nat => 'a word"  | 
| 24333 | 240  | 
"clearBit w n == set_bit w n False"  | 
241  | 
||
242  | 
||
| 24465 | 243  | 
subsection "Shift operations"  | 
244  | 
||
245  | 
constdefs  | 
|
246  | 
sshiftr1 :: "'a :: len word => 'a word"  | 
|
247  | 
"sshiftr1 w == word_of_int (bin_rest (sint w))"  | 
|
248  | 
||
249  | 
bshiftr1 :: "bool => 'a :: len word => 'a word"  | 
|
250  | 
"bshiftr1 b w == of_bl (b # butlast (to_bl w))"  | 
|
251  | 
||
252  | 
sshiftr :: "'a :: len word => nat => 'a word" (infixl ">>>" 55)  | 
|
253  | 
"w >>> n == (sshiftr1 ^ n) w"  | 
|
254  | 
||
255  | 
mask :: "nat => 'a::len word"  | 
|
256  | 
"mask n == (1 << n) - 1"  | 
|
257  | 
||
258  | 
revcast :: "'a :: len0 word => 'b :: len0 word"  | 
|
259  | 
  "revcast w ==  of_bl (takefill False (len_of TYPE('b)) (to_bl w))"
 | 
|
260  | 
||
261  | 
slice1 :: "nat => 'a :: len0 word => 'b :: len0 word"  | 
|
262  | 
"slice1 n w == of_bl (takefill False n (to_bl w))"  | 
|
263  | 
||
264  | 
slice :: "nat => 'a :: len0 word => 'b :: len0 word"  | 
|
265  | 
"slice n w == slice1 (size w - n) w"  | 
|
266  | 
||
267  | 
||
268  | 
subsection "Rotation"  | 
|
269  | 
||
270  | 
constdefs  | 
|
271  | 
rotater1 :: "'a list => 'a list"  | 
|
272  | 
"rotater1 ys ==  | 
|
273  | 
case ys of [] => [] | x # xs => last ys # butlast ys"  | 
|
274  | 
||
275  | 
rotater :: "nat => 'a list => 'a list"  | 
|
276  | 
"rotater n == rotater1 ^ n"  | 
|
277  | 
||
278  | 
word_rotr :: "nat => 'a :: len0 word => 'a :: len0 word"  | 
|
279  | 
"word_rotr n w == of_bl (rotater n (to_bl w))"  | 
|
280  | 
||
281  | 
word_rotl :: "nat => 'a :: len0 word => 'a :: len0 word"  | 
|
282  | 
"word_rotl n w == of_bl (rotate n (to_bl w))"  | 
|
283  | 
||
284  | 
word_roti :: "int => 'a :: len0 word => 'a :: len0 word"  | 
|
285  | 
"word_roti i w == if i >= 0 then word_rotr (nat i) w  | 
|
286  | 
else word_rotl (nat (- i)) w"  | 
|
287  | 
||
288  | 
||
289  | 
subsection "Split and cat operations"  | 
|
290  | 
||
291  | 
constdefs  | 
|
292  | 
word_cat :: "'a :: len0 word => 'b :: len0 word => 'c :: len0 word"  | 
|
293  | 
  "word_cat a b == word_of_int (bin_cat (uint a) (len_of TYPE ('b)) (uint b))"
 | 
|
294  | 
||
295  | 
  word_split :: "'a :: len0 word => ('b :: len0 word) * ('c :: len0 word)"
 | 
|
296  | 
"word_split a ==  | 
|
297  | 
   case bin_split (len_of TYPE ('c)) (uint a) of 
 | 
|
298  | 
(u, v) => (word_of_int u, word_of_int v)"  | 
|
299  | 
||
300  | 
word_rcat :: "'a :: len0 word list => 'b :: len0 word"  | 
|
301  | 
"word_rcat ws ==  | 
|
302  | 
  word_of_int (bin_rcat (len_of TYPE ('a)) (map uint ws))"
 | 
|
303  | 
||
304  | 
word_rsplit :: "'a :: len0 word => 'b :: len word list"  | 
|
305  | 
"word_rsplit w ==  | 
|
306  | 
  map word_of_int (bin_rsplit (len_of TYPE ('b)) (len_of TYPE ('a), uint w))"
 | 
|
307  | 
||
| 24333 | 308  | 
constdefs  | 
309  | 
-- "Largest representable machine integer."  | 
|
| 24465 | 310  | 
max_word :: "'a::len word"  | 
311  | 
  "max_word \<equiv> word_of_int (2^len_of TYPE('a) - 1)"
 | 
|
| 24333 | 312  | 
|
313  | 
consts  | 
|
| 24465 | 314  | 
of_bool :: "bool \<Rightarrow> 'a::len word"  | 
| 24333 | 315  | 
primrec  | 
316  | 
"of_bool False = 0"  | 
|
317  | 
"of_bool True = 1"  | 
|
318  | 
||
319  | 
||
| 24465 | 320  | 
lemmas of_nth_def = word_set_bits_def  | 
321  | 
||
| 24333 | 322  | 
lemmas word_size_gt_0 [iff] =  | 
| 25762 | 323  | 
xtr1 [OF word_size len_gt_0, standard]  | 
| 24465 | 324  | 
lemmas lens_gt_0 = word_size_gt_0 len_gt_0  | 
| 24333 | 325  | 
lemmas lens_not_0 [iff] = lens_gt_0 [THEN gr_implies_not0, standard]  | 
326  | 
||
327  | 
lemma uints_num: "uints n = {i. 0 \<le> i \<and> i < 2 ^ n}"
 | 
|
328  | 
by (simp add: uints_def range_bintrunc)  | 
|
329  | 
||
330  | 
lemma sints_num: "sints n = {i. - (2 ^ (n - 1)) \<le> i \<and> i < 2 ^ (n - 1)}"
 | 
|
331  | 
by (simp add: sints_def range_sbintrunc)  | 
|
332  | 
||
333  | 
lemmas atLeastLessThan_alt = atLeastLessThan_def [unfolded  | 
|
334  | 
atLeast_def lessThan_def Collect_conj_eq [symmetric]]  | 
|
335  | 
||
336  | 
lemma mod_in_reps: "m > 0 ==> y mod m : {0::int ..< m}"
 | 
|
337  | 
unfolding atLeastLessThan_alt by auto  | 
|
338  | 
||
339  | 
lemma  | 
|
340  | 
Rep_word_0:"0 <= Rep_word x" and  | 
|
| 24465 | 341  | 
  Rep_word_lt: "Rep_word (x::'a::len0 word) < 2 ^ len_of TYPE('a)"
 | 
| 24333 | 342  | 
by (auto simp: Rep_word [simplified])  | 
343  | 
||
344  | 
lemma Rep_word_mod_same:  | 
|
| 24465 | 345  | 
  "Rep_word x mod 2 ^ len_of TYPE('a) = Rep_word (x::'a::len0 word)"
 | 
| 24333 | 346  | 
by (simp add: int_mod_eq Rep_word_lt Rep_word_0)  | 
347  | 
||
348  | 
lemma td_ext_uint:  | 
|
| 24465 | 349  | 
  "td_ext (uint :: 'a word => int) word_of_int (uints (len_of TYPE('a::len0))) 
 | 
350  | 
    (%w::int. w mod 2 ^ len_of TYPE('a))"
 | 
|
| 24333 | 351  | 
apply (unfold td_ext_def')  | 
352  | 
apply (simp add: uints_num uint_def word_of_int_def bintrunc_mod2p)  | 
|
353  | 
apply (simp add: Rep_word_mod_same Rep_word_0 Rep_word_lt  | 
|
354  | 
word.Rep_word_inverse word.Abs_word_inverse int_mod_lem)  | 
|
355  | 
done  | 
|
356  | 
||
357  | 
lemmas int_word_uint = td_ext_uint [THEN td_ext.eq_norm, standard]  | 
|
358  | 
||
359  | 
interpretation word_uint:  | 
|
| 24465 | 360  | 
td_ext ["uint::'a::len0 word \<Rightarrow> int"  | 
| 24333 | 361  | 
word_of_int  | 
| 24465 | 362  | 
          "uints (len_of TYPE('a::len0))"
 | 
363  | 
          "\<lambda>w. w mod 2 ^ len_of TYPE('a::len0)"]
 | 
|
| 24333 | 364  | 
by (rule td_ext_uint)  | 
365  | 
||
366  | 
lemmas td_uint = word_uint.td_thm  | 
|
367  | 
||
368  | 
lemmas td_ext_ubin = td_ext_uint  | 
|
| 24465 | 369  | 
[simplified len_gt_0 no_bintr_alt1 [symmetric]]  | 
| 24333 | 370  | 
|
371  | 
interpretation word_ubin:  | 
|
| 24465 | 372  | 
td_ext ["uint::'a::len0 word \<Rightarrow> int"  | 
| 24333 | 373  | 
word_of_int  | 
| 24465 | 374  | 
          "uints (len_of TYPE('a::len0))"
 | 
375  | 
          "bintrunc (len_of TYPE('a::len0))"]
 | 
|
| 24333 | 376  | 
by (rule td_ext_ubin)  | 
377  | 
||
378  | 
lemma sint_sbintrunc':  | 
|
379  | 
"sint (word_of_int bin :: 'a word) =  | 
|
| 24465 | 380  | 
    (sbintrunc (len_of TYPE ('a :: len) - 1) bin)"
 | 
| 24333 | 381  | 
unfolding sint_uint  | 
382  | 
by (auto simp: word_ubin.eq_norm sbintrunc_bintrunc_lt)  | 
|
383  | 
||
384  | 
lemma uint_sint:  | 
|
| 24465 | 385  | 
  "uint w = bintrunc (len_of TYPE('a)) (sint (w :: 'a :: len word))"
 | 
| 24333 | 386  | 
unfolding sint_uint by (auto simp: bintrunc_sbintrunc_le)  | 
387  | 
||
388  | 
lemma bintr_uint':  | 
|
389  | 
"n >= size w ==> bintrunc n (uint w) = uint w"  | 
|
390  | 
apply (unfold word_size)  | 
|
391  | 
apply (subst word_ubin.norm_Rep [symmetric])  | 
|
392  | 
apply (simp only: bintrunc_bintrunc_min word_size min_def)  | 
|
393  | 
apply simp  | 
|
394  | 
done  | 
|
395  | 
||
396  | 
lemma wi_bintr':  | 
|
397  | 
"wb = word_of_int bin ==> n >= size wb ==>  | 
|
398  | 
word_of_int (bintrunc n bin) = wb"  | 
|
399  | 
unfolding word_size  | 
|
400  | 
by (clarsimp simp add : word_ubin.norm_eq_iff [symmetric] min_def)  | 
|
401  | 
||
402  | 
lemmas bintr_uint = bintr_uint' [unfolded word_size]  | 
|
403  | 
lemmas wi_bintr = wi_bintr' [unfolded word_size]  | 
|
404  | 
||
405  | 
lemma td_ext_sbin:  | 
|
| 24465 | 406  | 
  "td_ext (sint :: 'a word => int) word_of_int (sints (len_of TYPE('a::len))) 
 | 
407  | 
    (sbintrunc (len_of TYPE('a) - 1))"
 | 
|
| 24333 | 408  | 
apply (unfold td_ext_def' sint_uint)  | 
409  | 
apply (simp add : word_ubin.eq_norm)  | 
|
| 24465 | 410  | 
  apply (cases "len_of TYPE('a)")
 | 
| 24333 | 411  | 
apply (auto simp add : sints_def)  | 
412  | 
apply (rule sym [THEN trans])  | 
|
413  | 
apply (rule word_ubin.Abs_norm)  | 
|
414  | 
apply (simp only: bintrunc_sbintrunc)  | 
|
415  | 
apply (drule sym)  | 
|
416  | 
apply simp  | 
|
417  | 
done  | 
|
418  | 
||
419  | 
lemmas td_ext_sint = td_ext_sbin  | 
|
| 24465 | 420  | 
[simplified len_gt_0 no_sbintr_alt2 Suc_pred' [symmetric]]  | 
| 24333 | 421  | 
|
422  | 
(* We do sint before sbin, before sint is the user version  | 
|
423  | 
and interpretations do not produce thm duplicates. I.e.  | 
|
424  | 
we get the name word_sint.Rep_eqD, but not word_sbin.Req_eqD,  | 
|
425  | 
because the latter is the same thm as the former *)  | 
|
426  | 
interpretation word_sint:  | 
|
| 24465 | 427  | 
td_ext ["sint ::'a::len word => int"  | 
| 24333 | 428  | 
word_of_int  | 
| 24465 | 429  | 
          "sints (len_of TYPE('a::len))"
 | 
430  | 
          "%w. (w + 2^(len_of TYPE('a::len) - 1)) mod 2^len_of TYPE('a::len) -
 | 
|
431  | 
               2 ^ (len_of TYPE('a::len) - 1)"]
 | 
|
| 24333 | 432  | 
by (rule td_ext_sint)  | 
433  | 
||
434  | 
interpretation word_sbin:  | 
|
| 24465 | 435  | 
td_ext ["sint ::'a::len word => int"  | 
| 24333 | 436  | 
word_of_int  | 
| 24465 | 437  | 
          "sints (len_of TYPE('a::len))"
 | 
438  | 
          "sbintrunc (len_of TYPE('a::len) - 1)"]
 | 
|
| 24333 | 439  | 
by (rule td_ext_sbin)  | 
440  | 
||
441  | 
lemmas int_word_sint = td_ext_sint [THEN td_ext.eq_norm, standard]  | 
|
442  | 
||
443  | 
lemmas td_sint = word_sint.td  | 
|
444  | 
||
445  | 
lemma word_number_of_alt: "number_of b == word_of_int (number_of b)"  | 
|
446  | 
unfolding word_number_of_def by (simp add: number_of_eq)  | 
|
447  | 
||
448  | 
lemma word_no_wi: "number_of = word_of_int"  | 
|
449  | 
by (auto simp: word_number_of_def intro: ext)  | 
|
450  | 
||
| 24465 | 451  | 
lemma to_bl_def':  | 
452  | 
"(to_bl :: 'a :: len0 word => bool list) =  | 
|
453  | 
    bin_to_bl (len_of TYPE('a)) o uint"
 | 
|
454  | 
by (auto simp: to_bl_def intro: ext)  | 
|
455  | 
||
| 
25349
 
0d46bea01741
eliminated illegal schematic variables in where/of;
 
wenzelm 
parents: 
25149 
diff
changeset
 | 
456  | 
lemmas word_reverse_no_def [simp] = word_reverse_def [of "number_of w", standard]  | 
| 24465 | 457  | 
|
| 24333 | 458  | 
lemmas uints_mod = uints_def [unfolded no_bintr_alt1]  | 
459  | 
||
460  | 
lemma uint_bintrunc: "uint (number_of bin :: 'a word) =  | 
|
| 24465 | 461  | 
    number_of (bintrunc (len_of TYPE ('a :: len0)) bin)"
 | 
| 24333 | 462  | 
unfolding word_number_of_def number_of_eq  | 
463  | 
by (auto intro: word_ubin.eq_norm)  | 
|
464  | 
||
465  | 
lemma sint_sbintrunc: "sint (number_of bin :: 'a word) =  | 
|
| 24465 | 466  | 
    number_of (sbintrunc (len_of TYPE ('a :: len) - 1) bin)" 
 | 
| 24333 | 467  | 
unfolding word_number_of_def number_of_eq  | 
| 25149 | 468  | 
by (subst word_sbin.eq_norm) simp  | 
| 24333 | 469  | 
|
470  | 
lemma unat_bintrunc:  | 
|
| 24465 | 471  | 
"unat (number_of bin :: 'a :: len0 word) =  | 
472  | 
    number_of (bintrunc (len_of TYPE('a)) bin)"
 | 
|
| 24333 | 473  | 
unfolding unat_def nat_number_of_def  | 
474  | 
by (simp only: uint_bintrunc)  | 
|
475  | 
||
476  | 
(* WARNING - these may not always be helpful *)  | 
|
477  | 
declare  | 
|
478  | 
uint_bintrunc [simp]  | 
|
479  | 
sint_sbintrunc [simp]  | 
|
480  | 
unat_bintrunc [simp]  | 
|
481  | 
||
| 24465 | 482  | 
lemma size_0_eq: "size (w :: 'a :: len0 word) = 0 ==> v = w"  | 
| 24333 | 483  | 
apply (unfold word_size)  | 
484  | 
apply (rule word_uint.Rep_eqD)  | 
|
485  | 
apply (rule box_equals)  | 
|
486  | 
defer  | 
|
487  | 
apply (rule word_ubin.norm_Rep)+  | 
|
488  | 
apply simp  | 
|
489  | 
done  | 
|
490  | 
||
491  | 
lemmas uint_lem = word_uint.Rep [unfolded uints_num mem_Collect_eq]  | 
|
492  | 
lemmas sint_lem = word_sint.Rep [unfolded sints_num mem_Collect_eq]  | 
|
493  | 
lemmas uint_ge_0 [iff] = uint_lem [THEN conjunct1, standard]  | 
|
494  | 
lemmas uint_lt2p [iff] = uint_lem [THEN conjunct2, standard]  | 
|
495  | 
lemmas sint_ge = sint_lem [THEN conjunct1, standard]  | 
|
496  | 
lemmas sint_lt = sint_lem [THEN conjunct2, standard]  | 
|
497  | 
||
498  | 
lemma sign_uint_Pls [simp]:  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents: 
25762 
diff
changeset
 | 
499  | 
"bin_sign (uint x) = Int.Pls"  | 
| 24333 | 500  | 
by (simp add: sign_Pls_ge_0 number_of_eq)  | 
501  | 
||
502  | 
lemmas uint_m2p_neg = iffD2 [OF diff_less_0_iff_less uint_lt2p, standard]  | 
|
503  | 
lemmas uint_m2p_not_non_neg =  | 
|
504  | 
iffD2 [OF linorder_not_le uint_m2p_neg, standard]  | 
|
505  | 
||
506  | 
lemma lt2p_lem:  | 
|
| 24465 | 507  | 
  "len_of TYPE('a) <= n ==> uint (w :: 'a :: len0 word) < 2 ^ n"
 | 
| 24333 | 508  | 
by (rule xtr8 [OF _ uint_lt2p]) simp  | 
509  | 
||
510  | 
lemmas uint_le_0_iff [simp] =  | 
|
511  | 
uint_ge_0 [THEN leD, THEN linorder_antisym_conv1, standard]  | 
|
512  | 
||
513  | 
lemma uint_nat: "uint w == int (unat w)"  | 
|
514  | 
unfolding unat_def by auto  | 
|
515  | 
||
516  | 
lemma uint_number_of:  | 
|
| 24465 | 517  | 
  "uint (number_of b :: 'a :: len0 word) = number_of b mod 2 ^ len_of TYPE('a)"
 | 
| 24333 | 518  | 
unfolding word_number_of_alt  | 
519  | 
by (simp only: int_word_uint)  | 
|
520  | 
||
521  | 
lemma unat_number_of:  | 
|
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents: 
25762 
diff
changeset
 | 
522  | 
"bin_sign b = Int.Pls ==>  | 
| 24465 | 523  | 
  unat (number_of b::'a::len0 word) = number_of b mod 2 ^ len_of TYPE ('a)"
 | 
| 24333 | 524  | 
apply (unfold unat_def)  | 
525  | 
apply (clarsimp simp only: uint_number_of)  | 
|
526  | 
apply (rule nat_mod_distrib [THEN trans])  | 
|
527  | 
apply (erule sign_Pls_ge_0 [THEN iffD1])  | 
|
528  | 
apply (simp_all add: nat_power_eq)  | 
|
529  | 
done  | 
|
530  | 
||
| 24465 | 531  | 
lemma sint_number_of: "sint (number_of b :: 'a :: len word) = (number_of b +  | 
532  | 
    2 ^ (len_of TYPE('a) - 1)) mod 2 ^ len_of TYPE('a) -
 | 
|
533  | 
    2 ^ (len_of TYPE('a) - 1)"
 | 
|
| 24333 | 534  | 
unfolding word_number_of_alt by (rule int_word_sint)  | 
535  | 
||
536  | 
lemma word_of_int_bin [simp] :  | 
|
| 24465 | 537  | 
"(word_of_int (number_of bin) :: 'a :: len0 word) = (number_of bin)"  | 
| 24333 | 538  | 
unfolding word_number_of_alt by auto  | 
539  | 
||
540  | 
lemma word_int_case_wi:  | 
|
541  | 
"word_int_case f (word_of_int i :: 'b word) =  | 
|
| 24465 | 542  | 
    f (i mod 2 ^ len_of TYPE('b::len0))"
 | 
| 24333 | 543  | 
unfolding word_int_case_def by (simp add: word_uint.eq_norm)  | 
544  | 
||
545  | 
lemma word_int_split:  | 
|
546  | 
"P (word_int_case f x) =  | 
|
| 24465 | 547  | 
(ALL i. x = (word_of_int i :: 'b :: len0 word) &  | 
548  | 
      0 <= i & i < 2 ^ len_of TYPE('b) --> P (f i))"
 | 
|
| 24333 | 549  | 
unfolding word_int_case_def  | 
550  | 
by (auto simp: word_uint.eq_norm int_mod_eq')  | 
|
551  | 
||
552  | 
lemma word_int_split_asm:  | 
|
553  | 
"P (word_int_case f x) =  | 
|
| 24465 | 554  | 
(~ (EX n. x = (word_of_int n :: 'b::len0 word) &  | 
555  | 
      0 <= n & n < 2 ^ len_of TYPE('b::len0) & ~ P (f n)))"
 | 
|
| 24333 | 556  | 
unfolding word_int_case_def  | 
557  | 
by (auto simp: word_uint.eq_norm int_mod_eq')  | 
|
558  | 
||
559  | 
lemmas uint_range' =  | 
|
560  | 
word_uint.Rep [unfolded uints_num mem_Collect_eq, standard]  | 
|
561  | 
lemmas sint_range' = word_sint.Rep [unfolded One_nat_def  | 
|
562  | 
sints_num mem_Collect_eq, standard]  | 
|
563  | 
||
564  | 
lemma uint_range_size: "0 <= uint w & uint w < 2 ^ size w"  | 
|
565  | 
unfolding word_size by (rule uint_range')  | 
|
566  | 
||
567  | 
lemma sint_range_size:  | 
|
568  | 
"- (2 ^ (size w - Suc 0)) <= sint w & sint w < 2 ^ (size w - Suc 0)"  | 
|
569  | 
unfolding word_size by (rule sint_range')  | 
|
570  | 
||
571  | 
lemmas sint_above_size = sint_range_size  | 
|
572  | 
[THEN conjunct2, THEN [2] xtr8, folded One_nat_def, standard]  | 
|
573  | 
||
574  | 
lemmas sint_below_size = sint_range_size  | 
|
575  | 
[THEN conjunct1, THEN [2] order_trans, folded One_nat_def, standard]  | 
|
576  | 
||
| 24465 | 577  | 
lemma test_bit_eq_iff: "(test_bit (u::'a::len0 word) = test_bit v) = (u = v)"  | 
| 24333 | 578  | 
unfolding word_test_bit_def by (simp add: bin_nth_eq_iff)  | 
579  | 
||
| 24465 | 580  | 
lemma test_bit_size [rule_format] : "(w::'a::len0 word) !! n --> n < size w"  | 
| 24333 | 581  | 
apply (unfold word_test_bit_def)  | 
582  | 
apply (subst word_ubin.norm_Rep [symmetric])  | 
|
583  | 
apply (simp only: nth_bintr word_size)  | 
|
584  | 
apply fast  | 
|
585  | 
done  | 
|
586  | 
||
587  | 
lemma word_eqI [rule_format] :  | 
|
| 24465 | 588  | 
fixes u :: "'a::len0 word"  | 
| 24333 | 589  | 
shows "(ALL n. n < size u --> u !! n = v !! n) ==> u = v"  | 
590  | 
apply (rule test_bit_eq_iff [THEN iffD1])  | 
|
591  | 
apply (rule ext)  | 
|
592  | 
apply (erule allE)  | 
|
593  | 
apply (erule impCE)  | 
|
594  | 
prefer 2  | 
|
595  | 
apply assumption  | 
|
596  | 
apply (auto dest!: test_bit_size simp add: word_size)  | 
|
597  | 
done  | 
|
598  | 
||
599  | 
lemmas word_eqD = test_bit_eq_iff [THEN iffD2, THEN fun_cong, standard]  | 
|
600  | 
||
601  | 
lemma test_bit_bin': "w !! n = (n < size w & bin_nth (uint w) n)"  | 
|
602  | 
unfolding word_test_bit_def word_size  | 
|
603  | 
by (simp add: nth_bintr [symmetric])  | 
|
604  | 
||
605  | 
lemmas test_bit_bin = test_bit_bin' [unfolded word_size]  | 
|
606  | 
||
607  | 
lemma bin_nth_uint_imp': "bin_nth (uint w) n --> n < size w"  | 
|
608  | 
apply (unfold word_size)  | 
|
609  | 
apply (rule impI)  | 
|
610  | 
apply (rule nth_bintr [THEN iffD1, THEN conjunct1])  | 
|
611  | 
apply (subst word_ubin.norm_Rep)  | 
|
612  | 
apply assumption  | 
|
613  | 
done  | 
|
614  | 
||
615  | 
lemma bin_nth_sint':  | 
|
616  | 
"n >= size w --> bin_nth (sint w) n = bin_nth (sint w) (size w - 1)"  | 
|
617  | 
apply (rule impI)  | 
|
618  | 
apply (subst word_sbin.norm_Rep [symmetric])  | 
|
619  | 
apply (simp add : nth_sbintr word_size)  | 
|
620  | 
apply auto  | 
|
621  | 
done  | 
|
622  | 
||
623  | 
lemmas bin_nth_uint_imp = bin_nth_uint_imp' [rule_format, unfolded word_size]  | 
|
624  | 
lemmas bin_nth_sint = bin_nth_sint' [rule_format, unfolded word_size]  | 
|
625  | 
||
| 24465 | 626  | 
(* type definitions theorem for in terms of equivalent bool list *)  | 
627  | 
lemma td_bl:  | 
|
628  | 
"type_definition (to_bl :: 'a::len0 word => bool list)  | 
|
629  | 
of_bl  | 
|
630  | 
                   {bl. length bl = len_of TYPE('a)}"
 | 
|
631  | 
apply (unfold type_definition_def of_bl_def to_bl_def)  | 
|
632  | 
apply (simp add: word_ubin.eq_norm)  | 
|
633  | 
apply safe  | 
|
634  | 
apply (drule sym)  | 
|
635  | 
apply simp  | 
|
636  | 
done  | 
|
637  | 
||
638  | 
interpretation word_bl:  | 
|
639  | 
type_definition ["to_bl :: 'a::len0 word => bool list"  | 
|
640  | 
of_bl  | 
|
641  | 
                   "{bl. length bl = len_of TYPE('a::len0)}"]
 | 
|
642  | 
by (rule td_bl)  | 
|
643  | 
||
644  | 
lemma word_size_bl: "size w == size (to_bl w)"  | 
|
645  | 
unfolding word_size by auto  | 
|
646  | 
||
647  | 
lemma to_bl_use_of_bl:  | 
|
648  | 
"(to_bl w = bl) = (w = of_bl bl \<and> length bl = length (to_bl w))"  | 
|
649  | 
by (fastsimp elim!: word_bl.Abs_inverse [simplified])  | 
|
650  | 
||
651  | 
lemma to_bl_word_rev: "to_bl (word_reverse w) = rev (to_bl w)"  | 
|
652  | 
unfolding word_reverse_def by (simp add: word_bl.Abs_inverse)  | 
|
653  | 
||
654  | 
lemma word_rev_rev [simp] : "word_reverse (word_reverse w) = w"  | 
|
655  | 
unfolding word_reverse_def by (simp add : word_bl.Abs_inverse)  | 
|
656  | 
||
657  | 
lemma word_rev_gal: "word_reverse w = u ==> word_reverse u = w"  | 
|
658  | 
by auto  | 
|
659  | 
||
660  | 
lemmas word_rev_gal' = sym [THEN word_rev_gal, symmetric, standard]  | 
|
661  | 
||
662  | 
lemmas length_bl_gt_0 [iff] = xtr1 [OF word_bl.Rep' len_gt_0, standard]  | 
|
663  | 
lemmas bl_not_Nil [iff] =  | 
|
664  | 
length_bl_gt_0 [THEN length_greater_0_conv [THEN iffD1], standard]  | 
|
665  | 
lemmas length_bl_neq_0 [iff] = length_bl_gt_0 [THEN gr_implies_not0]  | 
|
666  | 
||
| 
25919
 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 
haftmann 
parents: 
25762 
diff
changeset
 | 
667  | 
lemma hd_bl_sign_sint: "hd (to_bl w) = (bin_sign (sint w) = Int.Min)"  | 
| 24465 | 668  | 
apply (unfold to_bl_def sint_uint)  | 
669  | 
apply (rule trans [OF _ bl_sbin_sign])  | 
|
670  | 
apply simp  | 
|
671  | 
done  | 
|
672  | 
||
673  | 
lemma of_bl_drop':  | 
|
674  | 
  "lend = length bl - len_of TYPE ('a :: len0) ==> 
 | 
|
675  | 
of_bl (drop lend bl) = (of_bl bl :: 'a word)"  | 
|
676  | 
apply (unfold of_bl_def)  | 
|
677  | 
apply (clarsimp simp add : trunc_bl2bin [symmetric])  | 
|
678  | 
done  | 
|
679  | 
||
680  | 
lemmas of_bl_no = of_bl_def [folded word_number_of_def]  | 
|
681  | 
||
682  | 
lemma test_bit_of_bl:  | 
|
683  | 
  "(of_bl bl::'a::len0 word) !! n = (rev bl ! n \<and> n < len_of TYPE('a) \<and> n < length bl)"
 | 
|
684  | 
apply (unfold of_bl_def word_test_bit_def)  | 
|
685  | 
apply (auto simp add: word_size word_ubin.eq_norm nth_bintr bin_nth_of_bl)  | 
|
686  | 
done  | 
|
687  | 
||
688  | 
lemma no_of_bl:  | 
|
689  | 
  "(number_of bin ::'a::len0 word) = of_bl (bin_to_bl (len_of TYPE ('a)) bin)"
 | 
|
690  | 
unfolding word_size of_bl_no by (simp add : word_number_of_def)  | 
|
691  | 
||
692  | 
lemma uint_bl: "to_bl w == bin_to_bl (size w) (uint w)"  | 
|
693  | 
unfolding word_size to_bl_def by auto  | 
|
694  | 
||
695  | 
lemma to_bl_bin: "bl_to_bin (to_bl w) = uint w"  | 
|
696  | 
unfolding uint_bl by (simp add : word_size)  | 
|
697  | 
||
698  | 
lemma to_bl_of_bin:  | 
|
699  | 
  "to_bl (word_of_int bin::'a::len0 word) = bin_to_bl (len_of TYPE('a)) bin"
 | 
|
700  | 
unfolding uint_bl by (clarsimp simp add: word_ubin.eq_norm word_size)  | 
|
701  | 
||
702  | 
lemmas to_bl_no_bin [simp] = to_bl_of_bin [folded word_number_of_def]  | 
|
703  | 
||
704  | 
lemma to_bl_to_bin [simp] : "bl_to_bin (to_bl w) = uint w"  | 
|
705  | 
unfolding uint_bl by (simp add : word_size)  | 
|
706  | 
||
707  | 
lemmas uint_bl_bin [simp] = trans [OF bin_bl_bin word_ubin.norm_Rep, standard]  | 
|
| 24333 | 708  | 
|
709  | 
lemmas num_AB_u [simp] = word_uint.Rep_inverse  | 
|
710  | 
[unfolded o_def word_number_of_def [symmetric], standard]  | 
|
711  | 
lemmas num_AB_s [simp] = word_sint.Rep_inverse  | 
|
712  | 
[unfolded o_def word_number_of_def [symmetric], standard]  | 
|
713  | 
||
714  | 
(* naturals *)  | 
|
715  | 
lemma uints_unats: "uints n = int ` unats n"  | 
|
716  | 
apply (unfold unats_def uints_num)  | 
|
717  | 
apply safe  | 
|
718  | 
apply (rule_tac image_eqI)  | 
|
719  | 
apply (erule_tac nat_0_le [symmetric])  | 
|
720  | 
apply auto  | 
|
721  | 
apply (erule_tac nat_less_iff [THEN iffD2])  | 
|
722  | 
apply (rule_tac [2] zless_nat_eq_int_zless [THEN iffD1])  | 
|
723  | 
apply (auto simp add : nat_power_eq int_power)  | 
|
724  | 
done  | 
|
725  | 
||
726  | 
lemma unats_uints: "unats n = nat ` uints n"  | 
|
| 
25349
 
0d46bea01741
eliminated illegal schematic variables in where/of;
 
wenzelm 
parents: 
25149 
diff
changeset
 | 
727  | 
by (auto simp add : uints_unats image_iff)  | 
| 24333 | 728  | 
|
729  | 
lemmas bintr_num = word_ubin.norm_eq_iff  | 
|
730  | 
[symmetric, folded word_number_of_def, standard]  | 
|
731  | 
lemmas sbintr_num = word_sbin.norm_eq_iff  | 
|
732  | 
[symmetric, folded word_number_of_def, standard]  | 
|
733  | 
||
734  | 
lemmas num_of_bintr = word_ubin.Abs_norm [folded word_number_of_def, standard]  | 
|
735  | 
lemmas num_of_sbintr = word_sbin.Abs_norm [folded word_number_of_def, standard];  | 
|
736  | 
||
737  | 
(* don't add these to simpset, since may want bintrunc n w to be simplified;  | 
|
738  | 
may want these in reverse, but loop as simp rules, so use following *)  | 
|
739  | 
||
740  | 
lemma num_of_bintr':  | 
|
| 24465 | 741  | 
  "bintrunc (len_of TYPE('a :: len0)) a = b ==> 
 | 
| 24333 | 742  | 
number_of a = (number_of b :: 'a word)"  | 
743  | 
apply safe  | 
|
744  | 
apply (rule_tac num_of_bintr [symmetric])  | 
|
745  | 
done  | 
|
746  | 
||
747  | 
lemma num_of_sbintr':  | 
|
| 24465 | 748  | 
  "sbintrunc (len_of TYPE('a :: len) - 1) a = b ==> 
 | 
| 24333 | 749  | 
number_of a = (number_of b :: 'a word)"  | 
750  | 
apply safe  | 
|
751  | 
apply (rule_tac num_of_sbintr [symmetric])  | 
|
752  | 
done  | 
|
753  | 
||
754  | 
lemmas num_abs_bintr = sym [THEN trans,  | 
|
| 25762 | 755  | 
OF num_of_bintr word_number_of_def, standard]  | 
| 24333 | 756  | 
lemmas num_abs_sbintr = sym [THEN trans,  | 
| 25762 | 757  | 
OF num_of_sbintr word_number_of_def, standard]  | 
| 24465 | 758  | 
|
| 24333 | 759  | 
(** cast - note, no arg for new length, as it's determined by type of result,  | 
760  | 
thus in "cast w = w, the type means cast to length of w! **)  | 
|
761  | 
||
762  | 
lemma ucast_id: "ucast w = w"  | 
|
763  | 
unfolding ucast_def by auto  | 
|
764  | 
||
765  | 
lemma scast_id: "scast w = w"  | 
|
766  | 
unfolding scast_def by auto  | 
|
767  | 
||
| 24465 | 768  | 
lemma ucast_bl: "ucast w == of_bl (to_bl w)"  | 
769  | 
unfolding ucast_def of_bl_def uint_bl  | 
|
770  | 
by (auto simp add : word_size)  | 
|
771  | 
||
| 24333 | 772  | 
lemma nth_ucast:  | 
| 24465 | 773  | 
  "(ucast w::'a::len0 word) !! n = (w !! n & n < len_of TYPE('a))"
 | 
| 24333 | 774  | 
apply (unfold ucast_def test_bit_bin)  | 
775  | 
apply (simp add: word_ubin.eq_norm nth_bintr word_size)  | 
|
776  | 
apply (fast elim!: bin_nth_uint_imp)  | 
|
777  | 
done  | 
|
778  | 
||
779  | 
(* for literal u(s)cast *)  | 
|
780  | 
||
781  | 
lemma ucast_bintr [simp]:  | 
|
| 24465 | 782  | 
"ucast (number_of w ::'a::len0 word) =  | 
783  | 
   number_of (bintrunc (len_of TYPE('a)) w)"
 | 
|
| 24333 | 784  | 
unfolding ucast_def by simp  | 
785  | 
||
786  | 
lemma scast_sbintr [simp]:  | 
|
| 24465 | 787  | 
"scast (number_of w ::'a::len word) =  | 
788  | 
   number_of (sbintrunc (len_of TYPE('a) - Suc 0) w)"
 | 
|
| 24333 | 789  | 
unfolding scast_def by simp  | 
790  | 
||
791  | 
lemmas source_size = source_size_def [unfolded Let_def word_size]  | 
|
792  | 
lemmas target_size = target_size_def [unfolded Let_def word_size]  | 
|
793  | 
lemmas is_down = is_down_def [unfolded source_size target_size]  | 
|
794  | 
lemmas is_up = is_up_def [unfolded source_size target_size]  | 
|
795  | 
||
796  | 
lemmas is_up_down =  | 
|
797  | 
trans [OF is_up [THEN meta_eq_to_obj_eq]  | 
|
798  | 
is_down [THEN meta_eq_to_obj_eq, symmetric],  | 
|
799  | 
standard]  | 
|
800  | 
||
801  | 
lemma down_cast_same': "uc = ucast ==> is_down uc ==> uc = scast"  | 
|
802  | 
apply (unfold is_down)  | 
|
803  | 
apply safe  | 
|
804  | 
apply (rule ext)  | 
|
805  | 
apply (unfold ucast_def scast_def uint_sint)  | 
|
806  | 
apply (rule word_ubin.norm_eq_iff [THEN iffD1])  | 
|
807  | 
apply simp  | 
|
808  | 
done  | 
|
809  | 
||
| 24465 | 810  | 
lemma word_rev_tf':  | 
811  | 
"r = to_bl (of_bl bl) ==> r = rev (takefill False (length r) (rev bl))"  | 
|
812  | 
unfolding of_bl_def uint_bl  | 
|
813  | 
by (clarsimp simp add: bl_bin_bl_rtf word_ubin.eq_norm word_size)  | 
|
814  | 
||
815  | 
lemmas word_rev_tf = refl [THEN word_rev_tf', unfolded word_bl.Rep', standard]  | 
|
816  | 
||
817  | 
lemmas word_rep_drop = word_rev_tf [simplified takefill_alt,  | 
|
818  | 
simplified, simplified rev_take, simplified]  | 
|
819  | 
||
820  | 
lemma to_bl_ucast:  | 
|
821  | 
"to_bl (ucast (w::'b::len0 word) ::'a::len0 word) =  | 
|
822  | 
   replicate (len_of TYPE('a) - len_of TYPE('b)) False @
 | 
|
823  | 
   drop (len_of TYPE('b) - len_of TYPE('a)) (to_bl w)"
 | 
|
824  | 
apply (unfold ucast_bl)  | 
|
825  | 
apply (rule trans)  | 
|
826  | 
apply (rule word_rep_drop)  | 
|
827  | 
apply simp  | 
|
828  | 
done  | 
|
829  | 
||
830  | 
lemma ucast_up_app':  | 
|
831  | 
"uc = ucast ==> source_size uc + n = target_size uc ==>  | 
|
832  | 
to_bl (uc w) = replicate n False @ (to_bl w)"  | 
|
833  | 
apply (auto simp add : source_size target_size to_bl_ucast)  | 
|
834  | 
apply (rule_tac f = "%n. replicate n False" in arg_cong)  | 
|
835  | 
apply simp  | 
|
836  | 
done  | 
|
837  | 
||
838  | 
lemma ucast_down_drop':  | 
|
839  | 
"uc = ucast ==> source_size uc = target_size uc + n ==>  | 
|
840  | 
to_bl (uc w) = drop n (to_bl w)"  | 
|
841  | 
by (auto simp add : source_size target_size to_bl_ucast)  | 
|
842  | 
||
843  | 
lemma scast_down_drop':  | 
|
844  | 
"sc = scast ==> source_size sc = target_size sc + n ==>  | 
|
845  | 
to_bl (sc w) = drop n (to_bl w)"  | 
|
846  | 
apply (subgoal_tac "sc = ucast")  | 
|
847  | 
apply safe  | 
|
848  | 
apply simp  | 
|
849  | 
apply (erule refl [THEN ucast_down_drop'])  | 
|
850  | 
apply (rule refl [THEN down_cast_same', symmetric])  | 
|
851  | 
apply (simp add : source_size target_size is_down)  | 
|
852  | 
done  | 
|
853  | 
||
| 24333 | 854  | 
lemma sint_up_scast':  | 
855  | 
"sc = scast ==> is_up sc ==> sint (sc w) = sint w"  | 
|
856  | 
apply (unfold is_up)  | 
|
857  | 
apply safe  | 
|
858  | 
apply (simp add: scast_def word_sbin.eq_norm)  | 
|
859  | 
apply (rule box_equals)  | 
|
860  | 
prefer 3  | 
|
861  | 
apply (rule word_sbin.norm_Rep)  | 
|
862  | 
apply (rule sbintrunc_sbintrunc_l)  | 
|
863  | 
defer  | 
|
864  | 
apply (subst word_sbin.norm_Rep)  | 
|
865  | 
apply (rule refl)  | 
|
866  | 
apply simp  | 
|
867  | 
done  | 
|
868  | 
||
869  | 
lemma uint_up_ucast':  | 
|
870  | 
"uc = ucast ==> is_up uc ==> uint (uc w) = uint w"  | 
|
871  | 
apply (unfold is_up)  | 
|
872  | 
apply safe  | 
|
873  | 
apply (rule bin_eqI)  | 
|
874  | 
apply (fold word_test_bit_def)  | 
|
875  | 
apply (auto simp add: nth_ucast)  | 
|
876  | 
apply (auto simp add: test_bit_bin)  | 
|
877  | 
done  | 
|
878  | 
||
879  | 
lemmas down_cast_same = refl [THEN down_cast_same']  | 
|
| 24465 | 880  | 
lemmas ucast_up_app = refl [THEN ucast_up_app']  | 
881  | 
lemmas ucast_down_drop = refl [THEN ucast_down_drop']  | 
|
882  | 
lemmas scast_down_drop = refl [THEN scast_down_drop']  | 
|
| 24333 | 883  | 
lemmas uint_up_ucast = refl [THEN uint_up_ucast']  | 
884  | 
lemmas sint_up_scast = refl [THEN sint_up_scast']  | 
|
885  | 
||
886  | 
lemma ucast_up_ucast': "uc = ucast ==> is_up uc ==> ucast (uc w) = ucast w"  | 
|
887  | 
apply (simp (no_asm) add: ucast_def)  | 
|
888  | 
apply (clarsimp simp add: uint_up_ucast)  | 
|
889  | 
done  | 
|
890  | 
||
891  | 
lemma scast_up_scast': "sc = scast ==> is_up sc ==> scast (sc w) = scast w"  | 
|
892  | 
apply (simp (no_asm) add: scast_def)  | 
|
893  | 
apply (clarsimp simp add: sint_up_scast)  | 
|
894  | 
done  | 
|
895  | 
||
| 24465 | 896  | 
lemma ucast_of_bl_up':  | 
897  | 
"w = of_bl bl ==> size bl <= size w ==> ucast w = of_bl bl"  | 
|
898  | 
by (auto simp add : nth_ucast word_size test_bit_of_bl intro!: word_eqI)  | 
|
899  | 
||
| 24333 | 900  | 
lemmas ucast_up_ucast = refl [THEN ucast_up_ucast']  | 
901  | 
lemmas scast_up_scast = refl [THEN scast_up_scast']  | 
|
| 24465 | 902  | 
lemmas ucast_of_bl_up = refl [THEN ucast_of_bl_up']  | 
| 24333 | 903  | 
|
904  | 
lemmas ucast_up_ucast_id = trans [OF ucast_up_ucast ucast_id]  | 
|
905  | 
lemmas scast_up_scast_id = trans [OF scast_up_scast scast_id]  | 
|
906  | 
||
907  | 
lemmas isduu = is_up_down [where c = "ucast", THEN iffD2]  | 
|
908  | 
lemmas isdus = is_up_down [where c = "scast", THEN iffD2]  | 
|
909  | 
lemmas ucast_down_ucast_id = isduu [THEN ucast_up_ucast_id]  | 
|
910  | 
lemmas scast_down_scast_id = isdus [THEN ucast_up_ucast_id]  | 
|
911  | 
||
912  | 
lemma up_ucast_surj:  | 
|
| 24465 | 913  | 
"is_up (ucast :: 'b::len0 word => 'a::len0 word) ==>  | 
| 24333 | 914  | 
surj (ucast :: 'a word => 'b word)"  | 
915  | 
by (rule surjI, erule ucast_up_ucast_id)  | 
|
916  | 
||
917  | 
lemma up_scast_surj:  | 
|
| 24465 | 918  | 
"is_up (scast :: 'b::len word => 'a::len word) ==>  | 
| 24333 | 919  | 
surj (scast :: 'a word => 'b word)"  | 
920  | 
by (rule surjI, erule scast_up_scast_id)  | 
|
921  | 
||
922  | 
lemma down_scast_inj:  | 
|
| 24465 | 923  | 
"is_down (scast :: 'b::len word => 'a::len word) ==>  | 
| 24333 | 924  | 
inj_on (ucast :: 'a word => 'b word) A"  | 
925  | 
by (rule inj_on_inverseI, erule scast_down_scast_id)  | 
|
926  | 
||
927  | 
lemma down_ucast_inj:  | 
|
| 24465 | 928  | 
"is_down (ucast :: 'b::len0 word => 'a::len0 word) ==>  | 
| 24333 | 929  | 
inj_on (ucast :: 'a word => 'b word) A"  | 
930  | 
by (rule inj_on_inverseI, erule ucast_down_ucast_id)  | 
|
931  | 
||
| 24465 | 932  | 
lemma of_bl_append_same: "of_bl (X @ to_bl w) = w"  | 
933  | 
by (rule word_bl.Rep_eqD) (simp add: word_rep_drop)  | 
|
| 24333 | 934  | 
|
935  | 
lemma ucast_down_no':  | 
|
936  | 
"uc = ucast ==> is_down uc ==> uc (number_of bin) = number_of bin"  | 
|
937  | 
apply (unfold word_number_of_def is_down)  | 
|
938  | 
apply (clarsimp simp add: ucast_def word_ubin.eq_norm)  | 
|
939  | 
apply (rule word_ubin.norm_eq_iff [THEN iffD1])  | 
|
940  | 
apply (erule bintrunc_bintrunc_ge)  | 
|
941  | 
done  | 
|
942  | 
||
943  | 
lemmas ucast_down_no = ucast_down_no' [OF refl]  | 
|
944  | 
||
| 24465 | 945  | 
lemma ucast_down_bl': "uc = ucast ==> is_down uc ==> uc (of_bl bl) = of_bl bl"  | 
946  | 
unfolding of_bl_no by clarify (erule ucast_down_no)  | 
|
947  | 
||
948  | 
lemmas ucast_down_bl = ucast_down_bl' [OF refl]  | 
|
949  | 
||
950  | 
lemmas slice_def' = slice_def [unfolded word_size]  | 
|
| 26559 | 951  | 
lemmas test_bit_def' = word_test_bit_def [THEN fun_cong]  | 
| 24465 | 952  | 
|
953  | 
lemmas word_log_defs = word_and_def word_or_def word_xor_def word_not_def  | 
|
954  | 
lemmas word_log_bin_defs = word_log_defs  | 
|
955  | 
||
| 24333 | 956  | 
end  |