| author | chaieb | 
| Wed, 22 Aug 2007 17:13:41 +0200 | |
| changeset 24403 | b7c3ee2ca184 | 
| parent 24397 | eaf37b780683 | 
| child 24408 | 058c5613a86f | 
| 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  | 
|
| 
24397
 
eaf37b780683
move bool list stuff from WordDefinition and WordArith to new WordBoolList theory
 
huffman 
parents: 
24378 
diff
changeset
 | 
11  | 
theory WordDefinition imports Size BinOperations TdThs begin  | 
| 24333 | 12  | 
|
13  | 
typedef (open word) 'a word  | 
|
14  | 
  = "{(0::int) ..< 2^len_of TYPE('a::len0)}" by auto
 | 
|
15  | 
||
16  | 
instance word :: (len0) number ..  | 
|
17  | 
instance word :: (type) minus ..  | 
|
18  | 
instance word :: (type) plus ..  | 
|
19  | 
instance word :: (type) one ..  | 
|
20  | 
instance word :: (type) zero ..  | 
|
21  | 
instance word :: (type) times ..  | 
|
22  | 
instance word :: (type) power ..  | 
|
23  | 
instance word :: (type) size ..  | 
|
24  | 
instance word :: (type) inverse ..  | 
|
25  | 
instance word :: (type) bit ..  | 
|
26  | 
||
27  | 
||
| 24350 | 28  | 
subsection "Type conversions and casting"  | 
| 24333 | 29  | 
|
30  | 
constdefs  | 
|
31  | 
  -- {* representation of words using unsigned or signed bins, 
 | 
|
32  | 
only difference in these is the type class *}  | 
|
33  | 
word_of_int :: "int => 'a :: len0 word"  | 
|
34  | 
  "word_of_int w == Abs_word (bintrunc (len_of TYPE ('a)) w)" 
 | 
|
35  | 
||
36  | 
  -- {* uint and sint cast a word to an integer,
 | 
|
37  | 
uint treats the word as unsigned,  | 
|
38  | 
sint treats the most-significant-bit as a sign bit *}  | 
|
39  | 
uint :: "'a :: len0 word => int"  | 
|
40  | 
"uint w == Rep_word w"  | 
|
41  | 
sint :: "'a :: len word => int"  | 
|
42  | 
  sint_uint: "sint w == sbintrunc (len_of TYPE ('a) - 1) (uint w)"
 | 
|
43  | 
unat :: "'a :: len0 word => nat"  | 
|
44  | 
"unat w == nat (uint w)"  | 
|
45  | 
||
46  | 
-- "the sets of integers representing the words"  | 
|
47  | 
uints :: "nat => int set"  | 
|
48  | 
"uints n == range (bintrunc n)"  | 
|
49  | 
sints :: "nat => int set"  | 
|
50  | 
"sints n == range (sbintrunc (n - 1))"  | 
|
51  | 
unats :: "nat => nat set"  | 
|
52  | 
  "unats n == {i. i < 2 ^ n}"
 | 
|
53  | 
norm_sint :: "nat => int => int"  | 
|
54  | 
"norm_sint n w == (w + 2 ^ (n - 1)) mod 2 ^ n - 2 ^ (n - 1)"  | 
|
55  | 
||
56  | 
defs (overloaded)  | 
|
57  | 
  word_size: "size (w :: 'a :: len0 word) == len_of TYPE('a)"
 | 
|
58  | 
word_number_of_def: "number_of w == word_of_int w"  | 
|
59  | 
||
60  | 
constdefs  | 
|
61  | 
  word_int_case :: "(int => 'b) => ('a :: len0 word) => 'b"
 | 
|
62  | 
"word_int_case f w == f (uint w)"  | 
|
63  | 
||
64  | 
syntax  | 
|
65  | 
of_int :: "int => 'a"  | 
|
66  | 
translations  | 
|
67  | 
"case x of of_int y => b" == "word_int_case (%y. b) x"  | 
|
68  | 
||
69  | 
||
| 24350 | 70  | 
subsection "Arithmetic operations"  | 
| 24333 | 71  | 
|
72  | 
defs (overloaded)  | 
|
73  | 
  word_1_wi: "(1 :: ('a :: len0) word) == word_of_int 1"
 | 
|
74  | 
  word_0_wi: "(0 :: ('a :: len0) word) == word_of_int 0"
 | 
