1(*
2 * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230)
3 *
4 * SPDX-License-Identifier: BSD-2-Clause
5 *)
6
7section "Lemmas with Generic Word Length"
8
9theory Word_Lemmas
10  imports
11    Complex_Main
12    Word_EqI
13    Word_Enum
14    "HOL-Library.Sublist"
15begin
16
17lemma word_plus_mono_left:
18  fixes x :: "'a :: len word"
19  shows "\<lbrakk>y \<le> z; x \<le> x + z\<rbrakk> \<Longrightarrow> y + x \<le> z + x"
20  by unat_arith
21
22lemma word_shiftl_add_distrib:
23  fixes x :: "'a :: len word"
24  shows "(x + y) << n = (x << n) + (y << n)"
25  by (simp add: shiftl_t2n ring_distribs)
26
27lemma less_Suc_unat_less_bound:
28  "n < Suc (unat (x :: 'a :: len word)) \<Longrightarrow> n < 2 ^ LENGTH('a)"
29  by (auto elim!: order_less_le_trans intro: Suc_leI)
30
31lemma up_ucast_inj:
32  "\<lbrakk> ucast x = (ucast y::'b::len word); LENGTH('a) \<le> len_of TYPE ('b) \<rbrakk> \<Longrightarrow> x = (y::'a::len word)"
33  by (subst (asm) bang_eq) (fastforce simp: nth_ucast word_size intro: word_eqI)
34
35lemmas ucast_up_inj = up_ucast_inj
36
37lemma up_ucast_inj_eq:
38  "LENGTH('a) \<le> len_of TYPE ('b) \<Longrightarrow> (ucast x = (ucast y::'b::len word)) = (x = (y::'a::len word))"
39  by (fastforce dest: up_ucast_inj)
40
41lemma no_plus_overflow_neg:
42  "(x :: 'a :: len word) < -y \<Longrightarrow> x \<le> x + y"
43  by (metis diff_minus_eq_add less_imp_le sub_wrap_lt)
44
45lemma ucast_ucast_eq:
46  "\<lbrakk> ucast x = (ucast (ucast y::'a word)::'c::len word); LENGTH('a) \<le> LENGTH('b);
47     LENGTH('b) \<le> LENGTH('c) \<rbrakk> \<Longrightarrow>
48   x = ucast y" for x :: "'a::len word" and y :: "'b::len word"
49  by (fastforce intro: word_eqI simp: bang_eq nth_ucast word_size)
50
51lemma ucast_0_I:
52  "x = 0 \<Longrightarrow> ucast x = 0"
53  by simp
54
55text \<open>right-padding a word to a certain length\<close>
56
57definition
58  "bl_pad_to bl sz \<equiv> bl @ (replicate (sz - length bl) False)"
59
60lemma bl_pad_to_length:
61  assumes lbl: "length bl \<le> sz"
62  shows   "length (bl_pad_to bl sz) = sz"
63  using lbl by (simp add: bl_pad_to_def)
64
65lemma bl_pad_to_prefix:
66  "prefix bl (bl_pad_to bl sz)"
67  by (simp add: bl_pad_to_def)
68
69lemma same_length_is_parallel:
70  assumes len: "\<forall>y \<in> set as. length y = x"
71  shows  "\<forall>x \<in> set as. \<forall>y \<in> set as - {x}. x \<parallel> y"
72proof (rule, rule)
73  fix x y
74  assume xi: "x \<in> set as" and yi: "y \<in> set as - {x}"
75  from len obtain q where len': "\<forall>y \<in> set as. length y = q" ..
76
77  show "x \<parallel> y"
78  proof (rule not_equal_is_parallel)
79    from xi yi show "x \<noteq> y" by auto
80    from xi yi len' show "length x = length y" by (auto dest: bspec)
81  qed
82qed
83
84text \<open>Lemmas about words\<close>
85
86lemmas and_bang = word_and_nth
87
88lemma of_drop_to_bl:
89  "of_bl (drop n (to_bl x)) = (x && mask (size x - n))"
90  by (simp add: of_bl_drop word_size_bl)
91
92lemma word_add_offset_less:
93  fixes x :: "'a :: len word"
94  assumes yv: "y < 2 ^ n"
95  and     xv: "x < 2 ^ m"
96  and     mnv: "sz < LENGTH('a :: len)"
97  and    xv': "x < 2 ^ (LENGTH('a :: len) - n)"
98  and     mn: "sz = m + n"
99  shows   "x * 2 ^ n + y < 2 ^ sz"
100proof (subst mn)
101  from mnv mn have nv: "n < LENGTH('a)" and mv: "m < LENGTH('a)"  by auto
102
103  have uy: "unat y < 2 ^ n"
104    by (rule order_less_le_trans [OF unat_mono [OF yv] order_eq_refl],
105        rule unat_power_lower[OF nv])
106
107  have ux: "unat x < 2 ^ m"
108    by (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl],
109        rule unat_power_lower[OF mv])
110
111  then show "x * 2 ^ n + y < 2 ^ (m + n)" using ux uy nv mnv xv'
112    apply (subst word_less_nat_alt)
113    apply (subst unat_word_ariths)+
114    apply (subst mod_less)
115     apply simp
116     apply (subst mult.commute)
117     apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]])
118     apply (rule order_less_le_trans [OF unat_mono [OF xv']])
119     apply (cases "n = 0"; simp)
120    apply (subst unat_power_lower[OF nv])
121    apply (subst mod_less)
122     apply (erule order_less_le_trans [OF nat_add_offset_less], assumption)
123      apply (rule mn)
124     apply simp
125    apply (simp add: mn mnv)
126    apply (erule nat_add_offset_less; simp)
127    done
128qed
129
130lemma word_less_power_trans:
131  fixes n :: "'a :: len word"
132  assumes nv: "n < 2 ^ (m - k)"
133  and     kv: "k \<le> m"
134  and     mv: "m < len_of TYPE ('a)"
135  shows "2 ^ k * n < 2 ^ m"
136  using nv kv mv
137  apply -
138  apply (subst word_less_nat_alt)
139  apply (subst unat_word_ariths)
140  apply (subst mod_less)
141   apply simp
142   apply (rule nat_less_power_trans)
143    apply (erule order_less_trans [OF unat_mono])
144    apply simp
145   apply simp
146  apply simp
147  apply (rule nat_less_power_trans)
148   apply (subst unat_power_lower[where 'a = 'a, symmetric])
149    apply simp
150   apply (erule unat_mono)
151  apply simp
152  done
153
154lemma Suc_unat_diff_1:
155  fixes x :: "'a :: len word"
156  assumes lt: "1 \<le> x"
157  shows "Suc (unat (x - 1)) = unat x"
158proof -
159  have "0 < unat x"
160    by (rule order_less_le_trans [where y = 1], simp, subst unat_1 [symmetric],
161        rule iffD1 [OF word_le_nat_alt lt])
162
163  then show ?thesis
164    by ((subst unat_sub [OF lt])+, simp only:  unat_1)
165qed
166
167lemma word_div_sub:
168  fixes x :: "'a :: len word"
169  assumes yx: "y \<le> x"
170  and     y0: "0 < y"
171  shows "(x - y) div y = x div y - 1"
172  apply (rule word_unat.Rep_eqD)
173  apply (subst unat_div)
174  apply (subst unat_sub [OF yx])
175  apply (subst unat_sub)
176   apply (subst word_le_nat_alt)
177   apply (subst unat_div)
178   apply (subst le_div_geq)
179     apply (rule order_le_less_trans [OF _ unat_mono [OF y0]])
180     apply simp
181    apply (subst word_le_nat_alt [symmetric], rule yx)
182   apply simp
183  apply (subst unat_div)
184  apply (subst le_div_geq [OF _ iffD1 [OF word_le_nat_alt yx]])
185   apply (rule order_le_less_trans [OF _ unat_mono [OF y0]])
186   apply simp
187  apply simp
188  done
189
190lemma word_mult_less_mono1:
191  fixes i :: "'a :: len word"
192  assumes ij: "i < j"
193  and    knz: "0 < k"
194  and    ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)"
195  shows  "i * k < j * k"
196proof -
197  from ij ujk knz have jk: "unat i * unat k < 2 ^ len_of TYPE ('a)"
198    by (auto intro: order_less_subst2 simp: word_less_nat_alt elim: mult_less_mono1)
199
200  then show ?thesis using ujk knz ij
201    by (auto simp: word_less_nat_alt iffD1 [OF unat_mult_lem])
202qed
203
204lemma word_mult_less_dest:
205  fixes i :: "'a :: len word"
206  assumes ij: "i * k < j * k"
207  and    uik: "unat i * unat k < 2 ^ len_of TYPE ('a)"
208  and    ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)"
209  shows  "i < j"
210  using uik ujk ij
211  by (auto simp: word_less_nat_alt iffD1 [OF unat_mult_lem] elim: mult_less_mono1)
212
213lemma word_mult_less_cancel:
214  fixes k :: "'a :: len word"
215  assumes knz: "0 < k"
216  and    uik: "unat i * unat k < 2 ^ len_of TYPE ('a)"
217  and    ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)"
218  shows "(i * k < j * k) = (i < j)"
219  by (rule iffI [OF word_mult_less_dest [OF _ uik ujk] word_mult_less_mono1 [OF _ knz ujk]])
220
221lemma Suc_div_unat_helper:
222  assumes szv: "sz < LENGTH('a :: len)"
223  and   usszv: "us \<le> sz"
224  shows "2 ^ (sz - us) = Suc (unat (((2::'a :: len word) ^ sz - 1) div 2 ^ us))"
225proof -
226  note usv = order_le_less_trans [OF usszv szv]
227
228  from usszv obtain q where qv: "sz = us + q" by (auto simp: le_iff_add)
229
230  have "Suc (unat (((2:: 'a word) ^ sz - 1) div 2 ^ us)) =
231    (2 ^ us + unat ((2:: 'a word) ^ sz - 1)) div 2 ^ us"
232    apply (subst unat_div unat_power_lower[OF usv])+
233    apply (subst div_add_self1, simp+)
234    done
235
236  also have "\<dots> = ((2 ^ us - 1) + 2 ^ sz) div 2 ^ us" using szv
237    by (simp add: unat_minus_one)
238
239  also have "\<dots> = 2 ^ q + ((2 ^ us - 1) div 2 ^ us)"
240    apply (subst qv)
241    apply (subst power_add)
242    apply (subst div_mult_self2; simp)
243    done
244
245  also have "\<dots> = 2 ^ (sz - us)" using qv by simp
246
247  finally show ?thesis ..
248qed
249
250
251lemma set_enum_word8_def:
252  "(set enum::word8 set) = {0, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19,
253                            20, 21, 22, 23, 24, 25, 26, 27, 28, 29, 30, 31, 32, 33, 34, 35, 36,
254                            37, 38, 39, 40, 41, 42, 43, 44, 45, 46, 47, 48, 49, 50, 51, 52, 53,
255                            54, 55, 56, 57, 58, 59, 60, 61, 62, 63, 64, 65, 66, 67, 68, 69, 70,
256                            71, 72, 73, 74, 75, 76, 77, 78, 79, 80, 81, 82, 83, 84, 85, 86, 87,
257                            88, 89, 90, 91, 92, 93, 94, 95, 96, 97, 98, 99, 100, 101, 102, 103,
258                            104, 105, 106, 107, 108, 109, 110, 111, 112, 113, 114, 115, 116, 117,
259                            118, 119, 120, 121, 122, 123, 124, 125, 126, 127, 128, 129, 130, 131,
260                            132, 133, 134, 135, 136, 137, 138, 139, 140, 141, 142, 143, 144, 145,
261                            146, 147, 148, 149, 150, 151, 152, 153, 154, 155, 156, 157, 158, 159,
262                            160, 161, 162, 163, 164, 165, 166, 167, 168, 169, 170, 171, 172, 173,
263                            174, 175, 176, 177, 178, 179, 180, 181, 182, 183, 184, 185, 186, 187,
264                            188, 189, 190, 191, 192, 193, 194, 195, 196, 197, 198, 199, 200, 201,
265                            202, 203, 204, 205, 206, 207, 208, 209, 210, 211, 212, 213, 214, 215,
266                            216, 217, 218, 219, 220, 221, 222, 223, 224, 225, 226, 227, 228, 229,
267                            230, 231, 232, 233, 234, 235, 236, 237, 238, 239, 240, 241, 242, 243,
268                            244, 245, 246, 247, 248, 249, 250, 251, 252, 253, 254, 255}"
269  by eval
270
271lemma set_strip_insert: "\<lbrakk> x \<in> insert a S; x \<noteq> a \<rbrakk> \<Longrightarrow> x \<in> S"
272  by simp
273
274lemma word8_exhaust:
275  fixes x :: word8
276  shows "\<lbrakk>x \<noteq> 0; x \<noteq> 1; x \<noteq> 2; x \<noteq> 3; x \<noteq> 4; x \<noteq> 5; x \<noteq> 6; x \<noteq> 7; x \<noteq> 8; x \<noteq> 9; x \<noteq> 10; x \<noteq> 11; x \<noteq>
277          12; x \<noteq> 13; x \<noteq> 14; x \<noteq> 15; x \<noteq> 16; x \<noteq> 17; x \<noteq> 18; x \<noteq> 19; x \<noteq> 20; x \<noteq> 21; x \<noteq> 22; x \<noteq>
278          23; x \<noteq> 24; x \<noteq> 25; x \<noteq> 26; x \<noteq> 27; x \<noteq> 28; x \<noteq> 29; x \<noteq> 30; x \<noteq> 31; x \<noteq> 32; x \<noteq> 33; x \<noteq>
279          34; x \<noteq> 35; x \<noteq> 36; x \<noteq> 37; x \<noteq> 38; x \<noteq> 39; x \<noteq> 40; x \<noteq> 41; x \<noteq> 42; x \<noteq> 43; x \<noteq> 44; x \<noteq>
280          45; x \<noteq> 46; x \<noteq> 47; x \<noteq> 48; x \<noteq> 49; x \<noteq> 50; x \<noteq> 51; x \<noteq> 52; x \<noteq> 53; x \<noteq> 54; x \<noteq> 55; x \<noteq>
281          56; x \<noteq> 57; x \<noteq> 58; x \<noteq> 59; x \<noteq> 60; x \<noteq> 61; x \<noteq> 62; x \<noteq> 63; x \<noteq> 64; x \<noteq> 65; x \<noteq> 66; x \<noteq>
282          67; x \<noteq> 68; x \<noteq> 69; x \<noteq> 70; x \<noteq> 71; x \<noteq> 72; x \<noteq> 73; x \<noteq> 74; x \<noteq> 75; x \<noteq> 76; x \<noteq> 77; x \<noteq>
283          78; x \<noteq> 79; x \<noteq> 80; x \<noteq> 81; x \<noteq> 82; x \<noteq> 83; x \<noteq> 84; x \<noteq> 85; x \<noteq> 86; x \<noteq> 87; x \<noteq> 88; x \<noteq>
284          89; x \<noteq> 90; x \<noteq> 91; x \<noteq> 92; x \<noteq> 93; x \<noteq> 94; x \<noteq> 95; x \<noteq> 96; x \<noteq> 97; x \<noteq> 98; x \<noteq> 99; x \<noteq>
285          100; x \<noteq> 101; x \<noteq> 102; x \<noteq> 103; x \<noteq> 104; x \<noteq> 105; x \<noteq> 106; x \<noteq> 107; x \<noteq> 108; x \<noteq> 109; x \<noteq>
286          110; x \<noteq> 111; x \<noteq> 112; x \<noteq> 113; x \<noteq> 114; x \<noteq> 115; x \<noteq> 116; x \<noteq> 117; x \<noteq> 118; x \<noteq> 119; x \<noteq>
287          120; x \<noteq> 121; x \<noteq> 122; x \<noteq> 123; x \<noteq> 124; x \<noteq> 125; x \<noteq> 126; x \<noteq> 127; x \<noteq> 128; x \<noteq> 129; x \<noteq>
288          130; x \<noteq> 131; x \<noteq> 132; x \<noteq> 133; x \<noteq> 134; x \<noteq> 135; x \<noteq> 136; x \<noteq> 137; x \<noteq> 138; x \<noteq> 139; x \<noteq>
289          140; x \<noteq> 141; x \<noteq> 142; x \<noteq> 143; x \<noteq> 144; x \<noteq> 145; x \<noteq> 146; x \<noteq> 147; x \<noteq> 148; x \<noteq> 149; x \<noteq>
290          150; x \<noteq> 151; x \<noteq> 152; x \<noteq> 153; x \<noteq> 154; x \<noteq> 155; x \<noteq> 156; x \<noteq> 157; x \<noteq> 158; x \<noteq> 159; x \<noteq>
291          160; x \<noteq> 161; x \<noteq> 162; x \<noteq> 163; x \<noteq> 164; x \<noteq> 165; x \<noteq> 166; x \<noteq> 167; x \<noteq> 168; x \<noteq> 169; x \<noteq>
292          170; x \<noteq> 171; x \<noteq> 172; x \<noteq> 173; x \<noteq> 174; x \<noteq> 175; x \<noteq> 176; x \<noteq> 177; x \<noteq> 178; x \<noteq> 179; x \<noteq>
293          180; x \<noteq> 181; x \<noteq> 182; x \<noteq> 183; x \<noteq> 184; x \<noteq> 185; x \<noteq> 186; x \<noteq> 187; x \<noteq> 188; x \<noteq> 189; x \<noteq>
294          190; x \<noteq> 191; x \<noteq> 192; x \<noteq> 193; x \<noteq> 194; x \<noteq> 195; x \<noteq> 196; x \<noteq> 197; x \<noteq> 198; x \<noteq> 199; x \<noteq>
295          200; x \<noteq> 201; x \<noteq> 202; x \<noteq> 203; x \<noteq> 204; x \<noteq> 205; x \<noteq> 206; x \<noteq> 207; x \<noteq> 208; x \<noteq> 209; x \<noteq>
296          210; x \<noteq> 211; x \<noteq> 212; x \<noteq> 213; x \<noteq> 214; x \<noteq> 215; x \<noteq> 216; x \<noteq> 217; x \<noteq> 218; x \<noteq> 219; x \<noteq>
297          220; x \<noteq> 221; x \<noteq> 222; x \<noteq> 223; x \<noteq> 224; x \<noteq> 225; x \<noteq> 226; x \<noteq> 227; x \<noteq> 228; x \<noteq> 229; x \<noteq>
298          230; x \<noteq> 231; x \<noteq> 232; x \<noteq> 233; x \<noteq> 234; x \<noteq> 235; x \<noteq> 236; x \<noteq> 237; x \<noteq> 238; x \<noteq> 239; x \<noteq>
299          240; x \<noteq> 241; x \<noteq> 242; x \<noteq> 243; x \<noteq> 244; x \<noteq> 245; x \<noteq> 246; x \<noteq> 247; x \<noteq> 248; x \<noteq> 249; x \<noteq>
300          250; x \<noteq> 251; x \<noteq> 252; x \<noteq> 253; x \<noteq> 254; x \<noteq> 255\<rbrakk> \<Longrightarrow> P"
301  apply (subgoal_tac "x \<in> set enum", subst (asm) set_enum_word8_def)
302    apply (drule set_strip_insert, assumption)+
303   apply (erule emptyE)
304  apply (subst enum_UNIV, rule UNIV_I)
305  done
306
307lemma upto_enum_red':
308  assumes lt: "1 \<le> X"
309  shows "[(0::'a :: len word) .e. X - 1] =  map of_nat [0 ..< unat X]"
310proof -
311  have lt': "unat X < 2 ^ LENGTH('a)"
312    by (rule unat_lt2p)
313
314  show ?thesis
315    apply (subst upto_enum_red)
316    apply (simp del: upt.simps)
317    apply (subst Suc_unat_diff_1 [OF lt])
318    apply (rule map_cong [OF refl])
319    apply (rule toEnum_of_nat)
320    apply simp
321    apply (erule order_less_trans [OF _ lt'])
322    done
323qed
324
325lemma upto_enum_red2:
326  assumes szv: "sz < LENGTH('a :: len)"
327  shows "[(0:: 'a :: len word) .e. 2 ^ sz - 1] =
328  map of_nat [0 ..< 2 ^ sz]" using szv
329  apply (subst unat_power_lower[OF szv, symmetric])
330  apply (rule upto_enum_red')
331  apply (subst word_le_nat_alt, simp)
332  done
333
334lemma upto_enum_step_red:
335  assumes szv: "sz < LENGTH('a)"
336  and   usszv: "us \<le> sz"
337  shows "[0 :: 'a :: len word , 2 ^ us .e. 2 ^ sz - 1] =
338  map (\<lambda>x. of_nat x * 2 ^ us) [0 ..< 2 ^ (sz - us)]" using szv
339  unfolding upto_enum_step_def
340  apply (subst if_not_P)
341   apply (rule leD)
342   apply (subst word_le_nat_alt)
343   apply (subst unat_minus_one)
344    apply simp
345   apply simp
346  apply simp
347  apply (subst upto_enum_red)
348  apply (simp del: upt.simps)
349  apply (subst Suc_div_unat_helper [where 'a = 'a, OF szv usszv, symmetric])
350  apply clarsimp
351  apply (subst toEnum_of_nat)
352   apply (erule order_less_trans)
353   using szv
354   apply simp
355  apply simp
356  done
357
358lemma upto_enum_word:
359  "[x .e. y] = map of_nat [unat x ..< Suc (unat y)]"
360  apply (subst upto_enum_red)
361  apply clarsimp
362  apply (subst toEnum_of_nat)
363   prefer 2
364   apply (rule refl)
365  apply (erule disjE, simp)
366  apply clarsimp
367  apply (erule order_less_trans)
368  apply simp
369  done
370
371lemma word_upto_Cons_eq:
372  "x < y \<Longrightarrow> [x::'a::len word .e. y] = x # [x + 1 .e. y]"
373  apply (subst upto_enum_red)
374  apply (subst upt_conv_Cons, unat_arith)
375  apply (simp only: list.map list.inject upto_enum_red to_from_enum simp_thms)
376  apply (rule map_cong[OF _ refl])
377  apply (rule arg_cong2[where f = "\<lambda>x y. [x ..< y]"], unat_arith)
378  apply (rule refl)
379  done
380
381lemma distinct_enum_upto:
382  "distinct [(0 :: 'a::len word) .e. b]"
383proof -
384  have "\<And>(b::'a word). [0 .e. b] = nths enum {..< Suc (fromEnum b)}"
385    apply (subst upto_enum_red)
386    apply (subst nths_upt_eq_take)
387    apply (subst enum_word_def)
388    apply (subst take_map)
389    apply (subst take_upt)
390     apply (simp only: add_0 fromEnum_unat)
391     apply (rule order_trans [OF _ order_eq_refl])
392      apply (rule Suc_leI [OF unat_lt2p])
393     apply simp
394    apply clarsimp
395    apply (rule toEnum_of_nat)
396    apply (erule order_less_trans [OF _ unat_lt2p])
397    done
398
399  then show ?thesis
400    by (rule ssubst) (rule distinct_nthsI, simp)
401qed
402
403lemma upto_enum_set_conv [simp]:
404  fixes a :: "'a :: len word"
405  shows "set [a .e. b] = {x. a \<le> x \<and> x \<le> b}"
406  apply (subst upto_enum_red)
407  apply (subst set_map)
408  apply safe
409    apply simp
410    apply clarsimp
411    apply (erule disjE)
412     apply simp
413     apply (erule iffD2 [OF word_le_nat_alt])
414    apply clarsimp
415    apply (erule word_unat.Rep_cases [OF unat_le [OF order_less_imp_le]])
416    apply simp
417    apply (erule iffD2 [OF word_le_nat_alt])
418   apply simp
419
420   apply clarsimp
421   apply (erule disjE)
422    apply simp
423   apply clarsimp
424   apply (rule word_unat.Rep_cases [OF unat_le  [OF order_less_imp_le]])
425    apply assumption
426   apply simp
427   apply (erule order_less_imp_le [OF iffD2 [OF word_less_nat_alt]])
428  apply clarsimp
429  apply (rule_tac x="fromEnum x" in image_eqI)
430   apply clarsimp
431  apply clarsimp
432  apply (rule conjI)
433   apply (subst word_le_nat_alt [symmetric])
434   apply simp
435  apply safe
436   apply (simp add: word_le_nat_alt [symmetric])
437  apply (simp add: word_less_nat_alt [symmetric])
438  done
439
440lemma upto_enum_less:
441  assumes xin: "x \<in> set [(a::'a::len word).e.2 ^ n - 1]"
442  and     nv:  "n < LENGTH('a::len)"
443  shows   "x < 2 ^ n"
444proof (cases n)
445  case 0
446  then show ?thesis using xin by simp
447next
448  case (Suc m)
449  show ?thesis using xin nv by simp
450qed
451
452lemma upto_enum_len_less:
453  "\<lbrakk> n \<le> length [a, b .e. c]; n \<noteq> 0 \<rbrakk> \<Longrightarrow> a \<le> c"
454  unfolding upto_enum_step_def
455  by (simp split: if_split_asm)
456
457lemma length_upto_enum_step:
458  fixes x :: "'a :: len word"
459  shows "x \<le> z \<Longrightarrow> length [x , y .e. z] = (unat ((z - x) div (y - x))) + 1"
460  unfolding upto_enum_step_def
461  by (simp add: upto_enum_red)
462
463lemma map_length_unfold_one:
464  fixes x :: "'a::len word"
465  assumes xv: "Suc (unat x) < 2 ^ LENGTH('a)"
466  and     ax: "a < x"
467  shows   "map f [a .e. x] = f a # map f [a + 1 .e. x]"
468  by (subst word_upto_Cons_eq, auto, fact+)
469
470lemma upto_enum_set_conv2:
471  fixes a :: "'a::len word"
472  shows "set [a .e. b] = {a .. b}"
473  by auto
474
475lemma of_nat_unat [simp]:
476  "of_nat \<circ> unat = id"
477  by (rule ext, simp)
478
479lemma Suc_unat_minus_one [simp]:
480  "x \<noteq> 0 \<Longrightarrow> Suc (unat (x - 1)) = unat x"
481  by (metis Suc_diff_1 unat_gt_0 unat_minus_one)
482
483lemma word_add_le_dest:
484  fixes i :: "'a :: len word"
485  assumes le: "i + k \<le> j + k"
486  and    uik: "unat i + unat k < 2 ^ len_of TYPE ('a)"
487  and    ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)"
488  shows  "i \<le> j"
489  using uik ujk le
490  by (auto simp: word_le_nat_alt iffD1 [OF unat_add_lem] elim: add_le_mono1)
491
492lemma mask_shift:
493  "(x && ~~ (mask y)) >> y = x >> y"
494  by word_eqI
495
496lemma word_add_le_mono1:
497  fixes i :: "'a :: len word"
498  assumes ij: "i \<le> j"
499  and    ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)"
500  shows  "i + k \<le> j + k"
501proof -
502  from ij ujk have jk: "unat i + unat k < 2 ^ len_of TYPE ('a)"
503    by (auto elim: order_le_less_subst2 simp: word_le_nat_alt elim: add_le_mono1)
504
505  then show ?thesis using ujk ij
506    by (auto simp: word_le_nat_alt iffD1 [OF unat_add_lem])
507qed
508
509lemma word_add_le_mono2:
510  fixes i :: "'a :: len word"
511  shows "\<lbrakk>i \<le> j; unat j + unat k < 2 ^ LENGTH('a)\<rbrakk> \<Longrightarrow> k + i \<le> k + j"
512  by (subst field_simps, subst field_simps, erule (1) word_add_le_mono1)
513
514lemma word_add_le_iff:
515  fixes i :: "'a :: len word"
516  assumes uik: "unat i + unat k < 2 ^ len_of TYPE ('a)"
517  and     ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)"
518  shows  "(i + k \<le> j + k) = (i \<le> j)"
519proof
520  assume "i \<le> j"
521  show "i + k \<le> j + k" by (rule word_add_le_mono1) fact+
522next
523  assume "i + k \<le> j + k"
524  show "i \<le> j" by (rule word_add_le_dest) fact+
525qed
526
527lemma word_add_less_mono1:
528  fixes i :: "'a :: len word"
529  assumes ij: "i < j"
530  and    ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)"
531  shows  "i + k < j + k"
532proof -
533  from ij ujk have jk: "unat i + unat k < 2 ^ len_of TYPE ('a)"
534    by (auto elim: order_le_less_subst2 simp: word_less_nat_alt elim: add_less_mono1)
535
536  then show ?thesis using ujk ij
537    by (auto simp: word_less_nat_alt iffD1 [OF unat_add_lem])
538qed
539
540lemma word_add_less_dest:
541  fixes i :: "'a :: len word"
542  assumes le: "i + k < j + k"
543  and    uik: "unat i + unat k < 2 ^ len_of TYPE ('a)"
544  and    ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)"
545  shows  "i < j"
546  using uik ujk le
547  by (auto simp: word_less_nat_alt iffD1 [OF unat_add_lem] elim: add_less_mono1)
548
549lemma word_add_less_iff:
550  fixes i :: "'a :: len word"
551  assumes uik: "unat i + unat k < 2 ^ len_of TYPE ('a)"
552  and     ujk: "unat j + unat k < 2 ^ len_of TYPE ('a)"
553  shows  "(i + k < j + k) = (i < j)"
554proof
555  assume "i < j"
556  show "i + k < j + k" by (rule word_add_less_mono1) fact+
557next
558  assume "i + k < j + k"
559  show "i < j" by (rule word_add_less_dest) fact+
560qed
561
562lemma shiftr_div_2n':
563  "unat (w >> n) = unat w div 2 ^ n"
564  apply (unfold unat_def)
565  apply (subst shiftr_div_2n)
566  apply (subst nat_div_distrib)
567   apply simp
568  apply (simp add: nat_power_eq)
569  done
570
571lemma shiftl_shiftr_id:
572  assumes nv: "n < LENGTH('a)"
573  and     xv: "x < 2 ^ (LENGTH('a) - n)"
574  shows "x << n >> n = (x::'a::len word)"
575  apply (simp add: shiftl_t2n)
576  apply (rule word_unat.Rep_eqD)
577  apply (subst shiftr_div_2n')
578  apply (cases n)
579   apply simp
580  apply (subst iffD1 [OF unat_mult_lem])+
581   apply (subst unat_power_lower[OF nv])
582   apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]])
583   apply (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl])
584   apply (rule unat_power_lower)
585   apply simp
586  apply (subst unat_power_lower[OF nv])
587  apply simp
588  done
589
590lemma ucast_shiftl_eq_0:
591  fixes w :: "'a :: len word"
592  shows "\<lbrakk> n \<ge> LENGTH('b) \<rbrakk> \<Longrightarrow> ucast (w << n) = (0 :: 'b :: len word)"
593  by (case_tac "size w \<le> n", clarsimp simp: shiftl_zero_size)
594     (clarsimp simp: not_le ucast_bl bl_shiftl bang_eq test_bit_of_bl rev_nth nth_append)
595
596lemma word_mult_less_iff:
597  fixes i :: "'a :: len word"
598  assumes knz: "0 < k"
599  and     uik: "unat i * unat k < 2 ^ len_of TYPE ('a)"
600  and     ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)"
601  shows  "(i * k < j * k) = (i < j)"
602  using assms by (rule word_mult_less_cancel)
603
604lemma word_le_imp_diff_le:
605  fixes n :: "'a::len word"
606  shows "\<lbrakk>k \<le> n; n \<le> m\<rbrakk> \<Longrightarrow> n - k \<le> m"
607  by (auto simp: unat_sub word_le_nat_alt)
608
609lemma word_less_imp_diff_less:
610  fixes n :: "'a::len word"
611  shows "\<lbrakk>k \<le> n; n < m\<rbrakk> \<Longrightarrow> n - k < m"
612  by (clarsimp simp: unat_sub word_less_nat_alt
613             intro!: less_imp_diff_less)
614
615lemma word_mult_le_mono1:
616  fixes i :: "'a :: len word"
617  assumes ij: "i \<le> j"
618  and    knz: "0 < k"
619  and    ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)"
620  shows  "i * k \<le> j * k"
621proof -
622  from ij ujk knz have jk: "unat i * unat k < 2 ^ len_of TYPE ('a)"
623    by (auto elim: order_le_less_subst2 simp: word_le_nat_alt elim: mult_le_mono1)
624
625  then show ?thesis using ujk knz ij
626    by (auto simp: word_le_nat_alt iffD1 [OF unat_mult_lem])
627qed
628
629lemma word_mult_le_iff:
630  fixes i :: "'a :: len word"
631  assumes knz: "0 < k"
632  and     uik: "unat i * unat k < 2 ^ len_of TYPE ('a)"
633  and     ujk: "unat j * unat k < 2 ^ len_of TYPE ('a)"
634  shows  "(i * k \<le> j * k) = (i \<le> j)"
635proof
636  assume "i \<le> j"
637  show "i * k \<le> j * k" by (rule word_mult_le_mono1) fact+
638next
639  assume p: "i * k \<le> j * k"
640
641  have "0 < unat k" using knz by (simp add: word_less_nat_alt)
642  then show "i \<le> j" using p
643    by (clarsimp simp: word_le_nat_alt iffD1 [OF unat_mult_lem uik]
644      iffD1 [OF unat_mult_lem ujk])
645qed
646
647lemma word_diff_less:
648  fixes n :: "'a :: len word"
649  shows "\<lbrakk>0 < n; 0 < m; n \<le> m\<rbrakk> \<Longrightarrow> m - n < m"
650  apply (subst word_less_nat_alt)
651  apply (subst unat_sub)
652   apply assumption
653  apply (rule diff_less)
654   apply (simp_all add: word_less_nat_alt)
655   done
656
657lemma MinI:
658  assumes fa: "finite A"
659  and     ne: "A \<noteq> {}"
660  and     xv: "m \<in> A"
661  and    min: "\<forall>y \<in> A. m \<le> y"
662  shows "Min A = m" using fa ne xv min
663proof (induct A arbitrary: m rule: finite_ne_induct)
664  case singleton then show ?case by simp
665next
666  case (insert y F)
667
668  from insert.prems have yx: "m \<le> y" and fx: "\<forall>y \<in> F. m \<le> y" by auto
669  have "m \<in> insert y F" by fact
670  then show ?case
671  proof
672    assume mv: "m = y"
673
674    have mlt: "m \<le> Min F"
675      by (rule iffD2 [OF Min_ge_iff [OF insert.hyps(1) insert.hyps(2)] fx])
676
677    show ?case
678      apply (subst Min_insert [OF insert.hyps(1) insert.hyps(2)])
679      apply (subst mv [symmetric])
680      apply (auto simp: min_def mlt)
681      done
682  next
683    assume "m \<in> F"
684    then have mf: "Min F = m"
685      by (rule insert.hyps(4) [OF _ fx])
686
687    show ?case
688      apply (subst Min_insert [OF insert.hyps(1) insert.hyps(2)])
689      apply (subst mf)
690      apply (rule iffD2 [OF _ yx])
691      apply (auto simp: min_def)
692      done
693  qed
694qed
695
696lemma length_upto_enum [simp]:
697  fixes a :: "'a :: len word"
698  shows "length [a .e. b] = Suc (unat b) - unat a"
699  apply (simp add: word_le_nat_alt upto_enum_red)
700  apply (clarsimp simp: Suc_diff_le)
701  done
702
703lemma length_upto_enum_cases:
704  fixes a :: "'a::len word"
705  shows "length [a .e. b] = (if a \<le> b then Suc (unat b) - unat a else 0)"
706  apply (case_tac "a \<le> b")
707   apply (clarsimp)
708  apply (clarsimp simp: upto_enum_def)
709  apply unat_arith
710  done
711
712lemma length_upto_enum_less_one:
713  "\<lbrakk>a \<le> b; b \<noteq> 0\<rbrakk>
714  \<Longrightarrow> length [a .e. b - 1] = unat (b - a)"
715  apply clarsimp
716  apply (subst unat_sub[symmetric], assumption)
717  apply clarsimp
718  done
719
720lemma drop_upto_enum:
721  "drop (unat n) [0 .e. m] = [n .e. m]"
722  apply (clarsimp simp: upto_enum_def)
723  apply (induct m, simp)
724  by (metis drop_map drop_upt plus_nat.add_0)
725
726lemma distinct_enum_upto' [simp]:
727  "distinct [a::'a::len word .e. b]"
728  apply (subst drop_upto_enum [symmetric])
729  apply (rule distinct_drop)
730  apply (rule distinct_enum_upto)
731  done
732
733lemma length_interval:
734  "\<lbrakk>set xs = {x. (a::'a::len word) \<le> x \<and> x \<le> b}; distinct xs\<rbrakk>
735  \<Longrightarrow> length xs = Suc (unat b) - unat a"
736  apply (frule distinct_card)
737  apply (subgoal_tac "set xs = set [a .e. b]")
738   apply (cut_tac distinct_card [where xs="[a .e. b]"])
739    apply (subst (asm) length_upto_enum)
740    apply clarsimp
741   apply (rule distinct_enum_upto')
742  apply simp
743  done
744
745lemma not_empty_eq:
746  "(S \<noteq> {}) = (\<exists>x. x \<in> S)"
747  by auto
748
749lemma range_subset_lower:
750  fixes c :: "'a ::linorder"
751  shows "\<lbrakk> {a..b} \<subseteq> {c..d}; x \<in> {a..b} \<rbrakk> \<Longrightarrow> c \<le> a"
752  apply (frule (1) subsetD)
753  apply (rule classical)
754  apply clarsimp
755  done
756
757lemma range_subset_upper:
758  fixes c :: "'a ::linorder"
759  shows "\<lbrakk> {a..b} \<subseteq> {c..d}; x \<in> {a..b} \<rbrakk> \<Longrightarrow> b \<le> d"
760  apply (frule (1) subsetD)
761  apply (rule classical)
762  apply clarsimp
763  done
764
765lemma range_subset_eq:
766  fixes a::"'a::linorder"
767  assumes non_empty: "a \<le> b"
768  shows "({a..b} \<subseteq> {c..d}) = (c \<le> a \<and> b \<le> d)"
769  apply (insert non_empty)
770  apply (rule iffI)
771   apply (frule range_subset_lower [where x=a], simp)
772   apply (drule range_subset_upper [where x=a], simp)
773   apply simp
774  apply auto
775  done
776
777lemma range_eq:
778  fixes a::"'a::linorder"
779  assumes non_empty: "a \<le> b"
780  shows "({a..b} = {c..d}) = (a = c \<and> b = d)"
781  by (metis atLeastatMost_subset_iff eq_iff non_empty)
782
783lemma range_strict_subset_eq:
784  fixes a::"'a::linorder"
785  assumes non_empty: "a \<le> b"
786  shows "({a..b} \<subset> {c..d}) = (c \<le> a \<and> b \<le> d \<and> (a = c \<longrightarrow> b \<noteq> d))"
787  apply (insert non_empty)
788  apply (subst psubset_eq)
789  apply (subst range_subset_eq, assumption+)
790  apply (subst range_eq, assumption+)
791  apply simp
792  done
793
794lemma range_subsetI:
795  fixes x :: "'a :: order"
796  assumes xX: "X \<le> x"
797  and     yY: "y \<le> Y"
798  shows   "{x .. y} \<subseteq> {X .. Y}"
799  using xX yY by auto
800
801lemma set_False [simp]:
802  "(set bs \<subseteq> {False}) = (True \<notin> set bs)" by auto
803
804declare of_nat_power [simp del]
805
806(* TODO: move to word *)
807lemma unat_of_bl_length:
808  "unat (of_bl xs :: 'a::len word) < 2 ^ (length xs)"
809proof (cases "length xs < LENGTH('a)")
810  case True
811  then have "(of_bl xs::'a::len word) < 2 ^ length xs"
812    by (simp add: of_bl_length_less)
813  with True
814  show ?thesis
815    by (simp add: word_less_nat_alt word_unat_power unat_of_nat)
816next
817  case False
818  have "unat (of_bl xs::'a::len word) < 2 ^ LENGTH('a)"
819    by (simp split: unat_split)
820  also
821  from False
822  have "LENGTH('a) \<le> length xs" by simp
823  then have "2 ^ LENGTH('a) \<le> (2::nat) ^ length xs"
824    by (rule power_increasing) simp
825  finally
826  show ?thesis .
827qed
828
829lemma is_aligned_0'[simp]:
830  "is_aligned 0 n"
831  by (simp add: is_aligned_def)
832
833lemma p_assoc_help:
834  fixes p :: "'a::{ring,power,numeral,one}"
835  shows "p + 2^sz - 1 = p + (2^sz - 1)"
836  by simp
837
838lemma word_add_increasing:
839  fixes x :: "'a :: len word"
840  shows "\<lbrakk> p + w \<le> x; p \<le> p + w \<rbrakk> \<Longrightarrow> p \<le> x"
841  by unat_arith
842
843lemma word_random:
844  fixes x :: "'a :: len word"
845  shows "\<lbrakk> p \<le> p + x'; x \<le> x' \<rbrakk> \<Longrightarrow> p \<le> p + x"
846  by unat_arith
847
848lemma word_sub_mono:
849  "\<lbrakk> a \<le> c; d \<le> b; a - b \<le> a; c - d \<le> c \<rbrakk>
850    \<Longrightarrow> (a - b) \<le> (c - d :: 'a :: len word)"
851  by unat_arith
852
853lemma power_not_zero:
854  "n < LENGTH('a::len) \<Longrightarrow> (2 :: 'a word) ^ n \<noteq> 0"
855  by (metis p2_gt_0 word_neq_0_conv)
856
857lemma word_gt_a_gt_0:
858  "a < n \<Longrightarrow> (0 :: 'a::len word) < n"
859  apply (case_tac "n = 0")
860   apply clarsimp
861  apply (clarsimp simp: word_neq_0_conv)
862  done
863
864lemma word_shift_nonzero:
865  "\<lbrakk> (x::'a::len word) \<le> 2 ^ m; m + n < LENGTH('a::len); x \<noteq> 0\<rbrakk>
866   \<Longrightarrow> x << n \<noteq> 0"
867  apply (simp only: word_neq_0_conv word_less_nat_alt
868                    shiftl_t2n mod_0 unat_word_ariths
869                    unat_power_lower word_le_nat_alt)
870  apply (subst mod_less)
871   apply (rule order_le_less_trans)
872    apply (erule mult_le_mono2)
873   apply (subst power_add[symmetric])
874   apply (rule power_strict_increasing)
875    apply simp
876   apply simp
877  apply simp
878  done
879
880lemma word_power_less_1 [simp]:
881  "sz < LENGTH('a::len) \<Longrightarrow> (2::'a word) ^ sz - 1 < 2 ^ sz"
882  apply (simp add: word_less_nat_alt)
883  apply (subst unat_minus_one)
884  apply (simp add: word_unat.Rep_inject [symmetric])
885  apply simp
886  done
887
888lemma nasty_split_lt:
889  "\<lbrakk> (x :: 'a:: len word) < 2 ^ (m - n); n \<le> m; m < LENGTH('a::len) \<rbrakk>
890     \<Longrightarrow> x * 2 ^ n + (2 ^ n - 1) \<le> 2 ^ m - 1"
891  apply (simp only: add_diff_eq)
892  apply (subst mult_1[symmetric], subst distrib_right[symmetric])
893  apply (rule word_sub_mono)
894     apply (rule order_trans)
895      apply (rule word_mult_le_mono1)
896        apply (rule inc_le)
897        apply assumption
898       apply (subst word_neq_0_conv[symmetric])
899       apply (rule power_not_zero)
900       apply simp
901      apply (subst unat_power_lower, simp)+
902      apply (subst power_add[symmetric])
903      apply (rule power_strict_increasing)
904       apply simp
905      apply simp
906     apply (subst power_add[symmetric])
907     apply simp
908    apply simp
909   apply (rule word_sub_1_le)
910   apply (subst mult.commute)
911   apply (subst shiftl_t2n[symmetric])
912   apply (rule word_shift_nonzero)
913     apply (erule inc_le)
914    apply simp
915   apply (unat_arith)
916  apply (drule word_power_less_1)
917  apply simp
918  done
919
920lemma nasty_split_less:
921  "\<lbrakk>m \<le> n; n \<le> nm; nm < LENGTH('a::len); x < 2 ^ (nm - n)\<rbrakk>
922   \<Longrightarrow> (x :: 'a word) * 2 ^ n + (2 ^ m - 1) < 2 ^ nm"
923  apply (simp only: word_less_sub_le[symmetric])
924  apply (rule order_trans [OF _ nasty_split_lt])
925     apply (rule word_plus_mono_right)
926      apply (rule word_sub_mono)
927         apply (simp add: word_le_nat_alt)
928        apply simp
929       apply (simp add: word_sub_1_le[OF power_not_zero])
930      apply (simp add: word_sub_1_le[OF power_not_zero])
931     apply (rule is_aligned_no_wrap')
932      apply (rule is_aligned_mult_triv2)
933     apply simp
934    apply (erule order_le_less_trans, simp)
935   apply simp+
936  done
937
938lemma int_not_emptyD:
939  "A \<inter> B \<noteq> {} \<Longrightarrow> \<exists>x. x \<in> A \<and> x \<in> B"
940  by (erule contrapos_np, clarsimp simp: disjoint_iff_not_equal)
941
942lemma unat_less_power:
943  fixes k :: "'a::len word"
944  assumes szv: "sz < LENGTH('a)"
945  and     kv:  "k < 2 ^ sz"
946  shows   "unat k < 2 ^ sz"
947  using szv unat_mono [OF kv] by simp
948
949lemma unat_mult_power_lem:
950  assumes kv: "k < 2 ^ (LENGTH('a::len) - sz)"
951  shows "unat (2 ^ sz * of_nat k :: (('a::len) word)) = 2 ^ sz * k"
952proof cases
953  assume szv: "sz < LENGTH('a::len)"
954  show ?thesis
955  proof (cases "sz = 0")
956    case True
957    then show ?thesis using kv szv
958     by (simp add: unat_of_nat)
959  next
960    case False
961    then have sne: "0 < sz" ..
962
963    have uk: "unat (of_nat k :: 'a word) = k"
964      apply (subst unat_of_nat)
965      apply (simp add: nat_mod_eq less_trans[OF kv] sne)
966      done
967
968    show ?thesis using szv
969      apply (subst iffD1 [OF unat_mult_lem])
970      apply (simp add: uk nat_less_power_trans[OF kv order_less_imp_le [OF szv]])+
971      done
972  qed
973next
974  assume "\<not> sz < LENGTH('a)"
975  with kv show ?thesis by (simp add: not_less power_overflow)
976qed
977
978lemma aligned_add_offset_no_wrap:
979  fixes off :: "('a::len) word"
980  and     x :: "'a word"
981  assumes al: "is_aligned x sz"
982  and   offv: "off < 2 ^ sz"
983  shows  "unat x + unat off < 2 ^ LENGTH('a)"
984proof cases
985  assume szv: "sz < LENGTH('a)"
986  from al obtain k where xv: "x = 2 ^ sz * (of_nat k)"
987    and kl: "k < 2 ^ (LENGTH('a) - sz)"
988    by (auto elim: is_alignedE)
989
990  show ?thesis using szv
991    apply (subst xv)
992    apply (subst unat_mult_power_lem[OF kl])
993    apply (subst mult.commute, rule nat_add_offset_less)
994      apply (rule less_le_trans[OF unat_mono[OF offv, simplified]])
995      apply (erule eq_imp_le[OF unat_power_lower])
996     apply (rule kl)
997    apply simp
998   done
999next
1000  assume "\<not> sz < LENGTH('a)"
1001  with offv show ?thesis by (simp add: not_less power_overflow )
1002qed
1003
1004lemma aligned_add_offset_mod:
1005  fixes x :: "('a::len) word"
1006  assumes al: "is_aligned x sz"
1007  and     kv: "k < 2 ^ sz"
1008  shows   "(x + k) mod 2 ^ sz = k"
1009proof cases
1010  assume szv: "sz < LENGTH('a)"
1011
1012  have ux: "unat x + unat k < 2 ^ LENGTH('a)"
1013    by (rule aligned_add_offset_no_wrap) fact+
1014
1015  show ?thesis using al szv
1016    apply -
1017    apply (erule is_alignedE)
1018    apply (subst word_unat.Rep_inject [symmetric])
1019    apply (subst unat_mod)
1020    apply (subst iffD1 [OF unat_add_lem], rule ux)
1021    apply simp
1022    apply (subst unat_mult_power_lem, assumption+)
1023    apply (simp)
1024    apply (rule mod_less[OF less_le_trans[OF unat_mono], OF kv])
1025    apply (erule eq_imp_le[OF unat_power_lower])
1026    done
1027next
1028  assume "\<not> sz < LENGTH('a)"
1029  with al show ?thesis
1030    by (simp add: not_less power_overflow is_aligned_mask mask_def
1031                  word_mod_by_0)
1032qed
1033
1034lemma word_plus_mcs_4:
1035  "\<lbrakk>v + x \<le> w + x; x \<le> v + x\<rbrakk> \<Longrightarrow> v \<le> (w::'a::len word)"
1036  by uint_arith
1037
1038lemma word_plus_mcs_3:
1039  "\<lbrakk>v \<le> w; x \<le> w + x\<rbrakk> \<Longrightarrow> v + x \<le> w + (x::'a::len word)"
1040  by unat_arith
1041
1042lemma aligned_neq_into_no_overlap:
1043  fixes x :: "'a::len word"
1044  assumes neq: "x \<noteq> y"
1045  and     alx: "is_aligned x sz"
1046  and     aly: "is_aligned y sz"
1047  shows  "{x .. x + (2 ^ sz - 1)} \<inter> {y .. y + (2 ^ sz - 1)} = {}"
1048proof cases
1049  assume szv: "sz < LENGTH('a)"
1050  show ?thesis
1051  proof (rule equals0I, clarsimp)
1052    fix z
1053    assume xb: "x \<le> z" and xt: "z \<le> x + (2 ^ sz - 1)"
1054      and yb: "y \<le> z" and yt: "z \<le> y + (2 ^ sz - 1)"
1055
1056    have rl: "\<And>(p::'a word) k w. \<lbrakk>uint p + uint k < 2 ^ LENGTH('a); w = p + k; w \<le> p + (2 ^ sz - 1) \<rbrakk>
1057      \<Longrightarrow> k < 2 ^ sz"
1058      apply -
1059      apply simp
1060      apply (subst (asm) add.commute, subst (asm) add.commute, drule word_plus_mcs_4)
1061      apply (subst add.commute, subst no_plus_overflow_uint_size)
1062      apply (simp add: word_size_bl)
1063      apply (erule iffD1 [OF word_less_sub_le[OF szv]])
1064      done
1065
1066    from xb obtain kx where
1067      kx: "z = x + kx" and
1068      kxl: "uint x + uint kx < 2 ^ LENGTH('a)"
1069      by (clarsimp dest!: word_le_exists')
1070
1071    from yb obtain ky where
1072      ky: "z = y + ky" and
1073      kyl: "uint y + uint ky < 2 ^ LENGTH('a)"
1074      by (clarsimp dest!: word_le_exists')
1075
1076    have "x = y"
1077    proof -
1078      have "kx = z mod 2 ^ sz"
1079      proof (subst kx, rule sym, rule aligned_add_offset_mod)
1080        show "kx < 2 ^ sz" by (rule rl) fact+
1081      qed fact+
1082
1083      also have "\<dots> = ky"
1084      proof (subst ky, rule aligned_add_offset_mod)
1085        show "ky < 2 ^ sz"
1086          using kyl ky yt by (rule rl)
1087      qed fact+
1088
1089      finally have kxky: "kx = ky" .
1090      moreover have "x + kx = y + ky" by (simp add: kx [symmetric] ky [symmetric])
1091      ultimately show ?thesis by simp
1092    qed
1093
1094    then show False using neq by simp
1095  qed
1096next
1097  assume "\<not> sz < LENGTH('a)"
1098  with neq alx aly
1099  have False by (simp add: is_aligned_mask mask_def power_overflow)
1100  then show ?thesis ..
1101qed
1102
1103lemma less_two_pow_divD:
1104  "\<lbrakk> (x :: nat) < 2 ^ n div 2 ^ m \<rbrakk>
1105    \<Longrightarrow> n \<ge> m \<and> (x < 2 ^ (n - m))"
1106  apply (rule context_conjI)
1107   apply (rule ccontr)
1108   apply (simp add: power_strict_increasing)
1109  apply (simp add: power_sub)
1110  done
1111
1112lemma less_two_pow_divI:
1113  "\<lbrakk> (x :: nat) < 2 ^ (n - m); m \<le> n \<rbrakk> \<Longrightarrow> x < 2 ^ n div 2 ^ m"
1114  by (simp add: power_sub)
1115
1116lemma word_less_two_pow_divI:
1117  "\<lbrakk> (x :: 'a::len word) < 2 ^ (n - m); m \<le> n; n < LENGTH('a) \<rbrakk> \<Longrightarrow> x < 2 ^ n div 2 ^ m"
1118  apply (simp add: word_less_nat_alt)
1119  apply (subst unat_word_ariths)
1120  apply (subst mod_less)
1121   apply (rule order_le_less_trans [OF div_le_dividend])
1122   apply (rule unat_lt2p)
1123  apply (simp add: power_sub)
1124  done
1125
1126lemma word_less_two_pow_divD:
1127  "\<lbrakk> (x :: 'a::len word) < 2 ^ n div 2 ^ m \<rbrakk>
1128     \<Longrightarrow> n \<ge> m \<and> (x < 2 ^ (n - m))"
1129  apply (cases "n < LENGTH('a)")
1130   apply (cases "m < LENGTH('a)")
1131    apply (simp add: word_less_nat_alt)
1132    apply (subst(asm) unat_word_ariths)
1133    apply (subst(asm) mod_less)
1134     apply (rule order_le_less_trans [OF div_le_dividend])
1135     apply (rule unat_lt2p)
1136    apply (clarsimp dest!: less_two_pow_divD)
1137   apply (simp add: power_overflow)
1138   apply (simp add: word_div_def)
1139  apply (simp add: power_overflow word_div_def)
1140  done
1141
1142lemma of_nat_less_two_pow_div_set:
1143  "\<lbrakk> n < LENGTH('a) \<rbrakk> \<Longrightarrow>
1144   {x. x < (2 ^ n div 2 ^ m :: 'a::len word)}
1145      = of_nat ` {k. k < 2 ^ n div 2 ^ m}"
1146  apply (simp add: image_def)
1147  apply (safe dest!: word_less_two_pow_divD less_two_pow_divD
1148             intro!: word_less_two_pow_divI)
1149   apply (rule_tac x="unat x" in exI)
1150   apply (simp add: power_sub[symmetric])
1151   apply (subst unat_power_lower[symmetric, where 'a='a])
1152    apply simp
1153   apply (erule unat_mono)
1154  apply (subst word_unat_power)
1155  apply (rule of_nat_mono_maybe)
1156   apply (rule power_strict_increasing)
1157    apply simp
1158   apply simp
1159  apply assumption
1160  done
1161
1162lemma  word_less_power_trans2:
1163  fixes n :: "'a::len word"
1164  shows "\<lbrakk>n < 2 ^ (m - k); k \<le> m; m < LENGTH('a)\<rbrakk> \<Longrightarrow> n * 2 ^ k < 2 ^ m"
1165  by (subst field_simps, rule word_less_power_trans)
1166
1167(* shadows the slightly weaker Word.nth_ucast *)
1168lemma nth_ucast:
1169  "(ucast (w::'a::len0 word)::'b::len0 word) !! n =
1170   (w !! n \<and> n < min LENGTH('a) LENGTH('b))"
1171  by (simp add: ucast_def test_bit_bin word_ubin.eq_norm nth_bintr word_size)
1172     (fast elim!: bin_nth_uint_imp)
1173
1174lemma ucast_less:
1175  "LENGTH('b) < LENGTH('a) \<Longrightarrow>
1176   (ucast (x :: 'b :: len word) :: ('a :: len word)) < 2 ^ LENGTH('b)"
1177  by (meson Word.nth_ucast test_bit_conj_lt le_def upper_bits_unset_is_l2p)
1178
1179lemma ucast_range_less:
1180  "LENGTH('a :: len) < LENGTH('b :: len) \<Longrightarrow>
1181   range (ucast :: 'a word \<Rightarrow> 'b word) = {x. x < 2 ^ len_of TYPE ('a)}"
1182  apply safe
1183   apply (erule ucast_less)
1184  apply (simp add: image_def)
1185  apply (rule_tac x="ucast x" in exI)
1186  by word_eqI_solve
1187
1188lemma word_power_less_diff:
1189  "\<lbrakk>2 ^ n * q < (2::'a::len word) ^ m; q < 2 ^ (LENGTH('a) - n)\<rbrakk> \<Longrightarrow> q < 2 ^ (m - n)"
1190  apply (case_tac "m \<ge> LENGTH('a)")
1191   apply (simp add: power_overflow)
1192  apply (case_tac "n \<ge> LENGTH('a)")
1193   apply (simp add: power_overflow)
1194  apply (cases "n = 0")
1195   apply simp
1196  apply (subst word_less_nat_alt)
1197  apply (subst unat_power_lower)
1198   apply simp
1199  apply (rule nat_power_less_diff)
1200  apply (simp add: word_less_nat_alt)
1201  apply (subst (asm) iffD1 [OF unat_mult_lem])
1202   apply (simp add:nat_less_power_trans)
1203  apply simp
1204  done
1205
1206lemmas word_diff_ls'' = word_diff_ls [where xa=x and x=x for x]
1207lemmas word_diff_ls' = word_diff_ls'' [simplified]
1208
1209lemmas word_l_diffs' = word_l_diffs [where xa=x and x=x for x]
1210lemmas word_l_diffs = word_l_diffs' [simplified]
1211
1212lemma is_aligned_diff:
1213  fixes m :: "'a::len word"
1214  assumes alm: "is_aligned m s1"
1215  and     aln: "is_aligned n s2"
1216  and    s2wb: "s2 < LENGTH('a)"
1217  and      nm: "m \<in> {n .. n + (2 ^ s2 - 1)}"
1218  and    s1s2: "s1 \<le> s2"
1219  and     s10: "0 < s1" (* Probably can be folded into the proof \<dots> *)
1220  shows  "\<exists>q. m - n = of_nat q * 2 ^ s1 \<and> q < 2 ^ (s2 - s1)"
1221proof -
1222  have rl: "\<And>m s. \<lbrakk> m < 2 ^ (LENGTH('a) - s); s < LENGTH('a) \<rbrakk> \<Longrightarrow> unat ((2::'a word) ^ s * of_nat m) = 2 ^ s * m"
1223  proof -
1224    fix m :: nat and  s
1225    assume m: "m < 2 ^ (LENGTH('a) - s)" and s: "s < LENGTH('a)"
1226    then have "unat ((of_nat m) :: 'a word) = m"
1227      apply (subst unat_of_nat)
1228      apply (subst mod_less)
1229       apply (erule order_less_le_trans)
1230       apply (rule power_increasing)
1231        apply simp_all
1232      done
1233
1234    then show "?thesis m s" using s m
1235      apply (subst iffD1 [OF unat_mult_lem])
1236      apply (simp add: nat_less_power_trans)+
1237      done
1238  qed
1239  have s1wb: "s1 < LENGTH('a)" using s2wb s1s2 by simp
1240  from alm obtain mq where mmq: "m = 2 ^ s1 * of_nat mq" and mq: "mq < 2 ^ (LENGTH('a) - s1)"
1241    by (auto elim: is_alignedE simp: field_simps)
1242  from aln obtain nq where nnq: "n = 2 ^ s2 * of_nat nq" and nq: "nq < 2 ^ (LENGTH('a) - s2)"
1243    by (auto elim: is_alignedE simp: field_simps)
1244  from s1s2 obtain sq where sq: "s2 = s1 + sq" by (auto simp: le_iff_add)
1245
1246  note us1 = rl [OF mq s1wb]
1247  note us2 = rl [OF nq s2wb]
1248
1249  from nm have "n \<le> m" by clarsimp
1250  then have "(2::'a word) ^ s2 * of_nat nq \<le> 2 ^ s1 * of_nat mq" using nnq mmq by simp
1251  then have "2 ^ s2 * nq \<le> 2 ^ s1 * mq" using s1wb s2wb
1252    by (simp add: word_le_nat_alt us1 us2)
1253  then have nqmq: "2 ^ sq * nq \<le> mq" using sq by (simp add: power_add)
1254
1255  have "m - n = 2 ^ s1 * of_nat mq - 2 ^ s2 * of_nat nq" using mmq nnq by simp
1256  also have "\<dots> = 2 ^ s1 * of_nat mq - 2 ^ s1 * 2 ^ sq * of_nat nq" using sq by (simp add: power_add)
1257  also have "\<dots> = 2 ^ s1 * (of_nat mq - 2 ^ sq * of_nat nq)" by (simp add: field_simps)
1258  also have "\<dots> = 2 ^ s1 * of_nat (mq - 2 ^ sq * nq)" using s1wb s2wb us1 us2 nqmq
1259    by (simp add:  word_unat_power)
1260  finally have mn: "m - n = of_nat (mq - 2 ^ sq * nq) * 2 ^ s1" by simp
1261  moreover
1262  from nm have "m - n \<le> 2 ^ s2 - 1"
1263    by - (rule word_diff_ls', (simp add: field_simps)+)
1264  then have "(2::'a word) ^ s1 * of_nat (mq - 2 ^ sq * nq) < 2 ^ s2" using mn s2wb by (simp add: field_simps)
1265  then have "of_nat (mq - 2 ^ sq * nq) < (2::'a word) ^ (s2 - s1)"
1266  proof (rule word_power_less_diff)
1267    have mm: "mq - 2 ^ sq * nq < 2 ^ (LENGTH('a) - s1)" using mq by simp
1268    moreover from s10 have "LENGTH('a) - s1 < LENGTH('a)"
1269      by (rule diff_less, simp)
1270    ultimately show "of_nat (mq - 2 ^ sq * nq) < (2::'a word) ^ (LENGTH('a) - s1)"
1271      apply (simp add: word_less_nat_alt)
1272      apply (subst unat_of_nat)
1273      apply (subst mod_less)
1274       apply (erule order_less_le_trans)
1275       apply simp+
1276      done
1277  qed
1278  then have "mq - 2 ^ sq * nq < 2 ^ (s2 - s1)" using mq s2wb
1279    apply (simp add: word_less_nat_alt)
1280    apply (subst (asm) unat_of_nat)
1281    apply (subst (asm) mod_less)
1282    apply (rule order_le_less_trans)
1283    apply (rule diff_le_self)
1284    apply (erule order_less_le_trans)
1285    apply simp
1286    apply assumption
1287    done
1288  ultimately show ?thesis by auto
1289qed
1290
1291lemma word_less_sub_1:
1292  "x < (y :: 'a :: len word) \<Longrightarrow> x \<le> y - 1"
1293  apply (erule udvd_minus_le')
1294   apply (simp add: udvd_def)+
1295  done
1296
1297lemma word_sub_mono2:
1298  "\<lbrakk> a + b \<le> c + d; c \<le> a; b \<le> a + b; d \<le> c + d \<rbrakk>
1299    \<Longrightarrow> b \<le> (d :: 'a :: len word)"
1300  apply (drule(1) word_sub_mono)
1301    apply simp
1302   apply simp
1303  apply simp
1304  done
1305
1306lemma word_not_le:
1307  "(\<not> x \<le> (y :: 'a :: len word)) = (y < x)"
1308  by fastforce
1309
1310lemma word_subset_less:
1311  "\<lbrakk> {x .. x + r - 1} \<subseteq> {y .. y + s - 1};
1312     x \<le> x + r - 1; y \<le> y + (s :: 'a :: len word) - 1;
1313     s \<noteq> 0 \<rbrakk>
1314     \<Longrightarrow> r \<le> s"
1315  apply (frule subsetD[where c=x])
1316   apply simp
1317  apply (drule subsetD[where c="x + r - 1"])
1318   apply simp
1319  apply (clarsimp simp: add_diff_eq[symmetric])
1320  apply (drule(1) word_sub_mono2)
1321    apply (simp_all add: olen_add_eqv[symmetric])
1322  apply (erule word_le_minus_cancel)
1323  apply (rule ccontr)
1324  apply (simp add: word_not_le)
1325  done
1326
1327lemma uint_power_lower:
1328  "n < LENGTH('a) \<Longrightarrow> uint (2 ^ n :: 'a :: len word) = (2 ^ n :: int)"
1329  by (rule uint_2p_alt)
1330
1331lemma power_le_mono:
1332  "\<lbrakk>2 ^ n \<le> (2::'a::len word) ^ m; n < LENGTH('a); m < LENGTH('a)\<rbrakk>
1333   \<Longrightarrow> n \<le> m"
1334  apply (clarsimp simp add: le_less)
1335  apply safe
1336  apply (simp add: word_less_nat_alt)
1337  apply (simp only: uint_arith_simps(3))
1338  apply (drule uint_power_lower)+
1339  apply simp
1340  done
1341
1342lemma sublist_equal_part:
1343  "prefix xs ys \<Longrightarrow> take (length xs) ys = xs"
1344  by (clarsimp simp: prefix_def)
1345
1346lemma two_power_eq:
1347  "\<lbrakk>n < LENGTH('a); m < LENGTH('a)\<rbrakk>
1348   \<Longrightarrow> ((2::'a::len word) ^ n = 2 ^ m) = (n = m)"
1349  apply safe
1350  apply (rule order_antisym)
1351   apply (simp add: power_le_mono[where 'a='a])+
1352  done
1353
1354lemma prefix_length_less:
1355  "strict_prefix xs ys \<Longrightarrow> length xs < length ys"
1356  apply (clarsimp simp: strict_prefix_def)
1357  apply (frule prefix_length_le)
1358  apply (rule ccontr, simp)
1359  apply (clarsimp simp: prefix_def)
1360  done
1361
1362lemmas take_less = take_strict_prefix
1363
1364lemma not_prefix_longer:
1365  "\<lbrakk> length xs > length ys \<rbrakk> \<Longrightarrow> \<not> prefix xs ys"
1366  by (clarsimp dest!: prefix_length_le)
1367
1368lemma of_bl_length:
1369  "length xs < LENGTH('a) \<Longrightarrow> of_bl xs < (2 :: 'a::len word) ^ length xs"
1370  by (simp add: of_bl_length_less)
1371
1372lemma unat_of_nat_eq:
1373  "x < 2 ^ LENGTH('a) \<Longrightarrow> unat (of_nat x ::'a::len word) = x"
1374  by (rule unat_of_nat_len)
1375
1376lemma unat_eq_of_nat:
1377  "n < 2 ^ LENGTH('a) \<Longrightarrow> (unat (x :: 'a::len word) = n) = (x = of_nat n)"
1378  by (subst unat_of_nat_eq[where x=n, symmetric], simp+)
1379
1380lemma unat_less_helper:
1381  "x < of_nat n \<Longrightarrow> unat x < n"
1382  apply (simp add: word_less_nat_alt)
1383  apply (erule order_less_le_trans)
1384  apply (simp add: unat_of_nat)
1385  done
1386
1387lemma nat_uint_less_helper:
1388  "nat (uint y) = z \<Longrightarrow> x < y \<Longrightarrow> nat (uint x) < z"
1389  apply (erule subst)
1390  apply (subst unat_def [symmetric])
1391  apply (subst unat_def [symmetric])
1392  by (simp add: unat_mono)
1393
1394lemma of_nat_0:
1395  "\<lbrakk>of_nat n = (0::'a::len word); n < 2 ^ LENGTH('a)\<rbrakk> \<Longrightarrow> n = 0"
1396  by (drule unat_of_nat_eq, simp)
1397
1398lemma word_leq_le_minus_one:
1399  "\<lbrakk> x \<le> y; x \<noteq> 0 \<rbrakk> \<Longrightarrow> x - 1 < (y :: 'a :: len word)"
1400  apply (simp add: word_less_nat_alt word_le_nat_alt)
1401  apply (subst unat_minus_one)
1402   apply assumption
1403  apply (cases "unat x")
1404   apply (simp add: unat_eq_zero)
1405  apply arith
1406  done
1407
1408lemma of_nat_inj:
1409  "\<lbrakk>x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\<rbrakk> \<Longrightarrow>
1410   (of_nat x = (of_nat y :: 'a :: len word)) = (x = y)"
1411  by (simp add: word_unat.norm_eq_iff [symmetric])
1412
1413lemma map_prefixI:
1414  "prefix xs ys \<Longrightarrow> prefix (map f xs) (map f ys)"
1415  by (clarsimp simp: prefix_def)
1416
1417lemma if_Some_None_eq_None:
1418  "((if P then Some v else None) = None) = (\<not> P)"
1419  by simp
1420
1421lemma CollectPairFalse [iff]:
1422  "{(a,b). False} = {}"
1423  by (simp add: split_def)
1424
1425lemma if_conj_dist:
1426  "((if b then w else x) \<and> (if b then y else z) \<and> X) =
1427  ((if b then w \<and> y else x \<and> z) \<and> X)"
1428  by simp
1429
1430lemma if_P_True1:
1431  "Q \<Longrightarrow> (if P then True else Q)"
1432  by simp
1433
1434lemma if_P_True2:
1435  "Q \<Longrightarrow> (if P then Q else True)"
1436  by simp
1437
1438lemma list_all2_induct [consumes 1, case_names Nil Cons]:
1439  assumes lall: "list_all2 Q xs ys"
1440  and     nilr: "P [] []"
1441  and    consr: "\<And>x xs y ys. \<lbrakk>list_all2 Q xs ys; Q x y; P xs ys\<rbrakk> \<Longrightarrow> P (x # xs) (y # ys)"
1442  shows  "P xs ys"
1443  using lall
1444proof (induct rule: list_induct2 [OF list_all2_lengthD [OF lall]])
1445  case 1 then show ?case by auto fact+
1446next
1447  case (2 x xs y ys)
1448
1449  show ?case
1450  proof (rule consr)
1451    from "2.prems" show "list_all2 Q xs ys" and "Q x y" by simp_all
1452    then show "P xs ys" by (intro "2.hyps")
1453  qed
1454qed
1455
1456lemma list_all2_induct_suffixeq [consumes 1, case_names Nil Cons]:
1457  assumes lall: "list_all2 Q as bs"
1458  and     nilr: "P [] []"
1459  and    consr: "\<And>x xs y ys.
1460  \<lbrakk>list_all2 Q xs ys; Q x y; P xs ys; suffix (x # xs) as; suffix (y # ys) bs\<rbrakk>
1461  \<Longrightarrow> P (x # xs) (y # ys)"
1462  shows  "P as bs"
1463proof -
1464  define as' where "as' == as"
1465  define bs' where "bs' == bs"
1466
1467  have "suffix as as' \<and> suffix bs bs'" unfolding as'_def bs'_def by simp
1468  then show ?thesis using lall
1469  proof (induct rule: list_induct2 [OF list_all2_lengthD [OF lall]])
1470    case 1 show ?case by fact
1471  next
1472    case (2 x xs y ys)
1473
1474    show ?case
1475    proof (rule consr)
1476      from "2.prems" show "list_all2 Q xs ys" and "Q x y" by simp_all
1477      then show "P xs ys" using "2.hyps" "2.prems" by (auto dest: suffix_ConsD)
1478      from "2.prems" show "suffix (x # xs) as" and "suffix (y # ys) bs"
1479      by (auto simp: as'_def bs'_def)
1480    qed
1481  qed
1482qed
1483
1484lemma upto_enum_step_shift:
1485  "\<lbrakk> is_aligned p n \<rbrakk> \<Longrightarrow>
1486  ([p , p + 2 ^ m .e. p + 2 ^ n - 1])
1487      = map ((+) p) [0, 2 ^ m .e. 2 ^ n - 1]"
1488  apply (erule is_aligned_get_word_bits)
1489   prefer 2
1490   apply (simp add: map_idI)
1491  apply (clarsimp simp: upto_enum_step_def)
1492  apply (frule is_aligned_no_overflow)
1493  apply (simp add: linorder_not_le [symmetric])
1494  done
1495
1496lemma upto_enum_step_shift_red:
1497  "\<lbrakk> is_aligned p sz; sz < LENGTH('a); us \<le> sz \<rbrakk>
1498     \<Longrightarrow> [p :: 'a :: len word, p + 2 ^ us .e. p + 2 ^ sz - 1]
1499          = map (\<lambda>x. p + of_nat x * 2 ^ us) [0 ..< 2 ^ (sz - us)]"
1500  apply (subst upto_enum_step_shift, assumption)
1501  apply (simp add: upto_enum_step_red)
1502  done
1503
1504lemma div_to_mult_word_lt:
1505  "\<lbrakk> (x :: 'a :: len word) \<le> y div z \<rbrakk> \<Longrightarrow> x * z \<le> y"
1506  apply (cases "z = 0")
1507   apply simp
1508  apply (simp add: word_neq_0_conv)
1509  apply (rule order_trans)
1510   apply (erule(1) word_mult_le_mono1)
1511   apply (simp add: unat_div)
1512   apply (rule order_le_less_trans [OF div_mult_le])
1513   apply simp
1514  apply (rule word_div_mult_le)
1515  done
1516
1517lemma upto_enum_step_subset:
1518  "set [x, y .e. z] \<subseteq> {x .. z}"
1519  apply (clarsimp simp: upto_enum_step_def linorder_not_less)
1520  apply (drule div_to_mult_word_lt)
1521  apply (rule conjI)
1522   apply (erule word_random[rotated])
1523   apply simp
1524  apply (rule order_trans)
1525   apply (erule word_plus_mono_right)
1526   apply simp
1527  apply simp
1528  done
1529
1530lemma shiftr_less_t2n':
1531  "\<lbrakk> x && mask (n + m) = x; m < LENGTH('a) \<rbrakk> \<Longrightarrow> x >> n < 2 ^ m" for x :: "'a :: len word"
1532  apply (simp add: word_size mask_eq_iff_w2p[symmetric])
1533  apply word_eqI
1534  apply (erule_tac x="na + n" in allE)
1535  apply fastforce
1536  done
1537
1538lemma shiftr_less_t2n:
1539  "x < 2 ^ (n + m) \<Longrightarrow> x >> n < 2 ^ m" for x :: "'a :: len word"
1540  apply (rule shiftr_less_t2n')
1541   apply (erule less_mask_eq)
1542  apply (rule ccontr)
1543  apply (simp add: not_less)
1544  apply (subst (asm) p2_eq_0[symmetric])
1545  apply (simp add: power_add)
1546  done
1547
1548lemma shiftr_eq_0:
1549  "n \<ge> LENGTH('a) \<Longrightarrow> ((w::'a::len word) >> n) = 0"
1550  apply (cut_tac shiftr_less_t2n'[of w n 0], simp)
1551   apply (simp add: mask_eq_iff)
1552  apply (simp add: lt2p_lem)
1553  apply simp
1554  done
1555
1556lemma shiftr_not_mask_0:
1557  "n+m \<ge> LENGTH('a :: len) \<Longrightarrow> ((w::'a::len word) >> n) && ~~ (mask m) = 0"
1558  apply (simp add: and_not_mask shiftr_less_t2n shiftr_shiftr)
1559  apply (subgoal_tac "w >> n + m = 0", simp)
1560  apply (simp add: le_mask_iff[symmetric] mask_def le_def)
1561  apply (subst (asm) p2_gt_0[symmetric])
1562  apply (simp add: power_add not_less)
1563  done
1564
1565lemma shiftl_less_t2n:
1566  fixes x :: "'a :: len word"
1567  shows "\<lbrakk> x < (2 ^ (m - n)); m < LENGTH('a) \<rbrakk> \<Longrightarrow> (x << n) < 2 ^ m"
1568  apply (simp add: word_size mask_eq_iff_w2p[symmetric])
1569  apply word_eqI
1570  apply (erule_tac x="na - n" in allE)
1571  apply auto
1572  done
1573
1574lemma shiftl_less_t2n':
1575  "(x::'a::len word) < 2 ^ m \<Longrightarrow> m+n < LENGTH('a) \<Longrightarrow> x << n < 2 ^ (m + n)"
1576  by (rule shiftl_less_t2n) simp_all
1577
1578lemma ucast_ucast_mask:
1579  "(ucast :: 'a :: len word \<Rightarrow> 'b :: len word) (ucast x) = x && mask (len_of TYPE ('a))"
1580  by word_eqI
1581
1582lemma ucast_ucast_len:
1583  "\<lbrakk> x < 2 ^ LENGTH('b) \<rbrakk> \<Longrightarrow> ucast (ucast x::'b::len word) = (x::'a::len word)"
1584  apply (subst ucast_ucast_mask)
1585  apply (erule less_mask_eq)
1586  done
1587
1588lemma ucast_ucast_id:
1589  "LENGTH('a) < LENGTH('b) \<Longrightarrow> ucast (ucast (x::'a::len word)::'b::len word) = x"
1590  by (auto intro: ucast_up_ucast_id simp: is_up_def source_size_def target_size_def word_size)
1591
1592lemma unat_ucast:
1593  "unat (ucast x :: ('a :: len0) word) = unat x mod 2 ^ (LENGTH('a))"
1594  apply (simp add: unat_def ucast_def)
1595  apply (subst word_uint.eq_norm)
1596  apply (subst nat_mod_distrib)
1597    apply simp
1598   apply simp
1599  apply (subst nat_power_eq)
1600   apply simp
1601  apply simp
1602  done
1603
1604lemma ucast_less_ucast:
1605  "LENGTH('a) \<le> LENGTH('b) \<Longrightarrow>
1606   (ucast x < ((ucast (y :: 'a::len word)) :: 'b::len word)) = (x < y)"
1607  apply (simp add: word_less_nat_alt unat_ucast)
1608  apply (subst mod_less)
1609   apply(rule less_le_trans[OF unat_lt2p], simp)
1610  apply (subst mod_less)
1611   apply(rule less_le_trans[OF unat_lt2p], simp)
1612  apply simp
1613  done
1614
1615\<comment> \<open>This weaker version was previously called @{text ucast_less_ucast}. We retain it to
1616    support existing proofs.\<close>
1617lemmas ucast_less_ucast_weak = ucast_less_ucast[OF order.strict_implies_order]
1618
1619lemma sints_subset:
1620  "m \<le> n \<Longrightarrow> sints m \<subseteq> sints n"
1621  apply (simp add: sints_num)
1622  apply clarsimp
1623  apply (rule conjI)
1624   apply (erule order_trans[rotated])
1625   apply simp
1626  apply (erule order_less_le_trans)
1627  apply simp
1628  done
1629
1630lemma up_scast_inj:
1631      "\<lbrakk> scast x = (scast y :: 'b :: len word); size x \<le> LENGTH('b) \<rbrakk>
1632         \<Longrightarrow> x = y"
1633  apply (simp add: scast_def)
1634  apply (subst(asm) word_sint.Abs_inject)
1635    apply (erule subsetD [OF sints_subset])
1636    apply (simp add: word_size)
1637   apply (erule subsetD [OF sints_subset])
1638   apply (simp add: word_size)
1639  apply simp
1640  done
1641
1642lemma up_scast_inj_eq:
1643  "LENGTH('a) \<le> len_of TYPE ('b) \<Longrightarrow>
1644  (scast x = (scast y::'b::len word)) = (x = (y::'a::len word))"
1645  by (fastforce dest: up_scast_inj simp: word_size)
1646
1647lemma nth_bounded:
1648  "\<lbrakk>(x :: 'a :: len word) !! n; x < 2 ^ m; m \<le> len_of TYPE ('a)\<rbrakk> \<Longrightarrow> n < m"
1649  apply (frule test_bit_size)
1650  apply (clarsimp simp: test_bit_bl word_size)
1651  apply (simp add: nth_rev)
1652  apply (subst(asm) is_aligned_add_conv[OF is_aligned_0',
1653                                        simplified add_0_left, rotated])
1654   apply assumption+
1655  apply (simp only: to_bl_0)
1656  apply (simp add: nth_append split: if_split_asm)
1657  done
1658
1659lemma is_aligned_add_or:
1660  "\<lbrakk>is_aligned p n; d < 2 ^ n\<rbrakk> \<Longrightarrow> p + d = p || d"
1661  by (rule word_plus_and_or_coroll, word_eqI) blast
1662
1663lemma two_power_increasing:
1664  "\<lbrakk> n \<le> m; m < LENGTH('a) \<rbrakk> \<Longrightarrow> (2 :: 'a :: len word) ^ n \<le> 2 ^ m"
1665  by (simp add: word_le_nat_alt)
1666
1667lemma is_aligned_add_less_t2n:
1668  "\<lbrakk>is_aligned (p::'a::len word) n; d < 2^n; n \<le> m; p < 2^m\<rbrakk> \<Longrightarrow> p + d < 2^m"
1669  apply (case_tac "m < LENGTH('a)")
1670   apply (subst mask_eq_iff_w2p[symmetric])
1671    apply (simp add: word_size)
1672   apply (simp add: is_aligned_add_or word_ao_dist less_mask_eq)
1673   apply (subst less_mask_eq)
1674    apply (erule order_less_le_trans)
1675    apply (erule(1) two_power_increasing)
1676   apply simp
1677  apply (simp add: power_overflow)
1678  done
1679
1680
1681lemma aligned_offset_non_zero:
1682  "\<lbrakk> is_aligned x n; y < 2 ^ n; x \<noteq> 0 \<rbrakk> \<Longrightarrow> x + y \<noteq> 0"
1683  apply (cases "y = 0")
1684   apply simp
1685  apply (subst word_neq_0_conv)
1686  apply (subst gt0_iff_gem1)
1687  apply (erule is_aligned_get_word_bits)
1688   apply (subst field_simps[symmetric], subst plus_le_left_cancel_nowrap)
1689     apply (rule is_aligned_no_wrap')
1690      apply simp
1691     apply (rule word_leq_le_minus_one)
1692      apply simp
1693     apply assumption
1694    apply (erule (1) is_aligned_no_wrap')
1695   apply (simp add: gt0_iff_gem1 [symmetric] word_neq_0_conv)
1696  apply simp
1697  done
1698
1699lemmas mask_inner_mask = mask_eqs(1)
1700
1701lemma mask_add_aligned:
1702  "is_aligned p n \<Longrightarrow> (p + q) && mask n = q && mask n"
1703  apply (simp add: is_aligned_mask)
1704  apply (subst mask_inner_mask [symmetric])
1705  apply simp
1706  done
1707
1708lemma take_prefix:
1709  "(take (length xs) ys = xs) = prefix xs ys"
1710proof (induct xs arbitrary: ys)
1711  case Nil then show ?case by simp
1712next
1713  case Cons then show ?case by (cases ys) auto
1714qed
1715
1716lemma cart_singleton_empty:
1717  "(S \<times> {e} = {}) = (S = {})"
1718  by blast
1719
1720lemma word_div_1:
1721  "(n :: 'a :: len word) div 1 = n"
1722  by (simp add: word_div_def)
1723
1724lemma word_minus_one_le:
1725  "-1 \<le> (x :: 'a :: len word) = (x = -1)"
1726  apply (insert word_n1_ge[where y=x])
1727  apply safe
1728  apply (erule(1) order_antisym)
1729  done
1730
1731lemma mask_out_sub_mask:
1732  "(x && ~~ (mask n)) = x - (x && (mask n))"
1733  by (simp add: field_simps word_plus_and_or_coroll2)
1734
1735lemma is_aligned_addD1:
1736  assumes al1: "is_aligned (x + y) n"
1737  and     al2: "is_aligned (x::'a::len word) n"
1738  shows "is_aligned y n"
1739  using al2
1740proof (rule is_aligned_get_word_bits)
1741  assume "x = 0" then show ?thesis using al1 by simp
1742next
1743  assume nv: "n < LENGTH('a)"
1744  from al1 obtain q1
1745    where xy: "x + y = 2 ^ n * of_nat q1" and "q1 < 2 ^ (LENGTH('a) - n)"
1746    by (rule is_alignedE)
1747  moreover from al2 obtain q2
1748    where x: "x = 2 ^ n * of_nat q2" and "q2 < 2 ^ (LENGTH('a) - n)"
1749    by (rule is_alignedE)
1750  ultimately have "y = 2 ^ n * (of_nat q1 - of_nat q2)"
1751    by (simp add: field_simps)
1752  then show ?thesis using nv by (simp add: is_aligned_mult_triv1)
1753qed
1754
1755lemmas is_aligned_addD2 =
1756       is_aligned_addD1[OF subst[OF add.commute,
1757                                 of "%x. is_aligned x n" for n]]
1758
1759lemma is_aligned_add:
1760  "\<lbrakk>is_aligned p n; is_aligned q n\<rbrakk> \<Longrightarrow> is_aligned (p + q) n"
1761  by (simp add: is_aligned_mask mask_add_aligned)
1762
1763lemma word_le_add:
1764  fixes x :: "'a :: len word"
1765  shows "x \<le> y \<Longrightarrow> \<exists>n. y = x + of_nat n"
1766  by (rule exI [where x = "unat (y - x)"]) simp
1767
1768lemma word_plus_mcs_4':
1769  fixes x :: "'a :: len word"
1770  shows "\<lbrakk>x + v \<le> x + w; x \<le> x + v\<rbrakk> \<Longrightarrow> v \<le> w"
1771  apply (rule word_plus_mcs_4)
1772   apply (simp add: add.commute)
1773  apply (simp add: add.commute)
1774  done
1775
1776lemma shiftl_mask_is_0[simp]:
1777  "(x << n) && mask n = 0"
1778  apply (rule iffD1 [OF is_aligned_mask])
1779  apply (rule is_aligned_shiftl_self)
1780  done
1781
1782definition
1783  sum_map :: "('a \<Rightarrow> 'b) \<Rightarrow> ('c \<Rightarrow> 'd) \<Rightarrow> 'a + 'c \<Rightarrow> 'b + 'd" where
1784 "sum_map f g x \<equiv> case x of Inl v \<Rightarrow> Inl (f v) | Inr v' \<Rightarrow> Inr (g v')"
1785
1786lemma sum_map_simps[simp]:
1787  "sum_map f g (Inl v) = Inl (f v)"
1788  "sum_map f g (Inr w) = Inr (g w)"
1789  by (simp add: sum_map_def)+
1790
1791lemma if_and_helper:
1792  "(If x v v') && v'' = If x (v && v'') (v' && v'')"
1793  by (rule if_distrib)
1794
1795lemma unat_Suc2:
1796  fixes n :: "'a :: len word"
1797  shows
1798  "n \<noteq> -1 \<Longrightarrow> unat (n + 1) = Suc (unat n)"
1799  apply (subst add.commute, rule unatSuc)
1800  apply (subst eq_diff_eq[symmetric], simp add: minus_equation_iff)
1801  done
1802
1803lemmas word_unat_Rep_inject1 = word_unat.Rep_inject[where y=1]
1804lemmas unat_eq_1 = unat_eq_0 word_unat_Rep_inject1[simplified]
1805
1806
1807lemma rshift_sub_mask_eq:
1808  "(a >> (size a - b)) && mask b = a >> (size a - b)"
1809  using shiftl_shiftr2[where a=a and b=0 and c="size a - b"]
1810  apply (cases "b < size a")
1811   apply simp
1812  apply (simp add: linorder_not_less mask_def word_size
1813                   p2_eq_0[THEN iffD2])
1814  done
1815
1816lemma shiftl_shiftr3:
1817  "b \<le> c \<Longrightarrow> a << b >> c = (a >> c - b) && mask (size a - c)"
1818  apply (cases "b = c")
1819   apply (simp add: shiftl_shiftr1)
1820  apply (simp add: shiftl_shiftr2)
1821  done
1822
1823lemma and_mask_shiftr_comm:
1824  "m\<le>size w \<Longrightarrow> (w && mask m) >> n = (w >> n) && mask (m-n)"
1825  by (simp add: and_mask shiftr_shiftr) (simp add: word_size shiftl_shiftr3)
1826
1827lemma and_mask_shiftl_comm:
1828  "m+n \<le> size w \<Longrightarrow> (w && mask m) << n = (w << n) && mask (m+n)"
1829  by (simp add: and_mask word_size shiftl_shiftl) (simp add: shiftl_shiftr1)
1830
1831lemma le_mask_shiftl_le_mask: "s = m + n \<Longrightarrow> x \<le> mask n \<Longrightarrow> x << m \<le> mask s"
1832  by (simp add: le_mask_iff shiftl_shiftr3)
1833
1834lemma and_not_mask_twice:
1835  "(w && ~~ (mask n)) && ~~ (mask m) = w && ~~ (mask (max m n))"
1836  apply (simp add: and_not_mask)
1837  apply (case_tac "n<m";
1838         simp add: shiftl_shiftr2 shiftl_shiftr1 not_less max_def shiftr_shiftr shiftl_shiftl)
1839   apply (cut_tac and_mask_shiftr_comm [where w=w and m="size w" and n=m, simplified,symmetric])
1840   apply (simp add: word_size mask_def)
1841  apply (cut_tac and_mask_shiftr_comm [where w=w and m="size w" and n=n, simplified,symmetric])
1842  apply (simp add: word_size mask_def)
1843  done
1844
1845lemma word_less_cases:
1846  "x < y \<Longrightarrow> x = y - 1 \<or> x < y - (1 ::'a::len word)"
1847  apply (drule word_less_sub_1)
1848  apply (drule order_le_imp_less_or_eq)
1849  apply auto
1850  done
1851
1852lemma eq_eqI:
1853  "a = b \<Longrightarrow> (a = x) = (b = x)"
1854  by simp
1855
1856lemma mask_and_mask:
1857  "mask a && mask b = mask (min a b)"
1858  by word_eqI
1859
1860lemma mask_eq_0_eq_x:
1861  "(x && w = 0) = (x && ~~ w = x)"
1862  using word_plus_and_or_coroll2[where x=x and w=w]
1863  by auto
1864
1865lemma mask_eq_x_eq_0:
1866  "(x && w = x) = (x && ~~ w = 0)"
1867  using word_plus_and_or_coroll2[where x=x and w=w]
1868  by auto
1869
1870definition
1871  "limited_and (x :: 'a :: len word) y = (x && y = x)"
1872
1873lemma limited_and_eq_0:
1874  "\<lbrakk> limited_and x z; y && ~~ z = y \<rbrakk> \<Longrightarrow> x && y = 0"
1875  unfolding limited_and_def
1876  apply (subst arg_cong2[where f="(&&)"])
1877    apply (erule sym)+
1878  apply (simp(no_asm) add: word_bw_assocs word_bw_comms word_bw_lcs)
1879  done
1880
1881lemma limited_and_eq_id:
1882  "\<lbrakk> limited_and x z; y && z = z \<rbrakk> \<Longrightarrow> x && y = x"
1883  unfolding limited_and_def
1884  by (erule subst, fastforce simp: word_bw_lcs word_bw_assocs word_bw_comms)
1885
1886lemma lshift_limited_and:
1887  "limited_and x z \<Longrightarrow> limited_and (x << n) (z << n)"
1888  unfolding limited_and_def
1889  by (simp add: shiftl_over_and_dist[symmetric])
1890
1891lemma rshift_limited_and:
1892  "limited_and x z \<Longrightarrow> limited_and (x >> n) (z >> n)"
1893  unfolding limited_and_def
1894  by (simp add: shiftr_over_and_dist[symmetric])
1895
1896lemmas limited_and_simps1 = limited_and_eq_0 limited_and_eq_id
1897
1898lemmas is_aligned_limited_and
1899    = is_aligned_neg_mask_eq[unfolded mask_def, folded limited_and_def]
1900
1901lemma compl_of_1: "~~ 1 = (-2 :: 'a :: len word)"
1902  by simp
1903
1904lemmas limited_and_simps = limited_and_simps1
1905       limited_and_simps1[OF is_aligned_limited_and]
1906       limited_and_simps1[OF lshift_limited_and]
1907       limited_and_simps1[OF rshift_limited_and]
1908       limited_and_simps1[OF rshift_limited_and, OF is_aligned_limited_and]
1909       compl_of_1 shiftl_shiftr1[unfolded word_size mask_def]
1910       shiftl_shiftr2[unfolded word_size mask_def]
1911
1912lemma split_word_eq_on_mask:
1913  "(x = y) = (x && m = y && m \<and> x && ~~ m = y && ~~ m)"
1914  by safe word_eqI_solve
1915
1916lemma map2_Cons_2_3:
1917  "(map2 f xs (y # ys) = (z # zs)) = (\<exists>x xs'. xs = x # xs' \<and> f x y = z \<and> map2 f xs' ys = zs)"
1918  by (case_tac xs, simp_all)
1919
1920lemma map2_xor_replicate_False:
1921  "map2 (\<lambda>x y. x \<longleftrightarrow> \<not> y) xs (replicate n False) = take n xs"
1922  apply (induct xs arbitrary: n, simp)
1923  apply (case_tac n; simp)
1924  done
1925
1926lemma word_and_1_shiftl:
1927  "x && (1 << n) = (if x !! n then (1 << n) else 0)" for x :: "'a :: len word"
1928  by word_eqI_solve
1929
1930lemmas word_and_1_shiftls'
1931    = word_and_1_shiftl[where n=0]
1932      word_and_1_shiftl[where n=1]
1933      word_and_1_shiftl[where n=2]
1934
1935lemmas word_and_1_shiftls = word_and_1_shiftls' [simplified]
1936
1937lemma word_and_mask_shiftl:
1938  "x && (mask n << m) = ((x >> m) && mask n) << m"
1939  by word_eqI_solve
1940
1941lemma plus_Collect_helper:
1942  "(+) x ` {xa. P (xa :: 'a :: len word)} = {xa. P (xa - x)}"
1943  by (fastforce simp add: image_def)
1944
1945lemma plus_Collect_helper2:
1946  "(+) (- x) ` {xa. P (xa :: 'a :: len word)} = {xa. P (x + xa)}"
1947  using plus_Collect_helper [of "- x" P] by (simp add: ac_simps)
1948
1949lemma word_FF_is_mask:
1950  "0xFF = mask 8"
1951  by (simp add: mask_def)
1952
1953lemma word_1FF_is_mask:
1954  "0x1FF = mask 9"
1955  by (simp add: mask_def)
1956
1957lemma ucast_of_nat_small:
1958  "x < 2 ^ LENGTH('a) \<Longrightarrow> ucast (of_nat x :: 'a :: len word) = (of_nat x :: 'b :: len word)"
1959  apply (rule sym, subst word_unat.inverse_norm)
1960  apply (simp add: ucast_def word_of_int[symmetric]
1961                   of_nat_nat[symmetric] unat_def[symmetric])
1962  apply (simp add: unat_of_nat)
1963  done
1964
1965lemma word_le_make_less:
1966  fixes x :: "'a :: len word"
1967  shows "y \<noteq> -1 \<Longrightarrow> (x \<le> y) = (x < (y + 1))"
1968  apply safe
1969  apply (erule plus_one_helper2)
1970  apply (simp add: eq_diff_eq[symmetric])
1971  done
1972
1973lemmas finite_word = finite [where 'a="'a::len word"]
1974
1975lemma word_to_1_set:
1976  "{0 ..< (1 :: 'a :: len word)} = {0}"
1977  by fastforce
1978
1979lemma range_subset_eq2:
1980  "{a :: 'a :: len word .. b} \<noteq> {} \<Longrightarrow> ({a .. b} \<subseteq> {c .. d}) = (c \<le> a \<and> b \<le> d)"
1981  by simp
1982
1983lemma word_leq_minus_one_le:
1984  fixes x :: "'a::len word"
1985  shows "\<lbrakk>y \<noteq> 0; x \<le> y - 1 \<rbrakk> \<Longrightarrow> x < y"
1986  using le_m1_iff_lt word_neq_0_conv by blast
1987
1988lemma word_count_from_top:
1989  "n \<noteq> 0 \<Longrightarrow> {0 ..< n :: 'a :: len word} = {0 ..< n - 1} \<union> {n - 1}"
1990  apply (rule set_eqI, rule iffI)
1991   apply simp
1992   apply (drule word_le_minus_one_leq)
1993   apply (rule disjCI)
1994   apply simp
1995  apply simp
1996  apply (erule word_leq_minus_one_le)
1997  apply fastforce
1998  done
1999
2000lemma word_minus_one_le_leq:
2001  "\<lbrakk> x - 1 < y \<rbrakk> \<Longrightarrow> x \<le> (y :: 'a :: len word)"
2002  apply (cases "x = 0")
2003   apply simp
2004  apply (simp add: word_less_nat_alt word_le_nat_alt)
2005  apply (subst(asm) unat_minus_one)
2006   apply (simp add: word_less_nat_alt)
2007  apply (cases "unat x")
2008   apply (simp add: unat_eq_zero)
2009  apply arith
2010  done
2011
2012lemma mod_mod_power:
2013  fixes k :: nat
2014  shows "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ (min m n)"
2015proof (cases "m \<le> n")
2016  case True
2017
2018  then have "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ m"
2019    apply -
2020    apply (subst mod_less [where n = "2 ^ n"])
2021    apply (rule order_less_le_trans [OF mod_less_divisor])
2022    apply simp+
2023    done
2024  also have "\<dots> = k mod  2 ^ (min m n)" using True by simp
2025  finally show ?thesis .
2026next
2027  case False
2028  then have "n < m" by simp
2029  then obtain d where md: "m = n + d"
2030    by (auto dest: less_imp_add_positive)
2031  then have "k mod 2 ^ m = 2 ^ n * (k div 2 ^ n mod 2 ^ d) + k mod 2 ^ n"
2032    by (simp add: mod_mult2_eq power_add)
2033  then have "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ n"
2034    by (simp add: mod_add_left_eq)
2035  then show ?thesis using False
2036    by simp
2037qed
2038
2039lemma word_div_less:
2040  "m < n \<Longrightarrow> m div n = 0" for m :: "'a :: len word"
2041  by (simp add: unat_mono word_arith_nat_defs(6))
2042
2043lemma word_must_wrap:
2044  "\<lbrakk> x \<le> n - 1; n \<le> x \<rbrakk> \<Longrightarrow> n = (0 :: 'a :: len word)"
2045  using dual_order.trans sub_wrap word_less_1 by blast
2046
2047lemma range_subset_card:
2048  "\<lbrakk> {a :: 'a :: len word .. b} \<subseteq> {c .. d}; b \<ge> a \<rbrakk> \<Longrightarrow> d \<ge> c \<and> d - c \<ge> b - a"
2049  using word_sub_le word_sub_mono by fastforce
2050
2051lemma less_1_simp:
2052  "n - 1 < m = (n \<le> (m :: 'a :: len word) \<and> n \<noteq> 0)"
2053  by unat_arith
2054
2055lemma alignUp_div_helper:
2056  fixes a :: "'a::len word"
2057  assumes kv: "k < 2 ^ (LENGTH('a) - n)"
2058  and     xk: "x = 2 ^ n * of_nat k"
2059  and    le: "a \<le> x"
2060  and    sz: "n < LENGTH('a)"
2061  and   anz: "a mod 2 ^ n \<noteq> 0"
2062  shows "a div 2 ^ n < of_nat k"
2063proof -
2064  have kn: "unat (of_nat k :: 'a word) * unat ((2::'a word) ^ n) < 2 ^ LENGTH('a)"
2065    using xk kv sz
2066    apply (subst unat_of_nat_eq)
2067     apply (erule order_less_le_trans)
2068     apply simp
2069    apply (subst unat_power_lower, simp)
2070    apply (subst mult.commute)
2071    apply (rule nat_less_power_trans)
2072     apply simp
2073    apply simp
2074    done
2075
2076  have "unat a div 2 ^ n * 2 ^ n \<noteq> unat a"
2077  proof -
2078    have "unat a = unat a div 2 ^ n * 2 ^ n + unat a mod 2 ^ n"
2079      by (simp add: div_mult_mod_eq)
2080    also have "\<dots> \<noteq> unat a div 2 ^ n * 2 ^ n" using sz anz
2081      by (simp add: unat_arith_simps)
2082    finally show ?thesis ..
2083  qed
2084
2085  then have "a div 2 ^ n * 2 ^ n < a" using sz anz
2086    apply (subst word_less_nat_alt)
2087    apply (subst unat_word_ariths)
2088    apply (subst unat_div)
2089    apply simp
2090    apply (rule order_le_less_trans [OF mod_less_eq_dividend])
2091    apply (erule order_le_neq_trans [OF div_mult_le])
2092    done
2093
2094  also from xk le have "\<dots> \<le> of_nat k * 2 ^ n" by (simp add: field_simps)
2095  finally show ?thesis using sz kv
2096    apply -
2097    apply (erule word_mult_less_dest [OF _ _ kn])
2098    apply (simp add: unat_div)
2099    apply (rule order_le_less_trans [OF div_mult_le])
2100    apply (rule unat_lt2p)
2101    done
2102qed
2103
2104lemma nat_mod_power_lem:
2105  fixes a :: nat
2106  shows "1 < a \<Longrightarrow> a ^ n mod a ^ m = (if m \<le> n then 0 else a ^ n)"
2107  apply (clarsimp)
2108  apply (clarsimp simp add: le_iff_add power_add)
2109  done
2110
2111lemma power_mod_div:
2112  fixes x :: "nat"
2113  shows "x mod 2 ^ n div 2 ^ m = x div 2 ^ m mod 2 ^ (n - m)" (is "?LHS = ?RHS")
2114proof (cases "n \<le> m")
2115  case True
2116  then have "?LHS = 0"
2117    apply -
2118    apply (rule div_less)
2119    apply (rule order_less_le_trans [OF mod_less_divisor]; simp)
2120    done
2121  also have "\<dots> = ?RHS" using True
2122    by simp
2123  finally show ?thesis .
2124next
2125  case False
2126  then have lt: "m < n" by simp
2127  then obtain q where nv: "n = m + q" and "0 < q"
2128    by (auto dest: less_imp_Suc_add)
2129
2130  then have "x mod 2 ^ n = 2 ^ m * (x div 2 ^ m mod 2 ^ q) + x mod 2 ^ m"
2131    by (simp add: power_add mod_mult2_eq)
2132
2133  then have "?LHS = x div 2 ^ m mod 2 ^ q"
2134    by (simp add: div_add1_eq)
2135
2136  also have "\<dots> = ?RHS" using nv
2137    by simp
2138
2139  finally show ?thesis .
2140qed
2141
2142lemma word_power_mod_div:
2143  fixes x :: "'a::len word"
2144  shows "\<lbrakk> n < LENGTH('a); m < LENGTH('a)\<rbrakk>
2145  \<Longrightarrow> x mod 2 ^ n div 2 ^ m = x div 2 ^ m mod 2 ^ (n - m)"
2146  apply (simp add: word_arith_nat_div unat_mod power_mod_div)
2147  apply (subst unat_arith_simps(3))
2148  apply (subst unat_mod)
2149  apply (subst unat_of_nat)+
2150  apply (simp add: mod_mod_power min.commute)
2151  done
2152
2153lemma word_range_minus_1':
2154  fixes a :: "'a :: len word"
2155  shows "a \<noteq> 0 \<Longrightarrow> {a - 1<..b} = {a..b}"
2156  by (simp add: greaterThanAtMost_def atLeastAtMost_def greaterThan_def atLeast_def less_1_simp)
2157
2158lemma word_range_minus_1:
2159  fixes a :: "'a :: len word"
2160  shows "b \<noteq> 0 \<Longrightarrow> {a..b - 1} = {a..<b}"
2161  apply (simp add: atLeastLessThan_def atLeastAtMost_def atMost_def lessThan_def)
2162  apply (rule arg_cong [where f = "\<lambda>x. {a..} \<inter> x"])
2163  apply rule
2164   apply clarsimp
2165   apply (erule contrapos_pp)
2166   apply (simp add: linorder_not_less linorder_not_le word_must_wrap)
2167  apply (clarsimp)
2168  apply (drule word_le_minus_one_leq)
2169  apply (auto simp: word_less_sub_1)
2170  done
2171
2172lemma ucast_nat_def:
2173  "of_nat (unat x) = (ucast :: 'a :: len word \<Rightarrow> 'b :: len word) x"
2174  by (simp add: ucast_def word_of_int_nat unat_def)
2175
2176lemmas if_fun_split = if_apply_def2
2177
2178lemma i_hate_words_helper:
2179  "i \<le> (j - k :: nat) \<Longrightarrow> i \<le> j"
2180  by simp
2181
2182lemma i_hate_words:
2183  "unat (a :: 'a word) \<le> unat (b :: 'a :: len word) - Suc 0
2184    \<Longrightarrow> a \<noteq> -1"
2185  apply (frule i_hate_words_helper)
2186  apply (subst(asm) word_le_nat_alt[symmetric])
2187  apply (clarsimp simp only: word_minus_one_le)
2188  apply (simp only: linorder_not_less[symmetric])
2189  apply (erule notE)
2190  apply (rule diff_Suc_less)
2191  apply (subst neq0_conv[symmetric])
2192  apply (subst unat_eq_0)
2193  apply (rule notI, drule arg_cong[where f="(+) 1"])
2194  apply simp
2195  done
2196
2197lemma overflow_plus_one_self:
2198  "(1 + p \<le> p) = (p = (-1 :: 'a :: len word))"
2199  apply (safe, simp_all)
2200  apply (rule ccontr)
2201  apply (drule plus_one_helper2)
2202   apply (rule notI)
2203   apply (drule arg_cong[where f="\<lambda>x. x - 1"])
2204   apply simp
2205  apply (simp add: field_simps)
2206  done
2207
2208lemma plus_1_less:
2209  "(x + 1 \<le> (x :: 'a :: len word)) = (x = -1)"
2210  apply (rule iffI)
2211   apply (rule ccontr)
2212   apply (cut_tac plus_one_helper2[where x=x, OF order_refl])
2213    apply simp
2214   apply clarsimp
2215   apply (drule arg_cong[where f="\<lambda>x. x - 1"])
2216   apply simp
2217  apply simp
2218  done
2219
2220lemma pos_mult_pos_ge:
2221  "[|x > (0::int); n>=0 |] ==> n * x >= n*1"
2222  apply (simp only: mult_left_mono)
2223  done
2224
2225lemma If_eq_obvious:
2226  "x \<noteq> z \<Longrightarrow> ((if P then x else y) = z) = (\<not> P \<and> y = z)"
2227  by simp
2228
2229lemma Some_to_the:
2230  "v = Some x \<Longrightarrow> x = the v"
2231  by simp
2232
2233lemma dom_if_Some:
2234  "dom (\<lambda>x. if P x then Some (f x) else g x) = {x. P x} \<union> dom g"
2235  by fastforce
2236
2237lemma dom_insert_absorb:
2238  "x \<in> dom f \<Longrightarrow> insert x (dom f) = dom f" by auto
2239
2240lemma emptyE2:
2241  "\<lbrakk> S = {}; x \<in> S \<rbrakk> \<Longrightarrow> P"
2242  by simp
2243
2244lemma mod_div_equality_div_eq:
2245  "a div b * b = (a - (a mod b) :: int)"
2246  by (simp add: field_simps)
2247
2248lemma zmod_helper:
2249  "n mod m = k \<Longrightarrow> ((n :: int) + a) mod m = (k + a) mod m"
2250  by (metis add.commute mod_add_right_eq)
2251
2252lemma int_div_sub_1:
2253  "\<lbrakk> m \<ge> 1 \<rbrakk> \<Longrightarrow> (n - (1 :: int)) div m = (if m dvd n then (n div m) - 1 else n div m)"
2254  apply (subgoal_tac "m = 0 \<or> (n - (1 :: int)) div m = (if m dvd n then (n div m) - 1 else n div m)")
2255   apply fastforce
2256  apply (subst mult_cancel_right[symmetric])
2257  apply (simp only: left_diff_distrib split: if_split)
2258  apply (simp only: mod_div_equality_div_eq)
2259  apply (clarsimp simp: field_simps)
2260  apply (clarsimp simp: dvd_eq_mod_eq_0)
2261  apply (cases "m = 1")
2262   apply simp
2263  apply (subst mod_diff_eq[symmetric], simp add: zmod_minus1)
2264  apply clarsimp
2265  apply (subst diff_add_cancel[where b=1, symmetric])
2266  apply (subst mod_add_eq[symmetric])
2267  apply (simp add: field_simps)
2268  apply (rule mod_pos_pos_trivial)
2269   apply (subst add_0_right[where a=0, symmetric])
2270   apply (rule add_mono)
2271    apply simp
2272   apply simp
2273  apply (cases "(n - 1) mod m = m - 1")
2274   apply (drule zmod_helper[where a=1])
2275   apply simp
2276  apply (subgoal_tac "1 + (n - 1) mod m \<le> m")
2277   apply simp
2278  apply (subst field_simps, rule zless_imp_add1_zle)
2279  apply simp
2280  done
2281
2282lemma ptr_add_image_multI:
2283  "\<lbrakk> \<And>x y. (x * val = y * val') = (x * val'' = y); x * val'' \<in> S \<rbrakk> \<Longrightarrow>
2284     ptr_add ptr (x * val) \<in> (\<lambda>p. ptr_add ptr (p * val')) ` S"
2285  apply (simp add: image_def)
2286  apply (erule rev_bexI)
2287  apply (rule arg_cong[where f="ptr_add ptr"])
2288  apply simp
2289  done
2290
2291lemma shift_times_fold:
2292  "(x :: 'a :: len word) * (2 ^ n) << m = x << (m + n)"
2293  by (simp add: shiftl_t2n ac_simps power_add)
2294
2295lemma word_plus_strict_mono_right:
2296  fixes x :: "'a :: len word"
2297  shows "\<lbrakk>y < z; x \<le> x + z\<rbrakk> \<Longrightarrow> x + y < x + z"
2298  by unat_arith
2299
2300lemma replicate_minus:
2301  "k < n \<Longrightarrow> replicate n False = replicate (n - k) False @ replicate k False"
2302  by (subst replicate_add [symmetric]) simp
2303
2304lemmas map_prod_split_imageI'
2305  = map_prod_imageI[where f="case_prod f" and g="case_prod g"
2306                    and a="(a, b)" and b="(c, d)" for a b c d f g]
2307lemmas map_prod_split_imageI = map_prod_split_imageI'[simplified]
2308
2309lemma word_div_mult:
2310  "0 < c \<Longrightarrow> a < b * c \<Longrightarrow> a div c < b" for a b c :: "'a::len word"
2311  by (rule classical)
2312     (use div_to_mult_word_lt [of b a c] in
2313      \<open>auto simp add: word_less_nat_alt word_le_nat_alt unat_div\<close>)
2314
2315lemma word_less_power_trans_ofnat:
2316  "\<lbrakk>n < 2 ^ (m - k); k \<le> m; m < LENGTH('a)\<rbrakk>
2317   \<Longrightarrow> of_nat n * 2 ^ k < (2::'a::len word) ^ m"
2318  apply (subst mult.commute)
2319  apply (rule word_less_power_trans)
2320    apply (simp add: word_less_nat_alt)
2321    apply (subst unat_of_nat_eq)
2322     apply (erule order_less_trans)
2323     apply simp+
2324    done
2325
2326lemma word_1_le_power:
2327  "n < LENGTH('a) \<Longrightarrow> (1 :: 'a :: len word) \<le> 2 ^ n"
2328  by (rule inc_le[where i=0, simplified], erule iffD2[OF p2_gt_0])
2329
2330lemma enum_word_div:
2331  fixes v :: "'a :: len word" shows
2332  "\<exists>xs ys. enum = xs @ [v] @ ys
2333             \<and> (\<forall>x \<in> set xs. x < v)
2334             \<and> (\<forall>y \<in> set ys. v < y)"
2335  apply (simp only: enum_word_def)
2336  apply (subst upt_add_eq_append'[where j="unat v"])
2337    apply simp
2338   apply (rule order_less_imp_le, simp)
2339  apply (simp add: upt_conv_Cons)
2340  apply (intro exI conjI)
2341    apply fastforce
2342   apply clarsimp
2343   apply (drule of_nat_mono_maybe[rotated, where 'a='a])
2344    apply simp
2345   apply simp
2346  apply (clarsimp simp: Suc_le_eq)
2347  apply (drule of_nat_mono_maybe[rotated, where 'a='a])
2348   apply simp
2349  apply simp
2350  done
2351
2352lemma of_bool_nth:
2353  "of_bool (x !! v) = (x >> v) && 1"
2354  by (word_eqI cong: rev_conj_cong)
2355
2356lemma unat_1_0:
2357  "1 \<le> (x::'a::len word) = (0 < unat x)"
2358  by (auto simp add: word_le_nat_alt)
2359
2360lemma x_less_2_0_1':
2361  fixes x :: "'a::len word"
2362  shows "\<lbrakk>LENGTH('a) \<noteq> 1; x < 2\<rbrakk> \<Longrightarrow> x = 0 \<or> x = 1"
2363  apply (induct x)
2364   apply clarsimp+
2365  by (metis Suc_eq_plus1 add_lessD1 less_irrefl one_add_one unatSuc word_less_nat_alt)
2366
2367lemmas word_add_le_iff2 = word_add_le_iff [folded no_olen_add_nat]
2368
2369lemma of_nat_power:
2370  shows "\<lbrakk> p < 2 ^ x; x < len_of TYPE ('a) \<rbrakk> \<Longrightarrow> of_nat p < (2 :: 'a :: len word) ^ x"
2371  apply (rule order_less_le_trans)
2372   apply (rule of_nat_mono_maybe)
2373    apply (erule power_strict_increasing)
2374    apply simp
2375   apply assumption
2376  apply (simp add: word_unat_power)
2377  done
2378
2379lemma of_nat_n_less_equal_power_2:
2380  "n < LENGTH('a::len) \<Longrightarrow> ((of_nat n)::'a word) < 2 ^ n"
2381  apply (induct n)
2382   apply clarsimp
2383  apply clarsimp
2384  apply (metis of_nat_power n_less_equal_power_2 of_nat_Suc power_Suc)
2385  done
2386
2387lemma eq_mask_less:
2388  fixes w :: "'a::len word"
2389  assumes eqm: "w = w && mask n"
2390  and      sz: "n < len_of TYPE ('a)"
2391  shows "w < (2::'a word) ^ n"
2392  by (subst eqm, rule and_mask_less' [OF sz])
2393
2394lemma of_nat_mono_maybe':
2395  fixes Y :: "nat"
2396  assumes xlt: "x < 2 ^ len_of TYPE ('a)"
2397  assumes ylt: "y < 2 ^ len_of TYPE ('a)"
2398  shows   "(y < x) = (of_nat y < (of_nat x :: 'a :: len word))"
2399  apply (subst word_less_nat_alt)
2400  apply (subst unat_of_nat)+
2401  apply (subst mod_less)
2402   apply (rule ylt)
2403  apply (subst mod_less)
2404   apply (rule xlt)
2405  apply simp
2406  done
2407
2408lemma shiftr_mask_eq:
2409  "(x >> n) && mask (size x - n) = x >> n" for x :: "'a :: len word"
2410  by word_eqI_solve
2411
2412lemma shiftr_mask_eq':
2413  "m = (size x - n) \<Longrightarrow> (x >> n) && mask m = x >> n" for x :: "'a :: len word"
2414  by (simp add: shiftr_mask_eq)
2415
2416lemma dom_if:
2417  "dom (\<lambda>a. if a \<in> addrs then Some (f a) else g a)  = addrs \<union> dom g"
2418  by (auto simp: dom_def split: if_split)
2419
2420lemma less_is_non_zero_p1:
2421  fixes a :: "'a :: len word"
2422  shows "a < k \<Longrightarrow> a + 1 \<noteq> 0"
2423  apply (erule contrapos_pn)
2424  apply (drule max_word_wrap)
2425  apply (simp add: not_less)
2426  done
2427
2428lemma of_nat_mono_maybe_le:
2429  "\<lbrakk>x < 2 ^ LENGTH('a); y < 2 ^ LENGTH('a)\<rbrakk> \<Longrightarrow>
2430  (y \<le> x) = ((of_nat y :: 'a :: len word) \<le> of_nat x)"
2431  apply (clarsimp simp: le_less)
2432  apply (rule disj_cong)
2433   apply (rule of_nat_mono_maybe', assumption+)
2434  apply (simp add: word_unat.norm_eq_iff [symmetric])
2435  done
2436
2437lemma mask_AND_NOT_mask:
2438  "(w && ~~ (mask n)) && mask n = 0"
2439  by word_eqI
2440
2441lemma AND_NOT_mask_plus_AND_mask_eq:
2442  "(w && ~~ (mask n)) + (w && mask n) = w"
2443  by (subst word_plus_and_or_coroll; word_eqI_solve)
2444
2445lemma mask_eqI:
2446  fixes x :: "'a :: len word"
2447  assumes m1: "x && mask n = y && mask n"
2448  and     m2: "x && ~~ (mask n) = y && ~~ (mask n)"
2449  shows "x = y"
2450proof (subst bang_eq, rule allI)
2451  fix m
2452
2453  show "x !! m = y !! m"
2454  proof (cases "m < n")
2455    case True
2456    then have "x !! m = ((x && mask n) !! m)"
2457      by (simp add: word_size test_bit_conj_lt)
2458    also have "\<dots> = ((y && mask n) !! m)" using m1 by simp
2459    also have "\<dots> = y !! m" using True
2460      by (simp add: word_size test_bit_conj_lt)
2461    finally show ?thesis .
2462  next
2463    case False
2464    then have "x !! m = ((x && ~~ (mask n)) !! m)"
2465      by (simp add: neg_mask_test_bit test_bit_conj_lt)
2466    also have "\<dots> = ((y && ~~ (mask n)) !! m)" using m2 by simp
2467    also have "\<dots> = y !! m" using False
2468      by (simp add: neg_mask_test_bit test_bit_conj_lt)
2469    finally show ?thesis .
2470  qed
2471qed
2472
2473lemma nat_less_power_trans2:
2474  fixes n :: nat
2475  shows "\<lbrakk>n < 2 ^ (m - k); k \<le> m\<rbrakk> \<Longrightarrow> n * 2 ^ k  < 2 ^ m"
2476  by (subst mult.commute, erule (1) nat_less_power_trans)
2477
2478lemma nat_move_sub_le: "(a::nat) + b \<le> c \<Longrightarrow> a \<le> c - b" by arith
2479
2480lemma neq_0_no_wrap:
2481  fixes x :: "'a :: len word"
2482  shows "\<lbrakk> x \<le> x + y; x \<noteq> 0 \<rbrakk> \<Longrightarrow> x + y \<noteq> 0"
2483  by clarsimp
2484
2485lemma plus_minus_one_rewrite:
2486  "v + (- 1 :: ('a :: {ring, one, uminus})) \<equiv> v - 1"
2487  by (simp add: field_simps)
2488
2489lemma power_minus_is_div:
2490  "b \<le> a \<Longrightarrow> (2 :: nat) ^ (a - b) = 2 ^ a div 2 ^ b"
2491  apply (induct a arbitrary: b)
2492   apply simp
2493  apply (erule le_SucE)
2494   apply (clarsimp simp:Suc_diff_le le_iff_add power_add)
2495  apply simp
2496  done
2497
2498lemma two_pow_div_gt_le:
2499  "v < 2 ^ n div (2 ^ m :: nat) \<Longrightarrow> m \<le> n"
2500  by (clarsimp dest!: less_two_pow_divD)
2501
2502lemma unatSuc2:
2503  fixes n :: "'a :: len word"
2504  shows "n + 1 \<noteq> 0 \<Longrightarrow> unat (n + 1) = Suc (unat n)"
2505  by (simp add: add.commute unatSuc)
2506
2507lemma word_of_nat_less:
2508  "\<lbrakk> n < unat x \<rbrakk> \<Longrightarrow> of_nat n < x"
2509  apply (simp add: word_less_nat_alt)
2510  apply (erule order_le_less_trans[rotated])
2511  apply (simp add: unat_of_nat)
2512  done
2513
2514lemma word_of_nat_le:
2515  "n \<le> unat x \<Longrightarrow> of_nat n \<le> x"
2516  apply (simp add: word_le_nat_alt unat_of_nat)
2517  apply (erule order_trans[rotated])
2518  apply simp
2519  done
2520
2521lemma word_unat_less_le:
2522   "a \<le> of_nat b \<Longrightarrow> unat a \<le> b"
2523   by (metis eq_iff le_cases le_unat_uoi word_of_nat_le)
2524
2525lemma and_eq_0_is_nth:
2526  fixes x :: "'a :: len word"
2527  shows "y = 1 << n \<Longrightarrow> ((x && y) = 0) = (\<not> (x !! n))"
2528  apply safe
2529   apply (drule_tac u="(x && (1 << n))" and x=n in word_eqD)
2530   apply (simp add: nth_w2p)
2531   apply (simp add: test_bit_bin)
2532  apply word_eqI
2533  done
2534
2535lemmas arg_cong_Not = arg_cong [where f=Not]
2536lemmas and_neq_0_is_nth = arg_cong_Not [OF and_eq_0_is_nth, simplified]
2537
2538lemma nth_is_and_neq_0:
2539  "(x::'a::len word) !! n = (x && 2 ^ n \<noteq> 0)"
2540  by (subst and_neq_0_is_nth; rule refl)
2541
2542lemma mask_Suc_0 : "mask (Suc 0) = 1"
2543  by (simp add: mask_def)
2544
2545lemma ucast_ucast_add:
2546  fixes x :: "'a :: len word"
2547  fixes y :: "'b :: len word"
2548  shows
2549  "LENGTH('b) \<ge> LENGTH('a) \<Longrightarrow>
2550    ucast (ucast x + y) = x + ucast y"
2551  apply (rule word_unat.Rep_eqD)
2552  apply (simp add: unat_ucast unat_word_ariths mod_mod_power
2553                   min.absorb2 unat_of_nat)
2554  apply (subst mod_add_left_eq[symmetric])
2555  apply (simp add: mod_mod_power min.absorb2)
2556  apply (subst mod_add_right_eq)
2557  apply simp
2558  done
2559
2560lemma word_shift_zero:
2561  "\<lbrakk> x << n = 0; x \<le> 2^m; m + n < LENGTH('a)\<rbrakk> \<Longrightarrow> (x::'a::len word) = 0"
2562  apply (rule ccontr)
2563  apply (drule (2) word_shift_nonzero)
2564  apply simp
2565  done
2566
2567lemma bool_mask':
2568  fixes x :: "'a :: len word"
2569  shows "2 < LENGTH('a) \<Longrightarrow> (0 < x && 1) = (x && 1 = 1)"
2570  apply (rule iffI)
2571   prefer 2
2572   apply simp
2573  apply (subgoal_tac "x && mask 1 < 2^1")
2574   prefer 2
2575   apply (rule and_mask_less_size)
2576   apply (simp add: word_size)
2577  apply (simp add: mask_def)
2578  apply (drule word_less_cases [where y=2])
2579  apply (erule disjE, simp)
2580  apply simp
2581  done
2582
2583lemma sint_eq_uint:
2584  "\<not> msb x \<Longrightarrow> sint x = uint x"
2585  apply (rule word_uint.Abs_eqD, subst word_sint.Rep_inverse)
2586    apply simp_all
2587  apply (cut_tac x=x in word_sint.Rep)
2588  apply (clarsimp simp add: uints_num sints_num)
2589  apply (rule conjI)
2590   apply (rule ccontr)
2591   apply (simp add: linorder_not_le word_msb_sint[symmetric])
2592  apply (erule order_less_le_trans)
2593  apply simp
2594  done
2595
2596lemma scast_eq_ucast:
2597  "\<not> msb x \<Longrightarrow> scast x = ucast x"
2598  by (simp add: scast_def ucast_def sint_eq_uint)
2599
2600lemma lt1_neq0:
2601  fixes x :: "'a :: len word"
2602  shows "(1 \<le> x) = (x \<noteq> 0)" by unat_arith
2603
2604lemma word_plus_one_nonzero:
2605  fixes x :: "'a :: len word"
2606  shows "\<lbrakk>x \<le> x + y; y \<noteq> 0\<rbrakk> \<Longrightarrow> x + 1 \<noteq> 0"
2607  apply (subst lt1_neq0 [symmetric])
2608  apply (subst olen_add_eqv [symmetric])
2609  apply (erule word_random)
2610  apply (simp add: lt1_neq0)
2611  done
2612
2613lemma word_sub_plus_one_nonzero:
2614  fixes n :: "'a :: len word"
2615  shows "\<lbrakk>n' \<le> n; n' \<noteq> 0\<rbrakk> \<Longrightarrow> (n - n') + 1 \<noteq> 0"
2616  apply (subst lt1_neq0 [symmetric])
2617  apply (subst olen_add_eqv [symmetric])
2618  apply (rule word_random [where x' = n'])
2619   apply simp
2620   apply (erule word_sub_le)
2621  apply (simp add: lt1_neq0)
2622  done
2623
2624lemma word_le_minus_mono_right:
2625  fixes x :: "'a :: len word"
2626  shows "\<lbrakk> z \<le> y; y \<le> x; z \<le> x \<rbrakk> \<Longrightarrow> x - y \<le> x - z"
2627  apply (rule word_sub_mono)
2628     apply simp
2629    apply assumption
2630   apply (erule word_sub_le)
2631  apply (erule word_sub_le)
2632  done
2633
2634lemma drop_append_miracle:
2635  "n = length xs \<Longrightarrow> drop n (xs @ ys) = ys"
2636  by simp
2637
2638lemma foldr_does_nothing_to_xf:
2639  "\<lbrakk> \<And>x s. x \<in> set xs \<Longrightarrow> xf (f x s) = xf s \<rbrakk> \<Longrightarrow> xf (foldr f xs s) = xf s"
2640  by (induct xs, simp_all)
2641
2642lemma nat_less_mult_monoish: "\<lbrakk> a < b; c < (d :: nat) \<rbrakk> \<Longrightarrow> (a + 1) * (c + 1) <= b * d"
2643  apply (drule Suc_leI)+
2644  apply (drule(1) mult_le_mono)
2645  apply simp
2646  done
2647
2648lemma word_0_sle_from_less[unfolded word_size]:
2649  "\<lbrakk> x < 2 ^ (size x - 1) \<rbrakk>  \<Longrightarrow> 0 <=s x"
2650  apply (clarsimp simp: word_sle_msb_le)
2651  apply (simp add: word_msb_nth)
2652  apply (subst (asm) word_test_bit_def [symmetric])
2653  apply (drule less_mask_eq)
2654  apply (drule_tac x="size x - 1" in word_eqD)
2655  apply (simp add: word_size)
2656  done
2657
2658lemma not_msb_from_less:
2659  "(v :: 'a word) < 2 ^ (LENGTH('a :: len) - 1) \<Longrightarrow> \<not> msb v"
2660  apply (clarsimp simp add: msb_nth)
2661  apply (drule less_mask_eq)
2662  apply (drule word_eqD, drule(1) iffD2)
2663  apply simp
2664  done
2665
2666lemma distinct_lemma: "f x \<noteq> f y \<Longrightarrow> x \<noteq> y" by auto
2667
2668
2669lemma ucast_sub_ucast:
2670  fixes x :: "'a::len word"
2671  assumes "y \<le> x"
2672  assumes T: "LENGTH('a) \<le> LENGTH('b)"
2673  shows "ucast (x - y) = (ucast x - ucast y :: 'b::len word)"
2674proof -
2675  from T
2676  have P: "unat x < 2 ^ LENGTH('b)" "unat y < 2 ^ LENGTH('b)"
2677    by (fastforce intro!: less_le_trans[OF unat_lt2p])+
2678  then show ?thesis
2679    by (simp add: unat_arith_simps unat_ucast assms[simplified unat_arith_simps])
2680qed
2681
2682lemma word_1_0:
2683  "\<lbrakk>a + (1::('a::len) word) \<le> b; a < of_nat x\<rbrakk> \<Longrightarrow> a < b"
2684  by unat_arith
2685
2686lemma unat_of_nat_less:"\<lbrakk> a < b; unat b = c \<rbrakk> \<Longrightarrow> a < of_nat c"
2687  by fastforce
2688
2689lemma word_le_plus_1: "\<lbrakk> (y::('a::len) word) < y + n; a < n \<rbrakk> \<Longrightarrow> y + a \<le> y + a + 1"
2690  by unat_arith
2691
2692lemma word_le_plus:"\<lbrakk>(a::('a::len) word) < a + b; c < b\<rbrakk> \<Longrightarrow> a \<le> a + c"
2693by (metis order_less_imp_le word_random)
2694
2695(*
2696 * Basic signed arithemetic properties.
2697 *)
2698
2699lemma sint_minus1 [simp]: "(sint x = -1) = (x = -1)"
2700  by (metis sint_n1 word_sint.Rep_inverse')
2701
2702lemma sint_0 [simp]: "(sint x = 0) = (x = 0)"
2703  by (metis sint_0 word_sint.Rep_inverse')
2704
2705(* It is not always that case that "sint 1 = 1", because of 1-bit word sizes.
2706 * This lemma produces the different cases. *)
2707lemma sint_1_cases:
2708  "\<lbrakk> \<lbrakk> len_of TYPE ('a::len) = 1; (a::'a word) = 0; sint a = 0 \<rbrakk> \<Longrightarrow> P;
2709     \<lbrakk> len_of TYPE ('a) = 1; a = 1; sint (1 :: 'a word) = -1 \<rbrakk> \<Longrightarrow> P;
2710      \<lbrakk> len_of TYPE ('a) > 1; sint (1 :: 'a word) = 1 \<rbrakk> \<Longrightarrow> P \<rbrakk>
2711                \<Longrightarrow> P"
2712   apply atomize_elim
2713  apply (case_tac "len_of TYPE ('a) = 1")
2714   apply clarsimp
2715   apply (subgoal_tac "(UNIV :: 'a word set) = {0, 1}")
2716    apply (metis UNIV_I insert_iff singletonE)
2717   apply (subst word_unat.univ)
2718   apply (clarsimp simp: unats_def image_def)
2719   apply (rule set_eqI, rule iffI)
2720    apply clarsimp
2721    apply (metis One_nat_def less_2_cases of_nat_1 semiring_1_class.of_nat_0)
2722   apply clarsimp
2723   apply (metis Abs_fnat_hom_0 Suc_1 lessI of_nat_1 zero_less_Suc)
2724  apply clarsimp
2725  apply (metis One_nat_def arith_is_1 le_def len_gt_0)
2726  done
2727
2728lemma sint_int_min:
2729  "sint (- (2 ^ (LENGTH('a) - Suc 0)) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))"
2730  apply (subst word_sint.Abs_inverse' [where r="- (2 ^ (LENGTH('a) - Suc 0))"])
2731    apply (clarsimp simp: sints_num)
2732   apply (clarsimp simp: wi_hom_syms word_of_int_2p)
2733  apply clarsimp
2734  done
2735
2736lemma sint_int_max_plus_1:
2737  "sint (2 ^ (LENGTH('a) - Suc 0) :: ('a::len) word) = - (2 ^ (LENGTH('a) - Suc 0))"
2738  apply (subst word_of_int_2p [symmetric])
2739  apply (subst int_word_sint)
2740  apply clarsimp
2741  apply (metis Suc_pred int_word_uint len_gt_0 power_Suc uint_eq_0 word_of_int_2p word_pow_0)
2742  done
2743
2744lemma sbintrunc_eq_in_range:
2745  "(sbintrunc n x = x) = (x \<in> range (sbintrunc n))"
2746  "(x = sbintrunc n x) = (x \<in> range (sbintrunc n))"
2747  apply (simp_all add: image_def)
2748  apply (metis sbintrunc_sbintrunc)+
2749  done
2750
2751lemma sbintrunc_If:
2752  "- 3 * (2 ^ n) \<le> x \<and> x < 3 * (2 ^ n)
2753    \<Longrightarrow> sbintrunc n x = (if x < - (2 ^ n) then x + 2 * (2 ^ n)
2754        else if x \<ge> 2 ^ n then x - 2 * (2 ^ n) else x)"
2755  apply (simp add: no_sbintr_alt2, safe)
2756   apply (simp add: mod_pos_geq)
2757  apply (subst mod_add_self1[symmetric], simp)
2758  done
2759
2760lemma signed_arith_eq_checks_to_ord:
2761  "(sint a + sint b = sint (a + b ))
2762    = ((a <=s a + b) = (0 <=s b))"
2763  "(sint a - sint b = sint (a - b ))
2764    = ((0 <=s a - b) = (b <=s a))"
2765  "(- sint a = sint (- a)) = (0 <=s (- a) = (a <=s 0))"
2766  using sint_range'[where x=a] sint_range'[where x=b]
2767  by (simp_all add: sint_word_ariths word_sle_def word_sless_alt sbintrunc_If)
2768
2769(* Basic proofs that signed word div/mod operations are
2770 * truncations of their integer counterparts. *)
2771
2772lemma signed_div_arith:
2773    "sint ((a::('a::len) word) sdiv b) = sbintrunc (LENGTH('a) - 1) (sint a sdiv sint b)"
2774  apply (subst word_sbin.norm_Rep [symmetric])
2775  apply (subst bin_sbin_eq_iff' [symmetric])
2776   apply simp
2777  apply (subst uint_sint [symmetric])
2778  apply (clarsimp simp: sdiv_int_def sdiv_word_def)
2779  apply (metis word_ubin.eq_norm)
2780  done
2781
2782lemma signed_mod_arith:
2783    "sint ((a::('a::len) word) smod b) = sbintrunc (LENGTH('a) - 1) (sint a smod sint b)"
2784  apply (subst word_sbin.norm_Rep [symmetric])
2785  apply (subst bin_sbin_eq_iff' [symmetric])
2786   apply simp
2787  apply (subst uint_sint [symmetric])
2788  apply (clarsimp simp: smod_int_def smod_word_def)
2789  apply (metis word_ubin.eq_norm)
2790  done
2791
2792(* Signed word arithmetic overflow constraints. *)
2793
2794lemma signed_arith_ineq_checks_to_eq:
2795  "((- (2 ^ (size a - 1)) \<le> (sint a + sint b)) \<and> (sint a + sint b \<le> (2 ^ (size a - 1) - 1)))
2796    = (sint a + sint b = sint (a + b ))"
2797  "((- (2 ^ (size a - 1)) \<le> (sint a - sint b)) \<and> (sint a - sint b \<le> (2 ^ (size a - 1) - 1)))
2798    = (sint a - sint b = sint (a - b))"
2799  "((- (2 ^ (size a - 1)) \<le> (- sint a)) \<and> (- sint a) \<le> (2 ^ (size a - 1) - 1))
2800    = ((- sint a) = sint (- a))"
2801  "((- (2 ^ (size a - 1)) \<le> (sint a * sint b)) \<and> (sint a * sint b \<le> (2 ^ (size a - 1) - 1)))
2802    = (sint a * sint b = sint (a * b))"
2803  "((- (2 ^ (size a - 1)) \<le> (sint a sdiv sint b)) \<and> (sint a sdiv sint b \<le> (2 ^ (size a - 1) - 1)))
2804    = (sint a sdiv sint b = sint (a sdiv b))"
2805  "((- (2 ^ (size a - 1)) \<le> (sint a smod sint b)) \<and> (sint a smod sint b \<le> (2 ^ (size a - 1) - 1)))
2806    = (sint a smod sint b = sint (a smod b))"
2807  by (auto simp: sint_word_ariths word_size signed_div_arith signed_mod_arith
2808                    sbintrunc_eq_in_range range_sbintrunc)
2809
2810lemma signed_arith_sint:
2811  "((- (2 ^ (size a - 1)) \<le> (sint a + sint b)) \<and> (sint a + sint b \<le> (2 ^ (size a - 1) - 1)))
2812    \<Longrightarrow> sint (a + b) = (sint a + sint b)"
2813  "((- (2 ^ (size a - 1)) \<le> (sint a - sint b)) \<and> (sint a - sint b \<le> (2 ^ (size a - 1) - 1)))
2814    \<Longrightarrow> sint (a - b) = (sint a - sint b)"
2815  "((- (2 ^ (size a - 1)) \<le> (- sint a)) \<and> (- sint a) \<le> (2 ^ (size a - 1) - 1))
2816    \<Longrightarrow> sint (- a) = (- sint a)"
2817  "((- (2 ^ (size a - 1)) \<le> (sint a * sint b)) \<and> (sint a * sint b \<le> (2 ^ (size a - 1) - 1)))
2818    \<Longrightarrow> sint (a * b) = (sint a * sint b)"
2819  "((- (2 ^ (size a - 1)) \<le> (sint a sdiv sint b)) \<and> (sint a sdiv sint b \<le> (2 ^ (size a - 1) - 1)))
2820    \<Longrightarrow> sint (a sdiv b) = (sint a sdiv sint b)"
2821  "((- (2 ^ (size a - 1)) \<le> (sint a smod sint b)) \<and> (sint a smod sint b \<le> (2 ^ (size a - 1) - 1)))
2822    \<Longrightarrow> sint (a smod b) = (sint a smod sint b)"
2823  by (subst (asm) signed_arith_ineq_checks_to_eq; simp)+
2824
2825lemma signed_mult_eq_checks_double_size:
2826  assumes mult_le: "(2 ^ (len_of TYPE ('a) - 1) + 1) ^ 2 \<le> (2 :: int) ^ (len_of TYPE ('b) - 1)"
2827           and le: "2 ^ (LENGTH('a) - 1) \<le> (2 :: int) ^ (len_of TYPE ('b) - 1)"
2828  shows "(sint (a :: 'a :: len word) * sint b = sint (a * b))
2829       = (scast a * scast b = (scast (a * b) :: 'b :: len word))"
2830proof -
2831  have P: "sbintrunc (size a - 1) (sint a * sint b) \<in> range (sbintrunc (size a - 1))"
2832    by simp
2833
2834  have abs: "!! x :: 'a word. abs (sint x) < 2 ^ (size a - 1) + 1"
2835    apply (cut_tac x=x in sint_range')
2836    apply (simp add: abs_le_iff word_size)
2837    done
2838  have abs_ab: "abs (sint a * sint b) < 2 ^ (LENGTH('b) - 1)"
2839    using abs_mult_less[OF abs[where x=a] abs[where x=b]] mult_le
2840    by (simp add: abs_mult power2_eq_square word_size)
2841  show ?thesis
2842    using P[unfolded range_sbintrunc] abs_ab le
2843    apply (simp add: sint_word_ariths scast_def)
2844    apply (simp add: wi_hom_mult)
2845    apply (subst word_sint.Abs_inject, simp_all)
2846     apply (simp add: sints_def range_sbintrunc abs_less_iff)
2847    apply clarsimp
2848    apply (simp add: sints_def range_sbintrunc word_size)
2849    apply (auto elim: order_less_le_trans order_trans[rotated])
2850    done
2851qed
2852
2853(* Properties about signed division. *)
2854
2855lemma int_sdiv_simps [simp]:
2856    "(a :: int) sdiv 1 = a"
2857    "(a :: int) sdiv 0 = 0"
2858    "(a :: int) sdiv -1 = -a"
2859  apply (auto simp: sdiv_int_def sgn_if)
2860  done
2861
2862lemma sgn_div_eq_sgn_mult:
2863    "a div b \<noteq> 0 \<Longrightarrow> sgn ((a :: int) div b) = sgn (a * b)"
2864  apply (clarsimp simp: sgn_if zero_le_mult_iff neg_imp_zdiv_nonneg_iff not_less)
2865  apply (metis less_le mult_le_0_iff neg_imp_zdiv_neg_iff not_less pos_imp_zdiv_neg_iff zdiv_eq_0_iff)
2866  done
2867
2868lemma sgn_sdiv_eq_sgn_mult:
2869  "a sdiv b \<noteq> 0 \<Longrightarrow> sgn ((a :: int) sdiv b) = sgn (a * b)"
2870  by (auto simp: sdiv_int_def sgn_div_eq_sgn_mult sgn_mult)
2871
2872lemma int_sdiv_same_is_1 [simp]:
2873    "a \<noteq> 0 \<Longrightarrow> ((a :: int) sdiv b = a) = (b = 1)"
2874  apply (rule iffI)
2875   apply (clarsimp simp: sdiv_int_def)
2876   apply (subgoal_tac "b > 0")
2877    apply (case_tac "a > 0")
2878     apply (clarsimp simp: sgn_if)
2879    apply (clarsimp simp: algebra_split_simps not_less)
2880    apply (metis int_div_same_is_1 le_neq_trans minus_minus neg_0_le_iff_le neg_equal_0_iff_equal)
2881   apply (case_tac "a > 0")
2882    apply (case_tac "b = 0")
2883     apply clarsimp
2884    apply (rule classical)
2885    apply (clarsimp simp: sgn_mult not_less)
2886    apply (metis le_less neg_0_less_iff_less not_less_iff_gr_or_eq pos_imp_zdiv_neg_iff)
2887   apply (rule classical)
2888   apply (clarsimp simp: algebra_split_simps sgn_mult not_less sgn_if split: if_splits)
2889   apply (metis antisym less_le neg_imp_zdiv_nonneg_iff)
2890  apply (clarsimp simp: sdiv_int_def sgn_if)
2891  done
2892
2893lemma int_sdiv_negated_is_minus1 [simp]:
2894    "a \<noteq> 0 \<Longrightarrow> ((a :: int) sdiv b = - a) = (b = -1)"
2895  apply (clarsimp simp: sdiv_int_def)
2896  apply (rule iffI)
2897   apply (subgoal_tac "b < 0")
2898    apply (case_tac "a > 0")
2899     apply (clarsimp simp: sgn_if algebra_split_simps not_less)
2900    apply (case_tac "sgn (a * b) = -1")
2901     apply (clarsimp simp: not_less algebra_split_simps)
2902    apply (clarsimp simp: algebra_split_simps not_less)
2903   apply (rule classical)
2904   apply (case_tac "b = 0")
2905    apply (clarsimp simp: not_less sgn_mult)
2906   apply (case_tac "a > 0")
2907    apply (clarsimp simp: not_less sgn_mult)
2908    apply (metis less_le neg_less_0_iff_less not_less_iff_gr_or_eq pos_imp_zdiv_neg_iff)
2909   apply (clarsimp simp: not_less sgn_mult)
2910   apply (metis antisym_conv div_minus_right neg_imp_zdiv_nonneg_iff neg_le_0_iff_le not_less)
2911  apply (clarsimp simp: sgn_if)
2912  done
2913
2914lemma sdiv_int_range:
2915    "(a :: int) sdiv b \<in> { - (abs a) .. (abs a) }"
2916  apply (unfold sdiv_int_def)
2917  apply (subgoal_tac "(abs a) div (abs b) \<le> (abs a)")
2918   apply (clarsimp simp: sgn_if)
2919   apply (meson abs_ge_zero neg_le_0_iff_le nonneg_mod_div order_trans)
2920  apply (metis abs_eq_0 abs_ge_zero div_by_0 zdiv_le_dividend zero_less_abs_iff)
2921  done
2922
2923lemma word_sdiv_div1 [simp]:
2924    "(a :: ('a::len) word) sdiv 1 = a"
2925  apply (rule sint_1_cases [where a=a])
2926    apply (clarsimp simp: sdiv_word_def sdiv_int_def)
2927   apply (clarsimp simp: sdiv_word_def sdiv_int_def simp del: sint_minus1)
2928  apply (clarsimp simp: sdiv_word_def)
2929  done
2930
2931lemma sdiv_int_div_0 [simp]:
2932  "(x :: int) sdiv 0 = 0"
2933  by (clarsimp simp: sdiv_int_def)
2934
2935lemma sdiv_int_0_div [simp]:
2936  "0 sdiv (x :: int) = 0"
2937  by (clarsimp simp: sdiv_int_def)
2938
2939lemma word_sdiv_div0 [simp]:
2940    "(a :: ('a::len) word) sdiv 0 = 0"
2941  apply (auto simp: sdiv_word_def sdiv_int_def sgn_if)
2942  done
2943
2944lemma word_sdiv_div_minus1 [simp]:
2945    "(a :: ('a::len) word) sdiv -1 = -a"
2946  apply (auto simp: sdiv_word_def sdiv_int_def sgn_if)
2947  apply (metis wi_hom_neg word_sint.Rep_inverse')
2948  done
2949
2950lemmas word_sdiv_0 = word_sdiv_div0
2951
2952lemma sdiv_word_min:
2953    "- (2 ^ (size a - 1)) \<le> sint (a :: ('a::len) word) sdiv sint (b :: ('a::len) word)"
2954  apply (clarsimp simp: word_size)
2955  apply (cut_tac sint_range' [where x=a])
2956  apply (cut_tac sint_range' [where x=b])
2957  apply clarsimp
2958  apply (insert sdiv_int_range [where a="sint a" and b="sint b"])
2959  apply (clarsimp simp: max_def abs_if split: if_split_asm)
2960  done
2961
2962lemma sdiv_word_max:
2963    "(sint (a :: ('a::len) word) sdiv sint (b :: ('a::len) word) < (2 ^ (size a - 1))) =
2964          ((a \<noteq> - (2 ^ (size a - 1)) \<or> (b \<noteq> -1)))"
2965    (is "?lhs = (\<not> ?a_int_min \<or> \<not> ?b_minus1)")
2966proof (rule classical)
2967  assume not_thesis: "\<not> ?thesis"
2968
2969  have not_zero: "b \<noteq> 0"
2970    using not_thesis
2971    by (clarsimp)
2972
2973  have result_range: "sint a sdiv sint b \<in> (sints (size a)) \<union> {2 ^ (size a - 1)}"
2974    apply (cut_tac sdiv_int_range [where a="sint a" and b="sint b"])
2975    apply (erule rev_subsetD)
2976    using sint_range' [where x=a]  sint_range' [where x=b]
2977    apply (auto simp: max_def abs_if word_size sints_num)
2978    done
2979
2980  have result_range_overflow: "(sint a sdiv sint b = 2 ^ (size a - 1)) = (?a_int_min \<and> ?b_minus1)"
2981    apply (rule iffI [rotated])
2982     apply (clarsimp simp: sdiv_int_def sgn_if word_size sint_int_min)
2983    apply (rule classical)
2984    apply (case_tac "?a_int_min")
2985     apply (clarsimp simp: word_size sint_int_min)
2986     apply (metis diff_0_right
2987              int_sdiv_negated_is_minus1 minus_diff_eq minus_int_code(2)
2988              power_eq_0_iff sint_minus1 zero_neq_numeral)
2989    apply (subgoal_tac "abs (sint a) < 2 ^ (size a - 1)")
2990     apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1]
2991     apply (clarsimp simp: word_size)
2992    apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1]
2993    apply (insert word_sint.Rep [where x="a"])[1]
2994    apply (clarsimp simp: minus_le_iff word_size abs_if sints_num split: if_split_asm)
2995    apply (metis minus_minus sint_int_min word_sint.Rep_inject)
2996    done
2997
2998  have result_range_simple: "(sint a sdiv sint b \<in> (sints (size a))) \<Longrightarrow> ?thesis"
2999    apply (insert sdiv_int_range [where a="sint a" and b="sint b"])
3000    apply (clarsimp simp: word_size sints_num sint_int_min)
3001    done
3002
3003  show ?thesis
3004    apply (rule UnE [OF result_range result_range_simple])
3005     apply simp
3006    apply (clarsimp simp: word_size)
3007    using result_range_overflow
3008    apply (clarsimp simp: word_size)
3009    done
3010qed
3011
3012lemmas sdiv_word_min' = sdiv_word_min [simplified word_size, simplified]
3013lemmas sdiv_word_max' = sdiv_word_max [simplified word_size, simplified]
3014
3015lemmas word_sdiv_numerals_lhs = sdiv_word_def[where a="numeral x" for x]
3016    sdiv_word_def[where a=0] sdiv_word_def[where a=1]
3017
3018lemmas word_sdiv_numerals = word_sdiv_numerals_lhs[where b="numeral y" for y]
3019    word_sdiv_numerals_lhs[where b=0] word_sdiv_numerals_lhs[where b=1]
3020
3021(*
3022 * Signed modulo properties.
3023 *)
3024
3025lemma smod_int_alt_def:
3026     "(a::int) smod b = sgn (a) * (abs a mod abs b)"
3027  apply (clarsimp simp: smod_int_def sdiv_int_def)
3028  apply (clarsimp simp: minus_div_mult_eq_mod [symmetric] abs_sgn sgn_mult sgn_if algebra_split_simps)
3029  done
3030
3031lemma smod_int_range:
3032  "b \<noteq> 0 \<Longrightarrow> (a::int) smod b \<in> { - abs b + 1 .. abs b - 1 }"
3033  apply (case_tac  "b > 0")
3034   apply (insert pos_mod_conj [where a=a and b=b])[1]
3035   apply (insert pos_mod_conj [where a="-a" and b=b])[1]
3036   apply (auto simp: smod_int_alt_def algebra_simps sgn_if
3037              abs_if not_less add1_zle_eq [simplified add.commute])[1]
3038    apply (metis add_nonneg_nonneg int_one_le_iff_zero_less le_less less_add_same_cancel2 not_le pos_mod_conj)
3039  apply (metis (full_types) add.inverse_inverse eucl_rel_int eucl_rel_int_iff le_less_trans neg_0_le_iff_le)
3040  apply (insert neg_mod_conj [where a=a and b="b"])[1]
3041  apply (insert neg_mod_conj [where a="-a" and b="b"])[1]
3042  apply (clarsimp simp: smod_int_alt_def algebra_simps sgn_if
3043            abs_if not_less add1_zle_eq [simplified add.commute])
3044  apply (metis neg_0_less_iff_less neg_mod_conj not_le not_less_iff_gr_or_eq order_trans pos_mod_conj)
3045  done
3046
3047lemma smod_int_compares:
3048   "\<lbrakk> 0 \<le> a; 0 < b \<rbrakk> \<Longrightarrow> (a :: int) smod b < b"
3049   "\<lbrakk> 0 \<le> a; 0 < b \<rbrakk> \<Longrightarrow> 0 \<le> (a :: int) smod b"
3050   "\<lbrakk> a \<le> 0; 0 < b \<rbrakk> \<Longrightarrow> -b < (a :: int) smod b"
3051   "\<lbrakk> a \<le> 0; 0 < b \<rbrakk> \<Longrightarrow> (a :: int) smod b \<le> 0"
3052   "\<lbrakk> 0 \<le> a; b < 0 \<rbrakk> \<Longrightarrow> (a :: int) smod b < - b"
3053   "\<lbrakk> 0 \<le> a; b < 0 \<rbrakk> \<Longrightarrow> 0 \<le> (a :: int) smod b"
3054   "\<lbrakk> a \<le> 0; b < 0 \<rbrakk> \<Longrightarrow> (a :: int) smod b \<le> 0"
3055   "\<lbrakk> a \<le> 0; b < 0 \<rbrakk> \<Longrightarrow> b \<le> (a :: int) smod b"
3056  apply (insert smod_int_range [where a=a and b=b])
3057  apply (auto simp: add1_zle_eq smod_int_alt_def sgn_if)
3058  done
3059
3060lemma smod_int_mod_0 [simp]:
3061  "x smod (0 :: int) = x"
3062  by (clarsimp simp: smod_int_def)
3063
3064lemma smod_int_0_mod [simp]:
3065  "0 smod (x :: int) = 0"
3066  by (clarsimp simp: smod_int_alt_def)
3067
3068lemma smod_word_mod_0 [simp]:
3069  "x smod (0 :: ('a::len) word) = x"
3070  by (clarsimp simp: smod_word_def)
3071
3072lemma smod_word_0_mod [simp]:
3073  "0 smod (x :: ('a::len) word) = 0"
3074  by (clarsimp simp: smod_word_def)
3075
3076lemma smod_word_max:
3077    "sint (a::'a word) smod sint (b::'a word) < 2 ^ (LENGTH('a::len) - Suc 0)"
3078  apply (case_tac "b = 0")
3079   apply (insert word_sint.Rep [where x=a, simplified sints_num])[1]
3080   apply (clarsimp)
3081  apply (insert word_sint.Rep [where x="b", simplified sints_num])[1]
3082  apply (insert smod_int_range [where a="sint a" and b="sint b"])
3083  apply (clarsimp simp: abs_if split: if_split_asm)
3084  done
3085
3086lemma smod_word_min:
3087    "- (2 ^ (LENGTH('a::len) - Suc 0)) \<le> sint (a::'a word) smod sint (b::'a word)"
3088  apply (case_tac "b = 0")
3089   apply (insert word_sint.Rep [where x=a, simplified sints_num])[1]
3090   apply clarsimp
3091  apply (insert word_sint.Rep [where x=b, simplified sints_num])[1]
3092  apply (insert smod_int_range [where a="sint a" and b="sint b"])
3093  apply (clarsimp simp: abs_if add1_zle_eq split: if_split_asm)
3094  done
3095
3096lemma smod_word_alt_def:
3097  "(a :: ('a::len) word) smod b = a - (a sdiv b) * b"
3098  apply (case_tac "a \<noteq> - (2 ^ (LENGTH('a) - 1)) \<or> b \<noteq> -1")
3099   apply (clarsimp simp: smod_word_def sdiv_word_def smod_int_def
3100             minus_word.abs_eq [symmetric] times_word.abs_eq [symmetric])
3101  apply (clarsimp simp: smod_word_def smod_int_def)
3102  done
3103
3104lemmas word_smod_numerals_lhs = smod_word_def[where a="numeral x" for x]
3105    smod_word_def[where a=0] smod_word_def[where a=1]
3106
3107lemmas word_smod_numerals = word_smod_numerals_lhs[where b="numeral y" for y]
3108    word_smod_numerals_lhs[where b=0] word_smod_numerals_lhs[where b=1]
3109
3110lemma sint_of_int_eq:
3111  "\<lbrakk> - (2 ^ (LENGTH('a) - 1)) \<le> x; x < 2 ^ (LENGTH('a) - 1) \<rbrakk> \<Longrightarrow> sint (of_int x :: ('a::len) word) = x"
3112  apply (clarsimp simp: word_of_int int_word_sint)
3113  apply (subst int_mod_eq')
3114    apply simp
3115   apply (subst (2) power_minus_simp)
3116    apply clarsimp
3117   apply clarsimp
3118  apply clarsimp
3119  done
3120
3121lemma of_int_sint [simp]:
3122    "of_int (sint a) = a"
3123  apply (insert word_sint.Rep [where x=a])
3124  apply (clarsimp simp: word_of_int)
3125  done
3126
3127lemma nth_w2p_scast [simp]:
3128  "((scast ((2::'a::len signed word) ^ n) :: 'a word) !! m)
3129         \<longleftrightarrow> ((((2::'a::len  word) ^ n) :: 'a word) !! m)"
3130  apply (subst nth_w2p)
3131  apply (case_tac "n \<ge> LENGTH('a)")
3132   apply (subst power_overflow, simp)
3133   apply clarsimp
3134  apply (metis nth_w2p scast_def test_bit_conj_lt
3135               len_signed nth_word_of_int word_sint.Rep_inverse)
3136  done
3137
3138lemma scast_2_power [simp]: "scast ((2 :: 'a::len signed word) ^ x) = ((2 :: 'a word) ^ x)"
3139  by (clarsimp simp: word_eq_iff)
3140
3141lemma scast_bit_test [simp]:
3142    "scast ((1 :: 'a::len signed word) << n) = (1 :: 'a word) << n"
3143  by (clarsimp simp: word_eq_iff)
3144
3145lemma ucast_nat_def':
3146  "of_nat (unat x) = (ucast :: 'a :: len word \<Rightarrow> ('b :: len) signed word) x"
3147  by (simp add: ucast_def word_of_int_nat unat_def)
3148
3149lemma mod_mod_power_int:
3150  fixes k :: int
3151  shows "k mod 2 ^ m mod 2 ^ n = k mod 2 ^ (min m n)"
3152  by (metis bintrunc_bintrunc_min bintrunc_mod2p min.commute)
3153
3154(* Normalise combinations of scast and ucast. *)
3155
3156lemma ucast_distrib:
3157  fixes M :: "'a::len word \<Rightarrow> 'a::len word \<Rightarrow> 'a::len word"
3158  fixes M' :: "'b::len word \<Rightarrow> 'b::len word \<Rightarrow> 'b::len word"
3159  fixes L :: "int \<Rightarrow> int \<Rightarrow> int"
3160  assumes lift_M: "\<And>x y. uint (M x y) = L (uint x) (uint y)  mod 2 ^ LENGTH('a)"
3161  assumes lift_M': "\<And>x y. uint (M' x y) = L (uint x) (uint y)  mod 2 ^ LENGTH('b)"
3162  assumes distrib: "\<And>x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b))
3163                               = (L x y) mod (2 ^ LENGTH('b))"
3164  assumes is_down: "is_down (ucast :: 'a word \<Rightarrow> 'b word)"
3165  shows "ucast (M a b) = M' (ucast a) (ucast b)"
3166  apply (clarsimp simp: word_of_int ucast_def)
3167  apply (subst lift_M)
3168  apply (subst of_int_uint [symmetric], subst lift_M')
3169  apply (subst (1 2) int_word_uint)
3170  apply (subst word_of_int)
3171  apply (subst word.abs_eq_iff)
3172  apply (subst (1 2) bintrunc_mod2p)
3173  apply (insert is_down)
3174  apply (unfold is_down_def)
3175  apply (clarsimp simp: target_size source_size)
3176  apply (clarsimp simp: mod_mod_power_int min_def)
3177  apply (rule distrib [symmetric])
3178  done
3179
3180lemma ucast_down_add:
3181    "is_down (ucast:: 'a word \<Rightarrow> 'b word) \<Longrightarrow>  ucast ((a :: 'a::len word) + b) = (ucast a + ucast b :: 'b::len word)"
3182  by (rule ucast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp)
3183
3184lemma ucast_down_minus:
3185    "is_down (ucast:: 'a word \<Rightarrow> 'b word) \<Longrightarrow>  ucast ((a :: 'a::len word) - b) = (ucast a - ucast b :: 'b::len word)"
3186  apply (rule ucast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+)
3187  apply (metis mod_diff_left_eq mod_diff_right_eq)
3188  apply simp
3189  done
3190
3191lemma ucast_down_mult:
3192    "is_down (ucast:: 'a word \<Rightarrow> 'b word) \<Longrightarrow>  ucast ((a :: 'a::len word) * b) = (ucast a * ucast b :: 'b::len word)"
3193  apply (rule ucast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+)
3194  apply (metis mod_mult_eq)
3195  apply simp
3196  done
3197
3198lemma scast_distrib:
3199  fixes M :: "'a::len word \<Rightarrow> 'a::len word \<Rightarrow> 'a::len word"
3200  fixes M' :: "'b::len word \<Rightarrow> 'b::len word \<Rightarrow> 'b::len word"
3201  fixes L :: "int \<Rightarrow> int \<Rightarrow> int"
3202  assumes lift_M: "\<And>x y. uint (M x y) = L (uint x) (uint y)  mod 2 ^ LENGTH('a)"
3203  assumes lift_M': "\<And>x y. uint (M' x y) = L (uint x) (uint y)  mod 2 ^ LENGTH('b)"
3204  assumes distrib: "\<And>x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b))
3205                               = (L x y) mod (2 ^ LENGTH('b))"
3206  assumes is_down: "is_down (scast :: 'a word \<Rightarrow> 'b word)"
3207  shows "scast (M a b) = M' (scast a) (scast b)"
3208  apply (subst (1 2 3) down_cast_same [symmetric])
3209   apply (insert is_down)
3210   apply (clarsimp simp: is_down_def target_size source_size is_down)
3211  apply (rule ucast_distrib [where L=L, OF lift_M lift_M' distrib])
3212  apply (insert is_down)
3213  apply (clarsimp simp: is_down_def target_size source_size is_down)
3214  done
3215
3216lemma scast_down_add:
3217    "is_down (scast:: 'a word \<Rightarrow> 'b word) \<Longrightarrow>  scast ((a :: 'a::len word) + b) = (scast a + scast b :: 'b::len word)"
3218  by (rule scast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp)
3219
3220lemma scast_down_minus:
3221    "is_down (scast:: 'a word \<Rightarrow> 'b word) \<Longrightarrow>  scast ((a :: 'a::len word) - b) = (scast a - scast b :: 'b::len word)"
3222  apply (rule scast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+)
3223  apply (metis mod_diff_left_eq mod_diff_right_eq)
3224  apply simp
3225  done
3226
3227lemma scast_down_mult:
3228    "is_down (scast:: 'a word \<Rightarrow> 'b word) \<Longrightarrow>  scast ((a :: 'a::len word) * b) = (scast a * scast b :: 'b::len word)"
3229  apply (rule scast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+)
3230  apply (metis mod_mult_eq)
3231  apply simp
3232  done
3233
3234lemma scast_ucast_1:
3235  "\<lbrakk> is_down (ucast :: 'a word \<Rightarrow> 'b word); is_down (ucast :: 'b word \<Rightarrow> 'c word) \<rbrakk> \<Longrightarrow>
3236         (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a"
3237  by (metis down_cast_same ucast_def ucast_down_wi)
3238
3239lemma scast_ucast_3:
3240  "\<lbrakk> is_down (ucast :: 'a word \<Rightarrow> 'c word); is_down (ucast :: 'b word \<Rightarrow> 'c word) \<rbrakk> \<Longrightarrow>
3241         (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a"
3242  by (metis down_cast_same ucast_def ucast_down_wi)
3243
3244lemma scast_ucast_4:
3245  "\<lbrakk> is_up (ucast :: 'a word \<Rightarrow> 'b word); is_down (ucast :: 'b word \<Rightarrow> 'c word) \<rbrakk> \<Longrightarrow>
3246         (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a"
3247  by (metis down_cast_same ucast_def ucast_down_wi)
3248
3249lemma scast_scast_b:
3250  "\<lbrakk> is_up (scast :: 'a word \<Rightarrow> 'b word) \<rbrakk> \<Longrightarrow>
3251     (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a"
3252  by (metis scast_def sint_up_scast)
3253
3254lemma ucast_scast_1:
3255  "\<lbrakk> is_down (scast :: 'a word \<Rightarrow> 'b word); is_down (ucast :: 'b word \<Rightarrow> 'c word) \<rbrakk> \<Longrightarrow>
3256            (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a"
3257  by (metis scast_def ucast_down_wi)
3258
3259lemma ucast_scast_3:
3260  "\<lbrakk> is_down (scast :: 'a word \<Rightarrow> 'c word); is_down (ucast :: 'b word \<Rightarrow> 'c word) \<rbrakk> \<Longrightarrow>
3261     (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a"
3262  by (metis scast_def ucast_down_wi)
3263
3264lemma ucast_scast_4:
3265  "\<lbrakk> is_up (scast :: 'a word \<Rightarrow> 'b word); is_down (ucast :: 'b word \<Rightarrow> 'c word) \<rbrakk> \<Longrightarrow>
3266     (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a"
3267  by (metis down_cast_same scast_def sint_up_scast)
3268
3269lemma ucast_ucast_a:
3270  "\<lbrakk> is_down (ucast :: 'b word \<Rightarrow> 'c word) \<rbrakk> \<Longrightarrow>
3271        (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a"
3272  by (metis down_cast_same ucast_def ucast_down_wi)
3273
3274lemma ucast_ucast_b:
3275  "\<lbrakk> is_up (ucast :: 'a word \<Rightarrow> 'b word) \<rbrakk> \<Longrightarrow>
3276     (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a"
3277  by (metis ucast_up_ucast)
3278
3279lemma scast_scast_a:
3280  "\<lbrakk> is_down (scast :: 'b word \<Rightarrow> 'c word) \<rbrakk> \<Longrightarrow>
3281            (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a"
3282  apply (clarsimp simp: scast_def)
3283  apply (metis down_cast_same is_up_down scast_def ucast_down_wi)
3284  done
3285
3286lemma scast_down_wi [OF refl]:
3287  "uc = scast \<Longrightarrow> is_down uc \<Longrightarrow> uc (word_of_int x) = word_of_int x"
3288  by (metis down_cast_same is_up_down ucast_down_wi)
3289
3290lemmas cast_simps =
3291  is_down is_up
3292  scast_down_add scast_down_minus scast_down_mult
3293  ucast_down_add ucast_down_minus ucast_down_mult
3294  scast_ucast_1 scast_ucast_3 scast_ucast_4
3295  ucast_scast_1 ucast_scast_3 ucast_scast_4
3296  ucast_ucast_a ucast_ucast_b
3297  scast_scast_a scast_scast_b
3298  ucast_down_bl
3299  ucast_down_wi scast_down_wi
3300  ucast_of_nat scast_of_nat
3301  uint_up_ucast sint_up_scast
3302  up_scast_surj up_ucast_surj
3303
3304lemma smod_mod_positive:
3305    "\<lbrakk> 0 \<le> (a :: int); 0 \<le> b \<rbrakk> \<Longrightarrow> a smod b = a mod b"
3306  by (clarsimp simp: smod_int_alt_def zsgn_def)
3307
3308lemma nat_mult_power_less_eq:
3309  "b > 0 \<Longrightarrow> (a * b ^ n < (b :: nat) ^ m) = (a < b ^ (m - n))"
3310  using mult_less_cancel2[where m = a and k = "b ^ n" and n="b ^ (m - n)"]
3311        mult_less_cancel2[where m="a * b ^ (n - m)" and k="b ^ m" and n=1]
3312  apply (simp only: power_add[symmetric] nat_minus_add_max)
3313  apply (simp only: power_add[symmetric] nat_minus_add_max ac_simps)
3314  apply (simp add: max_def split: if_split_asm)
3315  done
3316
3317lemma signed_shift_guard_to_word:
3318  "\<lbrakk> n < len_of TYPE ('a); n > 0 \<rbrakk>
3319    \<Longrightarrow> (unat (x :: 'a :: len word) * 2 ^ y < 2 ^ n)
3320    = (x = 0 \<or> x < (1 << n >> y))"
3321  apply (simp only: nat_mult_power_less_eq)
3322  apply (cases "y \<le> n")
3323   apply (simp only: shiftl_shiftr1)
3324   apply (subst less_mask_eq)
3325    apply (simp add: word_less_nat_alt word_size)
3326    apply (rule order_less_le_trans[rotated], rule power_increasing[where n=1])
3327      apply simp
3328     apply simp
3329    apply simp
3330   apply (simp add: nat_mult_power_less_eq word_less_nat_alt word_size)
3331   apply auto[1]
3332  apply (simp only: shiftl_shiftr2, simp add: unat_eq_0)
3333  done
3334
3335lemma sint_ucast_eq_uint:
3336    "\<lbrakk> \<not> is_down (ucast :: ('a::len word \<Rightarrow> 'b::len word)) \<rbrakk>
3337            \<Longrightarrow> sint ((ucast :: ('a::len word \<Rightarrow> 'b::len word)) x) = uint x"
3338  apply (subst sint_eq_uint)
3339   apply (clarsimp simp: msb_nth nth_ucast is_down)
3340   apply (metis Suc_leI Suc_pred len_gt_0)
3341  apply (clarsimp simp: uint_up_ucast is_up is_down)
3342  done
3343
3344lemma word_less_nowrapI':
3345  "(x :: 'a :: len0 word) \<le> z - k \<Longrightarrow> k \<le> z \<Longrightarrow> 0 < k \<Longrightarrow> x < x + k"
3346  by uint_arith
3347
3348lemma mask_plus_1:
3349  "mask n + 1 = 2 ^ n"
3350  by (clarsimp simp: mask_def)
3351
3352lemma unat_inj: "inj unat"
3353  by (metis eq_iff injI word_le_nat_alt)
3354
3355lemma unat_ucast_upcast:
3356  "is_up (ucast :: 'b word \<Rightarrow> 'a word)
3357      \<Longrightarrow> unat (ucast x :: ('a::len) word) = unat (x :: ('b::len) word)"
3358  unfolding ucast_def unat_def
3359  apply (subst int_word_uint)
3360  apply (subst mod_pos_pos_trivial)
3361    apply simp
3362   apply (rule lt2p_lem)
3363   apply (clarsimp simp: is_up)
3364  apply simp
3365  done
3366
3367lemma ucast_mono:
3368  "\<lbrakk> (x :: 'b :: len word) < y; y < 2 ^ LENGTH('a) \<rbrakk>
3369   \<Longrightarrow> ucast x < ((ucast y) :: 'a :: len word)"
3370  apply (simp add: ucast_nat_def [symmetric])
3371  apply (rule of_nat_mono_maybe)
3372  apply (rule unat_less_helper)
3373  apply (simp add: Power.of_nat_power)
3374  apply (simp add: word_less_nat_alt)
3375  done
3376
3377lemma ucast_mono_le:
3378  "\<lbrakk>x \<le> y; y < 2 ^ LENGTH('b)\<rbrakk> \<Longrightarrow> (ucast (x :: 'a :: len word) :: 'b :: len word) \<le> ucast y"
3379  apply (simp add: ucast_nat_def [symmetric])
3380  apply (subst of_nat_mono_maybe_le[symmetric])
3381    apply (rule unat_less_helper)
3382    apply (simp add: Power.of_nat_power)
3383   apply (rule unat_less_helper)
3384   apply (erule le_less_trans)
3385   apply (simp add: Power.of_nat_power)
3386  apply (simp add: word_le_nat_alt)
3387  done
3388
3389lemma ucast_mono_le':
3390  "\<lbrakk> unat y < 2 ^ LENGTH('b); LENGTH('b::len) < LENGTH('a::len); x \<le> y \<rbrakk>
3391   \<Longrightarrow> UCAST('a \<rightarrow> 'b) x \<le> UCAST('a \<rightarrow> 'b) y"
3392  by (auto simp: word_less_nat_alt intro: ucast_mono_le)
3393
3394lemma zero_sle_ucast_up:
3395  "\<not> is_down (ucast :: 'a word \<Rightarrow> 'b signed word) \<Longrightarrow>
3396          (0 <=s ((ucast (b::('a::len) word)) :: ('b::len) signed word))"
3397  apply (subgoal_tac "\<not> msb (ucast b :: 'b signed word)")
3398   apply (clarsimp simp: word_sle_msb_le)
3399  apply (clarsimp simp: is_down not_le msb_nth nth_ucast)
3400  apply (subst (asm) test_bit_conj_lt [symmetric])
3401  apply clarsimp
3402  apply arith
3403  done
3404
3405lemma word_le_ucast_sless:
3406  "\<lbrakk> x \<le> y; y \<noteq> -1; LENGTH('a) < LENGTH('b) \<rbrakk> \<Longrightarrow>
3407    UCAST (('a :: len) \<rightarrow> ('b :: len) signed) x <s ucast (y + 1)"
3408  by (clarsimp simp: word_le_make_less word_sless_alt sint_ucast_eq_uint is_down word_less_def)
3409
3410lemma msb_ucast_eq:
3411    "LENGTH('a) = LENGTH('b) \<Longrightarrow>
3412         msb (ucast x :: ('a::len) word) = msb (x :: ('b::len) word)"
3413  apply (clarsimp simp: word_msb_alt)
3414  apply (subst ucast_down_drop [where n=0])
3415   apply (clarsimp simp: source_size_def target_size_def word_size)
3416  apply clarsimp
3417  done
3418
3419lemma msb_big:
3420     "msb (a :: ('a::len) word) = (a \<ge> 2 ^ (LENGTH('a)  - Suc 0))"
3421  apply (rule iffI)
3422   apply (clarsimp simp: msb_nth)
3423   apply (drule bang_is_le)
3424   apply simp
3425  apply (rule ccontr)
3426  apply (subgoal_tac "a = a && mask (LENGTH('a) - Suc 0)")
3427   apply (cut_tac and_mask_less' [where w=a and n="LENGTH('a) - Suc 0"])
3428    apply (clarsimp simp: word_not_le [symmetric])
3429   apply clarsimp
3430  apply (rule sym, subst and_mask_eq_iff_shiftr_0)
3431  apply (clarsimp simp: msb_shift)
3432  done
3433
3434lemma zero_sle_ucast:
3435  "(0 <=s ((ucast (b::('a::len) word)) :: ('a::len) signed word))
3436                = (uint b < 2 ^ (LENGTH('a) - 1))"
3437  apply (case_tac "msb b")
3438   apply (clarsimp simp: word_sle_msb_le not_less msb_ucast_eq del: notI)
3439   apply (clarsimp simp: msb_big word_le_def uint_2p_alt)
3440  apply (clarsimp simp: word_sle_msb_le not_less msb_ucast_eq del: notI)
3441  apply (clarsimp simp: msb_big word_le_def uint_2p_alt)
3442  done
3443
3444
3445(* to_bool / from_bool. *)
3446
3447definition
3448  from_bool :: "bool \<Rightarrow> 'a::len word" where
3449  "from_bool b \<equiv> case b of True \<Rightarrow> of_nat 1
3450                         | False \<Rightarrow> of_nat 0"
3451
3452lemma from_bool_0:
3453  "(from_bool x = 0) = (\<not> x)"
3454  by (simp add: from_bool_def split: bool.split)
3455
3456definition
3457  to_bool :: "'a::len word \<Rightarrow> bool" where
3458  "to_bool \<equiv> (\<noteq>) 0"
3459
3460lemma to_bool_and_1:
3461  "to_bool (x && 1) = (x !! 0)"
3462  apply (simp add: to_bool_def)
3463  apply (rule iffI)
3464   apply (rule classical, erule notE, word_eqI)
3465   apply (case_tac n; simp)
3466  apply (rule notI, drule word_eqD[where x=0])
3467  apply simp
3468  done
3469
3470lemma to_bool_from_bool [simp]:
3471  "to_bool (from_bool r) = r"
3472  unfolding from_bool_def to_bool_def
3473  by (simp split: bool.splits)
3474
3475lemma from_bool_neq_0 [simp]:
3476  "(from_bool b \<noteq> 0) = b"
3477  by (simp add: from_bool_def split: bool.splits)
3478
3479lemma from_bool_mask_simp [simp]:
3480  "(from_bool r :: 'a::len word) && 1 = from_bool r"
3481  unfolding from_bool_def
3482  by (clarsimp split: bool.splits)
3483
3484lemma from_bool_1 [simp]:
3485  "(from_bool P = 1) = P"
3486  by (simp add: from_bool_def split: bool.splits)
3487
3488lemma ge_0_from_bool [simp]:
3489  "(0 < from_bool P) = P"
3490  by (simp add: from_bool_def split: bool.splits)
3491
3492lemma limited_and_from_bool:
3493  "limited_and (from_bool b) 1"
3494  by (simp add: from_bool_def limited_and_def split: bool.split)
3495
3496lemma to_bool_1 [simp]: "to_bool 1" by (simp add: to_bool_def)
3497lemma to_bool_0 [simp]: "\<not>to_bool 0" by (simp add: to_bool_def)
3498
3499lemma from_bool_eq_if:
3500  "(from_bool Q = (if P then 1 else 0)) = (P = Q)"
3501  by (simp add: case_bool_If from_bool_def split: if_split)
3502
3503lemma to_bool_eq_0:
3504  "(\<not> to_bool x) = (x = 0)"
3505  by (simp add: to_bool_def)
3506
3507lemma to_bool_neq_0:
3508  "(to_bool x) = (x \<noteq> 0)"
3509  by (simp add: to_bool_def)
3510
3511lemma from_bool_all_helper:
3512  "(\<forall>bool. from_bool bool = val \<longrightarrow> P bool)
3513      = ((\<exists>bool. from_bool bool = val) \<longrightarrow> P (val \<noteq> 0))"
3514  by (auto simp: from_bool_0)
3515
3516lemma from_bool_to_bool_iff:
3517  "w = from_bool b \<longleftrightarrow> to_bool w = b \<and> (w = 0 \<or> w = 1)"
3518  by (cases b) (auto simp: from_bool_def to_bool_def)
3519
3520lemma from_bool_eqI:
3521  "from_bool x = from_bool y \<Longrightarrow> x = y"
3522  unfolding from_bool_def
3523  by (auto split: bool.splits)
3524
3525lemma word_rsplit_upt:
3526  "\<lbrakk> size x = LENGTH('a :: len) * n; n \<noteq> 0 \<rbrakk>
3527    \<Longrightarrow> word_rsplit x = map (\<lambda>i. ucast (x >> i * len_of TYPE ('a)) :: 'a word) (rev [0 ..< n])"
3528  apply (subgoal_tac "length (word_rsplit x :: 'a word list) = n")
3529   apply (rule nth_equalityI, simp)
3530   apply (intro allI word_eqI impI)
3531   apply (simp add: test_bit_rsplit_alt word_size)
3532   apply (simp add: nth_ucast nth_shiftr nth_rev field_simps)
3533  apply (simp add: length_word_rsplit_exp_size)
3534  apply (metis mult.commute given_quot_alt word_size word_size_gt_0)
3535  done
3536
3537lemma aligned_shift:
3538  "\<lbrakk>x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \<le> LENGTH('a)\<rbrakk>
3539   \<Longrightarrow> x + y >> n = y >> n"
3540  by (subst word_plus_and_or_coroll; word_eqI, blast)
3541
3542lemma aligned_shift':
3543  "\<lbrakk>x < 2 ^ n; is_aligned (y :: 'a :: len word) n;n \<le> LENGTH('a)\<rbrakk>
3544   \<Longrightarrow> y + x >> n = y >> n"
3545  by (subst word_plus_and_or_coroll; word_eqI, blast)
3546
3547lemma neg_mask_add_mask:
3548  "((x:: 'a :: len word) && ~~ (mask n)) + (2 ^ n - 1) = x || mask n"
3549  unfolding mask_2pm1[symmetric]
3550  by (subst word_plus_and_or_coroll; word_eqI_solve)
3551
3552lemma subtract_mask:
3553  "p - (p && mask n) = (p && ~~ (mask n))"
3554  "p - (p && ~~ (mask n)) = (p && mask n)"
3555  by (simp add: field_simps word_plus_and_or_coroll2)+
3556
3557lemma and_neg_mask_plus_mask_mono: "(p && ~~ (mask n)) + mask n \<ge> p"
3558  apply (rule word_le_minus_cancel[where x = "p && ~~ (mask n)"])
3559   apply (clarsimp simp: subtract_mask)
3560   using word_and_le1[where a = "mask n" and y = p]
3561   apply (clarsimp simp: mask_def word_le_less_eq)
3562  apply (rule is_aligned_no_overflow'[folded mask_2pm1])
3563  apply (clarsimp simp: is_aligned_neg_mask)
3564  done
3565
3566lemma word_neg_and_le:
3567  "ptr \<le> (ptr && ~~ (mask n)) + (2 ^ n - 1)"
3568  by (simp add: and_neg_mask_plus_mask_mono mask_2pm1[symmetric])
3569
3570lemma aligned_less_plus_1:
3571  "\<lbrakk> is_aligned x n; n > 0 \<rbrakk> \<Longrightarrow> x < x + 1"
3572  apply (rule plus_one_helper2)
3573   apply (rule order_refl)
3574  apply (clarsimp simp: field_simps)
3575  apply (drule arg_cong[where f="\<lambda>x. x - 1"])
3576  apply (clarsimp simp: is_aligned_mask)
3577  apply (drule word_eqD[where x=0])
3578  apply simp
3579  done
3580
3581lemma aligned_add_offset_less:
3582  "\<lbrakk>is_aligned x n; is_aligned y n; x < y; z < 2 ^ n\<rbrakk> \<Longrightarrow> x + z < y"
3583  apply (cases "y = 0")
3584   apply simp
3585  apply (erule is_aligned_get_word_bits[where p=y], simp_all)
3586  apply (cases "z = 0", simp_all)
3587  apply (drule(2) aligned_at_least_t2n_diff[rotated -1])
3588  apply (drule plus_one_helper2)
3589   apply (rule less_is_non_zero_p1)
3590   apply (rule aligned_less_plus_1)
3591    apply (erule aligned_sub_aligned[OF _ _ order_refl],
3592           simp_all add: is_aligned_triv)[1]
3593   apply (cases n, simp_all)[1]
3594  apply (simp only: trans[OF diff_add_eq diff_diff_eq2[symmetric]])
3595  apply (drule word_less_add_right)
3596   apply (rule ccontr, simp add: linorder_not_le)
3597   apply (drule aligned_small_is_0, erule order_less_trans)
3598    apply (clarsimp simp: power_overflow)
3599   apply simp
3600  apply (erule order_le_less_trans[rotated],
3601         rule word_plus_mono_right)
3602   apply (erule word_le_minus_one_leq)
3603  apply (simp add: is_aligned_no_wrap' is_aligned_no_overflow field_simps)
3604  done
3605
3606lemma is_aligned_add_helper:
3607  "\<lbrakk> is_aligned p n; d < 2 ^ n \<rbrakk>
3608     \<Longrightarrow> (p + d && mask n = d) \<and> (p + d && (~~ (mask n)) = p)"
3609  apply (subst(asm) is_aligned_mask)
3610  apply (drule less_mask_eq)
3611  apply (rule context_conjI)
3612   apply (subst word_plus_and_or_coroll; word_eqI; blast)
3613  using word_plus_and_or_coroll2[where x="p + d" and w="mask n"] by simp
3614
3615lemma is_aligned_sub_helper:
3616  "\<lbrakk> is_aligned (p - d) n; d < 2 ^ n \<rbrakk>
3617     \<Longrightarrow> (p && mask n = d) \<and> (p && (~~ (mask n)) = p - d)"
3618  by (drule(1) is_aligned_add_helper, simp)
3619
3620lemma mask_twice:
3621  "(x && mask n) && mask m = x && mask (min m n)"
3622  by word_eqI_solve
3623
3624lemma is_aligned_after_mask:
3625  "\<lbrakk>is_aligned k m;m\<le> n\<rbrakk> \<Longrightarrow> is_aligned (k && mask n) m"
3626  by (rule is_aligned_andI1)
3627
3628lemma and_mask_plus:
3629  "\<lbrakk>is_aligned ptr m; m \<le> n; a < 2 ^ m\<rbrakk>
3630   \<Longrightarrow> ptr + a && mask n = (ptr && mask n) + a"
3631  apply (rule mask_eqI[where n = m])
3632   apply (simp add:mask_twice min_def)
3633    apply (simp add:is_aligned_add_helper)
3634    apply (subst is_aligned_add_helper[THEN conjunct1])
3635      apply (erule is_aligned_after_mask)
3636     apply simp
3637    apply simp
3638   apply simp
3639  apply (subgoal_tac "(ptr + a && mask n) && ~~ (mask m)
3640     = (ptr + a && ~~ (mask m) ) && mask n")
3641   apply (simp add:is_aligned_add_helper)
3642   apply (subst is_aligned_add_helper[THEN conjunct2])
3643     apply (simp add:is_aligned_after_mask)
3644    apply simp
3645   apply simp
3646  apply (simp add:word_bw_comms word_bw_lcs)
3647  done
3648
3649lemma le_step_down_word:"\<lbrakk>(i::('a::len) word) \<le> n; i = n \<longrightarrow> P; i \<le> n - 1 \<longrightarrow> P\<rbrakk> \<Longrightarrow> P"
3650  by unat_arith
3651
3652lemma le_step_down_word_2:
3653  fixes x :: "'a::len word"
3654  shows "\<lbrakk>x \<le>  y; x \<noteq> y\<rbrakk> \<Longrightarrow> x \<le> y - 1"
3655  by (subst (asm) word_le_less_eq,
3656      clarsimp,
3657      simp add: word_le_minus_one_leq)
3658
3659lemma NOT_mask_AND_mask[simp]: "(w && mask n) && ~~ (mask n) = 0"
3660  apply (clarsimp simp:mask_def)
3661  by (metis word_bool_alg.conj_cancel_right word_bool_alg.conj_zero_right word_bw_comms(1) word_bw_lcs(1))
3662
3663lemma and_and_not[simp]:"(a && b) && ~~ b = 0"
3664  apply (subst word_bw_assocs(1))
3665  apply clarsimp
3666  done
3667
3668lemma mask_shift_and_negate[simp]:"(w && mask n << m) && ~~ (mask n << m) = 0"
3669  apply (clarsimp simp:mask_def)
3670  by (metis (erased, hide_lams) mask_eq_x_eq_0 shiftl_over_and_dist word_bool_alg.conj_absorb word_bw_assocs(1))
3671
3672lemma le_step_down_nat:"\<lbrakk>(i::nat) \<le> n; i = n \<longrightarrow> P; i \<le> n - 1 \<longrightarrow> P\<rbrakk> \<Longrightarrow> P"
3673  by arith
3674
3675lemma le_step_down_int:"\<lbrakk>(i::int) \<le> n; i = n \<longrightarrow> P; i \<le> n - 1 \<longrightarrow> P\<rbrakk> \<Longrightarrow> P"
3676  by arith
3677
3678lemma ex_mask_1[simp]: "(\<exists>x. mask x = 1)"
3679  apply (rule_tac x=1 in exI)
3680  apply (simp add:mask_def)
3681  done
3682
3683lemma not_switch:"~~ a = x \<Longrightarrow> a = ~~ x"
3684  by auto
3685
3686(* The seL4 bitfield generator produces functions containing mask and shift operations, such that
3687 * invoking two of them consecutively can produce something like the following.
3688 *)
3689lemma bitfield_op_twice:
3690  "(x && ~~ (mask n << m) || ((y && mask n) << m)) && ~~ (mask n << m) = x && ~~ (mask n << m)"
3691  by (induct n arbitrary: m) (auto simp: word_ao_dist)
3692
3693lemma bitfield_op_twice'':
3694  "\<lbrakk>~~ a = b << c; \<exists>x. b = mask x\<rbrakk> \<Longrightarrow> (x && a || (y && b << c)) && a = x && a"
3695  apply clarsimp
3696  apply (cut_tac n=xa and m=c and x=x and y=y in bitfield_op_twice)
3697  apply (clarsimp simp:mask_def)
3698  apply (drule not_switch)
3699  apply clarsimp
3700  done
3701
3702lemma bit_twiddle_min:
3703  "(y::'a::len word) xor (((x::'a::len word) xor y) && (if x < y then -1 else 0)) = min x y"
3704  by (metis (mono_tags) min_def word_bitwise_m1_simps(2) word_bool_alg.xor_left_commute
3705             word_bool_alg.xor_self word_bool_alg.xor_zero_right word_bw_comms(1)
3706             word_le_less_eq word_log_esimps(7))
3707
3708lemma bit_twiddle_max:
3709  "(x::'a::len word) xor (((x::'a::len word) xor y) && (if x < y then -1 else 0)) = max x y"
3710  by (metis (mono_tags) max_def word_bitwise_m1_simps(2) word_bool_alg.xor_left_self
3711            word_bool_alg.xor_zero_right word_bw_comms(1) word_le_less_eq word_log_esimps(7))
3712
3713lemma swap_with_xor:
3714  "\<lbrakk>(x::'a::len word) = a xor b; y = b xor x; z = x xor y\<rbrakk> \<Longrightarrow> z = b \<and> y = a"
3715  by (metis word_bool_alg.xor_assoc word_bool_alg.xor_commute word_bool_alg.xor_self
3716            word_bool_alg.xor_zero_right)
3717
3718lemma scast_nop1:
3719  "((scast ((of_int x)::('a::len) word))::'a sword) = of_int x"
3720  apply (clarsimp simp:scast_def word_of_int)
3721  by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse)
3722
3723lemma scast_nop2:
3724  "((scast ((of_int x)::('a::len) sword))::'a word) = of_int x"
3725  apply (clarsimp simp:scast_def word_of_int)
3726  by (metis len_signed sint_sbintrunc' word_sint.Rep_inverse)
3727
3728lemmas scast_nop[simp] = scast_nop1 scast_nop2 scast_id
3729
3730lemma le_mask_imp_and_mask:
3731  "(x::'a::len word) \<le> mask n \<Longrightarrow> x && mask n = x"
3732  by (metis and_mask_eq_iff_le_mask)
3733
3734lemma or_not_mask_nop:
3735  "((x::'a::len word) || ~~ (mask n)) && mask n = x && mask n"
3736  by (metis word_and_not word_ao_dist2 word_bw_comms(1) word_log_esimps(3))
3737
3738lemma mask_subsume:
3739  "\<lbrakk>n \<le> m\<rbrakk> \<Longrightarrow> ((x::'a::len word) || y && mask n) && ~~ (mask m) = x && ~~ (mask m)"
3740  apply (subst word_ao_dist)
3741  apply (subgoal_tac "(y && mask n) && ~~ (mask m) = 0")
3742   apply simp
3743  by (metis (no_types, hide_lams) is_aligned_mask is_aligned_weaken word_and_not
3744            word_bool_alg.conj_zero_right word_bw_comms(1) word_bw_lcs(1))
3745
3746lemma and_mask_0_iff_le_mask:
3747  fixes w :: "'a::len word"
3748  shows "(w && ~~(mask n) = 0) = (w \<le> mask n)"
3749  by (simp add: mask_eq_0_eq_x le_mask_imp_and_mask and_mask_eq_iff_le_mask)
3750
3751lemma mask_twice2:
3752  "n \<le> m \<Longrightarrow> ((x::'a::len word) && mask m) && mask n = x && mask n"
3753  by (metis mask_twice min_def)
3754
3755lemma uint_2_id:
3756  "LENGTH('a) \<ge> 2 \<Longrightarrow> uint (2::('a::len) word) = 2"
3757  apply clarsimp
3758  apply (subgoal_tac "2 \<in> uints (LENGTH('a))")
3759   apply (subst (asm) Word.word_ubin.set_iff_norm)
3760   apply simp
3761  apply (subst word_uint.set_iff_norm)
3762  apply clarsimp
3763  apply (rule int_mod_eq')
3764   apply simp
3765  apply (rule pow_2_gt)
3766  apply simp
3767  done
3768
3769lemma bintrunc_id:
3770  "\<lbrakk>m \<le> of_nat n; 0 < m\<rbrakk> \<Longrightarrow> bintrunc n m = m"
3771  by (simp add: bintrunc_mod2p le_less_trans)
3772
3773lemma shiftr1_unfold: "shiftr1 x = x >> 1"
3774  by (metis One_nat_def comp_apply funpow.simps(1) funpow.simps(2) id_apply shiftr_def)
3775
3776lemma shiftr1_is_div_2: "(x::('a::len) word) >> 1 = x div 2"
3777  apply (case_tac "LENGTH('a) = 1")
3778   apply simp
3779 apply (subgoal_tac "x = 0 \<or> x = 1")
3780    apply (erule disjE)
3781     apply (clarsimp simp:word_div_def)+
3782   apply (metis One_nat_def less_irrefl_nat sint_1_cases)
3783  apply clarsimp
3784  apply (subst word_div_def)
3785  apply clarsimp
3786  apply (subst bintrunc_id)
3787    apply (subgoal_tac "2 \<le> LENGTH('a)")
3788     apply simp
3789    apply (metis (no_types) le_0_eq le_SucE lens_not_0(2) not_less_eq_eq numeral_2_eq_2)
3790   apply simp
3791  apply (subst bin_rest_def[symmetric])
3792  apply (subst shiftr1_def[symmetric])
3793  apply (clarsimp simp:shiftr1_unfold)
3794  done
3795
3796lemma shiftl1_is_mult: "(x << 1) = (x :: 'a::len word) * 2"
3797  by (metis One_nat_def mult_2 mult_2_right one_add_one
3798        power_0 power_Suc shiftl_t2n)
3799
3800lemma div_of_0_id[simp]:"(0::('a::len) word) div n = 0"
3801  by (simp add: word_div_def)
3802
3803lemma degenerate_word:"LENGTH('a) = 1 \<Longrightarrow> (x::('a::len) word) = 0 \<or> x = 1"
3804  by (metis One_nat_def less_irrefl_nat sint_1_cases)
3805
3806lemma div_by_0_word:"(x::('a::len) word) div 0 = 0"
3807  by (metis div_0 div_by_0 unat_0 word_arith_nat_defs(6) word_div_1)
3808
3809lemma div_less_dividend_word:"\<lbrakk>x \<noteq> 0; n \<noteq> 1\<rbrakk> \<Longrightarrow> (x::('a::len) word) div n < x"
3810  apply (case_tac "n = 0")
3811   apply clarsimp
3812   apply (subst div_by_0_word)
3813   apply (simp add:word_neq_0_conv)
3814  apply (subst word_arith_nat_div)
3815  apply (rule word_of_nat_less)
3816  apply (rule div_less_dividend)
3817   apply (metis (poly_guards_query) One_nat_def less_one nat_neq_iff unat_eq_1(2) unat_eq_zero)
3818  apply (simp add:unat_gt_0)
3819  done
3820
3821lemma shiftr1_lt:"x \<noteq> 0 \<Longrightarrow> (x::('a::len) word) >> 1 < x"
3822  apply (subst shiftr1_is_div_2)
3823  apply (rule div_less_dividend_word)
3824   apply simp+
3825  done
3826
3827lemma word_less_div:
3828  fixes x :: "('a::len) word"
3829    and y :: "('a::len) word"
3830  shows "x div y = 0 \<Longrightarrow> y = 0 \<or> x < y"
3831  apply (case_tac "y = 0", clarsimp+)
3832  by (metis One_nat_def Suc_le_mono le0 le_div_geq not_less unat_0 unat_div unat_gt_0 word_less_nat_alt zero_less_one)
3833
3834lemma not_degenerate_imp_2_neq_0:"LENGTH('a) > 1 \<Longrightarrow> (2::('a::len) word) \<noteq> 0"
3835  by (metis numerals(1) power_not_zero power_zero_numeral)
3836
3837lemma shiftr1_0_or_1:"(x::('a::len) word) >> 1 = 0 \<Longrightarrow> x = 0 \<or> x = 1"
3838  apply (subst (asm) shiftr1_is_div_2)
3839  apply (drule word_less_div)
3840  apply (case_tac "LENGTH('a) = 1")
3841   apply (simp add:degenerate_word)
3842  apply (erule disjE)
3843   apply (subgoal_tac "(2::'a word) \<noteq> 0")
3844    apply simp
3845   apply (rule not_degenerate_imp_2_neq_0)
3846   apply (subgoal_tac "LENGTH('a) \<noteq> 0")
3847    apply arith
3848   apply simp
3849  apply (rule x_less_2_0_1', simp+)
3850  done
3851
3852lemma word_overflow:"(x::('a::len) word) + 1 > x \<or> x + 1 = 0"
3853  apply clarsimp
3854  by (metis diff_0 eq_diff_eq less_x_plus_1 max_word_max plus_1_less)
3855
3856lemma word_overflow_unat:"unat ((x::('a::len) word) + 1) = unat x + 1 \<or> x + 1 = 0"
3857  by (metis Suc_eq_plus1 add.commute unatSuc)
3858
3859lemma even_word_imp_odd_next:"even (unat (x::('a::len) word)) \<Longrightarrow> x + 1 = 0 \<or> odd (unat (x + 1))"
3860  apply (cut_tac x=x in word_overflow_unat)
3861  apply clarsimp
3862  done
3863
3864lemma odd_word_imp_even_next:"odd (unat (x::('a::len) word)) \<Longrightarrow> x + 1 = 0 \<or> even (unat (x + 1))"
3865  apply (cut_tac x=x in word_overflow_unat)
3866  apply clarsimp
3867  done
3868
3869lemma overflow_imp_lsb:"(x::('a::len) word) + 1 = 0 \<Longrightarrow> x !! 0"
3870  by (metis add.commute add_left_cancel max_word_max not_less word_and_1 word_bool_alg.conj_one_right word_bw_comms(1) word_overflow zero_neq_one)
3871
3872lemma word_lsb_nat:"lsb w = (unat w mod 2 = 1)"
3873  unfolding word_lsb_def bin_last_def
3874  by (metis (no_types, hide_lams) nat_mod_distrib nat_numeral not_mod_2_eq_1_eq_0 numeral_One uint_eq_0 uint_nonnegative unat_0 unat_def zero_le_numeral)
3875
3876lemma odd_iff_lsb:"odd (unat (x::('a::len) word)) = x !! 0"
3877  apply (simp add:even_iff_mod_2_eq_zero)
3878  apply (subst word_lsb_nat[unfolded One_nat_def, symmetric])
3879  apply (rule word_lsb_alt)
3880  done
3881
3882lemma of_nat_neq_iff_word:
3883      "x mod 2 ^ LENGTH('a) \<noteq> y mod 2 ^ LENGTH('a) \<Longrightarrow>
3884         (((of_nat x)::('a::len) word) \<noteq> of_nat y) = (x \<noteq> y)"
3885  apply (rule iffI)
3886   apply (case_tac "x = y")
3887   apply (subst (asm) of_nat_eq_iff[symmetric])
3888   apply simp+
3889  apply (case_tac "((of_nat x)::('a::len) word) = of_nat y")
3890   apply (subst (asm) word_unat.norm_eq_iff[symmetric])
3891   apply simp+
3892  done
3893
3894lemma shiftr1_irrelevant_lsb:"(x::('a::len) word) !! 0 \<or> x >> 1 = (x + 1) >> 1"
3895  apply (case_tac "LENGTH('a) = 1")
3896   apply clarsimp
3897   apply (drule_tac x=x in degenerate_word[unfolded One_nat_def])
3898   apply (erule disjE)
3899    apply clarsimp+
3900  apply (subst (asm) shiftr1_is_div_2[unfolded One_nat_def])+
3901  apply (subst (asm) word_arith_nat_div)+
3902  apply clarsimp
3903  apply (subst (asm) bintrunc_id)
3904    apply (subgoal_tac "LENGTH('a) > 0")
3905     apply linarith
3906    apply clarsimp+
3907  apply (subst (asm) bintrunc_id)
3908    apply (subgoal_tac "LENGTH('a) > 0")
3909     apply linarith
3910    apply clarsimp+
3911  apply (case_tac "x + 1 = 0")
3912   apply (clarsimp simp:overflow_imp_lsb)
3913  apply (cut_tac x=x in word_overflow_unat)
3914  apply clarsimp
3915  apply (case_tac "even (unat x)")
3916   apply (subgoal_tac "unat x div 2 = Suc (unat x) div 2")
3917    apply metis
3918   apply (subst numeral_2_eq_2)+
3919   apply simp
3920  apply (simp add:odd_iff_lsb)
3921  done
3922
3923lemma shiftr1_0_imp_only_lsb:"((x::('a::len) word) + 1) >> 1 = 0 \<Longrightarrow> x = 0 \<or> x + 1 = 0"
3924  by (metis One_nat_def shiftr1_0_or_1 word_less_1 word_overflow)
3925
3926lemma shiftr1_irrelevant_lsb':"\<not>((x::('a::len) word) !! 0) \<Longrightarrow> x >> 1 = (x + 1) >> 1"
3927  by (metis shiftr1_irrelevant_lsb)
3928
3929lemma lsb_this_or_next:"\<not>(((x::('a::len) word) + 1) !! 0) \<Longrightarrow> x !! 0"
3930  by (metis (poly_guards_query) even_word_imp_odd_next odd_iff_lsb overflow_imp_lsb)
3931
3932(* Perhaps this one should be a simp lemma, but it seems a little dangerous. *)
3933lemma cast_chunk_assemble_id:
3934  "\<lbrakk>n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\<rbrakk> \<Longrightarrow>
3935  (((ucast ((ucast (x::'b word))::'a word))::'b word) || (((ucast ((ucast (x >> n))::'a word))::'b word) << n)) = x"
3936  apply (subgoal_tac "((ucast ((ucast (x >> n))::'a word))::'b word) = x >> n")
3937   apply clarsimp
3938   apply (subst and_not_mask[symmetric])
3939   apply (subst ucast_ucast_mask)
3940   apply (subst word_ao_dist2[symmetric])
3941   apply clarsimp
3942  apply (rule ucast_ucast_len)
3943  apply (rule shiftr_less_t2n')
3944   apply (subst and_mask_eq_iff_le_mask)
3945   apply (clarsimp simp:mask_def)
3946   apply (metis max_word_eq max_word_max mult_2_right)
3947  apply (metis add_diff_cancel_left' diff_less len_gt_0 mult_2_right)
3948  done
3949
3950lemma cast_chunk_scast_assemble_id:
3951  "\<lbrakk>n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\<rbrakk> \<Longrightarrow>
3952  (((ucast ((scast (x::'b word))::'a word))::'b word) ||
3953   (((ucast ((scast (x >> n))::'a word))::'b word) << n)) = x"
3954  apply (subgoal_tac "((scast x)::'a word) = ((ucast x)::'a word)")
3955   apply (subgoal_tac "((scast (x >> n))::'a word) = ((ucast (x >> n))::'a word)")
3956    apply (simp add:cast_chunk_assemble_id)
3957   apply (subst down_cast_same[symmetric], subst is_down, arith, simp)+
3958  done
3959
3960lemma mask_or_not_mask:
3961  "x && mask n || x && ~~ (mask n) = x"
3962  apply (subst word_oa_dist, simp)
3963  apply (subst word_oa_dist2, simp)
3964  done
3965
3966lemma is_aligned_add_not_aligned:
3967  "\<lbrakk>is_aligned (p::'a::len word) n; \<not> is_aligned (q::'a::len word) n\<rbrakk> \<Longrightarrow> \<not> is_aligned (p + q) n"
3968  by (metis is_aligned_addD1)
3969
3970lemma word_gr0_conv_Suc: "(m::'a::len word) > 0 \<Longrightarrow> \<exists>n. m = n + 1"
3971  by (metis add.commute add_minus_cancel)
3972
3973lemma neg_mask_add_aligned:
3974  "\<lbrakk> is_aligned p n; q < 2 ^ n \<rbrakk> \<Longrightarrow> (p + q) && ~~ (mask n) = p && ~~ (mask n)"
3975  by (metis is_aligned_add_helper is_aligned_neg_mask_eq)
3976
3977lemma word_sless_sint_le:"x <s y \<Longrightarrow> sint x \<le> sint y - 1"
3978  by (metis word_sless_alt zle_diff1_eq)
3979
3980
3981lemma upper_trivial:
3982  fixes x :: "'a::len word"
3983  shows "x \<noteq> 2 ^ LENGTH('a) - 1 \<Longrightarrow> x < 2 ^ LENGTH('a) - 1"
3984  by (cut_tac n=x and 'a='a in max_word_max,
3985      clarsimp simp:max_word_def,
3986      simp add: less_le)
3987
3988lemma constraint_expand:
3989  fixes x :: "'a::len word"
3990  shows "x \<in> {y. lower \<le> y \<and> y \<le> upper} = (lower \<le> x \<and> x \<le> upper)"
3991  by (rule mem_Collect_eq)
3992
3993lemma card_map_elide:
3994  "card ((of_nat :: nat \<Rightarrow> 'a::len word) ` {0..<n}) = card {0..<n}"
3995    if "n \<le> CARD('a::len word)"
3996proof -
3997  let ?of_nat = "of_nat :: nat \<Rightarrow> 'a word"
3998  from word_unat.Abs_inj_on
3999  have "inj_on ?of_nat {i. i < CARD('a word)}"
4000    by (simp add: unats_def card_word)
4001  moreover have "{0..<n} \<subseteq> {i. i < CARD('a word)}"
4002    using that by auto
4003  ultimately have "inj_on ?of_nat {0..<n}"
4004    by (rule inj_on_subset)
4005  then show ?thesis
4006    by (simp add: card_image)
4007qed
4008
4009lemma card_map_elide2:
4010  "n \<le> CARD('a::len word) \<Longrightarrow> card ((of_nat::nat \<Rightarrow> 'a::len word) ` {0..<n}) = n"
4011  by (subst card_map_elide) clarsimp+
4012
4013lemma le_max_word_ucast_id:
4014  assumes "(x::'a::len word) \<le> ucast (max_word::'b::len word)"
4015  shows "ucast ((ucast x)::'b word) = x"
4016proof -
4017  {
4018    assume a1: "x \<le> word_of_int (uint (word_of_int (2 ^ LENGTH('b) - 1)::'b word))"
4019    have f2: "((\<exists>i ia. (0::int) \<le> i \<and> \<not> 0 \<le> i + - 1 * ia \<and> i mod ia \<noteq> i) \<or>
4020              \<not> (0::int) \<le> - 1 + 2 ^ LENGTH('b) \<or> (0::int) \<le> - 1 + 2 ^ LENGTH('b) + - 1 * 2 ^ LENGTH('b) \<or>
4021              (- (1::int) + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b) =
4022                - 1 + 2 ^ LENGTH('b)) = ((\<exists>i ia. (0::int) \<le> i \<and> \<not> 0 \<le> i + - 1 * ia \<and> i mod ia \<noteq> i) \<or>
4023              \<not> (1::int) \<le> 2 ^ LENGTH('b) \<or>
4024              2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)) = 1)"
4025      by force
4026    have f3: "\<forall>i ia. \<not> (0::int) \<le> i \<or> 0 \<le> i + - 1 * ia \<or> i mod ia = i"
4027      using mod_pos_pos_trivial by force
4028    have "(1::int) \<le> 2 ^ LENGTH('b)"
4029      by simp
4030    then have "2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ len_of TYPE ('b)) = 1"
4031      using f3 f2 by blast
4032    then have f4: "- (1::int) + 2 ^ LENGTH('b) = (- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)"
4033      by linarith
4034    have f5: "x \<le> word_of_int (uint (word_of_int (- 1 + 2 ^ LENGTH('b))::'b word))"
4035      using a1 by force
4036    have f6: "2 ^ LENGTH('b) + - (1::int) = - 1 + 2 ^ LENGTH('b)"
4037      by force
4038    have f7: "- (1::int) * 1 = - 1"
4039      by auto
4040    have "\<forall>x0 x1. (x1::int) - x0 = x1 + - 1 * x0"
4041      by force
4042    then have "x \<le> 2 ^ LENGTH('b) - 1"
4043      using f7 f6 f5 f4 by (metis uint_word_of_int wi_homs(2) word_arith_wis(8) word_of_int_2p)
4044  }
4045  with assms show ?thesis
4046    unfolding ucast_def
4047    apply (subst word_ubin.eq_norm)
4048    apply (subst and_mask_bintr[symmetric])
4049    apply (subst and_mask_eq_iff_le_mask)
4050    apply (clarsimp simp: max_word_def mask_def)
4051    done
4052qed
4053
4054lemma remdups_enum_upto:
4055  fixes s::"'a::len word"
4056  shows "remdups [s .e. e] = [s .e. e]"
4057  by simp
4058
4059lemma card_enum_upto:
4060  fixes s::"'a::len word"
4061  shows "card (set [s .e. e]) = Suc (unat e) - unat s"
4062  by (subst List.card_set) (simp add: remdups_enum_upto)
4063
4064lemma unat_mask:
4065  "unat (mask n :: 'a :: len word) = 2 ^ (min n (LENGTH('a))) - 1"
4066  apply (subst min.commute)
4067  apply (simp add: mask_def not_less min_def  split: if_split_asm)
4068  apply (intro conjI impI)
4069   apply (simp add: unat_sub_if_size)
4070   apply (simp add: power_overflow word_size)
4071  apply (simp add: unat_sub_if_size)
4072  done
4073
4074lemma word_shiftr_lt:
4075  fixes w :: "'a::len word"
4076  shows "unat (w >> n) < (2 ^ (LENGTH('a) - n))"
4077  apply (subst shiftr_div_2n')
4078  by (metis nat_mod_lem nat_zero_less_power_iff power_mod_div
4079      word_unat.Rep_inverse word_unat.eq_norm zero_less_numeral)
4080
4081lemma complement_nth_w2p:
4082  shows "n' < LENGTH('a) \<Longrightarrow> (~~ (2 ^ n :: 'a::len word)) !! n' = (n' \<noteq> n)"
4083  by (fastforce simp: word_ops_nth_size word_size nth_w2p)
4084
4085lemma word_unat_and_lt:
4086  "unat x < n \<or> unat y < n \<Longrightarrow> unat (x && y) < n"
4087  by (meson le_less_trans word_and_le1 word_and_le2 word_le_nat_alt)
4088
4089lemma word_unat_mask_lt:
4090  "m \<le> size w \<Longrightarrow> unat ((w::'a::len word) && mask m) < 2 ^ m"
4091  by (rule word_unat_and_lt) (simp add: unat_mask word_size)
4092
4093lemma unat_shiftr_less_t2n:
4094  fixes x :: "'a :: len word"
4095  shows "unat x < 2 ^ (n + m) \<Longrightarrow> unat (x >> n) < 2 ^ m"
4096  by (simp add: shiftr_div_2n' power_add mult.commute td_gal_lt)
4097
4098lemma le_or_mask:
4099  "w \<le> w' \<Longrightarrow> w || mask x \<le> w' || mask x"
4100  by (metis neg_mask_add_mask add.commute le_word_or1 mask_2pm1 neg_mask_mono_le word_plus_mono_left)
4101
4102lemma le_shiftr1':
4103  "\<lbrakk> shiftr1 u \<le> shiftr1 v ; shiftr1 u \<noteq> shiftr1 v \<rbrakk> \<Longrightarrow> u \<le> v"
4104  apply (unfold word_le_def shiftr1_def word_ubin.eq_norm)
4105  apply (unfold bin_rest_trunc_i
4106                trans [OF bintrunc_bintrunc_l word_ubin.norm_Rep,
4107                          unfolded word_ubin.norm_Rep,
4108                       OF order_refl [THEN le_SucI]])
4109  apply (case_tac "uint u" rule: bin_exhaust)
4110  apply (case_tac "uint v" rule: bin_exhaust)
4111  by (smt Bit_def bin_rest_BIT)
4112
4113lemma le_shiftr':
4114  "\<lbrakk> u >> n \<le> v >> n ; u >> n \<noteq> v >> n \<rbrakk> \<Longrightarrow> (u::'a::len0 word) \<le> v"
4115  apply (induct n; simp add: shiftr_def)
4116  apply (case_tac "(shiftr1 ^^ n) u = (shiftr1 ^^ n) v", simp)
4117  apply (fastforce dest: le_shiftr1')
4118  done
4119
4120lemma word_log2_nth_same:
4121  "w \<noteq> 0 \<Longrightarrow> w !! word_log2 w"
4122  unfolding word_log2_def
4123  using nth_length_takeWhile[where P=Not and xs="to_bl w"]
4124  apply (simp add: word_clz_def word_size to_bl_nth)
4125  apply (fastforce simp: linorder_not_less eq_zero_set_bl
4126                   dest: takeWhile_take_has_property)
4127  done
4128
4129lemma word_log2_nth_not_set:
4130  "\<lbrakk> word_log2 w < i ; i < size w \<rbrakk> \<Longrightarrow> \<not> w !! i"
4131  unfolding word_log2_def word_clz_def
4132  using takeWhile_take_has_property_nth[where P=Not and xs="to_bl w" and n="size w - Suc i"]
4133  by (fastforce simp add: to_bl_nth word_size)
4134
4135lemma word_log2_highest:
4136  assumes a: "w !! i"
4137  shows "i \<le> word_log2 w"
4138proof -
4139  from a have "i < size w" by - (rule test_bit_size)
4140  with a show ?thesis
4141    by - (rule ccontr, simp add: word_log2_nth_not_set)
4142qed
4143
4144lemma word_log2_max:
4145  "word_log2 w < size w"
4146  unfolding word_log2_def word_clz_def
4147  by simp
4148
4149lemma word_clz_0[simp]:
4150  "word_clz (0::'a::len word) = LENGTH('a)"
4151  unfolding word_clz_def
4152  by (simp add: takeWhile_replicate)
4153
4154lemma word_clz_minus_one[simp]:
4155  "word_clz (-1::'a::len word) = 0"
4156  unfolding word_clz_def
4157  by (simp add: max_word_bl[simplified max_word_minus] takeWhile_replicate)
4158
4159lemma word_add_no_overflow:"(x::'a::len word) < max_word \<Longrightarrow> x < x + 1"
4160  using less_x_plus_1 order_less_le by blast
4161
4162lemma lt_plus_1_le_word:
4163  fixes x :: "'a::len word"
4164  assumes bound:"n < unat (maxBound::'a word)"
4165  shows "x < 1 + of_nat n = (x \<le> of_nat n)"
4166  by (metis add.commute bound max_word_max word_Suc_leq word_not_le word_of_nat_less)
4167
4168lemma unat_ucast_up_simp:
4169  fixes x :: "'a::len word"
4170  assumes "LENGTH('a) \<le> LENGTH('b)"
4171  shows "unat (ucast x :: 'b::len word) = unat x"
4172  unfolding ucast_def unat_def
4173  apply (subst int_word_uint)
4174  apply (subst mod_pos_pos_trivial; simp?)
4175  apply (rule lt2p_lem)
4176  apply (simp add: assms)
4177  done
4178
4179lemma unat_ucast_less_no_overflow:
4180  "\<lbrakk>n < 2 ^ LENGTH('a); unat f < n\<rbrakk> \<Longrightarrow> (f::('a::len) word) < of_nat n"
4181  by (erule (1)  order_le_less_trans[OF _ of_nat_mono_maybe,rotated]) simp
4182
4183lemma unat_ucast_less_no_overflow_simp:
4184  "n < 2 ^ LENGTH('a) \<Longrightarrow> (unat f < n) = ((f::('a::len) word) < of_nat n)"
4185  using unat_less_helper unat_ucast_less_no_overflow by blast
4186
4187lemma unat_ucast_no_overflow_le:
4188  assumes no_overflow: "unat b < (2 :: nat) ^ LENGTH('a)"
4189  and upward_cast: "LENGTH('a) < LENGTH('b)"
4190  shows "(ucast (f::'a::len word) < (b :: 'b :: len word)) = (unat f < unat b)"
4191proof -
4192  have LR: "ucast f < b \<Longrightarrow> unat f < unat b"
4193    apply (rule unat_less_helper)
4194    apply (simp add:ucast_nat_def)
4195    apply (rule_tac 'b1 = 'b in  ucast_less_ucast[OF order.strict_implies_order, THEN iffD1])
4196     apply (rule upward_cast)
4197    apply (simp add: ucast_ucast_mask less_mask_eq word_less_nat_alt
4198                     unat_power_lower[OF upward_cast] no_overflow)
4199    done
4200  have RL: "unat f < unat b \<Longrightarrow> ucast f < b"
4201  proof-
4202    assume ineq: "unat f < unat b"
4203    have "ucast (f::'a::len word) < ((ucast (ucast b ::'a::len word)) :: 'b :: len word)"
4204      apply (simp add: ucast_less_ucast[OF order.strict_implies_order] upward_cast)
4205      apply (simp add: ucast_nat_def[symmetric])
4206      apply (rule unat_ucast_less_no_overflow[OF no_overflow ineq])
4207      done
4208    then show ?thesis
4209      apply (rule order_less_le_trans)
4210      apply (simp add:ucast_ucast_mask word_and_le2)
4211      done
4212  qed
4213  then show ?thesis by (simp add:RL LR iffI)
4214qed
4215
4216lemmas ucast_up_mono = ucast_less_ucast[THEN iffD2]
4217
4218(* casting a long word to a shorter word and casting back to the long word
4219   is equal to the original long word -- if the word is small enough.
4220  'l is the longer word.
4221  's is the shorter word.
4222*)
4223lemma bl_cast_long_short_long_ingoreLeadingZero_generic:
4224  "\<lbrakk> length (dropWhile Not (to_bl w)) \<le> LENGTH('s); LENGTH('s) \<le> LENGTH('l) \<rbrakk> \<Longrightarrow>
4225   (of_bl :: _ \<Rightarrow> 'l::len word) (to_bl ((of_bl::_ \<Rightarrow> 's::len word) (to_bl w))) = w"
4226  by (rule word_uint_eqI) (simp add: uint_of_bl_is_bl_to_bin uint_of_bl_is_bl_to_bin_drop)
4227
4228(*
4229 Casting between longer and shorter word.
4230  'l is the longer word.
4231  's is the shorter word.
4232 For example: 'l::len word is 128 word (full ipv6 address)
4233              's::len word is 16 word (address piece of ipv6 address in colon-text-representation)
4234*)
4235corollary ucast_short_ucast_long_ingoreLeadingZero:
4236  "\<lbrakk> length (dropWhile Not (to_bl w)) \<le> LENGTH('s); LENGTH('s) \<le> LENGTH('l) \<rbrakk> \<Longrightarrow>
4237   (ucast:: 's::len word \<Rightarrow> 'l::len word) ((ucast:: 'l::len word \<Rightarrow> 's::len word) w) = w"
4238  apply (subst Word.ucast_bl)+
4239  apply (rule bl_cast_long_short_long_ingoreLeadingZero_generic; simp)
4240  done
4241
4242lemma length_drop_mask:
4243  fixes w::"'a::len word"
4244  shows "length (dropWhile Not (to_bl (w AND mask n))) \<le> n"
4245proof -
4246  have "length (takeWhile Not (replicate n False @ ls)) = n + length (takeWhile Not ls)"
4247    for ls n by(subst takeWhile_append2) simp+
4248  then show ?thesis
4249    unfolding bl_and_mask by (simp add: dropWhile_eq_drop)
4250qed
4251
4252lemma minus_one_word:
4253  "(-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1"
4254  by simp
4255
4256lemma mask_exceed:
4257  "n \<ge> LENGTH('a) \<Longrightarrow> (x::'a::len word) && ~~ (mask n) = 0"
4258  by (simp add: and_not_mask shiftr_eq_0)
4259
4260lemma two_power_strict_part_mono:
4261  "strict_part_mono {..LENGTH('a) - 1} (\<lambda>x. (2 :: 'a :: len word) ^ x)"
4262proof -
4263  { fix n
4264    have "n < LENGTH('a) \<Longrightarrow> strict_part_mono {..n} (\<lambda>x. (2 :: 'a :: len word) ^ x)"
4265    proof (induct n)
4266      case 0 then show ?case by simp
4267    next
4268      case (Suc n)
4269      from Suc.prems
4270      have "2 ^ n < (2 :: 'a :: len word) ^ Suc n"
4271        using power_strict_increasing unat_power_lower word_less_nat_alt by fastforce
4272      with Suc
4273      show ?case by (subst strict_part_mono_by_steps) simp
4274    qed
4275  }
4276  then show ?thesis by simp
4277qed
4278
4279lemma word_shift_by_2:
4280  "x * 4 = (x::'a::len word) << 2"
4281  by (simp add: shiftl_t2n)
4282
4283lemma le_2p_upper_bits:
4284  "\<lbrakk> (p::'a::len word) \<le> 2^n - 1; n < LENGTH('a) \<rbrakk> \<Longrightarrow>
4285  \<forall>n'\<ge>n. n' < LENGTH('a) \<longrightarrow> \<not> p !! n'"
4286  by (subst upper_bits_unset_is_l2p; simp)
4287
4288lemma le2p_bits_unset:
4289  "p \<le> 2 ^ n - 1 \<Longrightarrow> \<forall>n'\<ge>n. n' < LENGTH('a) \<longrightarrow> \<not> (p::'a::len word) !! n'"
4290  using upper_bits_unset_is_l2p [where p=p]
4291  by (cases "n < LENGTH('a)") auto
4292
4293lemma ucast_less_shiftl_helper:
4294  "\<lbrakk> LENGTH('b) + 2 < LENGTH('a); 2 ^ (LENGTH('b) + 2) \<le> n\<rbrakk>
4295    \<Longrightarrow> (ucast (x :: 'b::len word) << 2) < (n :: 'a::len word)"
4296  apply (erule order_less_le_trans[rotated])
4297  using ucast_less[where x=x and 'a='a]
4298  apply (simp only: shiftl_t2n field_simps)
4299  apply (rule word_less_power_trans2; simp)
4300  done
4301
4302lemma word_power_nonzero:
4303  "\<lbrakk> (x :: 'a::len word) < 2 ^ (LENGTH('a) - n); n < LENGTH('a); x \<noteq> 0 \<rbrakk>
4304  \<Longrightarrow> x * 2 ^ n \<noteq> 0"
4305  by (metis and_mask_eq_iff_shiftr_0 less_mask_eq p2_gt_0 semiring_normalization_rules(7)
4306            shiftl_shiftr_id shiftl_t2n)
4307
4308lemma less_1_helper:
4309  "n \<le> m \<Longrightarrow> (n - 1 :: int) < m"
4310  by arith
4311
4312lemma div_power_helper:
4313  "\<lbrakk> x \<le> y; y < LENGTH('a) \<rbrakk> \<Longrightarrow> (2 ^ y - 1) div (2 ^ x :: 'a::len word) = 2 ^ (y - x) - 1"
4314  apply (rule word_uint.Rep_eqD)
4315  apply (simp only: uint_word_ariths uint_div uint_power_lower)
4316  apply (subst mod_pos_pos_trivial, fastforce, fastforce)+
4317  apply (subst mod_pos_pos_trivial)
4318    apply (simp add: le_diff_eq uint_2p_alt)
4319   apply (rule less_1_helper)
4320   apply (rule power_increasing; simp)
4321  apply (subst mod_pos_pos_trivial)
4322    apply (simp add: uint_2p_alt)
4323   apply (rule less_1_helper)
4324   apply (rule power_increasing; simp)
4325  apply (subst int_div_sub_1; simp add: uint_2p_alt)
4326  apply (subst power_0[symmetric])
4327  apply (simp add: uint_2p_alt le_imp_power_dvd power_sub_int)
4328  done
4329
4330
4331lemma word_add_power_off:
4332  fixes a :: "'a :: len word"
4333  assumes ak: "a < k"
4334  and kw: "k < 2 ^ (LENGTH('a) - m)"
4335  and mw: "m < LENGTH('a)"
4336  and off: "off < 2 ^ m"
4337  shows "(a * 2 ^ m) + off < k * 2 ^ m"
4338proof (cases "m = 0")
4339  case True
4340  then show ?thesis using off ak by simp
4341next
4342  case False
4343
4344  from ak have ak1: "a + 1 \<le> k" by (rule inc_le)
4345  then have "(a + 1) * 2 ^ m \<noteq> 0"
4346    apply -
4347    apply (rule word_power_nonzero)
4348    apply (erule order_le_less_trans  [OF _ kw])
4349    apply (rule mw)
4350    apply (rule less_is_non_zero_p1 [OF ak])
4351    done
4352  then have "(a * 2 ^ m) + off < ((a + 1) * 2 ^ m)" using kw mw
4353    apply -
4354    apply (simp add: distrib_right)
4355    apply (rule word_plus_strict_mono_right [OF off])
4356    apply (rule is_aligned_no_overflow'')
4357    apply (rule is_aligned_mult_triv2)
4358    apply assumption
4359    done
4360  also have "\<dots> \<le> k * 2 ^ m" using ak1 mw kw False
4361    apply -
4362    apply (erule word_mult_le_mono1)
4363    apply (simp add: p2_gt_0)
4364    apply (simp add: word_less_nat_alt)
4365    apply (rule nat_less_power_trans2[simplified])
4366    apply (simp add: word_less_nat_alt)
4367    apply simp
4368    done
4369  finally show ?thesis .
4370qed
4371
4372lemma offset_not_aligned:
4373  "\<lbrakk> is_aligned (p::'a::len word) n; i > 0; i < 2 ^ n; n < LENGTH('a)\<rbrakk> \<Longrightarrow>
4374   \<not> is_aligned (p + of_nat i) n"
4375  apply (erule is_aligned_add_not_aligned)
4376  unfolding is_aligned_def
4377  by (metis le_unat_uoi nat_dvd_not_less order_less_imp_le unat_power_lower)
4378
4379lemma length_upto_enum_one:
4380  fixes x :: "'a :: len word"
4381  assumes lt1: "x < y" and lt2: "z < y" and lt3: "x \<le> z"
4382  shows "[x , y .e. z] = [x]"
4383  unfolding upto_enum_step_def
4384proof (subst upto_enum_red, subst if_not_P [OF leD [OF lt3]], clarsimp, rule conjI)
4385  show "unat ((z - x) div (y - x)) = 0"
4386  proof (subst unat_div, rule div_less)
4387    have syx: "unat (y - x) = unat y - unat x"
4388      by (rule unat_sub [OF order_less_imp_le]) fact
4389    moreover have "unat (z - x) = unat z - unat x"
4390      by (rule unat_sub) fact
4391
4392    ultimately show "unat (z - x) < unat (y - x)"
4393      using lt2 lt3 unat_mono word_less_minus_mono_left by blast
4394  qed
4395
4396  then show "(z - x) div (y - x) * (y - x) = 0"
4397    by (metis mult_zero_left unat_0 word_unat.Rep_eqD)
4398qed
4399
4400lemma max_word_mask:
4401  "(max_word :: 'a::len word) = mask LENGTH('a)"
4402  unfolding mask_def by (metis max_word_eq shiftl_1)
4403
4404lemmas mask_len_max = max_word_mask[symmetric]
4405
4406lemma is_aligned_alignUp[simp]:
4407  "is_aligned (alignUp p n) n"
4408  by (simp add: alignUp_def complement_def is_aligned_mask mask_def word_bw_assocs)
4409
4410lemma alignUp_le[simp]:
4411  "alignUp p n \<le> p + 2 ^ n - 1"
4412  unfolding alignUp_def by (rule word_and_le2)
4413
4414lemma complement_mask:
4415  "complement (2 ^ n - 1) = ~~ (mask n)"
4416  unfolding complement_def mask_def by simp
4417
4418lemma alignUp_idem:
4419  fixes a :: "'a::len word"
4420  assumes "is_aligned a n" "n < LENGTH('a)"
4421  shows "alignUp a n = a"
4422  using assms unfolding alignUp_def
4423  by (metis complement_mask is_aligned_add_helper p_assoc_help power_2_ge_iff)
4424
4425lemma alignUp_not_aligned_eq:
4426  fixes a :: "'a :: len word"
4427  assumes al: "\<not> is_aligned a n"
4428  and     sz: "n < LENGTH('a)"
4429  shows   "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n"
4430proof -
4431  have anz: "a mod 2 ^ n \<noteq> 0"
4432    by (rule not_aligned_mod_nz) fact+
4433
4434  then have um: "unat (a mod 2 ^ n - 1) div 2 ^ n = 0" using sz
4435    by (meson Euclidean_Division.div_eq_0_iff le_m1_iff_lt measure_unat order_less_trans
4436              unat_less_power word_less_sub_le word_mod_less_divisor)
4437
4438  have "a + 2 ^ n - 1 = (a div 2 ^ n) * 2 ^ n + (a mod 2 ^ n) + 2 ^ n - 1"
4439    by (simp add: word_mod_div_equality)
4440  also have "\<dots> = (a mod 2 ^ n - 1) + (a div 2 ^ n + 1) * 2 ^ n"
4441    by (simp add: field_simps)
4442  finally show "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n" using sz
4443    unfolding alignUp_def
4444    apply (subst complement_mask)
4445    apply (erule ssubst)
4446    apply (subst neg_mask_is_div)
4447    apply (simp add: word_arith_nat_div)
4448    apply (subst unat_word_ariths(1) unat_word_ariths(2))+
4449    apply (subst uno_simps)
4450    apply (subst unat_1)
4451    apply (subst mod_add_right_eq)
4452    apply simp
4453    apply (subst power_mod_div)
4454    apply (subst div_mult_self1)
4455     apply simp
4456    apply (subst um)
4457    apply simp
4458    apply (subst mod_mod_power)
4459    apply simp
4460    apply (subst word_unat_power, subst Abs_fnat_hom_mult)
4461    apply (subst mult_mod_left)
4462    apply (subst power_add [symmetric])
4463    apply simp
4464    apply (subst Abs_fnat_hom_1)
4465    apply (subst Abs_fnat_hom_add)
4466    apply (subst word_unat_power, subst Abs_fnat_hom_mult)
4467    apply (subst word_unat.Rep_inverse[symmetric], subst Abs_fnat_hom_mult)
4468    apply simp
4469    done
4470qed
4471
4472lemma alignUp_ge:
4473  fixes a :: "'a :: len word"
4474  assumes sz: "n < LENGTH('a)"
4475  and nowrap: "alignUp a n \<noteq> 0"
4476  shows "a \<le> alignUp a n"
4477proof (cases "is_aligned a n")
4478  case True
4479  then show ?thesis using sz
4480    by (subst alignUp_idem, simp_all)
4481next
4482  case False
4483
4484  have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz
4485    by (metis shiftr_div_2n' word_shiftr_lt)
4486
4487  have"2 ^ n * (unat a div 2 ^ n + 1) \<le> 2 ^ LENGTH('a)" using sz
4488    by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right lt0 nat_le_power_trans nat_less_le)
4489  moreover have "2 ^ n * (unat a div 2 ^ n + 1) \<noteq> 2 ^ LENGTH('a)" using nowrap sz
4490    apply -
4491    apply (erule contrapos_nn)
4492    apply (subst alignUp_not_aligned_eq [OF False sz])
4493    apply (subst unat_arith_simps)
4494    apply (subst unat_word_ariths)
4495    apply (subst unat_word_ariths)
4496    apply simp
4497    apply (subst mult_mod_left)
4498    apply (simp add: unat_div field_simps power_add[symmetric] mod_mod_power min.absorb2)
4499    done
4500  ultimately have lt: "2 ^ n * (unat a div 2 ^ n + 1) < 2 ^ LENGTH('a)" by simp
4501
4502  have "a = a div 2 ^ n * 2 ^ n + a mod 2 ^ n" by (rule word_mod_div_equality [symmetric])
4503  also have "\<dots> < (a div 2 ^ n + 1) * 2 ^ n" using sz lt
4504    apply (simp add: field_simps)
4505    apply (rule word_add_less_mono1)
4506    apply (rule word_mod_less_divisor)
4507    apply (simp add: word_less_nat_alt)
4508    apply (subst unat_word_ariths)
4509    apply (simp add: unat_div)
4510    done
4511  also have "\<dots> =  alignUp a n"
4512    by (rule alignUp_not_aligned_eq [symmetric]) fact+
4513  finally show ?thesis by (rule order_less_imp_le)
4514qed
4515
4516lemma alignUp_le_greater_al:
4517  fixes x :: "'a :: len word"
4518  assumes le: "a \<le> x"
4519  and     sz: "n < LENGTH('a)"
4520  and     al: "is_aligned x n"
4521  shows   "alignUp a n \<le> x"
4522proof (cases "is_aligned a n")
4523  case True
4524  then show ?thesis using sz le by (simp add: alignUp_idem)
4525next
4526  case False
4527
4528  then have anz: "a mod 2 ^ n \<noteq> 0"
4529    by (rule not_aligned_mod_nz)
4530
4531  from al obtain k where xk: "x = 2 ^ n * of_nat k" and kv: "k < 2 ^ (LENGTH('a) - n)"
4532    by (auto elim!: is_alignedE)
4533
4534  then have kn: "unat (of_nat k :: 'a word) * unat ((2::'a word) ^ n) < 2 ^ LENGTH('a)"
4535    using sz
4536    apply (subst unat_of_nat_eq)
4537     apply (erule order_less_le_trans)
4538     apply simp
4539    apply (subst mult.commute)
4540    apply simp
4541    apply (rule nat_less_power_trans)
4542     apply simp
4543    apply simp
4544    done
4545
4546  have au: "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n"
4547    by (rule alignUp_not_aligned_eq) fact+
4548  also have "\<dots> \<le> of_nat k * 2 ^ n"
4549  proof (rule word_mult_le_mono1 [OF inc_le _ kn])
4550    show "a div 2 ^ n < of_nat k" using kv xk le sz anz
4551      by (simp add: alignUp_div_helper)
4552
4553    show "(0:: 'a word) < 2 ^ n" using sz by (simp add: p2_gt_0 sz)
4554  qed
4555
4556  finally show ?thesis using xk by (simp add: field_simps)
4557qed
4558
4559lemma alignUp_is_aligned_nz:
4560  fixes a :: "'a :: len word"
4561  assumes al: "is_aligned x n"
4562  and     sz: "n < LENGTH('a)"
4563  and     ax: "a \<le> x"
4564  and     az: "a \<noteq> 0"
4565  shows   "alignUp (a::'a :: len word) n \<noteq> 0"
4566proof (cases "is_aligned a n")
4567  case True
4568  then have "alignUp a n = a" using sz by (simp add: alignUp_idem)
4569  then show ?thesis using az by simp
4570next
4571  case False
4572  then have anz: "a mod 2 ^ n \<noteq> 0"
4573    by (rule not_aligned_mod_nz)
4574
4575  {
4576    assume asm: "alignUp a n = 0"
4577
4578    have lt0: "unat a div 2 ^ n < 2 ^ (LENGTH('a) - n)" using sz
4579      by (metis shiftr_div_2n' word_shiftr_lt)
4580
4581    have leq: "2 ^ n * (unat a div 2 ^ n + 1) \<le> 2 ^ LENGTH('a)" using sz
4582      by (metis One_nat_def Suc_leI add.right_neutral add_Suc_right lt0 nat_le_power_trans
4583                order_less_imp_le)
4584
4585    from al obtain k where  kv: "k < 2 ^ (LENGTH('a) - n)" and xk: "x = 2 ^ n * of_nat k"
4586      by (auto elim!: is_alignedE)
4587
4588    then have "a div 2 ^ n < of_nat k" using ax sz anz
4589      by (rule alignUp_div_helper)
4590
4591    then have r: "unat a div 2 ^ n < k" using sz
4592      by (metis unat_div unat_less_helper unat_power_lower)
4593
4594    have "alignUp a n = (a div 2 ^ n + 1) * 2 ^ n"
4595      by (rule alignUp_not_aligned_eq) fact+
4596
4597    then have "\<dots> = 0" using asm by simp
4598    then have "2 ^ LENGTH('a) dvd 2 ^ n * (unat a div 2 ^ n + 1)"
4599      using sz by (simp add: unat_arith_simps ac_simps)
4600                  (simp add: unat_word_ariths mod_simps mod_eq_0_iff_dvd)
4601    with leq have "2 ^ n * (unat a div 2 ^ n + 1) = 2 ^ LENGTH('a)"
4602      by (force elim!: le_SucE)
4603    then have "unat a div 2 ^ n = 2 ^ LENGTH('a) div 2 ^ n - 1"
4604      by (metis (no_types, hide_lams) Groups.add_ac(2) add.right_neutral
4605                add_diff_cancel_left' div_le_dividend div_mult_self4 gr_implies_not0
4606                le_neq_implies_less power_eq_0_iff unat_def zero_neq_numeral)
4607    then have "unat a div 2 ^ n = 2 ^ (LENGTH('a) - n) - 1"
4608      using sz by (simp add: power_sub)
4609    then have "2 ^ (LENGTH('a) - n) - 1 < k" using r
4610      by simp
4611    then have False using kv by simp
4612  } then show ?thesis by clarsimp
4613qed
4614
4615lemma alignUp_ar_helper:
4616  fixes a :: "'a :: len word"
4617  assumes al: "is_aligned x n"
4618  and     sz: "n < LENGTH('a)"
4619  and    sub: "{x..x + 2 ^ n - 1} \<subseteq> {a..b}"
4620  and    anz: "a \<noteq> 0"
4621  shows "a \<le> alignUp a n \<and> alignUp a n + 2 ^ n - 1 \<le> b"
4622proof
4623  from al have xl: "x \<le> x + 2 ^ n - 1" by (simp add: is_aligned_no_overflow)
4624
4625  from xl sub have ax: "a \<le> x"
4626    by (clarsimp elim!: range_subset_lower [where x = x])
4627
4628  show "a \<le> alignUp a n"
4629  proof (rule alignUp_ge)
4630    show "alignUp a n \<noteq> 0" using al sz ax anz
4631      by (rule alignUp_is_aligned_nz)
4632  qed fact+
4633
4634  show "alignUp a n + 2 ^ n - 1 \<le> b"
4635  proof (rule order_trans)
4636    from xl show tp: "x + 2 ^ n - 1 \<le> b" using sub
4637      by (clarsimp elim!: range_subset_upper [where x = x])
4638
4639    from ax have "alignUp a n \<le> x"
4640      by (rule alignUp_le_greater_al) fact+
4641    then have "alignUp a n + (2 ^ n - 1) \<le> x + (2 ^ n - 1)"
4642      using xl al is_aligned_no_overflow' olen_add_eqv word_plus_mcs_3 by blast
4643    then show "alignUp a n + 2 ^ n - 1 \<le> x + 2 ^ n - 1"
4644      by (simp add: field_simps)
4645  qed
4646qed
4647
4648lemma alignUp_def2:
4649  "alignUp a sz = a + 2 ^ sz - 1 && ~~ (mask sz)"
4650  unfolding alignUp_def[unfolded complement_def]
4651  by (simp add:mask_def[symmetric,unfolded shiftl_t2n,simplified])
4652
4653lemma mask_out_first_mask_some:
4654  "\<lbrakk> x && ~~ (mask n) = y; n \<le> m \<rbrakk> \<Longrightarrow> x && ~~ (mask m) = y && ~~ (mask m)"
4655  by word_eqI_solve
4656
4657lemma gap_between_aligned:
4658  "\<lbrakk>a < (b :: 'a ::len word); is_aligned a n; is_aligned b n; n < LENGTH('a) \<rbrakk>
4659  \<Longrightarrow> a + (2^n - 1) < b"
4660  by (simp add: aligned_add_offset_less)
4661
4662lemma mask_out_add_aligned:
4663  assumes al: "is_aligned p n"
4664  shows "p + (q && ~~ (mask n)) = (p + q) && ~~ (mask n)"
4665  using mask_add_aligned [OF al]
4666  by (simp add: mask_out_sub_mask)
4667
4668lemma alignUp_def3:
4669  "alignUp a sz = 2^ sz + (a - 1 && ~~ (mask sz))"
4670  by (simp add: alignUp_def2 is_aligned_triv field_simps mask_out_add_aligned)
4671
4672lemma  alignUp_plus:
4673  "is_aligned w us \<Longrightarrow> alignUp (w + a) us  = w + alignUp a us"
4674  by (clarsimp simp: alignUp_def2 mask_out_add_aligned field_simps)
4675
4676lemma mask_lower_twice:
4677  "n \<le> m \<Longrightarrow> (x && ~~ (mask n)) && ~~ (mask m) = x && ~~ (mask m)"
4678  by word_eqI_solve
4679
4680lemma mask_lower_twice2:
4681  "(a && ~~ (mask n)) && ~~ (mask m) = a && ~~ (mask (max n m))"
4682  by word_eqI_solve
4683
4684lemma ucast_and_neg_mask:
4685  "ucast (x && ~~ (mask n)) = ucast x && ~~ (mask n)"
4686  by word_eqI_solve
4687
4688lemma ucast_and_mask:
4689  "ucast (x && mask n) = ucast x && mask n"
4690  by word_eqI_solve
4691
4692lemma ucast_mask_drop:
4693  "LENGTH('a :: len) \<le> n \<Longrightarrow> (ucast (x && mask n) :: 'a word) = ucast x"
4694  by word_eqI
4695
4696lemma alignUp_distance:
4697  "alignUp (q :: 'a :: len word) sz - q \<le> mask sz"
4698  by (metis (no_types) add.commute add_diff_cancel_left alignUp_def2 diff_add_cancel
4699                       mask_2pm1 subtract_mask(2) word_and_le1 word_sub_le_iff)
4700
4701lemma is_aligned_diff_neg_mask:
4702  "is_aligned p sz \<Longrightarrow> (p - q && ~~ (mask sz)) = (p - ((alignUp q sz) && ~~ (mask sz)))"
4703  apply (clarsimp simp only:word_and_le2 diff_conv_add_uminus)
4704  apply (subst mask_out_add_aligned[symmetric]; simp)
4705  apply (rule sum_to_zero)
4706  apply (subst add.commute)
4707  by (simp add: alignUp_distance and_mask_0_iff_le_mask is_aligned_neg_mask_eq mask_out_add_aligned)
4708
4709lemma map_bits_rev_to_bl:
4710  "map ((!!) x) [0..<size x] = rev (to_bl x)"
4711  by (auto simp: list_eq_iff_nth_eq test_bit_bl word_size)
4712
4713(* negating a mask which has been shifted to the very left *)
4714lemma NOT_mask_shifted_lenword:
4715  "~~ (mask len << (LENGTH('a) - len) ::'a::len word) = mask (LENGTH('a) - len)"
4716  by (rule Word.word_bool_alg.compl_unique, simp) word_eqI_solve
4717
4718(* Comparisons between different word sizes. *)
4719
4720lemma eq_ucast_ucast_eq:
4721  "LENGTH('b) \<le> LENGTH('a) \<Longrightarrow> x = ucast y \<Longrightarrow> ucast x = y"
4722  for x :: "'a::len word" and y :: "'b::len word"
4723  by (simp add: is_down ucast_id ucast_ucast_a)
4724
4725lemma le_ucast_ucast_le:
4726  "x \<le> ucast y \<Longrightarrow> ucast x \<le> y"
4727  for x :: "'a::len word" and y :: "'b::len word"
4728  by (smt le_unat_uoi linorder_not_less order_less_imp_le ucast_nat_def unat_arith_simps(1))
4729
4730lemma less_ucast_ucast_less:
4731  "LENGTH('b) \<le> LENGTH('a) \<Longrightarrow> x < ucast y \<Longrightarrow> ucast x < y"
4732  for x :: "'a::len word" and y :: "'b::len word"
4733  by (metis ucast_nat_def unat_mono unat_ucast_up_simp word_of_nat_less)
4734
4735lemma ucast_le_ucast:
4736  "LENGTH('a) \<le> LENGTH('b) \<Longrightarrow> (ucast x \<le> (ucast y::'b::len word)) = (x \<le> y)"
4737  for x :: "'a::len word"
4738  by (simp add: unat_arith_simps(1) unat_ucast_up_simp)
4739
4740lemmas ucast_up_mono_le = ucast_le_ucast[THEN iffD2]
4741
4742lemma ucast_le_ucast_eq:
4743  fixes x y :: "'a::len word"
4744  assumes x: "x < 2 ^ n"
4745  assumes y: "y < 2 ^ n"
4746  assumes n: "n = LENGTH('b::len)"
4747  shows "(UCAST('a \<rightarrow> 'b) x \<le> UCAST('a \<rightarrow> 'b) y) = (x \<le> y)"
4748  apply (rule iffI)
4749   apply (cases "LENGTH('b) < LENGTH('a)")
4750    apply (subst less_mask_eq[OF x, symmetric])
4751    apply (subst less_mask_eq[OF y, symmetric])
4752    apply (unfold n)
4753    apply (subst ucast_ucast_mask[symmetric])+
4754    apply (simp add: ucast_le_ucast)+
4755  apply (erule ucast_mono_le[OF _ y[unfolded n]])
4756  done
4757
4758lemma word_le_not_less:
4759  "((b::'a::len word) \<le> a) = (\<not>(a < b))"
4760  by fastforce
4761
4762lemma ucast_or_distrib:
4763  fixes x :: "'a::len word"
4764  fixes y :: "'a::len word"
4765  shows "(ucast (x || y) :: ('b::len) word) = ucast x || ucast y"
4766  unfolding ucast_def Word.bitOR_word.abs_eq uint_or
4767  by blast
4768
4769lemma shiftr_less:
4770  "(w::'a::len word) < k \<Longrightarrow> w >> n < k"
4771  by (metis div_le_dividend le_less_trans shiftr_div_2n' unat_arith_simps(2))
4772
4773lemma word_and_notzeroD:
4774  "w && w' \<noteq> 0 \<Longrightarrow> w \<noteq> 0 \<and> w' \<noteq> 0"
4775  by auto
4776
4777lemma word_clz_max:
4778  "word_clz w \<le> size (w::'a::len word)"
4779  unfolding word_clz_def
4780  by (metis length_takeWhile_le word_size_bl)
4781
4782lemma word_clz_nonzero_max:
4783  fixes w :: "'a::len word"
4784  assumes nz: "w \<noteq> 0"
4785  shows "word_clz w < size (w::'a::len word)"
4786proof -
4787  {
4788    assume a: "word_clz w = size (w::'a::len word)"
4789    hence "length (takeWhile Not (to_bl w)) = length (to_bl w)"
4790      by (simp add: word_clz_def word_size)
4791    hence allj: "\<forall>j\<in>set(to_bl w). \<not> j"
4792      by (metis a length_takeWhile_less less_irrefl_nat word_clz_def)
4793    hence "to_bl w = replicate (length (to_bl w)) False"
4794      by (fastforce intro!: list_of_false)
4795    hence "w = 0"
4796      by (metis to_bl_0 word_bl.Rep_eqD word_bl_Rep')
4797    with nz have False by simp
4798  }
4799  thus ?thesis using word_clz_max
4800    by (fastforce intro: le_neq_trans)
4801qed
4802
4803lemma unat_add_lem':
4804  "(unat x + unat y < 2 ^ LENGTH('a)) \<Longrightarrow>
4805    (unat (x + y :: 'a :: len word) = unat x + unat y)"
4806  by (subst unat_add_lem[symmetric], assumption)
4807
4808lemma from_bool_eq_if':
4809  "((if P then 1 else 0) = from_bool Q) = (P = Q)"
4810  by (simp add: case_bool_If from_bool_def split: if_split)
4811
4812lemma word_exists_nth:
4813  "(w::'a::len word) \<noteq> 0 \<Longrightarrow> \<exists>i. w !! i"
4814  using word_log2_nth_same by blast
4815
4816lemma shiftr_le_0:
4817  "unat (w::'a::len word) < 2 ^ n \<Longrightarrow> w >> n = (0::'a::len word)"
4818  by (rule word_unat.Rep_eqD) (simp add: shiftr_div_2n')
4819
4820lemma of_nat_shiftl:
4821  "(of_nat x << n) = (of_nat (x * 2 ^ n) :: ('a::len) word)"
4822proof -
4823  have "(of_nat x::'a word) << n = of_nat (2 ^ n) * of_nat x"
4824    using shiftl_t2n by (metis word_unat_power)
4825  thus ?thesis by simp
4826qed
4827
4828lemma shiftl_1_not_0:
4829  "n < LENGTH('a) \<Longrightarrow> (1::'a::len word) << n \<noteq> 0"
4830  by (simp add: shiftl_t2n)
4831
4832lemma max_word_not_0[simp]:
4833  "max_word \<noteq> 0"
4834  by (simp add: max_word_minus)
4835
4836lemma ucast_zero_is_aligned:
4837  "UCAST('a::len \<rightarrow> 'b::len) w = 0 \<Longrightarrow> n \<le> LENGTH('b) \<Longrightarrow> is_aligned w n"
4838  by (clarsimp simp: is_aligned_mask word_eq_iff word_size nth_ucast)
4839
4840lemma unat_ucast_eq_unat_and_mask:
4841  "unat (UCAST('b::len \<rightarrow> 'a::len) w) = unat (w && mask LENGTH('a))"
4842  proof -
4843    have "unat (UCAST('b \<rightarrow> 'a) w) = unat (UCAST('a \<rightarrow> 'b) (UCAST('b \<rightarrow> 'a) w))"
4844      by (cases "LENGTH('a) < LENGTH('b)"; simp add: is_down ucast_ucast_a unat_ucast_up_simp)
4845    thus ?thesis using ucast_ucast_mask by simp
4846  qed
4847
4848lemma unat_max_word_pos[simp]: "0 < unat max_word"
4849  by (auto simp: unat_gt_0)
4850
4851
4852(* Miscellaneous conditional injectivity rules. *)
4853
4854lemma mult_pow2_inj:
4855  assumes ws: "m + n \<le> LENGTH('a)"
4856  assumes le: "x \<le> mask m" "y \<le> mask m"
4857  assumes eq: "x * 2^n = y * (2^n::'a::len word)"
4858  shows "x = y"
4859proof (cases n)
4860  case 0 thus ?thesis using eq by simp
4861next
4862  case (Suc n')
4863  have m_lt: "m < LENGTH('a)" using Suc ws by simp
4864  have xylt: "x < 2^m" "y < 2^m" using le m_lt unfolding mask_2pm1 by auto
4865  have lenm: "n \<le> LENGTH('a) - m" using ws by simp
4866  show ?thesis
4867    using eq xylt
4868    apply (fold shiftl_t2n[where n=n, simplified mult.commute])
4869    apply (simp only: word_bl.Rep_inject[symmetric] bl_shiftl)
4870    apply (erule ssubst[OF less_is_drop_replicate])+
4871    apply (clarsimp elim!: drop_eq_mono[OF lenm])
4872    done
4873qed
4874
4875lemma word_of_nat_inj:
4876  assumes bounded: "x < 2 ^ LENGTH('a)" "y < 2 ^ LENGTH('a)"
4877  assumes of_nats: "of_nat x = (of_nat y :: 'a::len word)"
4878  shows "x = y"
4879  by (rule contrapos_pp[OF of_nats]; cases "x < y"; cases "y < x")
4880     (auto dest: bounded[THEN of_nat_mono_maybe])
4881
4882(* Sign extension from bit n. *)
4883
4884lemma sign_extend_bitwise_if:
4885  "i < size w \<Longrightarrow> sign_extend e w !! i \<longleftrightarrow> (if i < e then w !! i else w !! e)"
4886  by (simp add: sign_extend_def neg_mask_test_bit word_size)
4887
4888lemma sign_extend_bitwise_disj:
4889  "i < size w \<Longrightarrow> sign_extend e w !! i \<longleftrightarrow> i \<le> e \<and> w !! i \<or> e \<le> i \<and> w !! e"
4890  by (auto simp: sign_extend_bitwise_if)
4891
4892lemma sign_extend_bitwise_cases:
4893  "i < size w \<Longrightarrow> sign_extend e w !! i \<longleftrightarrow> (i \<le> e \<longrightarrow> w !! i) \<and> (e \<le> i \<longrightarrow> w !! e)"
4894  by (auto simp: sign_extend_bitwise_if)
4895
4896lemmas sign_extend_bitwise_if'[word_eqI_simps] = sign_extend_bitwise_if[simplified word_size]
4897lemmas sign_extend_bitwise_disj' = sign_extend_bitwise_disj[simplified word_size]
4898lemmas sign_extend_bitwise_cases' = sign_extend_bitwise_cases[simplified word_size]
4899
4900(* Often, it is easier to reason about an operation which does not overwrite
4901   the bit which determines which mask operation to apply. *)
4902lemma sign_extend_def':
4903  "sign_extend n w = (if w !! n then w || ~~ (mask (Suc n)) else w && mask (Suc n))"
4904  by word_eqI (auto dest: less_antisym)
4905
4906lemma sign_extended_sign_extend:
4907  "sign_extended n (sign_extend n w)"
4908  by (clarsimp simp: sign_extended_def word_size sign_extend_bitwise_if')
4909
4910lemma sign_extended_iff_sign_extend:
4911  "sign_extended n w \<longleftrightarrow> sign_extend n w = w"
4912  apply (rule iffI)
4913   apply (word_eqI, rename_tac i)
4914   apply (case_tac "n < i"; simp add: sign_extended_def word_size)
4915  apply (erule subst, rule sign_extended_sign_extend)
4916  done
4917
4918lemma sign_extended_weaken:
4919  "sign_extended n w \<Longrightarrow> n \<le> m \<Longrightarrow> sign_extended m w"
4920  unfolding sign_extended_def by (cases "n < m") auto
4921
4922lemma sign_extend_sign_extend_eq:
4923  "sign_extend m (sign_extend n w) = sign_extend (min m n) w"
4924  by word_eqI
4925
4926lemma sign_extended_high_bits:
4927  "\<lbrakk> sign_extended e p; j < size p; e \<le> i; i < j \<rbrakk> \<Longrightarrow> p !! i = p !! j"
4928  by (drule (1) sign_extended_weaken; simp add: sign_extended_def)
4929
4930lemma sign_extend_eq:
4931  "w && mask (Suc n) = v && mask (Suc n) \<Longrightarrow> sign_extend n w = sign_extend n v"
4932  by word_eqI_solve
4933
4934lemma sign_extended_add:
4935  assumes p: "is_aligned p n"
4936  assumes f: "f < 2 ^ n"
4937  assumes e: "n \<le> e"
4938  assumes "sign_extended e p"
4939  shows "sign_extended e (p + f)"
4940proof (cases "e < size p")
4941  case True
4942  note and_or = is_aligned_add_or[OF p f]
4943  have "\<not> f !! e"
4944    using True e less_2p_is_upper_bits_unset[THEN iffD1, OF f]
4945    by (fastforce simp: word_size)
4946  hence i: "(p + f) !! e = p !! e"
4947    by (simp add: and_or)
4948  have fm: "f && mask e = f"
4949    by (fastforce intro: subst[where P="\<lambda>f. f && mask e = f", OF less_mask_eq[OF f]]
4950                  simp: mask_twice e)
4951  show ?thesis
4952    using assms
4953     apply (simp add: sign_extended_iff_sign_extend sign_extend_def i)
4954     apply (simp add: and_or word_bw_comms[of p f])
4955     apply (clarsimp simp: word_ao_dist fm word_bw_assocs split: if_splits)
4956    done
4957next
4958  case False thus ?thesis
4959    by (simp add: sign_extended_def word_size)
4960qed
4961
4962lemma sign_extended_neq_mask:
4963  "\<lbrakk>sign_extended n ptr; m \<le> n\<rbrakk> \<Longrightarrow> sign_extended n (ptr && ~~ (mask m))"
4964  by (fastforce simp: sign_extended_def word_size neg_mask_test_bit)
4965
4966(* Uints *)
4967
4968lemma uints_mono_iff: "uints l \<subseteq> uints m \<longleftrightarrow> l \<le> m"
4969  using power_increasing_iff[of "2::int" l m]
4970  apply (auto simp: uints_num subset_iff simp del: power_increasing_iff)
4971  by (meson less_irrefl not_less zle2p)
4972
4973lemmas uints_monoI = uints_mono_iff[THEN iffD2]
4974
4975lemma Bit_in_uints_Suc: "w BIT c \<in> uints (Suc m)" if "w \<in> uints m"
4976  using that
4977  by (auto simp: uints_num Bit_def)
4978
4979lemma Bit_in_uintsI: "w BIT c \<in> uints m" if "w \<in> uints (m - 1)" "m > 0"
4980  using Bit_in_uints_Suc[OF that(1)] that(2)
4981  by auto
4982
4983lemma bin_cat_in_uintsI: "bin_cat a n b \<in> uints m" if "a \<in> uints l" "m \<ge> l + n"
4984  using that
4985proof (induction n arbitrary: b m)
4986  case 0
4987  then have "uints l \<subseteq> uints m"
4988    by (intro uints_monoI) auto
4989  then show ?case using 0 by auto
4990next
4991  case (Suc n)
4992  then show ?case
4993    by (auto intro!: Bit_in_uintsI)
4994qed
4995
4996lemma bin_cat_cong: "bin_cat a n b = bin_cat c m d"
4997  if "n = m" "a = c" "bintrunc m b = bintrunc m d"
4998  using that(3) unfolding that(1,2)
4999proof (induction m arbitrary: b d)
5000  case (Suc m)
5001  show ?case
5002    using Suc.prems by (auto intro: Suc.IH)
5003qed simp
5004
5005lemma bin_cat_eqD1: "bin_cat a n b = bin_cat c n d \<Longrightarrow> a = c"
5006  by (induction n arbitrary: b d) auto
5007
5008lemma bin_cat_eqD2: "bin_cat a n b = bin_cat c n d \<Longrightarrow> bintrunc n b = bintrunc n d"
5009  by (induction n arbitrary: b d) auto
5010
5011lemma bin_cat_inj: "(bin_cat a n b) = bin_cat c n d \<longleftrightarrow> a = c \<and> bintrunc n b = bintrunc n d"
5012  by (auto intro: bin_cat_cong bin_cat_eqD1 bin_cat_eqD2)
5013
5014lemma word_of_int_bin_cat_eq_iff:
5015  "(word_of_int (bin_cat (uint a) LENGTH('b) (uint b))::'c::len0 word) =
5016  word_of_int (bin_cat (uint c) LENGTH('b) (uint d)) \<longleftrightarrow> b = d \<and> a = c"
5017  if "LENGTH('a) + LENGTH('b) \<le> LENGTH('c)"
5018  for a::"'a::len0 word" and b::"'b::len0 word"
5019  by (subst word_uint.Abs_inject)
5020     (auto simp: bin_cat_inj intro!: that bin_cat_in_uintsI)
5021
5022lemma word_cat_inj: "(word_cat a b::'c::len0 word) = word_cat c d \<longleftrightarrow> a = c \<and> b = d"
5023  if "LENGTH('a) + LENGTH('b) \<le> LENGTH('c)"
5024  for a::"'a::len0 word" and b::"'b::len0 word"
5025  by (auto simp: word_cat_def word_uint.Abs_inject word_of_int_bin_cat_eq_iff that)
5026
5027lemma p2_eq_1: "2 ^ n = (1::'a::len word) \<longleftrightarrow> n = 0"
5028proof -
5029  have "2 ^ n = (1::'a word) \<Longrightarrow> n = 0"
5030    by (metis One_nat_def not_less one_less_numeral_iff p2_eq_0 p2_gt_0 power_0 power_0
5031        power_inject_exp semiring_norm(76) unat_power_lower zero_neq_one)
5032  then show ?thesis by auto
5033qed
5034
5035(* usually: x,y = (len_of TYPE ('a)) *)
5036lemma bitmagic_zeroLast_leq_or1Last:
5037  "(a::('a::len) word) AND (mask len << x - len) \<le> a OR mask (y - len)"
5038  by (meson le_word_or2 order_trans word_and_le2)
5039
5040
5041lemma zero_base_lsb_imp_set_eq_as_bit_operation:
5042  fixes base ::"'a::len word"
5043  assumes valid_prefix: "mask (LENGTH('a) - len) AND base = 0"
5044  shows "(base = NOT (mask (LENGTH('a) - len)) AND a) \<longleftrightarrow>
5045         (a \<in> {base .. base OR mask (LENGTH('a) - len)})"
5046proof
5047  have helper3: "x OR y = x OR y AND NOT x" for x y ::"'a::len word" by (simp add: word_oa_dist2)
5048  from assms show "base = NOT (mask (LENGTH('a) - len)) AND a \<Longrightarrow>
5049                    a \<in> {base..base OR mask (LENGTH('a) - len)}"
5050    apply(simp add: word_and_le1)
5051    apply(metis helper3 le_word_or2 word_bw_comms(1) word_bw_comms(2))
5052  done
5053next
5054  assume "a \<in> {base..base OR mask (LENGTH('a) - len)}"
5055  hence a: "base \<le> a \<and> a \<le> base OR mask (LENGTH('a) - len)" by simp
5056  show "base = NOT (mask (LENGTH('a) - len)) AND a"
5057  proof -
5058    have f2: "\<forall>x\<^sub>0. base AND NOT (mask x\<^sub>0) \<le> a AND NOT (mask x\<^sub>0)"
5059      using a neg_mask_mono_le by blast
5060    have f3: "\<forall>x\<^sub>0. a AND NOT (mask x\<^sub>0) \<le> (base OR mask (LENGTH('a) - len)) AND NOT (mask x\<^sub>0)"
5061      using a neg_mask_mono_le by blast
5062    have f4: "base = base AND NOT (mask (LENGTH('a) - len))"
5063      using valid_prefix by (metis mask_eq_0_eq_x word_bw_comms(1))
5064    hence f5: "\<forall>x\<^sub>6. (base OR x\<^sub>6) AND NOT (mask (LENGTH('a) - len)) =
5065                      base OR x\<^sub>6 AND NOT (mask (LENGTH('a) - len))"
5066      using word_ao_dist by (metis)
5067    have f6: "\<forall>x\<^sub>2 x\<^sub>3. a AND NOT (mask x\<^sub>2) \<le> x\<^sub>3 \<or>
5068                      \<not> (base OR mask (LENGTH('a) - len)) AND NOT (mask x\<^sub>2) \<le> x\<^sub>3"
5069      using f3 dual_order.trans by auto
5070    have "base = (base OR mask (LENGTH('a) - len)) AND NOT (mask (LENGTH('a) - len))"
5071      using f5 by auto
5072    hence "base = a AND NOT (mask (LENGTH('a) - len))"
5073      using f2 f4 f6 by (metis eq_iff)
5074    thus "base = NOT (mask (LENGTH('a) - len)) AND a"
5075      by (metis word_bw_comms(1))
5076  qed
5077qed
5078
5079lemma unat_minus_one_word:
5080  "unat (-1 :: 'a :: len word) = 2 ^ LENGTH('a) - 1"
5081  by (subst minus_one_word)
5082     (subst unat_sub_if', clarsimp)
5083
5084lemma of_nat_eq_signed_scast:
5085  "(of_nat x = (y :: ('a::len) signed word))
5086   = (of_nat x = (scast y :: 'a word))"
5087  by (metis scast_of_nat scast_scast_id(2))
5088
5089lemma word_ctz_le:
5090  "word_ctz (w :: ('a::len word)) \<le> LENGTH('a)"
5091  apply (clarsimp simp: word_ctz_def)
5092  apply (rule nat_le_Suc_less_imp[where y="LENGTH('a) + 1" , simplified])
5093  apply (rule order_le_less_trans[OF List.length_takeWhile_le])
5094  apply simp
5095  done
5096
5097lemma word_ctz_less:
5098  "w \<noteq> 0 \<Longrightarrow> word_ctz (w :: ('a::len word)) < LENGTH('a)"
5099  apply (clarsimp simp: word_ctz_def eq_zero_set_bl)
5100  apply (rule order_less_le_trans[OF length_takeWhile_less])
5101   apply fastforce+
5102  done
5103
5104lemma word_ctz_not_minus_1:
5105  "1 < LENGTH('a) \<Longrightarrow> of_nat (word_ctz (w :: 'a :: len word)) \<noteq> (- 1 :: 'a::len word)"
5106  by (metis (mono_tags) One_nat_def add.right_neutral add_Suc_right le_diff_conv le_less_trans
5107                        n_less_equal_power_2 not_le suc_le_pow_2 unat_minus_one_word unat_of_nat_len
5108                        word_ctz_le)
5109
5110lemma word_aligned_add_no_wrap_bounded:
5111  "\<lbrakk> w + 2^n \<le> x; w + 2^n \<noteq> 0; is_aligned w n \<rbrakk> \<Longrightarrow> (w::'a::len word) < x"
5112  by (blast dest: is_aligned_no_overflow le_less_trans word_leq_le_minus_one)
5113
5114lemma mask_Suc:
5115  "mask (Suc n) = 2^n + mask n"
5116  by (simp add: mask_def)
5117
5118lemma is_aligned_no_overflow_mask:
5119  "is_aligned x n \<Longrightarrow> x \<le> x + mask n"
5120  by (simp add: mask_def) (erule is_aligned_no_overflow')
5121
5122lemma is_aligned_mask_offset_unat:
5123  fixes off :: "('a::len) word"
5124  and     x :: "'a word"
5125  assumes al: "is_aligned x sz"
5126  and   offv: "off \<le> mask sz"
5127  shows  "unat x + unat off < 2 ^ LENGTH('a)"
5128proof cases
5129  assume szv: "sz < LENGTH('a)"
5130  from al obtain k where xv: "x = 2 ^ sz * (of_nat k)"
5131    and kl: "k < 2 ^ (LENGTH('a) - sz)"
5132    by (auto elim: is_alignedE)
5133
5134  from offv szv have offv': "unat off < 2 ^ sz"
5135    by (simp add: mask_2pm1 unat_less_power)
5136
5137  show ?thesis using szv
5138    using al is_aligned_no_wrap''' offv' by blast
5139next
5140  assume "\<not> sz < LENGTH('a)"
5141  with al have "x = 0" by - word_eqI
5142  thus ?thesis by simp
5143qed
5144
5145lemma of_bl_max:
5146  "(of_bl xs :: 'a::len word) \<le> mask (length xs)"
5147  apply (induct xs)
5148   apply simp
5149  apply (simp add: of_bl_Cons mask_Suc)
5150  apply (rule conjI; clarsimp)
5151   apply (erule word_plus_mono_right)
5152   apply (rule is_aligned_no_overflow_mask)
5153   apply (rule is_aligned_triv)
5154  apply (simp add: word_le_nat_alt)
5155  apply (subst unat_add_lem')
5156   apply (rule is_aligned_mask_offset_unat)
5157    apply (rule is_aligned_triv)
5158   apply (simp add: mask_def)
5159  apply simp
5160  done
5161
5162lemma mask_over_length:
5163  "LENGTH('a) \<le> n \<Longrightarrow> mask n = (-1::'a::len word)"
5164  by (simp add: mask_def)
5165
5166lemma is_aligned_over_length:
5167  "\<lbrakk> is_aligned p n; LENGTH('a) \<le> n \<rbrakk> \<Longrightarrow> (p::'a::len word) = 0"
5168  by (simp add: is_aligned_mask mask_over_length)
5169
5170lemma Suc_2p_unat_mask:
5171  "n < LENGTH('a) \<Longrightarrow> Suc (2 ^ n * k + unat (mask n :: 'a::len word)) = 2 ^ n * (k+1)"
5172  by (simp add: unat_mask)
5173
5174lemma is_aligned_add_step_le:
5175  "\<lbrakk> is_aligned (a::'a::len word) n; is_aligned b n; a < b; b \<le> a + mask n \<rbrakk> \<Longrightarrow> False"
5176  apply (simp flip: not_le)
5177  apply (erule notE)
5178  apply (cases "LENGTH('a) \<le> n")
5179   apply (drule (1) is_aligned_over_length)+
5180   apply (drule mask_over_length)
5181   apply clarsimp
5182  apply (clarsimp simp: word_le_nat_alt not_less)
5183  apply (subst (asm) unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask)
5184  apply (clarsimp simp: is_aligned_def dvd_def word_le_nat_alt)
5185  apply (drule le_imp_less_Suc)
5186  apply (simp add: Suc_2p_unat_mask)
5187  by (metis Groups.mult_ac(2) Suc_leI linorder_not_less mult_le_mono order_refl times_nat.simps(2))
5188
5189lemma power_2_mult_step_le:
5190  "\<lbrakk>n' \<le> n; 2 ^ n' * k' < 2 ^ n * k\<rbrakk> \<Longrightarrow> 2 ^ n' * (k' + 1) \<le> 2 ^ n * (k::nat)"
5191  apply (cases "n'=n", simp)
5192   apply (metis Suc_leI le_refl mult_Suc_right mult_le_mono semiring_normalization_rules(7))
5193  apply (drule (1) le_neq_trans)
5194  apply clarsimp
5195  apply (subgoal_tac "\<exists>m. n = n' + m")
5196   prefer 2
5197   apply (simp add: le_Suc_ex)
5198  apply (clarsimp simp: power_add)
5199  by (metis Suc_leI mult.assoc mult_Suc_right nat_mult_le_cancel_disj)
5200
5201lemma aligned_mask_step:
5202  "\<lbrakk> n' \<le> n; p' \<le> p + mask n; is_aligned p n; is_aligned p' n' \<rbrakk> \<Longrightarrow>
5203   (p'::'a::len word) + mask n' \<le> p + mask n"
5204  apply (cases "LENGTH('a) \<le> n")
5205   apply (frule (1) is_aligned_over_length)
5206   apply (drule mask_over_length)
5207   apply clarsimp
5208  apply (simp add: not_le)
5209  apply (simp add: word_le_nat_alt unat_plus_simple)
5210  apply (subst unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask)+
5211  apply (subst (asm) unat_plus_simple[THEN iffD1], erule is_aligned_no_overflow_mask)
5212  apply (clarsimp simp: is_aligned_def dvd_def)
5213  apply (rename_tac k k')
5214  apply (thin_tac "unat p = x" for p x)+
5215  apply (subst Suc_le_mono[symmetric])
5216  apply (simp only: Suc_2p_unat_mask)
5217  apply (drule le_imp_less_Suc, subst (asm) Suc_2p_unat_mask, assumption)
5218  apply (erule (1) power_2_mult_step_le)
5219  done
5220
5221lemma mask_mono:
5222  "sz' \<le> sz \<Longrightarrow> mask sz' \<le> (mask sz :: 'a::len word)"
5223  by (simp add: le_mask_iff shiftr_mask_le)
5224
5225lemma aligned_mask_disjoint:
5226  "\<lbrakk> is_aligned (a :: 'a :: len word) n; b \<le> mask n \<rbrakk> \<Longrightarrow> a && b = 0"
5227  by word_eqI_solve
5228
5229lemma word_and_or_mask_aligned:
5230  "\<lbrakk> is_aligned a n; b \<le> mask n \<rbrakk> \<Longrightarrow> a + b = a || b"
5231  by (simp add: aligned_mask_disjoint word_plus_and_or_coroll)
5232
5233lemmas word_and_or_mask_aligned2 =
5234  word_and_or_mask_aligned[where a=b and b=a for a b,
5235                           simplified add.commute word_bool_alg.disj.commute]
5236
5237lemma is_aligned_ucastI:
5238  "is_aligned w n \<Longrightarrow> is_aligned (ucast w) n"
5239  by (clarsimp simp: word_eqI_simps)
5240
5241lemma ucast_le_maskI:
5242  "a \<le> mask n \<Longrightarrow> UCAST('a::len \<rightarrow> 'b::len) a \<le> mask n"
5243  by (metis and_mask_eq_iff_le_mask ucast_and_mask)
5244
5245lemma ucast_add_mask_aligned:
5246  "\<lbrakk> a \<le> mask n; is_aligned b n \<rbrakk> \<Longrightarrow> UCAST ('a::len \<rightarrow> 'b::len) (a + b) = ucast a + ucast b"
5247  by (metis is_aligned_ucastI ucast_le_maskI ucast_or_distrib word_and_or_mask_aligned2)
5248
5249lemma ucast_shiftl:
5250  "LENGTH('b) \<le> LENGTH ('a) \<Longrightarrow> UCAST ('a::len \<rightarrow> 'b::len) x << n = ucast (x << n)"
5251  by word_eqI_solve
5252
5253lemma ucast_leq_mask:
5254  "LENGTH('a) \<le> n \<Longrightarrow> ucast (x::'a::len0 word) \<le> mask n"
5255  by (clarsimp simp: le_mask_high_bits word_size nth_ucast)
5256
5257lemma shiftl_inj:
5258  "\<lbrakk> x << n = y << n; x \<le> mask (LENGTH('a)-n); y \<le> mask (LENGTH('a)-n) \<rbrakk> \<Longrightarrow>
5259   x = (y :: 'a :: len word)"
5260  apply word_eqI
5261  apply (rename_tac n')
5262  apply (case_tac "LENGTH('a) - n \<le> n'", simp)
5263  by (metis add.commute add.right_neutral diff_add_inverse le_diff_conv linorder_not_less zero_order(1))
5264
5265lemma distinct_word_add_ucast_shift_inj:
5266  "\<lbrakk> p + (UCAST('a::len \<rightarrow> 'b::len) off << n) = p' + (ucast off' << n);
5267     is_aligned p n'; is_aligned p' n'; n' = n + LENGTH('a); n' < LENGTH('b) \<rbrakk>
5268   \<Longrightarrow> p' = p \<and> off' = off"
5269  apply (simp add: word_and_or_mask_aligned le_mask_shiftl_le_mask[where n="LENGTH('a)"]
5270                   ucast_leq_mask)
5271  apply (simp add: is_aligned_nth)
5272  apply (rule conjI; word_eqI)
5273   apply (metis add.commute test_bit_conj_lt diff_add_inverse le_diff_conv nat_less_le)
5274  apply (rename_tac i)
5275  apply (erule_tac x="i+n" in allE)
5276  apply simp
5277  done
5278
5279lemma aligned_add_mask_lessD:
5280  "\<lbrakk> x + mask n < y; is_aligned x n \<rbrakk> \<Longrightarrow> x < y" for y::"'a::len word"
5281  by (metis is_aligned_no_overflow' mask_2pm1 order_le_less_trans)
5282
5283lemma aligned_add_mask_less_eq:
5284  "\<lbrakk> is_aligned x n; is_aligned y n;  n < LENGTH('a) \<rbrakk> \<Longrightarrow> (x + mask n < y) = (x < y)"
5285  for y::"'a::len word"
5286  using aligned_add_mask_lessD is_aligned_add_step_le word_le_not_less by blast
5287
5288lemma word_upto_Nil:
5289  "y < x \<Longrightarrow> [x .e. y ::'a::len word] = []"
5290  by (simp add: upto_enum_red not_le word_less_nat_alt)
5291
5292lemma word_enum_decomp_elem:
5293  assumes "[x .e. (y ::'a::len word)] = as @ a # bs"
5294  shows "x \<le> a \<and> a \<le> y"
5295proof -
5296  have "set as \<subseteq> set [x .e. y] \<and> a \<in> set [x .e. y]"
5297    using assms by (auto dest: arg_cong[where f=set])
5298  then show ?thesis by auto
5299qed
5300
5301lemma max_word_not_less[simp]:
5302   "\<not> max_word < x"
5303  by (simp add: not_less)
5304
5305lemma word_enum_prefix:
5306  "[x .e. (y ::'a::len word)] = as @ a # bs \<Longrightarrow> as = (if x < a then [x .e. a - 1] else [])"
5307  apply (induct as arbitrary: x; clarsimp)
5308   apply (case_tac "x < y")
5309    prefer 2
5310    apply (case_tac "x = y", simp)
5311    apply (simp add: not_less)
5312    apply (drule (1) dual_order.not_eq_order_implies_strict)
5313    apply (simp add: word_upto_Nil)
5314   apply (simp add: word_upto_Cons_eq)
5315  apply (case_tac "x < y")
5316   prefer 2
5317   apply (case_tac "x = y", simp)
5318   apply (simp add: not_less)
5319   apply (drule (1) dual_order.not_eq_order_implies_strict)
5320   apply (simp add: word_upto_Nil)
5321  apply (clarsimp simp: word_upto_Cons_eq)
5322  apply (frule word_enum_decomp_elem)
5323  apply clarsimp
5324  apply (rule conjI)
5325   prefer 2
5326   apply (subst word_Suc_le[symmetric]; clarsimp)
5327  apply (drule meta_spec)
5328  apply (drule (1) meta_mp)
5329  apply clarsimp
5330  apply (rule conjI; clarsimp)
5331  apply (subst (2) word_upto_Cons_eq)
5332   apply unat_arith
5333  apply simp
5334  done
5335
5336lemma word_enum_decomp_set:
5337  "[x .e. (y ::'a::len word)] = as @ a # bs \<Longrightarrow> a \<notin> set as"
5338  by (metis distinct_append distinct_enum_upto' not_distinct_conv_prefix)
5339
5340lemma word_enum_decomp:
5341  assumes "[x .e. (y ::'a::len word)] = as @ a # bs"
5342  shows "x \<le> a \<and> a \<le> y \<and> a \<notin> set as \<and> (\<forall>z \<in> set as. x \<le> z \<and> z \<le> y)"
5343proof -
5344  from assms
5345  have "set as \<subseteq> set [x .e. y] \<and> a \<in> set [x .e. y]"
5346    by (auto dest: arg_cong[where f=set])
5347  with word_enum_decomp_set[OF assms]
5348  show ?thesis by auto
5349qed
5350
5351lemma of_nat_unat_le_mask_ucast:
5352  "\<lbrakk>of_nat (unat t) = w; t \<le> mask LENGTH('a)\<rbrakk> \<Longrightarrow> t = UCAST('a::len \<rightarrow> 'b::len) w"
5353  by (clarsimp simp: ucast_nat_def ucast_ucast_mask simp flip: and_mask_eq_iff_le_mask)
5354
5355lemma fold_eq_0_to_bool:
5356  "(v = 0) = (\<not> to_bool v)"
5357  by (simp add: to_bool_def)
5358
5359lemma less_diff_gt0:
5360  "a < b \<Longrightarrow> (0 :: 'a :: len word) < b - a"
5361  by unat_arith
5362
5363lemma unat_plus_gt:
5364  "unat ((a :: 'a :: len word) + b) \<le> unat a + unat b"
5365  by (clarsimp simp: unat_plus_if_size)
5366
5367lemma const_less:
5368  "\<lbrakk> (a :: 'a :: len word) - 1 < b; a \<noteq> b \<rbrakk> \<Longrightarrow> a < b"
5369  by (metis less_1_simp word_le_less_eq)
5370
5371lemma add_mult_aligned_neg_mask:
5372  "m && (2 ^ n - 1) = 0 \<Longrightarrow> (x + y * m) && ~~(mask n) = (x && ~~(mask n)) + y * m"
5373  by (metis Groups.add_ac(2) is_aligned_mask mask_def mask_eqs(5) mask_out_add_aligned
5374            mult_zero_right shiftl_1 word_bw_comms(1) word_log_esimps(1))
5375
5376lemma unat_of_nat_minus_1:
5377  "\<lbrakk> n < 2 ^ LENGTH('a); n \<noteq> 0 \<rbrakk> \<Longrightarrow> unat ((of_nat n:: 'a :: len word) - 1) = n - 1"
5378  by (simp add: unat_eq_of_nat)
5379
5380lemma word_eq_zeroI:
5381  "a \<le> a - 1 \<Longrightarrow> a = 0" for a :: "'a :: len word"
5382  by (simp add: word_must_wrap)
5383
5384lemma word_add_format:
5385  "(-1 :: 'a :: len  word) + b + c = b + (c - 1)"
5386  by simp
5387
5388lemma upto_enum_word_nth:
5389  "\<lbrakk> i \<le> j; k \<le> unat (j - i) \<rbrakk> \<Longrightarrow> [i .e. j] ! k = i + of_nat k"
5390  apply (clarsimp simp: upto_enum_def nth_append)
5391  apply (clarsimp simp: word_le_nat_alt[symmetric])
5392  apply (rule conjI, clarsimp)
5393   apply (subst toEnum_of_nat, unat_arith)
5394   apply unat_arith
5395  apply (clarsimp simp: not_less unat_sub[symmetric])
5396  apply unat_arith
5397  done
5398
5399lemma upto_enum_step_nth:
5400  "\<lbrakk> a \<le> c; n \<le> unat ((c - a) div (b - a)) \<rbrakk>
5401   \<Longrightarrow> [a, b .e. c] ! n = a + of_nat n * (b - a)"
5402  by (clarsimp simp: upto_enum_step_def not_less[symmetric] upto_enum_word_nth)
5403
5404lemma upto_enum_inc_1_len:
5405  "a < - 1 \<Longrightarrow> [(0 :: 'a :: len word) .e. 1 + a] = [0 .e. a] @ [1 + a]"
5406  apply (simp add: upto_enum_word)
5407  apply (subgoal_tac "unat (1+a) = 1 + unat a")
5408   apply simp
5409  apply (subst unat_plus_simple[THEN iffD1])
5410   apply (metis add.commute no_plus_overflow_neg olen_add_eqv)
5411  apply unat_arith
5412  done
5413
5414lemma neg_mask_add:
5415  "y && mask n = 0 \<Longrightarrow> x + y && ~~(mask n) = (x && ~~(mask n)) + y"
5416  by (clarsimp simp: mask_out_sub_mask mask_eqs(7)[symmetric] mask_twice)
5417
5418lemma shiftr_shiftl_shiftr[simp]:
5419  "(x :: 'a :: len word)  >> a << a >> a = x >> a"
5420  by word_eqI_solve
5421
5422lemma add_right_shift:
5423  "\<lbrakk> x && mask n = 0; y && mask n = 0; x \<le> x + y \<rbrakk>
5424   \<Longrightarrow> (x + y :: ('a :: len) word) >> n = (x >> n) + (y >> n)"
5425  apply (simp add: no_olen_add_nat is_aligned_mask[symmetric])
5426  apply (simp add: unat_arith_simps shiftr_div_2n' split del: if_split)
5427  apply (subst if_P)
5428   apply (erule order_le_less_trans[rotated])
5429   apply (simp add: add_mono)
5430  apply (simp add: shiftr_div_2n' is_aligned_def)
5431  done
5432
5433lemma sub_right_shift:
5434  "\<lbrakk> x && mask n = 0; y && mask n = 0; y \<le> x \<rbrakk>
5435   \<Longrightarrow> (x - y) >> n = (x >> n :: 'a :: len word) - (y >> n)"
5436  using add_right_shift[where x="x - y" and y=y and n=n]
5437  by (simp add: aligned_sub_aligned is_aligned_mask[symmetric] word_sub_le)
5438
5439lemma and_and_mask_simple:
5440  "y && mask n = mask n \<Longrightarrow> (x && y) && mask n = x && mask n"
5441  by (simp add: word_bool_alg.conj.assoc)
5442
5443lemma and_and_mask_simple_not:
5444  "y && mask n = 0 \<Longrightarrow> (x && y) && mask n = 0"
5445  by (simp add: word_bool_alg.conj.assoc)
5446
5447lemma word_and_le':
5448  "b \<le> c \<Longrightarrow> (a :: 'a :: len word) && b \<le> c"
5449  by (metis word_and_le1 order_trans)
5450
5451lemma word_and_less':
5452  "b < c \<Longrightarrow> (a :: 'a :: len word) && b < c"
5453  by (metis word_and_le1 xtr7)
5454
5455lemma shiftr_w2p:
5456  "x < LENGTH('a) \<Longrightarrow> 2 ^ x = (2 ^ (LENGTH('a) - 1) >> (LENGTH('a) - 1 - x) :: 'a :: len word)"
5457  by word_eqI_solve
5458
5459lemma t2p_shiftr:
5460  "\<lbrakk> b \<le> a; a < LENGTH('a) \<rbrakk> \<Longrightarrow> (2 :: 'a :: len word) ^ a >> b = 2 ^ (a - b)"
5461  by word_eqI_solve
5462
5463lemma scast_1[simp]:
5464  "scast (1 :: 'a :: len signed word) = (1 :: 'a word)"
5465  by simp
5466
5467lemma ucast_ucast_mask_eq:
5468  "\<lbrakk> UCAST('a::len \<rightarrow> 'b::len) x = y; x && mask LENGTH('b) = x \<rbrakk> \<Longrightarrow> x = ucast y"
5469  by word_eqI_solve
5470
5471lemma ucast_up_eq:
5472  "\<lbrakk> ucast x = (ucast y::'b::len word); LENGTH('a) \<le> LENGTH ('b) \<rbrakk>
5473   \<Longrightarrow> ucast x = (ucast y::'a::len word)"
5474  by word_eqI_solve
5475
5476lemma ucast_up_neq:
5477  "\<lbrakk> ucast x \<noteq> (ucast y::'b::len word); LENGTH('b) \<le> LENGTH ('a) \<rbrakk>
5478   \<Longrightarrow> ucast x \<noteq> (ucast y::'a::len word)"
5479  by (fastforce dest: ucast_up_eq)
5480
5481lemma mask_AND_less_0:
5482  "\<lbrakk> x && mask n = 0; m \<le> n \<rbrakk> \<Longrightarrow> x && mask m = 0"
5483  by (metis mask_twice2 word_and_notzeroD)
5484
5485lemma mask_len_id [simp]:
5486  "(x :: 'a :: len word) && mask LENGTH('a) = x"
5487  using uint_lt2p [of x] by (simp add: mask_eq_iff)
5488
5489lemma scast_ucast_down_same:
5490  "LENGTH('b) \<le> LENGTH('a) \<Longrightarrow> SCAST('a \<rightarrow> 'b) = UCAST('a::len \<rightarrow> 'b::len)"
5491  by (simp add: down_cast_same is_down)
5492
5493lemma word_aligned_0_sum:
5494  "\<lbrakk> a + b = 0; is_aligned (a :: 'a :: len word) n; b \<le> mask n; n < LENGTH('a) \<rbrakk>
5495   \<Longrightarrow> a = 0 \<and> b = 0"
5496  by (simp add: word_plus_and_or_coroll aligned_mask_disjoint word_or_zero)
5497
5498lemma mask_eq1_nochoice:
5499  "\<lbrakk> LENGTH('a) > 1; (x :: 'a :: len word) && 1 = x \<rbrakk> \<Longrightarrow> x = 0 \<or> x = 1"
5500  by (metis word_and_1)
5501
5502lemma pow_mono_leq_imp_lt:
5503  "x \<le> y \<Longrightarrow> x < 2 ^ y"
5504  by (simp add: le_less_trans)
5505
5506lemma unat_of_nat_ctz_mw:
5507  "unat (of_nat (word_ctz (w :: 'a :: len word)) :: 'a :: len word) = word_ctz w"
5508  using word_ctz_le[where w=w, simplified] unat_of_nat_eq[where x="word_ctz w" and 'a="'a"]
5509        pow_mono_leq_imp_lt
5510  by simp
5511
5512lemma unat_of_nat_ctz_smw:
5513  "unat (of_nat (word_ctz (w :: 'a :: len word)) :: 'a :: len sword) = word_ctz w"
5514  using word_ctz_le[where w=w, simplified] unat_of_nat_eq[where x="word_ctz w" and 'a="'a"]
5515        pow_mono_leq_imp_lt
5516  by (metis le_unat_uoi le_unat_uoi linorder_neqE_nat nat_less_le scast_of_nat
5517            word_unat.Rep_inverse)
5518
5519lemma shiftr_and_eq_shiftl:
5520  "(w >> n) && x = y \<Longrightarrow> w && (x << n) = (y << n)" for y :: "'a:: len word"
5521  by (smt and_not_mask is_aligned_neg_mask_eq is_aligned_shift shift_over_ao_dists(4)
5522          word_bool_alg.conj_assoc word_bw_comms(1))
5523
5524lemma neg_mask_combine:
5525  "~~(mask a) && ~~(mask b) = ~~(mask (max a b))"
5526  by (auto simp: word_ops_nth_size word_size intro!: word_eqI)
5527
5528lemma neg_mask_twice:
5529  "x && ~~(mask n) && ~~(mask m) = x && ~~(mask (max n m))"
5530  by (metis neg_mask_combine)
5531
5532lemma multiple_mask_trivia:
5533  "n \<ge> m \<Longrightarrow> (x && ~~(mask n)) + (x && mask n && ~~(mask m)) = x && ~~(mask m)"
5534  apply (rule trans[rotated], rule_tac w="mask n" in word_plus_and_or_coroll2)
5535  apply (simp add: word_bw_assocs word_bw_comms word_bw_lcs neg_mask_twice
5536                   max_absorb2)
5537  done
5538
5539lemma add_mask_lower_bits':
5540  "\<lbrakk> len = LENGTH('a); is_aligned (x :: 'a :: len word) n;
5541     \<forall>n' \<ge> n. n' < len \<longrightarrow> \<not> p !! n' \<rbrakk>
5542   \<Longrightarrow> x + p && ~~(mask n) = x"
5543  using add_mask_lower_bits by auto
5544
5545lemma neg_mask_in_mask_range:
5546  "is_aligned ptr bits \<Longrightarrow> (ptr' && ~~(mask bits) = ptr) = (ptr' \<in> mask_range ptr bits)"
5547  apply (erule is_aligned_get_word_bits)
5548   apply (rule iffI)
5549    apply (drule sym)
5550    apply (simp add: word_and_le2)
5551    apply (subst word_plus_and_or_coroll, word_eqI_solve)
5552    apply (metis le_word_or2 neg_mask_add_mask word_bool_alg.conj.right_idem)
5553   apply clarsimp
5554   apply (smt add.right_neutral eq_iff is_aligned_neg_mask_eq mask_out_add_aligned neg_mask_mono_le
5555              word_and_not)
5556  apply (simp add: power_overflow mask_def)
5557  done
5558
5559lemma aligned_offset_in_range:
5560  "\<lbrakk> is_aligned (x :: 'a :: len word) m; y < 2 ^ m; is_aligned p n; n \<ge> m; n < LENGTH('a) \<rbrakk>
5561   \<Longrightarrow> (x + y \<in> {p .. p + mask n}) = (x \<in> mask_range p n)"
5562  apply (simp only: is_aligned_add_or flip: neg_mask_in_mask_range)
5563  by (metis less_mask_eq mask_subsume)
5564
5565lemma mask_range_to_bl':
5566  "\<lbrakk> is_aligned (ptr :: 'a :: len word) bits; bits < LENGTH('a) \<rbrakk>
5567   \<Longrightarrow> mask_range ptr bits
5568       = {x. take (LENGTH('a) - bits) (to_bl x) = take (LENGTH('a) - bits) (to_bl ptr)}"
5569  apply (rule set_eqI, rule iffI)
5570   apply clarsimp
5571   apply (subgoal_tac "\<exists>y. x = ptr + y \<and> y < 2 ^ bits")
5572    apply clarsimp
5573    apply (subst is_aligned_add_conv)
5574       apply assumption
5575      apply simp
5576    apply simp
5577   apply (rule_tac x="x - ptr" in exI)
5578   apply (simp add: add_diff_eq[symmetric])
5579   apply (simp only: word_less_sub_le[symmetric])
5580   apply (rule word_diff_ls')
5581    apply (simp add: field_simps mask_def)
5582   apply assumption
5583  apply simp
5584  apply (subgoal_tac "\<exists>y. y < 2 ^ bits \<and> to_bl (ptr + y) = to_bl x")
5585   apply clarsimp
5586   apply (rule conjI)
5587    apply (erule(1) is_aligned_no_wrap')
5588   apply (simp only: add_diff_eq[symmetric] mask_def)
5589   apply (rule word_plus_mono_right)
5590    apply simp
5591   apply (erule is_aligned_no_wrap')
5592   apply simp
5593  apply (rule_tac x="of_bl (drop (LENGTH('a) - bits) (to_bl x))" in exI)
5594  apply (rule context_conjI)
5595   apply (rule order_less_le_trans [OF of_bl_length])
5596    apply simp
5597   apply simp
5598  apply (subst is_aligned_add_conv)
5599     apply assumption
5600    apply simp
5601  apply (drule sym)
5602  apply (simp add: word_rep_drop)
5603  done
5604
5605lemma mask_range_to_bl:
5606  "is_aligned (ptr :: 'a :: len word) bits
5607   \<Longrightarrow> mask_range ptr bits
5608        = {x. take (LENGTH('a) - bits) (to_bl x) = take (LENGTH('a) - bits) (to_bl ptr)}"
5609  apply (erule is_aligned_get_word_bits)
5610   apply (erule(1) mask_range_to_bl')
5611  apply (rule set_eqI)
5612  apply (simp add: power_overflow mask_def)
5613  done
5614
5615lemma aligned_mask_range_cases:
5616  "\<lbrakk> is_aligned (p :: 'a :: len word) n; is_aligned (p' :: 'a :: len word) n' \<rbrakk>
5617   \<Longrightarrow> mask_range p n \<inter> mask_range p' n' = {} \<or>
5618       mask_range p n \<subseteq> mask_range p' n' \<or>
5619       mask_range p n \<supseteq> mask_range p' n'"
5620  apply (simp add: mask_range_to_bl)
5621  apply (rule Meson.disj_comm, rule disjCI)
5622  apply (erule nonemptyE)
5623  apply simp
5624  apply (subgoal_tac "(\<exists>n''. LENGTH('a) - n = (LENGTH('a) - n') + n'')
5625                    \<or> (\<exists>n''. LENGTH('a) - n' = (LENGTH('a) - n) + n'')")
5626   apply (fastforce simp: take_add)
5627  apply arith
5628  done
5629
5630lemma aligned_mask_range_offset_subset:
5631  assumes al: "is_aligned (ptr :: 'a :: len word) sz" and al': "is_aligned x sz'"
5632  and szv: "sz' \<le> sz"
5633  and xsz: "x < 2 ^ sz"
5634  shows "mask_range (ptr+x) sz' \<subseteq> mask_range ptr sz"
5635  using al
5636proof (rule is_aligned_get_word_bits)
5637  assume p0: "ptr = 0" and szv': "LENGTH ('a) \<le> sz"
5638  then have "(2 ::'a word) ^ sz = 0" by simp
5639  show ?thesis using p0
5640    by (simp add: \<open>2 ^ sz = 0\<close> mask_def)
5641next
5642  assume szv': "sz < LENGTH('a)"
5643
5644  hence blah: "2 ^ (sz - sz') < (2 :: nat) ^ LENGTH('a)"
5645    using szv by auto
5646  show ?thesis using szv szv'
5647    apply (intro range_subsetI)
5648     apply (rule is_aligned_no_wrap' [OF al xsz])
5649    apply (simp only: flip: add_diff_eq add_mask_fold)
5650    apply (subst add.assoc, rule word_plus_mono_right)
5651     using al' is_aligned_add_less_t2n xsz
5652     apply fastforce
5653    apply (simp add: field_simps szv al is_aligned_no_overflow)
5654    done
5655qed
5656
5657lemma aligned_mask_diff:
5658  "\<lbrakk> is_aligned (dest :: 'a :: len word) bits; is_aligned (ptr :: 'a :: len word) sz;
5659     bits \<le> sz; sz < LENGTH('a); dest < ptr \<rbrakk>
5660   \<Longrightarrow> mask bits + dest < ptr"
5661  apply (frule_tac p' = ptr in aligned_mask_range_cases, assumption)
5662  apply (elim disjE)
5663    apply (drule_tac is_aligned_no_overflow_mask, simp)+
5664    apply (simp add: algebra_split_simps word_le_not_less)
5665   apply (drule is_aligned_no_overflow_mask; fastforce)
5666  by (simp add: aligned_add_mask_less_eq is_aligned_weaken algebra_split_simps)
5667
5668lemma aligned_mask_ranges_disjoint:
5669  "\<lbrakk> is_aligned (p :: 'a :: len word) n; is_aligned (p' :: 'a :: len word) n';
5670     p && ~~(mask n') \<noteq> p'; p' && ~~(mask n) \<noteq> p \<rbrakk>
5671   \<Longrightarrow> mask_range p n \<inter> mask_range p' n' = {}"
5672  using aligned_mask_range_cases
5673  by (auto simp: neg_mask_in_mask_range)
5674
5675lemma aligned_mask_ranges_disjoint2:
5676  "\<lbrakk> is_aligned p n; is_aligned ptr bits; n \<ge> m; n < size p; m \<le> bits;
5677     (\<forall>y < 2 ^ (n - m). p + (y << m) \<notin> mask_range ptr bits) \<rbrakk>
5678   \<Longrightarrow> mask_range p n \<inter> mask_range ptr bits = {}"
5679  apply safe
5680  apply (simp only: flip: neg_mask_in_mask_range)
5681  apply (drule_tac x="x && mask n >> m" in spec)
5682  apply (clarsimp simp: shiftr_less_t2n and_mask_less_size wsst_TYs multiple_mask_trivia
5683                        word_bw_assocs neg_mask_twice max_absorb2 shiftr_shiftl1)
5684  done
5685
5686lemma leq_mask_shift:
5687  "(x :: 'a :: len word) \<le> mask (low_bits + high_bits) \<Longrightarrow> (x >> low_bits) \<le> mask high_bits"
5688  by (simp add: le_mask_iff shiftr_shiftr)
5689
5690lemma ucast_ucast_eq_mask_shift:
5691  "(x :: 'a :: len word) \<le> mask (low_bits + LENGTH('b))
5692   \<Longrightarrow> ucast((ucast (x >> low_bits)) :: 'b :: len word) = x >> low_bits"
5693  by (meson and_mask_eq_iff_le_mask eq_ucast_ucast_eq not_le_imp_less shiftr_less_t2n'
5694            ucast_ucast_len)
5695
5696lemma const_le_unat:
5697  "\<lbrakk> b < 2 ^ LENGTH('a); of_nat b \<le> a \<rbrakk> \<Longrightarrow> b \<le> unat (a :: 'a :: len word)"
5698  by (clarsimp simp: word_le_def uint_nat of_nat_inverse)
5699
5700lemma upt_enum_offset_trivial:
5701  "\<lbrakk> x < 2 ^ LENGTH('a) - 1 ; n \<le> unat x \<rbrakk>
5702   \<Longrightarrow> ([(0 :: 'a :: len word) .e. x] ! n) = of_nat n"
5703  apply (induct x arbitrary: n)
5704    apply simp
5705  by (simp add: upto_enum_word_nth)
5706
5707lemma word_le_mask_out_plus_2sz:
5708  "x \<le> (x && ~~(mask sz)) + 2 ^ sz - 1"
5709  by (metis add_diff_eq word_neg_and_le)
5710
5711lemma ucast_add:
5712  "ucast (a + (b :: 'a :: len word)) = ucast a + (ucast b :: ('a signed word))"
5713  apply (case_tac "LENGTH('a) = 1")
5714   apply (clarsimp simp: ucast_def)
5715   apply (metis (hide_lams, mono_tags) One_nat_def len_signed plus_word.abs_eq
5716                                       uint_word_arith_bintrs(1) word_ubin.Abs_norm)
5717  apply (clarsimp simp: ucast_def)
5718  apply (metis le_refl len_signed plus_word.abs_eq uint_word_arith_bintrs(1) wi_bintr)
5719  done
5720
5721lemma ucast_minus:
5722  "ucast (a - (b :: 'a :: len word)) = ucast a - (ucast b :: ('a signed word))"
5723  apply (insert ucast_add[where a=a and b="-b"])
5724  apply (metis (no_types, hide_lams) add_diff_eq diff_add_cancel ucast_add)
5725  done
5726
5727lemma scast_ucast_add_one [simp]:
5728  "scast (ucast (x :: 'a::len word) + (1 :: 'a signed word)) = x + 1"
5729  apply (subst ucast_1[symmetric])
5730  apply (subst ucast_add[symmetric])
5731  apply clarsimp
5732  done
5733
5734lemma word_and_le_plus_one:
5735  "a > 0 \<Longrightarrow> (x :: 'a :: len word) && (a - 1) < a"
5736  by (simp add: gt0_iff_gem1 word_and_less')
5737
5738lemma unat_of_ucast_then_shift_eq_unat_of_shift[simp]:
5739  "LENGTH('b) \<ge> LENGTH('a)
5740   \<Longrightarrow> unat ((ucast (x :: 'a :: len word) :: 'b :: len word) >> n) = unat (x >> n)"
5741  by (simp add: shiftr_div_2n' unat_ucast_up_simp)
5742
5743lemma unat_of_ucast_then_mask_eq_unat_of_mask[simp]:
5744  "LENGTH('b) \<ge> LENGTH('a)
5745   \<Longrightarrow> unat ((ucast (x :: 'a :: len word) :: 'b :: len word) && mask m) = unat (x && mask m)"
5746  by (metis ucast_and_mask unat_ucast_up_simp)
5747
5748lemma small_powers_of_2:
5749  "x \<ge> 3 \<Longrightarrow> x < 2 ^ (x - 1)"
5750  by (induct x; simp add: suc_le_pow_2)
5751
5752lemma word_clz_sint_upper[simp]:
5753  "LENGTH('a) \<ge> 3 \<Longrightarrow> sint (of_nat (word_clz (w :: 'a :: len word)) :: 'a sword) \<le> LENGTH('a)"
5754  using small_powers_of_2
5755  by (smt One_nat_def diff_less le_less_trans len_gt_0 len_signed lessI n_less_equal_power_2
5756           not_msb_from_less of_nat_mono sint_eq_uint uint_nat unat_of_nat_eq unat_power_lower
5757           word_clz_max word_of_nat_less wsst_TYs(3))
5758
5759lemma word_clz_sint_lower[simp]:
5760  "LENGTH('a) \<ge> 3
5761   \<Longrightarrow> - sint (of_nat (word_clz (w :: 'a :: len word)) :: 'a signed word) \<le> LENGTH('a)"
5762  apply (subst sint_eq_uint)
5763   using small_powers_of_2 uint_nat
5764   apply (simp add: order_le_less_trans[OF word_clz_max] not_msb_from_less word_of_nat_less
5765                    word_size)
5766  by (simp add: uint_nat)
5767
5768lemma shiftr_less_t2n3:
5769  "\<lbrakk> (2 :: 'a word) ^ (n + m) = 0; m < LENGTH('a) \<rbrakk>
5770   \<Longrightarrow> (x :: 'a :: len word) >> n < 2 ^ m"
5771  by (fastforce intro: shiftr_less_t2n' simp: mask_def power_overflow)
5772
5773lemma unat_shiftr_le_bound:
5774  "\<lbrakk> 2 ^ (LENGTH('a :: len) - n) - 1 \<le> bnd; 0 < n \<rbrakk>
5775   \<Longrightarrow> unat ((x :: 'a word) >> n) \<le> bnd"
5776  using less_not_refl3 le_step_down_nat le_trans less_or_eq_imp_le word_shiftr_lt
5777  by (metis (no_types, lifting))
5778
5779lemma shiftr_eqD:
5780  "\<lbrakk> x >> n = y >> n; is_aligned x n; is_aligned y n \<rbrakk>
5781   \<Longrightarrow> x = y"
5782  by (metis is_aligned_shiftr_shiftl)
5783
5784lemma word_shiftr_shiftl_shiftr_eq_shiftr:
5785  "a \<ge> b \<Longrightarrow> (x :: 'a :: len word) >> a << b >> b = x >> a"
5786  by (simp add: mask_shift multi_shift_simps(5) shiftr_shiftr)
5787
5788lemma of_int_uint_ucast:
5789   "of_int (uint (x :: 'a::len word)) = (ucast x :: 'b::len word)"
5790  by (simp add: ucast_def word_of_int)
5791
5792lemma mod_mask_drop:
5793  "\<lbrakk> m = 2 ^ n; 0 < m; mask n && msk = mask n \<rbrakk>
5794   \<Longrightarrow> (x mod m) && msk = x mod m"
5795  by (simp add: word_mod_2p_is_mask word_bw_assocs)
5796
5797lemma mask_eq_ucast_eq:
5798  "\<lbrakk> x && mask LENGTH('a) = (x :: ('c :: len word));
5799     LENGTH('a) \<le> LENGTH('b)\<rbrakk>
5800    \<Longrightarrow> ucast (ucast x :: ('a :: len word)) = (ucast x :: ('b :: len word))"
5801  by (metis ucast_and_mask ucast_id ucast_ucast_mask ucast_up_eq)
5802
5803lemma of_nat_less_t2n:
5804  "of_nat i < (2 :: ('a :: len) word) ^ n \<Longrightarrow> n < LENGTH('a) \<and> unat (of_nat i :: 'a word) < 2 ^ n"
5805  by (metis order_less_trans p2_gt_0 unat_less_power word_neq_0_conv)
5806
5807lemma two_power_increasing_less_1:
5808  "\<lbrakk> n \<le> m; m \<le> LENGTH('a) \<rbrakk> \<Longrightarrow> (2 :: 'a :: len word) ^ n - 1 \<le> 2 ^ m - 1"
5809  by (metis diff_diff_cancel le_m1_iff_lt less_imp_diff_less p2_gt_0 two_power_increasing
5810            word_1_le_power word_le_minus_mono_left word_less_sub_1)
5811
5812lemma word_sub_mono4:
5813  "\<lbrakk> y + x \<le> z + x; y \<le> y + x; z \<le> z + x \<rbrakk> \<Longrightarrow> y \<le> z" for y :: "'a :: len word"
5814  by (simp add: word_add_le_iff2)
5815
5816lemma eq_or_less_helperD:
5817  "\<lbrakk> n = unat (2 ^ m - 1 :: 'a :: len word) \<or> n < unat (2 ^ m - 1 :: 'a word); m < LENGTH('a) \<rbrakk>
5818   \<Longrightarrow> n < 2 ^ m"
5819  by (meson le_less_trans nat_less_le unat_less_power word_power_less_1)
5820
5821lemma mask_sub:
5822  "n \<le> m \<Longrightarrow> mask m - mask n = mask m && ~~(mask n)"
5823  by (metis (no_types) AND_NOT_mask_plus_AND_mask_eq add_right_imp_eq diff_add_cancel
5824                       and_mask_eq_iff_shiftr_0 shiftr_mask_le word_bool_alg.conj.commute)
5825
5826lemma neg_mask_diff_bound:
5827  "sz'\<le> sz \<Longrightarrow> (ptr && ~~(mask sz')) - (ptr && ~~(mask sz)) \<le> 2 ^ sz - 2 ^ sz'"
5828  (is "_ \<Longrightarrow> ?lhs \<le> ?rhs")
5829proof -
5830  assume lt: "sz' \<le> sz"
5831  hence "?lhs = ptr && (mask sz && ~~(mask sz'))"
5832    by (metis add_diff_cancel_left' multiple_mask_trivia)
5833  also have "\<dots> \<le> ?rhs" using lt
5834    by (metis (mono_tags) add_diff_eq diff_eq_eq eq_iff mask_2pm1 mask_sub word_and_le')
5835  finally show ?thesis by simp
5836qed
5837
5838lemma mask_range_subsetD:
5839  "\<lbrakk> p' \<in> mask_range p n; x' \<in> mask_range p' n'; n' \<le> n; is_aligned p n; is_aligned p' n' \<rbrakk> \<Longrightarrow>
5840   x' \<in> mask_range p n"
5841  using aligned_mask_step by fastforce
5842
5843lemma add_mult_in_mask_range:
5844  "\<lbrakk> is_aligned (base :: 'a :: len word) n; n < LENGTH('a); bits \<le> n; x < 2 ^ (n - bits) \<rbrakk>
5845   \<Longrightarrow> base + x * 2^bits \<in> mask_range base n"
5846  by (simp add: is_aligned_no_wrap' mask_2pm1 nasty_split_lt word_less_power_trans2
5847                word_plus_mono_right)
5848
5849lemma of_bl_length2:
5850  "length xs + c < LENGTH('a) \<Longrightarrow> of_bl xs * 2^c < (2::'a::len word) ^ (length xs + c)"
5851  by (simp add: of_bl_length word_less_power_trans2)
5852
5853lemma mask_out_eq_0:
5854  "\<lbrakk> idx < 2 ^ sz; sz < LENGTH('a) \<rbrakk> \<Longrightarrow> (of_nat idx :: 'a :: len word) && ~~(mask sz) = 0"
5855  by (simp add: Word_Lemmas.of_nat_power less_mask_eq mask_eq_0_eq_x)
5856
5857lemma is_aligned_neg_mask_eq':
5858  "is_aligned ptr sz = (ptr && ~~(mask sz) = ptr)"
5859  using is_aligned_mask mask_eq_0_eq_x by blast
5860
5861lemma neg_mask_mask_unat:
5862  "sz < LENGTH('a)
5863   \<Longrightarrow> unat ((ptr :: 'a :: len word) && ~~(mask sz)) + unat (ptr && mask sz) = unat ptr"
5864  by (metis AND_NOT_mask_plus_AND_mask_eq unat_plus_simple word_and_le2)
5865
5866lemma unat_pow_le_intro:
5867  "LENGTH('a) \<le> n \<Longrightarrow> unat (x :: 'a :: len word) < 2 ^ n"
5868  by (metis lt2p_lem not_le of_nat_le_iff of_nat_numeral semiring_1_class.of_nat_power uint_nat)
5869
5870lemma unat_shiftl_less_t2n:
5871  "\<lbrakk> unat (x :: 'a :: len word) < 2 ^ (m - n); m < LENGTH('a) \<rbrakk> \<Longrightarrow> unat (x << n) < 2 ^ m"
5872  by (metis (no_types) Word_Lemmas.of_nat_power diff_le_self le_less_trans shiftl_less_t2n
5873                       unat_less_power word_unat.Rep_inverse)
5874
5875lemma unat_is_aligned_add:
5876  "\<lbrakk> is_aligned p n; unat d < 2 ^ n \<rbrakk>
5877   \<Longrightarrow> unat (p + d && mask n) = unat d \<and> unat (p + d && ~~(mask n)) = unat p"
5878  by (metis add.right_neutral and_mask_eq_iff_le_mask and_not_mask le_mask_iff mask_add_aligned
5879            mask_out_add_aligned mult_zero_right shiftl_t2n shiftr_le_0)
5880
5881lemma unat_shiftr_shiftl_mask_zero:
5882  "\<lbrakk> c + a \<ge> LENGTH('a) + b ; c < LENGTH('a) \<rbrakk>
5883   \<Longrightarrow> unat (((q :: 'a :: len word) >> a << b) && ~~(mask c)) = 0"
5884  by (fastforce intro: unat_is_aligned_add[where p=0 and n=c, simplified, THEN conjunct2]
5885                       unat_shiftl_less_t2n unat_shiftr_less_t2n unat_pow_le_intro)
5886
5887lemmas of_nat_ucast = ucast_of_nat[symmetric]
5888
5889lemma shift_then_mask_eq_shift_low_bits:
5890  "x \<le> mask (low_bits + high_bits) \<Longrightarrow> (x >> low_bits) && mask high_bits = x >> low_bits"
5891  by (simp add: leq_mask_shift le_mask_imp_and_mask)
5892
5893lemma leq_low_bits_iff_zero:
5894  "\<lbrakk> x \<le> mask (low bits + high bits); x >> low_bits = 0 \<rbrakk> \<Longrightarrow> (x && mask low_bits = 0) = (x = 0)"
5895  using and_mask_eq_iff_shiftr_0 by force
5896
5897lemma unat_less_iff:
5898  "\<lbrakk> unat (a :: 'a :: len word) = b; c < 2 ^ LENGTH('a) \<rbrakk> \<Longrightarrow> (a < of_nat c) = (b < c)"
5899  using unat_ucast_less_no_overflow_simp by blast
5900
5901lemma is_aligned_no_overflow3:
5902 "\<lbrakk> is_aligned (a :: 'a :: len word) n; n < LENGTH('a); b < 2 ^ n; c \<le> 2 ^ n; b < c \<rbrakk>
5903  \<Longrightarrow> a + b \<le> a + (c - 1)"
5904  by (meson is_aligned_no_wrap' le_m1_iff_lt not_le word_less_sub_1 word_plus_mono_right)
5905
5906lemma mask_add_aligned_right:
5907  "is_aligned p n \<Longrightarrow> (q + p) && mask n = q && mask n"
5908  by (simp add: mask_add_aligned add.commute)
5909
5910lemma leq_high_bits_shiftr_low_bits_leq_bits_mask:
5911  "x \<le> mask high_bits \<Longrightarrow> (x :: 'a :: len word) << low_bits \<le> mask (low_bits + high_bits)"
5912  by (metis le_mask_shiftl_le_mask)
5913
5914lemma from_to_bool_last_bit:
5915  "from_bool (to_bool (x && 1)) = x && 1"
5916  by (metis from_bool_to_bool_iff word_and_1)
5917
5918lemma word_two_power_neg_ineq:
5919  "2 ^ m \<noteq> (0 :: 'a word) \<Longrightarrow> 2 ^ n \<le> - (2 ^ m :: 'a :: len word)"
5920  apply (cases "n < LENGTH('a)"; simp add: power_overflow)
5921  apply (cases "m < LENGTH('a)"; simp add: power_overflow)
5922  apply (simp add: word_le_nat_alt unat_minus word_size)
5923  apply (cases "LENGTH('a)"; simp)
5924  apply (simp add: less_Suc_eq_le)
5925  apply (drule power_increasing[where a=2 and n=n] power_increasing[where a=2 and n=m], simp)+
5926  apply (drule(1) add_le_mono)
5927  apply simp
5928  done
5929
5930lemma unat_shiftl_absorb:
5931  "\<lbrakk> x \<le> 2 ^ p; p + k < LENGTH('a) \<rbrakk> \<Longrightarrow> unat (x :: 'a :: len word) * 2 ^ k = unat (x * 2 ^ k)"
5932  by (smt add_diff_cancel_right' add_lessD1 le_add2 le_less_trans mult.commute nat_le_power_trans
5933          unat_lt2p unat_mult_lem unat_power_lower word_le_nat_alt)
5934
5935lemma word_plus_mono_right_split:
5936  "\<lbrakk> unat ((x :: 'a :: len word) && mask sz) + unat z < 2 ^ sz; sz < LENGTH('a) \<rbrakk>
5937   \<Longrightarrow> x \<le> x + z"
5938  apply (subgoal_tac "(x && ~~(mask sz)) + (x && mask sz) \<le> (x && ~~(mask sz)) + ((x && mask sz) + z)")
5939   apply (simp add:word_plus_and_or_coroll2 field_simps)
5940  apply (rule word_plus_mono_right)
5941   apply (simp add: less_le_trans no_olen_add_nat)
5942  using Word_Lemmas.of_nat_power is_aligned_no_wrap' by force
5943
5944lemma mul_not_mask_eq_neg_shiftl:
5945  "~~(mask n) = -1 << n"
5946  by (simp add: NOT_mask shiftl_t2n)
5947
5948lemma shiftr_mul_not_mask_eq_and_not_mask:
5949  "(x >> n) * ~~(mask n) = - (x && ~~(mask n))"
5950  by (metis NOT_mask and_not_mask mult_minus_left semiring_normalization_rules(7) shiftl_t2n)
5951
5952lemma mask_eq_n1_shiftr:
5953  "n \<le> LENGTH('a) \<Longrightarrow> (mask n :: 'a :: len word) = -1 >> (LENGTH('a) - n)"
5954  by (metis diff_diff_cancel eq_refl mask_full shiftr_mask2)
5955
5956lemma is_aligned_mask_out_add_eq:
5957  "is_aligned p n \<Longrightarrow> (p + x) && ~~(mask n) = p + (x && ~~(mask n))"
5958  by (simp add: mask_out_sub_mask mask_add_aligned)
5959
5960lemmas is_aligned_mask_out_add_eq_sub
5961    = is_aligned_mask_out_add_eq[where x="a - b" for a b, simplified field_simps]
5962
5963lemma aligned_bump_down:
5964  "is_aligned x n \<Longrightarrow> (x - 1) && ~~(mask n) = x - 2 ^ n"
5965  by (drule is_aligned_mask_out_add_eq[where x="-1"]) (simp add: NOT_mask)
5966
5967lemma unat_2tp_if:
5968  "unat (2 ^ n :: ('a :: len) word) = (if n < LENGTH ('a) then 2 ^ n else 0)"
5969  by (split if_split, simp_all add: power_overflow)
5970
5971lemma mask_of_mask:
5972  "mask (n::nat) && mask (m::nat) = mask (min m n)"
5973  by word_eqI_solve
5974
5975lemma unat_signed_ucast_less_ucast:
5976  "LENGTH('a) \<le> LENGTH('b) \<Longrightarrow> unat (ucast (x :: 'a :: len word) :: 'b :: len signed word) = unat x"
5977  by (simp add: unat_ucast_up_simp)
5978
5979lemma toEnum_of_ucast:
5980  "LENGTH('b) \<le> LENGTH('a) \<Longrightarrow>
5981   (toEnum (unat (b::'b :: len word))::'a :: len word) = of_nat (unat b)"
5982  by (simp add: unat_pow_le_intro)
5983
5984lemmas unat_ucast_mask = unat_ucast_eq_unat_and_mask[where w=a for a]
5985
5986lemma t2n_mask_eq_if:
5987  "2 ^ n && mask m = (if n < m then 2 ^ n else 0)"
5988  by (rule word_eqI, auto simp add: word_size nth_w2p split: if_split)
5989
5990lemma unat_ucast_le:
5991  "unat (ucast (x :: 'a :: len word) :: 'b :: len word) \<le> unat x"
5992  by (simp add: ucast_nat_def word_unat_less_le)
5993
5994lemma ucast_le_up_down_iff:
5995  "\<lbrakk> LENGTH('a) \<le> LENGTH('b); (x :: 'b :: len word) \<le> ucast (max_word :: 'a :: len word) \<rbrakk>
5996   \<Longrightarrow> (ucast x \<le> (y :: 'a word)) = (x \<le> ucast y)"
5997  using le_max_word_ucast_id ucast_le_ucast by metis
5998
5999lemma ucast_ucast_mask_shift:
6000  "a \<le> LENGTH('a) + b
6001   \<Longrightarrow> ucast (ucast (p && mask a >> b) :: 'a :: len word) = p && mask a >> b"
6002  by (metis add.commute le_mask_iff shiftr_mask_le ucast_ucast_eq_mask_shift word_and_le')
6003
6004lemma unat_ucast_mask_shift:
6005  "a \<le> LENGTH('a) + b
6006   \<Longrightarrow> unat (ucast (p && mask a >> b) :: 'a :: len word) = unat (p && mask a >> b)"
6007  by (metis linear ucast_ucast_mask_shift unat_ucast_up_simp)
6008
6009lemma mask_overlap_zero:
6010  "a \<le> b \<Longrightarrow> (p && mask a) && ~~(mask b) = 0"
6011  by (metis NOT_mask_AND_mask mask_lower_twice2 max_def)
6012
6013lemma mask_shifl_overlap_zero:
6014  "a + c \<le> b \<Longrightarrow> (p && mask a << c) && ~~(mask b) = 0"
6015  by (metis and_not_mask le_mask_iff mask_shiftl_decompose shiftl_0 shiftl_over_and_dist
6016            shiftr_mask_le word_and_le' word_bool_alg.conj_commute)
6017
6018lemma mask_overlap_zero':
6019  "a \<ge> b \<Longrightarrow> (p && ~~(mask a)) && mask b = 0"
6020  using mask_AND_NOT_mask mask_AND_less_0 by blast
6021
6022lemma mask_rshift_mult_eq_rshift_lshift:
6023  "((a :: 'a :: len word) >> b) * (1 << c) = (a >> b << c)"
6024  by (simp add: shiftl_t2n)
6025
6026lemma shift_alignment:
6027  "a \<ge> b \<Longrightarrow> is_aligned (p >> a << a) b"
6028  using is_aligned_shift is_aligned_weaken by blast
6029
6030lemma mask_split_sum_twice:
6031  "a \<ge> b \<Longrightarrow> (p && ~~(mask a)) + ((p && mask a) && ~~(mask b)) + (p && mask b) = p"
6032  by (simp add: add.commute multiple_mask_trivia word_bool_alg.conj_commute
6033                word_bool_alg.conj_left_commute word_plus_and_or_coroll2)
6034
6035lemma mask_shift_eq_mask_mask:
6036  "(p && mask a >> b << b) = (p && mask a) && ~~(mask b)"
6037  by (simp add: and_not_mask)
6038
6039lemma mask_shift_sum:
6040  "\<lbrakk> a \<ge> b; unat n = unat (p && mask b) \<rbrakk>
6041   \<Longrightarrow> (p && ~~(mask a)) + (p && mask a >> b) * (1 << b) + n = (p :: 'a :: len word)"
6042  by (metis and_not_mask mask_rshift_mult_eq_rshift_lshift mask_split_sum_twice word_unat.Rep_eqD)
6043
6044lemma is_up_compose:
6045  "\<lbrakk> is_up uc; is_up uc' \<rbrakk> \<Longrightarrow> is_up (uc' \<circ> uc)"
6046  unfolding is_up_def by (simp add: Word.target_size Word.source_size)
6047
6048lemma of_int_sint_scast:
6049  "of_int (sint (x :: 'a :: len word)) = (scast x :: 'b :: len word)"
6050  by (metis scast_def word_of_int)
6051
6052lemma scast_of_nat_to_signed [simp]:
6053  "scast (of_nat x :: 'a :: len word) = (of_nat x :: 'a signed word)"
6054  by (metis cast_simps(23) scast_scast_id(2))
6055
6056lemma scast_of_nat_signed_to_unsigned_add:
6057  "scast (of_nat x + of_nat y :: 'a :: len signed word) = (of_nat x + of_nat y :: 'a :: len word)"
6058  by (metis of_nat_add scast_of_nat)
6059
6060lemma scast_of_nat_unsigned_to_signed_add:
6061  "(scast (of_nat x + of_nat y :: 'a :: len word)) = (of_nat x + of_nat y :: 'a :: len signed word)"
6062  by (metis Abs_fnat_hom_add scast_of_nat_to_signed)
6063
6064lemma and_mask_cases:
6065  fixes x :: "'a :: len word"
6066  assumes len: "n < LENGTH('a)"
6067  shows "x && mask n \<in> of_nat ` set [0 ..< 2 ^ n]"
6068proof -
6069  have "x && mask n \<in> {0 .. 2 ^ n - 1}"
6070    by (simp add: mask_def word_and_le1)
6071  also
6072  have "... = of_nat ` {0 .. 2 ^ n - 1}"
6073    apply (rule set_eqI, rule iffI)
6074     apply (clarsimp simp: image_iff)
6075     apply (rule_tac x="unat x" in bexI; simp)
6076     using len
6077     apply (simp add: word_le_nat_alt unat_2tp_if unat_minus_one)
6078    using len
6079    apply (clarsimp simp: word_le_nat_alt unat_2tp_if unat_minus_one)
6080    apply (subst unat_of_nat_eq; simp add: nat_le_Suc_less)
6081    apply (erule less_le_trans)
6082    apply simp
6083    done
6084  also have "{0::nat .. 2^n - 1} = set [0 ..< 2^n]" by (auto simp: nat_le_Suc_less)
6085  finally show ?thesis .
6086qed
6087
6088lemma sint_of_nat_ge_zero:
6089  "x < 2 ^ (LENGTH('a) - 1) \<Longrightarrow> sint (of_nat x :: 'a :: len word) \<ge> 0"
6090  by (simp add: Word_Lemmas.of_nat_power not_msb_from_less sint_eq_uint)
6091
6092lemma sint_eq_uint_2pl:
6093  "\<lbrakk> (a :: 'a :: len word) < 2 ^ (LENGTH('a) - 1) \<rbrakk>
6094   \<Longrightarrow> sint a = uint a"
6095  by (simp add: not_msb_from_less sint_eq_uint word_2p_lem word_size)
6096
6097lemma sint_of_nat_le:
6098  "\<lbrakk> b < 2 ^ (LENGTH('a) - 1); a \<le> b \<rbrakk>
6099   \<Longrightarrow> sint (of_nat a :: 'a :: len word) \<le> sint (of_nat b :: 'a :: len word)"
6100  by (smt Word_Lemmas.of_nat_power diff_less le_less_trans len_gt_0 len_of_numeral_defs(2)
6101          nat_power_minus_less of_nat_le_iff sint_eq_uint_2pl uint_nat unat_of_nat_len)
6102
6103lemma int_eq_sint:
6104  "x < 2 ^ (LENGTH('a) - 1) \<Longrightarrow> sint (of_nat x :: 'a :: len word) = int x"
6105  by (smt Word_Lemmas.of_nat_power diff_less le_less_trans len_gt_0 len_of_numeral_defs(2)
6106          nat_less_le sint_eq_uint_2pl uint_nat unat_lt2p unat_of_nat_len unat_power_lower)
6107
6108lemma sint_ctz:
6109  "LENGTH('a) > 2
6110   \<Longrightarrow> 0 \<le> sint (of_nat (word_ctz (x :: 'a :: len word)) :: 'a signed word)
6111        \<and> sint (of_nat (word_ctz x) :: 'a signed word) \<le> LENGTH('a)"
6112  apply (subgoal_tac "LENGTH('a) < 2 ^ (LENGTH('a) - 1)")
6113   apply (rule conjI)
6114    apply (metis len_signed order_le_less_trans sint_of_nat_ge_zero word_ctz_le)
6115   apply (metis int_eq_sint len_signed sint_of_nat_le word_ctz_le)
6116  by (rule small_powers_of_2, simp)
6117
6118lemma pow_sub_less:
6119  "\<lbrakk> a + b \<le> LENGTH('a); unat (x :: 'a :: len word) = 2 ^ a \<rbrakk>
6120   \<Longrightarrow> unat (x * 2 ^ b - 1) < 2 ^ (a + b)"
6121  by (metis (mono_tags) eq_or_less_helperD not_less of_nat_numeral power_add
6122                        semiring_1_class.of_nat_power unat_pow_le_intro word_unat.Rep_inverse)
6123
6124lemma sle_le_2pl:
6125  "\<lbrakk> (b :: 'a :: len word) < 2 ^ (LENGTH('a) - 1); a \<le> b \<rbrakk> \<Longrightarrow> a <=s b"
6126  by (simp add: not_msb_from_less word_sle_msb_le)
6127
6128lemma sless_less_2pl:
6129  "\<lbrakk> (b :: 'a :: len word) < 2 ^ (LENGTH('a) - 1); a < b \<rbrakk> \<Longrightarrow> a <s b"
6130  using not_msb_from_less word_sless_msb_less by blast
6131
6132lemma and_mask2:
6133  "w << n >> n = w && mask (size w - n)"
6134  by (cases "n \<le> size w"; clarsimp simp: word_and_le2 and_mask shiftl_zero_size)
6135
6136lemma unat_of_nat_word_log2:
6137  "LENGTH('a) < 2 ^ LENGTH('b)
6138   \<Longrightarrow> unat (of_nat (word_log2 (n :: 'a :: len word)) :: 'b :: len word) = word_log2 n"
6139  by (metis less_trans unat_of_nat_eq word_log2_max word_size)
6140
6141lemma aligned_sub_aligned_simple:
6142  "\<lbrakk> is_aligned a n; is_aligned b n \<rbrakk> \<Longrightarrow> is_aligned (a - b) n"
6143  by (simp add: aligned_sub_aligned)
6144
6145lemma minus_one_shift:
6146  "- (1 << n) = (-1 << n :: 'a::len word)"
6147  by (simp add: mask_def NOT_eq flip: mul_not_mask_eq_neg_shiftl)
6148
6149lemma ucast_eq_mask:
6150  "(UCAST('a::len \<rightarrow> 'b::len) x = UCAST('a \<rightarrow> 'b) y) =
6151   (x && mask LENGTH('b) = y && mask LENGTH('b))"
6152  by (rule iffI; word_eqI_solve)
6153
6154context
6155  fixes w :: "'a::len word"
6156begin
6157
6158private lemma sbintrunc_uint_ucast:
6159  assumes "Suc n = LENGTH('b::len)"
6160  shows "sbintrunc n (uint (ucast w :: 'b word)) = sbintrunc n (uint w)"
6161  by (metis assms sbintrunc_bintrunc ucast_def word_ubin.eq_norm)
6162
6163private lemma test_bit_sbintrunc:
6164  assumes "i < LENGTH('a)"
6165  shows "(word_of_int (sbintrunc n (uint w)) :: 'a word) !! i
6166           = (if n < i then w !! n else w !! i)"
6167  using assms by (simp add: nth_sbintr)
6168                 (simp add: test_bit_bin)
6169
6170private lemma test_bit_sbintrunc_ucast:
6171  assumes len_a: "i < LENGTH('a)"
6172  shows "(word_of_int (sbintrunc (LENGTH('b) - 1) (uint (ucast w :: 'b word))) :: 'a word) !! i
6173          = (if LENGTH('b::len) \<le> i then w !! (LENGTH('b) - 1) else w !! i)"
6174  apply (subst sbintrunc_uint_ucast)
6175   apply simp
6176  apply (subst test_bit_sbintrunc)
6177   apply (rule len_a)
6178  apply (rule if_cong[OF _ refl refl])
6179  using leD less_linear by fastforce
6180
6181lemma scast_ucast_high_bits:
6182  "scast (ucast w :: 'b::len word) = w
6183     \<longleftrightarrow> (\<forall> i \<in> {LENGTH('b) ..< size w}. w !! i = w !! (LENGTH('b) - 1))"
6184  unfolding scast_def sint_uint word_size
6185  apply (subst word_eq_iff)
6186  apply (rule iffI)
6187   apply (rule ballI)
6188   apply (drule_tac x=i in spec)
6189   apply (subst (asm) test_bit_sbintrunc_ucast; simp)
6190  apply (rule allI)
6191  apply (case_tac "n < LENGTH('a)")
6192   apply (subst test_bit_sbintrunc_ucast)
6193    apply simp
6194   apply (case_tac "n \<ge> LENGTH('b)")
6195    apply (drule_tac x=n in bspec)
6196  by auto
6197
6198lemma scast_ucast_mask_compare:
6199  "scast (ucast w :: 'b::len word) = w
6200   \<longleftrightarrow> (w \<le> mask (LENGTH('b) - 1) \<or> ~~(mask (LENGTH('b) - 1)) \<le> w)"
6201  apply (clarsimp simp: le_mask_high_bits neg_mask_le_high_bits scast_ucast_high_bits word_size)
6202  apply (rule iffI; clarsimp)
6203   apply (rename_tac i j; case_tac "i = LENGTH('b) - 1"; case_tac "j = LENGTH('b) - 1")
6204  by auto
6205
6206lemma ucast_less_shiftl_helper':
6207  "\<lbrakk> LENGTH('b) + (a::nat) < LENGTH('a); 2 ^ (LENGTH('b) + a) \<le> n\<rbrakk>
6208   \<Longrightarrow> (ucast (x :: 'b::len word) << a) < (n :: 'a::len word)"
6209  apply (erule order_less_le_trans[rotated])
6210  using ucast_less[where x=x and 'a='a]
6211  apply (simp only: shiftl_t2n field_simps)
6212  apply (rule word_less_power_trans2; simp)
6213  done
6214
6215end
6216
6217lemma ucast_ucast_mask2:
6218  "is_down (UCAST ('a \<rightarrow> 'b)) \<Longrightarrow>
6219   UCAST ('b::len \<rightarrow> 'c::len) (UCAST ('a::len \<rightarrow> 'b::len) x) = UCAST ('a \<rightarrow> 'c) (x && mask LENGTH('b))"
6220  by word_eqI_solve
6221
6222lemma ucast_NOT:
6223  "ucast (~~x) = ~~(ucast x) && mask (LENGTH('a))" for x::"'a::len word"
6224  by word_eqI
6225
6226lemma ucast_NOT_down:
6227  "is_down UCAST('a::len \<rightarrow> 'b::len) \<Longrightarrow> UCAST('a \<rightarrow> 'b) (~~x) = ~~(UCAST('a \<rightarrow> 'b) x)"
6228  by word_eqI
6229
6230lemma of_bl_mult_and_not_mask_eq:
6231  "\<lbrakk>is_aligned (a :: 'a::len word) n; length b + m \<le> n\<rbrakk>
6232   \<Longrightarrow> a + of_bl b * (2^m) && ~~(mask n) = a"
6233  by (smt add.left_neutral add_diff_cancel_right' add_mask_lower_bits and_mask_plus is_aligned_mask
6234          is_aligned_weaken le_less_trans of_bl_length2 subtract_mask(1))
6235
6236lemma bin_to_bl_of_bl_eq:
6237  "\<lbrakk>is_aligned (a::'a::len word) n; length b + c \<le> n; length b + c < LENGTH('a)\<rbrakk>
6238  \<Longrightarrow> bin_to_bl (length b) (uint ((a + of_bl b * 2^c) >> c)) = b"
6239  apply (subst word_plus_and_or_coroll)
6240   apply (erule is_aligned_get_word_bits)
6241    apply (rule is_aligned_AND_less_0)
6242     apply (simp add: is_aligned_mask)
6243    apply (rule order_less_le_trans)
6244     apply (rule of_bl_length2)
6245     apply simp
6246    apply (simp add: two_power_increasing)
6247   apply simp
6248  apply (rule nth_equalityI)
6249   apply (simp only: len_bin_to_bl)
6250  apply (clarsimp simp only: len_bin_to_bl nth_bin_to_bl word_test_bit_def[symmetric])
6251  apply (simp add: nth_shiftr nth_shiftl
6252                   shiftl_t2n[where n=c, simplified mult.commute, simplified, symmetric])
6253  apply (simp add: is_aligned_nth[THEN iffD1, rule_format] test_bit_of_bl nth_rev)
6254  apply arith
6255  done
6256
6257end
6258