|
75  | 
||
76  | 
constdefs  | 
|
77  | 
word_succ :: "'a :: len0 word => 'a word"  | 
|
78  | 
"word_succ a == word_of_int (Numeral.succ (uint a))"  | 
|
79  | 
||
80  | 
word_pred :: "'a :: len0 word => 'a word"  | 
|
81  | 
"word_pred a == word_of_int (Numeral.pred (uint a))"  | 
|
82  | 
||
83  | 
consts  | 
|
84  | 
word_power :: "'a :: len0 word => nat => 'a word"  | 
|
85  | 
primrec  | 
|
86  | 
"word_power a 0 = 1"  | 
|
87  | 
"word_power a (Suc n) = a * word_power a n"  | 
|
88  | 
||
89  | 
defs (overloaded)  | 
|
90  | 
word_pow: "power == word_power"  | 
|
91  | 
word_add_def: "a + b == word_of_int (uint a + uint b)"  | 
|
92  | 
word_sub_wi: "a - b == word_of_int (uint a - uint b)"  | 
|
93  | 
word_minus_def: "- a == word_of_int (- uint a)"  | 
|
94  | 
word_mult_def: "a * b == word_of_int (uint a * uint b)"  | 
|
95  | 
||
96  | 
||
| 24350 | 97  | 
subsection "Bit-wise operations"  | 
| 24333 | 98  | 
|
99  | 
defs (overloaded)  | 
|
100  | 
word_and_def:  | 
|
| 24353 | 101  | 
"(a::'a::len0 word) AND b == word_of_int (uint a AND uint b)"  | 
| 24333 | 102  | 
|
103  | 
word_or_def:  | 
|
| 24353 | 104  | 
"(a::'a::len0 word) OR b == word_of_int (uint a OR uint b)"  | 
| 24333 | 105  | 
|
106  | 
word_xor_def:  | 
|
| 24353 | 107  | 
"(a::'a::len0 word) XOR b == word_of_int (uint a XOR uint b)"  | 
| 24333 | 108  | 
|
109  | 
word_not_def:  | 
|
| 24353 | 110  | 
"NOT (a::'a::len0 word) == word_of_int (NOT (uint a))"  | 
| 24333 | 111  | 
|
112  | 
word_test_bit_def:  | 
|
113  | 
"test_bit (a::'a::len0 word) == bin_nth (uint a)"  | 
|
114  | 
||
115  | 
word_set_bit_def:  | 
|
116  | 
"set_bit (a::'a::len0 word) n x ==  | 
|
117  | 
word_of_int (bin_sc n (If x bit.B1 bit.B0) (uint a))"  | 
|
118  | 
||
119  | 
word_lsb_def:  | 
|
120  | 
"lsb (a::'a::len0 word) == bin_last (uint a) = bit.B1"  | 
|
121  | 
||
122  | 
word_msb_def:  | 
|
123  | 
"msb (a::'a::len word) == bin_sign (sint a) = Numeral.Min"  | 
|
124  | 
||
125  | 
||
126  | 
constdefs  | 
|
127  | 
setBit :: "'a :: len0 word => nat => 'a word"  | 
|
128  | 
"setBit w n == set_bit w n True"  | 
|
129  | 
||
130  | 
clearBit :: "'a :: len0 word => nat => 'a word"  | 
|
131  | 
"clearBit w n == set_bit w n False"  | 
|
132  | 
||
133  | 
||
134  | 
constdefs  | 
|
135  | 
-- "Largest representable machine integer."  | 
|
136  | 
max_word :: "'a::len word"  | 
|
137  | 
  "max_word \<equiv> word_of_int (2^len_of TYPE('a) - 1)"
 | 
|
138  | 
||
139  | 
consts  | 
|
140  | 
of_bool :: "bool \<Rightarrow> 'a::len word"  | 
|
141  | 
primrec  | 
|
142  | 
"of_bool False = 0"  | 
|
143  | 
"of_bool True = 1"  | 
|
144  | 
||
145  | 
||
146  | 
||
147  | 
lemmas word_size_gt_0 [iff] =  | 
|
148  | 
xtr1 [OF word_size [THEN meta_eq_to_obj_eq] len_gt_0, standard]  | 
|
149  | 
lemmas lens_gt_0 = word_size_gt_0 len_gt_0  | 
|
150  | 
lemmas lens_not_0 [iff] = lens_gt_0 [THEN gr_implies_not0, standard]  | 
|
151  | 
||
152  | 
lemma uints_num: "uints n = {i. 0 \<le> i \<and> i < 2 ^ n}"
 | 
|
153  | 
by (simp add: uints_def range_bintrunc)  | 
|
154  | 
||
155  | 
lemma sints_num: "sints n = {i. - (2 ^ (n - 1)) \<le> i \<and> i < 2 ^ (n - 1)}"
 | 
|
156  | 
by (simp add: sints_def range_sbintrunc)  | 
|
157  | 
||
158  | 
lemmas atLeastLessThan_alt = atLeastLessThan_def [unfolded  | 
|
159  | 
atLeast_def lessThan_def Collect_conj_eq [symmetric]]  | 
|
160  | 
||
161  | 
lemma mod_in_reps: "m > 0 ==> y mod m : {0::int ..< m}"
 | 
|
162  | 
unfolding atLeastLessThan_alt by auto  | 
|
163  | 
||
164  | 
lemma  | 
|
165  | 
Rep_word_0:"0 <= Rep_word x" and  | 
|
166  | 
  Rep_word_lt: "Rep_word (x::'a::len0 word) < 2 ^ len_of TYPE('a)"
 | 
|
167  | 
by (auto simp: Rep_word [simplified])  | 
|
168  | 
||
169  | 
lemma Rep_word_mod_same:  | 
|
170  | 
  "Rep_word x mod 2 ^ len_of TYPE('a) = Rep_word (x::'a::len0 word)"
 | 
|
171  | 
by (simp add: int_mod_eq Rep_word_lt Rep_word_0)  | 
|
172  | 
||
173  | 
lemma td_ext_uint:  | 
|
174  | 
  "td_ext (uint :: 'a word => int) word_of_int (uints (len_of TYPE('a::len0))) 
 | 
|
175  | 
    (%w::int. w mod 2 ^ len_of TYPE('a))"
 | 
|
176  | 
apply (unfold td_ext_def')  | 
|
177  | 
apply (simp add: uints_num uint_def word_of_int_def bintrunc_mod2p)  | 
|
178  | 
apply (simp add: Rep_word_mod_same Rep_word_0 Rep_word_lt  | 
|
179  | 
word.Rep_word_inverse word.Abs_word_inverse int_mod_lem)  | 
|
180  | 
done  | 
|
181  | 
||
182  | 
lemmas int_word_uint = td_ext_uint [THEN td_ext.eq_norm, standard]  | 
|
183  | 
||
184  | 
interpretation word_uint:  | 
|
185  | 
td_ext ["uint::'a::len0 word \<Rightarrow> int"  | 
|
186  | 
word_of_int  | 
|
187  | 
          "uints (len_of TYPE('a::len0))"
 | 
|
188  | 
          "\<lambda>w. w mod 2 ^ len_of TYPE('a::len0)"]
 | 
|
189  | 
by (rule td_ext_uint)  | 
|
190  | 
||
191  | 
lemmas td_uint = word_uint.td_thm  | 
|
192  | 
||
193  | 
lemmas td_ext_ubin = td_ext_uint  | 
|
194  | 
[simplified len_gt_0 no_bintr_alt1 [symmetric]]  | 
|
195  | 
||
196  | 
interpretation word_ubin:  | 
|
197  | 
td_ext ["uint::'a::len0 word \<Rightarrow> int"  | 
|
198  | 
word_of_int  | 
|
199  | 
          "uints (len_of TYPE('a::len0))"
 | 
|
200  | 
          "bintrunc (len_of TYPE('a::len0))"]
 | 
|
201  | 
by (rule td_ext_ubin)  | 
|
202  | 
||
203  | 
lemma sint_sbintrunc':  | 
|
204  | 
"sint (word_of_int bin :: 'a word) =  | 
|
205  | 
    (sbintrunc (len_of TYPE ('a :: len) - 1) bin)"
 | 
|
206  | 
unfolding sint_uint  | 
|
207  | 
by (auto simp: word_ubin.eq_norm sbintrunc_bintrunc_lt)  | 
|
208  | 
||
209  | 
lemma uint_sint:  | 
|
210  | 
  "uint w = bintrunc (len_of TYPE('a)) (sint (w :: 'a :: len word))"
 | 
|
211  | 
unfolding sint_uint by (auto simp: bintrunc_sbintrunc_le)  | 
|
212  | 
||
213  | 
lemma bintr_uint':  | 
|
214  | 
"n >= size w ==> bintrunc n (uint w) = uint w"  | 
|
215  | 
apply (unfold word_size)  | 
|
216  | 
apply (subst word_ubin.norm_Rep [symmetric])  | 
|
217  | 
apply (simp only: bintrunc_bintrunc_min word_size min_def)  | 
|
218  | 
apply simp  | 
|
219  | 
done  | 
|
220  | 
||
221  | 
lemma wi_bintr':  | 
|
222  | 
"wb = word_of_int bin ==> n >= size wb ==>  | 
|
223  | 
word_of_int (bintrunc n bin) = wb"  | 
|
224  | 
unfolding word_size  | 
|
225  | 
by (clarsimp simp add : word_ubin.norm_eq_iff [symmetric] min_def)  | 
|
226  | 
||
227  | 
lemmas bintr_uint = bintr_uint' [unfolded word_size]  | 
|
228  | 
lemmas wi_bintr = wi_bintr' [unfolded word_size]  | 
|
229  | 
||
230  | 
lemma td_ext_sbin:  | 
|
231  | 
  "td_ext (sint :: 'a word => int) word_of_int (sints (len_of TYPE('a::len))) 
 | 
|
232  | 
    (sbintrunc (len_of TYPE('a) - 1))"
 | 
|
233  | 
apply (unfold td_ext_def' sint_uint)  | 
|
234  | 
apply (simp add : word_ubin.eq_norm)  | 
|
235  | 
  apply (cases "len_of TYPE('a)")
 | 
|
236  | 
apply (auto simp add : sints_def)  | 
|
237  | 
apply (rule sym [THEN trans])  | 
|
238  | 
apply (rule word_ubin.Abs_norm)  | 
|
239  | 
apply (simp only: bintrunc_sbintrunc)  | 
|
240  | 
apply (drule sym)  | 
|
241  | 
apply simp  | 
|
242  | 
done  | 
|
243  | 
||
244  | 
lemmas td_ext_sint = td_ext_sbin  | 
|
245  | 
[simplified len_gt_0 no_sbintr_alt2 Suc_pred' [symmetric]]  | 
|
246  | 
||
247  | 
(* We do sint before sbin, before sint is the user version  | 
|
248  | 
and interpretations do not produce thm duplicates. I.e.  | 
|
249  | 
we get the name word_sint.Rep_eqD, but not word_sbin.Req_eqD,  | 
|
250  | 
because the latter is the same thm as the former *)  | 
|
251  | 
interpretation word_sint:  | 
|
252  | 
td_ext ["sint ::'a::len word => int"  | 
|
253  | 
word_of_int  | 
|
254  | 
          "sints (len_of TYPE('a::len))"
 | 
|
255  | 
          "%w. (w + 2^(len_of TYPE('a::len) - 1)) mod 2^len_of TYPE('a::len) -
 | 
|
256  | 
               2 ^ (len_of TYPE('a::len) - 1)"]
 | 
|
257  | 
by (rule td_ext_sint)  | 
|
258  | 
||
259  | 
interpretation word_sbin:  | 
|
260  | 
td_ext ["sint ::'a::len word => int"  | 
|
261  | 
word_of_int  | 
|
262  | 
          "sints (len_of TYPE('a::len))"
 | 
|
263  | 
          "sbintrunc (len_of TYPE('a::len) - 1)"]
 | 
|
264  | 
by (rule td_ext_sbin)  | 
|
265  | 
||
266  | 
lemmas int_word_sint = td_ext_sint [THEN td_ext.eq_norm, standard]  | 
|
267  | 
||
268  | 
lemmas td_sint = word_sint.td  | 
|
269  | 
||
270  | 
lemma word_number_of_alt: "number_of b == word_of_int (number_of b)"  | 
|
271  | 
unfolding word_number_of_def by (simp add: number_of_eq)  | 
|
272  | 
||
273  | 
lemma word_no_wi: "number_of = word_of_int"  | 
|
274  | 
by (auto simp: word_number_of_def intro: ext)  | 
|
275  | 
||
276  | 
lemmas uints_mod = uints_def [unfolded no_bintr_alt1]  | 
|
277  | 
||
278  | 
lemma uint_bintrunc: "uint (number_of bin :: 'a word) =  | 
|
279  | 
    number_of (bintrunc (len_of TYPE ('a :: len0)) bin)"
 | 
|
280  | 
unfolding word_number_of_def number_of_eq  | 
|
281  | 
by (auto intro: word_ubin.eq_norm)  | 
|
282  | 
||
283  | 
lemma sint_sbintrunc: "sint (number_of bin :: 'a word) =  | 
|
284  | 
    number_of (sbintrunc (len_of TYPE ('a :: len) - 1) bin)" 
 | 
|
285  | 
unfolding word_number_of_def number_of_eq  | 
|
286  | 
by (auto intro!: word_sbin.eq_norm simp del: one_is_Suc_zero)  | 
|
287  | 
||
288  | 
lemma unat_bintrunc:  | 
|
289  | 
"unat (number_of bin :: 'a :: len0 word) =  | 
|
290  | 
    number_of (bintrunc (len_of TYPE('a)) bin)"
 | 
|
291  | 
unfolding unat_def nat_number_of_def  | 
|
292  | 
by (simp only: uint_bintrunc)  | 
|
293  | 
||
294  | 
(* WARNING - these may not always be helpful *)  | 
|
295  | 
declare  | 
|
296  | 
uint_bintrunc [simp]  | 
|
297  | 
sint_sbintrunc [simp]  | 
|
298  | 
unat_bintrunc [simp]  | 
|
299  | 
||
300  | 
lemma size_0_eq: "size (w :: 'a :: len0 word) = 0 ==> v = w"  | 
|
301  | 
apply (unfold word_size)  | 
|
302  | 
apply (rule word_uint.Rep_eqD)  | 
|
303  | 
apply (rule box_equals)  | 
|
304  | 
defer  | 
|
305  | 
apply (rule word_ubin.norm_Rep)+  | 
|
306  | 
apply simp  | 
|
307  | 
done  | 
|
308  | 
||
309  | 
lemmas uint_lem = word_uint.Rep [unfolded uints_num mem_Collect_eq]  | 
|
310  | 
lemmas sint_lem = word_sint.Rep [unfolded sints_num mem_Collect_eq]  | 
|
311  | 
lemmas uint_ge_0 [iff] = uint_lem [THEN conjunct1, standard]  | 
|
312  | 
lemmas uint_lt2p [iff] = uint_lem [THEN conjunct2, standard]  | 
|
313  | 
lemmas sint_ge = sint_lem [THEN conjunct1, standard]  | 
|
314  | 
lemmas sint_lt = sint_lem [THEN conjunct2, standard]  | 
|
315  | 
||
316  | 
lemma sign_uint_Pls [simp]:  | 
|
317  | 
"bin_sign (uint x) = Numeral.Pls"  | 
|
318  | 
by (simp add: sign_Pls_ge_0 number_of_eq)  | 
|
319  | 
||
320  | 
lemmas uint_m2p_neg = iffD2 [OF diff_less_0_iff_less uint_lt2p, standard]  | 
|
321  | 
lemmas uint_m2p_not_non_neg =  | 
|
322  | 
iffD2 [OF linorder_not_le uint_m2p_neg, standard]  | 
|
323  | 
||
324  | 
lemma lt2p_lem:  | 
|
325  | 
  "len_of TYPE('a) <= n ==> uint (w :: 'a :: len0 word) < 2 ^ n"
 | 
|
326  | 
by (rule xtr8 [OF _ uint_lt2p]) simp  | 
|
327  | 
||
328  | 
lemmas uint_le_0_iff [simp] =  | 
|
329  | 
uint_ge_0 [THEN leD, THEN linorder_antisym_conv1, standard]  | 
|
330  | 
||
331  | 
lemma uint_nat: "uint w == int (unat w)"  | 
|
332  | 
unfolding unat_def by auto  | 
|
333  | 
||
334  | 
lemma uint_number_of:  | 
|
335  | 
  "uint (number_of b :: 'a :: len0 word) = number_of b mod 2 ^ len_of TYPE('a)"
 | 
|
336  | 
unfolding word_number_of_alt  | 
|
337  | 
by (simp only: int_word_uint)  | 
|
338  | 
||
339  | 
lemma unat_number_of:  | 
|
340  | 
"bin_sign b = Numeral.Pls ==>  | 
|
341  | 
  unat (number_of b::'a::len0 word) = number_of b mod 2 ^ len_of TYPE ('a)"
 | 
|
342  | 
apply (unfold unat_def)  | 
|
343  | 
apply (clarsimp simp only: uint_number_of)  | 
|
344  | 
apply (rule nat_mod_distrib [THEN trans])  | 
|
345  | 
apply (erule sign_Pls_ge_0 [THEN iffD1])  | 
|
346  | 
apply (simp_all add: nat_power_eq)  | 
|
347  | 
done  | 
|
348  | 
||
349  | 
lemma sint_number_of: "sint (number_of b :: 'a :: len word) = (number_of b +  | 
|
350  | 
    2 ^ (len_of TYPE('a) - 1)) mod 2 ^ len_of TYPE('a) -
 | 
|
351  | 
    2 ^ (len_of TYPE('a) - 1)"
 | 
|
352  | 
unfolding word_number_of_alt by (rule int_word_sint)  | 
|
353  | 
||
354  | 
lemma word_of_int_bin [simp] :  | 
|
355  | 
"(word_of_int (number_of bin) :: 'a :: len0 word) = (number_of bin)"  | 
|
356  | 
unfolding word_number_of_alt by auto  | 
|
357  | 
||
358  | 
lemma word_int_case_wi:  | 
|
359  | 
"word_int_case f (word_of_int i :: 'b word) =  | 
|
360  | 
    f (i mod 2 ^ len_of TYPE('b::len0))"
 | 
|
361  | 
unfolding word_int_case_def by (simp add: word_uint.eq_norm)  | 
|
362  | 
||
363  | 
lemma word_int_split:  | 
|
364  | 
"P (word_int_case f x) =  | 
|
365  | 
(ALL i. x = (word_of_int i :: 'b :: len0 word) &  | 
|
366  | 
      0 <= i & i < 2 ^ len_of TYPE('b) --> P (f i))"
 | 
|
367  | 
unfolding word_int_case_def  | 
|
368  | 
by (auto simp: word_uint.eq_norm int_mod_eq')  | 
|
369  | 
||
370  | 
lemma word_int_split_asm:  | 
|
371  | 
"P (word_int_case f x) =  | 
|
372  | 
(~ (EX n. x = (word_of_int n :: 'b::len0 word) &  | 
|
373  | 
      0 <= n & n < 2 ^ len_of TYPE('b::len0) & ~ P (f n)))"
 | 
|
374  | 
unfolding word_int_case_def  | 
|
375  | 
by (auto simp: word_uint.eq_norm int_mod_eq')  | 
|
376  | 
||
377  | 
lemmas uint_range' =  | 
|
378  | 
word_uint.Rep [unfolded uints_num mem_Collect_eq, standard]  | 
|
379  | 
lemmas sint_range' = word_sint.Rep [unfolded One_nat_def  | 
|
380  | 
sints_num mem_Collect_eq, standard]  | 
|
381  | 
||
382  | 
lemma uint_range_size: "0 <= uint w & uint w < 2 ^ size w"  | 
|
383  | 
unfolding word_size by (rule uint_range')  | 
|
384  | 
||
385  | 
lemma sint_range_size:  | 
|
386  | 
"- (2 ^ (size w - Suc 0)) <= sint w & sint w < 2 ^ (size w - Suc 0)"  | 
|
387  | 
unfolding word_size by (rule sint_range')  | 
|
388  | 
||
389  | 
lemmas sint_above_size = sint_range_size  | 
|
390  | 
[THEN conjunct2, THEN [2] xtr8, folded One_nat_def, standard]  | 
|
391  | 
||
392  | 
lemmas sint_below_size = sint_range_size  | 
|
393  | 
[THEN conjunct1, THEN [2] order_trans, folded One_nat_def, standard]  | 
|
394  | 
||
395  | 
lemma test_bit_eq_iff: "(test_bit (u::'a::len0 word) = test_bit v) = (u = v)"  | 
|
396  | 
unfolding word_test_bit_def by (simp add: bin_nth_eq_iff)  | 
|
397  | 
||
398  | 
lemma test_bit_size [rule_format] : "(w::'a::len0 word) !! n --> n < size w"  | 
|
399  | 
apply (unfold word_test_bit_def)  | 
|
400  | 
apply (subst word_ubin.norm_Rep [symmetric])  | 
|
401  | 
apply (simp only: nth_bintr word_size)  | 
|
402  | 
apply fast  | 
|
403  | 
done  | 
|
404  | 
||
405  | 
lemma word_eqI [rule_format] :  | 
|
406  | 
fixes u :: "'a::len0 word"  | 
|
407  | 
shows "(ALL n. n < size u --> u !! n = v !! n) ==> u = v"  | 
|
408  | 
apply (rule test_bit_eq_iff [THEN iffD1])  | 
|
409  | 
apply (rule ext)  | 
|
410  | 
apply (erule allE)  | 
|
411  | 
apply (erule impCE)  | 
|
412  | 
prefer 2  | 
|
413  | 
apply assumption  | 
|
414  | 
apply (auto dest!: test_bit_size simp add: word_size)  | 
|
415  | 
done  | 
|
416  | 
||
417  | 
lemmas word_eqD = test_bit_eq_iff [THEN iffD2, THEN fun_cong, standard]  | 
|
418  | 
||
419  | 
lemma test_bit_bin': "w !! n = (n < size w & bin_nth (uint w) n)"  | 
|
420  | 
unfolding word_test_bit_def word_size  | 
|
421  | 
by (simp add: nth_bintr [symmetric])  | 
|
422  | 
||
423  | 
lemmas test_bit_bin = test_bit_bin' [unfolded word_size]  | 
|
424  | 
||
425  | 
lemma bin_nth_uint_imp': "bin_nth (uint w) n --> n < size w"  | 
|
426  | 
apply (unfold word_size)  | 
|
427  | 
apply (rule impI)  | 
|
428  | 
apply (rule nth_bintr [THEN iffD1, THEN conjunct1])  | 
|
429  | 
apply (subst word_ubin.norm_Rep)  | 
|
430  | 
apply assumption  | 
|
431  | 
done  | 
|
432  | 
||
433  | 
lemma bin_nth_sint':  | 
|
434  | 
"n >= size w --> bin_nth (sint w) n = bin_nth (sint w) (size w - 1)"  | 
|
435  | 
apply (rule impI)  | 
|
436  | 
apply (subst word_sbin.norm_Rep [symmetric])  | 
|
437  | 
apply (simp add : nth_sbintr word_size)  | 
|
438  | 
apply auto  | 
|
439  | 
done  | 
|
440  | 
||
441  | 
lemmas bin_nth_uint_imp = bin_nth_uint_imp' [rule_format, unfolded word_size]  | 
|
442  | 
lemmas bin_nth_sint = bin_nth_sint' [rule_format, unfolded word_size]  | 
|
443  | 
||
444  | 
||
445  | 
lemmas num_AB_u [simp] = word_uint.Rep_inverse  | 
|
446  | 
[unfolded o_def word_number_of_def [symmetric], standard]  | 
|
447  | 
lemmas num_AB_s [simp] = word_sint.Rep_inverse  | 
|
448  | 
[unfolded o_def word_number_of_def [symmetric], standard]  | 
|
449  | 
||
450  | 
(* naturals *)  | 
|
451  | 
lemma uints_unats: "uints n = int ` unats n"  | 
|
452  | 
apply (unfold unats_def uints_num)  | 
|
453  | 
apply safe  | 
|
454  | 
apply (rule_tac image_eqI)  | 
|
455  | 
apply (erule_tac nat_0_le [symmetric])  | 
|
456  | 
apply auto  | 
|
457  | 
apply (erule_tac nat_less_iff [THEN iffD2])  | 
|
458  | 
apply (rule_tac [2] zless_nat_eq_int_zless [THEN iffD1])  | 
|
459  | 
apply (auto simp add : nat_power_eq int_power)  | 
|
460  | 
done  | 
|
461  | 
||
462  | 
lemma unats_uints: "unats n = nat ` uints n"  | 
|
463  | 
apply (auto simp add : uints_unats image_iff)  | 
|
464  | 
done  | 
|
465  | 
||
466  | 
lemmas bintr_num = word_ubin.norm_eq_iff  | 
|
467  | 
[symmetric, folded word_number_of_def, standard]  | 
|
468  | 
lemmas sbintr_num = word_sbin.norm_eq_iff  | 
|
469  | 
[symmetric, folded word_number_of_def, standard]  | 
|
470  | 
||
471  | 
lemmas num_of_bintr = word_ubin.Abs_norm [folded word_number_of_def, standard]  | 
|
472  | 
lemmas num_of_sbintr = word_sbin.Abs_norm [folded word_number_of_def, standard];  | 
|
473  | 
||
474  | 
(* don't add these to simpset, since may want bintrunc n w to be simplified;  | 
|
475  | 
may want these in reverse, but loop as simp rules, so use following *)  | 
|
476  | 
||
477  | 
lemma num_of_bintr':  | 
|
478  | 
  "bintrunc (len_of TYPE('a :: len0)) a = b ==> 
 | 
|
479  | 
number_of a = (number_of b :: 'a word)"  | 
|
480  | 
apply safe  | 
|
481  | 
apply (rule_tac num_of_bintr [symmetric])  | 
|
482  | 
done  | 
|
483  | 
||
484  | 
lemma num_of_sbintr':  | 
|
485  | 
  "sbintrunc (len_of TYPE('a :: len) - 1) a = b ==> 
 | 
|
486  | 
number_of a = (number_of b :: 'a word)"  | 
|
487  | 
apply safe  | 
|
488  | 
apply (rule_tac num_of_sbintr [symmetric])  | 
|
489  | 
done  | 
|
490  | 
||
491  | 
lemmas num_abs_bintr = sym [THEN trans,  | 
|
492  | 
OF num_of_bintr word_number_of_def [THEN meta_eq_to_obj_eq], standard]  | 
|
493  | 
lemmas num_abs_sbintr = sym [THEN trans,  | 
|
494  | 
OF num_of_sbintr word_number_of_def [THEN meta_eq_to_obj_eq], standard]  | 
|
| 24375 | 495  | 
|
496  | 
lemmas test_bit_def' = word_test_bit_def [THEN meta_eq_to_obj_eq, THEN fun_cong]  | 
|
497  | 
||
498  | 
lemmas word_log_defs = word_and_def word_or_def word_xor_def word_not_def  | 
|
499  | 
lemmas word_log_bin_defs = word_log_defs  | 
|
500  | 
||
501  | 
||
502  | 
subsection {* Casting words to different lengths *}
 | 
|
503  | 
||
504  | 
constdefs  | 
|
505  | 
-- "cast a word to a different length"  | 
|
506  | 
scast :: "'a :: len word => 'b :: len word"  | 
|
507  | 
"scast w == word_of_int (sint w)"  | 
|
508  | 
ucast :: "'a :: len0 word => 'b :: len0 word"  | 
|
509  | 
"ucast w == word_of_int (uint w)"  | 
|
510  | 
||
511  | 
-- "whether a cast (or other) function is to a longer or shorter length"  | 
|
512  | 
  source_size :: "('a :: len0 word => 'b) => nat"
 | 
|
513  | 
"source_size c == let arb = arbitrary ; x = c arb in size arb"  | 
|
514  | 
  target_size :: "('a => 'b :: len0 word) => nat"
 | 
|
515  | 
"target_size c == size (c arbitrary)"  | 
|
516  | 
  is_up :: "('a :: len0 word => 'b :: len0 word) => bool"
 | 
|
517  | 
"is_up c == source_size c <= target_size c"  | 
|
518  | 
  is_down :: "('a :: len0 word => 'b :: len0 word) => bool"
 | 
|
519  | 
"is_down c == target_size c <= source_size c"  | 
|
520  | 
||
| 24333 | 521  | 
(** cast - note, no arg for new length, as it's determined by type of result,  | 
522  | 
thus in "cast w = w, the type means cast to length of w! **)  | 
|
523  | 
||
524  | 
lemma ucast_id: "ucast w = w"  | 
|
525  | 
unfolding ucast_def by auto  | 
|
526  | 
||
527  | 
lemma scast_id: "scast w = w"  | 
|
528  | 
unfolding scast_def by auto  | 
|
529  | 
||
530  | 
lemma nth_ucast:  | 
|
531  | 
  "(ucast w::'a::len0 word) !! n = (w !! n & n < len_of TYPE('a))"
 | 
|
532  | 
apply (unfold ucast_def test_bit_bin)  | 
|
533  | 
apply (simp add: word_ubin.eq_norm nth_bintr word_size)  | 
|
534  | 
apply (fast elim!: bin_nth_uint_imp)  | 
|
535  | 
done  | 
|
536  | 
||
537  | 
(* for literal u(s)cast *)  | 
|
538  | 
||
539  | 
lemma ucast_bintr [simp]:  | 
|
540  | 
"ucast (number_of w ::'a::len0 word) =  | 
|
541  | 
   number_of (bintrunc (len_of TYPE('a)) w)"
 | 
|
542  | 
unfolding ucast_def by simp  | 
|
543  | 
||
544  | 
lemma scast_sbintr [simp]:  | 
|
545  | 
"scast (number_of w ::'a::len word) =  | 
|
546  | 
   number_of (sbintrunc (len_of TYPE('a) - Suc 0) w)"
 | 
|
547  | 
unfolding scast_def by simp  | 
|
548  | 
||
549  | 
lemmas source_size = source_size_def [unfolded Let_def word_size]  | 
|
550  | 
lemmas target_size = target_size_def [unfolded Let_def word_size]  | 
|
551  | 
lemmas is_down = is_down_def [unfolded source_size target_size]  | 
|
552  | 
lemmas is_up = is_up_def [unfolded source_size target_size]  | 
|
553  | 
||
554  | 
lemmas is_up_down =  | 
|
555  | 
trans [OF is_up [THEN meta_eq_to_obj_eq]  | 
|
556  | 
is_down [THEN meta_eq_to_obj_eq, symmetric],  | 
|
557  | 
standard]  | 
|
558  | 
||
559  | 
lemma down_cast_same': "uc = ucast ==> is_down uc ==> uc = scast"  | 
|
560  | 
apply (unfold is_down)  | 
|
561  | 
apply safe  | 
|
562  | 
apply (rule ext)  | 
|
563  | 
apply (unfold ucast_def scast_def uint_sint)  | 
|
564  | 
apply (rule word_ubin.norm_eq_iff [THEN iffD1])  | 
|
565  | 
apply simp  | 
|
566  | 
done  | 
|
567  | 
||
568  | 
lemma sint_up_scast':  | 
|
569  | 
"sc = scast ==> is_up sc ==> sint (sc w) = sint w"  | 
|
570  | 
apply (unfold is_up)  | 
|
571  | 
apply safe  | 
|
572  | 
apply (simp add: scast_def word_sbin.eq_norm)  | 
|
573  | 
apply (rule box_equals)  | 
|
574  | 
prefer 3  | 
|
575  | 
apply (rule word_sbin.norm_Rep)  | 
|
576  | 
apply (rule sbintrunc_sbintrunc_l)  | 
|
577  | 
defer  | 
|
578  | 
apply (subst word_sbin.norm_Rep)  | 
|
579  | 
apply (rule refl)  | 
|
580  | 
apply simp  | 
|
581  | 
done  | 
|
582  | 
||
583  | 
lemma uint_up_ucast':  | 
|
584  | 
"uc = ucast ==> is_up uc ==> uint (uc w) = uint w"  | 
|
585  | 
apply (unfold is_up)  | 
|
586  | 
apply safe  | 
|
587  | 
apply (rule bin_eqI)  | 
|
588  | 
apply (fold word_test_bit_def)  | 
|
589  | 
apply (auto simp add: nth_ucast)  | 
|
590  | 
apply (auto simp add: test_bit_bin)  | 
|
591  | 
done  | 
|
592  | 
||
593  | 
lemmas down_cast_same = refl [THEN down_cast_same']  | 
|
594  | 
lemmas uint_up_ucast = refl [THEN uint_up_ucast']  | 
|
595  | 
lemmas sint_up_scast = refl [THEN sint_up_scast']  | 
|
596  | 
||
597  | 
lemma ucast_up_ucast': "uc = ucast ==> is_up uc ==> ucast (uc w) = ucast w"  | 
|
598  | 
apply (simp (no_asm) add: ucast_def)  | 
|
599  | 
apply (clarsimp simp add: uint_up_ucast)  | 
|
600  | 
done  | 
|
601  | 
||
602  | 
lemma scast_up_scast': "sc = scast ==> is_up sc ==> scast (sc w) = scast w"  | 
|
603  | 
apply (simp (no_asm) add: scast_def)  | 
|
604  | 
apply (clarsimp simp add: sint_up_scast)  | 
|
605  | 
done  | 
|
606  | 
||
607  | 
lemmas ucast_up_ucast = refl [THEN ucast_up_ucast']  | 
|
608  | 
lemmas scast_up_scast = refl [THEN scast_up_scast']  | 
|
609  | 
||
610  | 
lemmas ucast_up_ucast_id = trans [OF ucast_up_ucast ucast_id]  | 
|
611  | 
lemmas scast_up_scast_id = trans [OF scast_up_scast scast_id]  | 
|
612  | 
||
613  | 
lemmas isduu = is_up_down [where c = "ucast", THEN iffD2]  | 
|
614  | 
lemmas isdus = is_up_down [where c = "scast", THEN iffD2]  | 
|
615  | 
lemmas ucast_down_ucast_id = isduu [THEN ucast_up_ucast_id]  | 
|
616  | 
lemmas scast_down_scast_id = isdus [THEN ucast_up_ucast_id]  | 
|
617  | 
||
618  | 
lemma up_ucast_surj:  | 
|
619  | 
"is_up (ucast :: 'b::len0 word => 'a::len0 word) ==>  | 
|
620  | 
surj (ucast :: 'a word => 'b word)"  | 
|
621  | 
by (rule surjI, erule ucast_up_ucast_id)  | 
|
622  | 
||
623  | 
lemma up_scast_surj:  | 
|
624  | 
"is_up (scast :: 'b::len word => 'a::len word) ==>  | 
|
625  | 
surj (scast :: 'a word => 'b word)"  | 
|
626  | 
by (rule surjI, erule scast_up_scast_id)  | 
|
627  | 
||
628  | 
lemma down_scast_inj:  | 
|
629  | 
"is_down (scast :: 'b::len word => 'a::len word) ==>  | 
|
630  | 
inj_on (ucast :: 'a word => 'b word) A"  | 
|
631  | 
by (rule inj_on_inverseI, erule scast_down_scast_id)  | 
|
632  | 
||
633  | 
lemma down_ucast_inj:  | 
|
634  | 
"is_down (ucast :: 'b::len0 word => 'a::len0 word) ==>  | 
|
635  | 
inj_on (ucast :: 'a word => 'b word) A"  | 
|
636  | 
by (rule inj_on_inverseI, erule ucast_down_ucast_id)  | 
|
637  | 
||
638  | 
||
639  | 
lemma ucast_down_no':  | 
|
640  | 
"uc = ucast ==> is_down uc ==> uc (number_of bin) = number_of bin"  | 
|
641  | 
apply (unfold word_number_of_def is_down)  | 
|
642  | 
apply (clarsimp simp add: ucast_def word_ubin.eq_norm)  | 
|
643  | 
apply (rule word_ubin.norm_eq_iff [THEN iffD1])  | 
|
644  | 
apply (erule bintrunc_bintrunc_ge)  | 
|
645  | 
done  | 
|
646  | 
||
647  | 
lemmas ucast_down_no = ucast_down_no' [OF refl]  | 
|
648  | 
||
649  | 
end  |