author | huffman |
Fri, 24 Aug 2007 00:37:12 +0200 | |
changeset 24419 | 737622204802 |
parent 24417 | 5c3858b71f80 |
child 24465 | 70f0214b3ecc |
permissions | -rw-r--r-- |
24333 | 1 |
(* |
2 |
ID: $Id$ |
|
3 |
Author: Jeremy Dawson, NICTA |
|
4 |
||
5 |
contains basic definition to do with integers |
|
6 |
expressed using Pls, Min, BIT and important resulting theorems, |
|
7 |
in particular, bin_rec and related work |
|
8 |
*) |
|
9 |
||
24350 | 10 |
header {* Basic Definitions for Binary Integers *} |
11 |
||
24413 | 12 |
theory BinGeneral imports BinInduct Num_Lemmas |
24333 | 13 |
|
14 |
begin |
|
15 |
||
24384
0002537695df
move BIT datatype stuff from Num_Lemmas to BinGeneral
huffman
parents:
24383
diff
changeset
|
16 |
subsection {* BIT as a datatype constructor *} |
0002537695df
move BIT datatype stuff from Num_Lemmas to BinGeneral
huffman
parents:
24383
diff
changeset
|
17 |
|
0002537695df
move BIT datatype stuff from Num_Lemmas to BinGeneral
huffman
parents:
24383
diff
changeset
|
18 |
(** ways in which type Bin resembles a datatype **) |
0002537695df
move BIT datatype stuff from Num_Lemmas to BinGeneral
huffman
parents:
24383
diff
changeset
|
19 |
|
24413 | 20 |
lemmas BIT_eqI = conjI [THEN BIT_eq_iff [THEN iffD2]] |
24384
0002537695df
move BIT datatype stuff from Num_Lemmas to BinGeneral
huffman
parents:
24383
diff
changeset
|
21 |
|
0002537695df
move BIT datatype stuff from Num_Lemmas to BinGeneral
huffman
parents:
24383
diff
changeset
|
22 |
lemma neB1E [elim!]: |
0002537695df
move BIT datatype stuff from Num_Lemmas to BinGeneral
huffman
parents:
24383
diff
changeset
|
23 |
assumes ne: "y \<noteq> bit.B1" |
0002537695df
move BIT datatype stuff from Num_Lemmas to BinGeneral
huffman
parents:
24383
diff
changeset
|
24 |
assumes y: "y = bit.B0 \<Longrightarrow> P" |
0002537695df
move BIT datatype stuff from Num_Lemmas to BinGeneral
huffman
parents:
24383
diff
changeset
|
25 |
shows "P" |
0002537695df
move BIT datatype stuff from Num_Lemmas to BinGeneral
huffman
parents:
24383
diff
changeset
|
26 |
apply (rule y) |
0002537695df
move BIT datatype stuff from Num_Lemmas to BinGeneral
huffman
parents:
24383
diff
changeset
|
27 |
apply (cases y rule: bit.exhaust, simp) |
0002537695df
move BIT datatype stuff from Num_Lemmas to BinGeneral
huffman
parents:
24383
diff
changeset
|
28 |
apply (simp add: ne) |
0002537695df
move BIT datatype stuff from Num_Lemmas to BinGeneral
huffman
parents:
24383
diff
changeset
|
29 |
done |
0002537695df
move BIT datatype stuff from Num_Lemmas to BinGeneral
huffman
parents:
24383
diff
changeset
|
30 |
|
0002537695df
move BIT datatype stuff from Num_Lemmas to BinGeneral
huffman
parents:
24383
diff
changeset
|
31 |
lemma bin_exhaust: |
0002537695df
move BIT datatype stuff from Num_Lemmas to BinGeneral
huffman
parents:
24383
diff
changeset
|
32 |
assumes Q: "\<And>x b. bin = x BIT b \<Longrightarrow> Q" |
0002537695df
move BIT datatype stuff from Num_Lemmas to BinGeneral
huffman
parents:
24383
diff
changeset
|
33 |
shows "Q" |
24413 | 34 |
by (rule BIT_cases, rule Q) |
24384
0002537695df
move BIT datatype stuff from Num_Lemmas to BinGeneral
huffman
parents:
24383
diff
changeset
|
35 |
|
24350 | 36 |
subsection {* Recursion combinator for binary integers *} |
37 |
||
24333 | 38 |
function |
39 |
bin_rec' :: "int * 'a * 'a * (int => bit => 'a => 'a) => 'a" |
|
40 |
where |
|
41 |
"bin_rec' (bin, f1, f2, f3) = (if bin = Numeral.Pls then f1 |
|
42 |
else if bin = Numeral.Min then f2 |
|
24413 | 43 |
else f3 (bin_rest bin) (bin_last bin) |
44 |
(bin_rec' (bin_rest bin, f1, f2, f3)))" |
|
24333 | 45 |
by pat_completeness auto |
46 |
||
47 |
termination |
|
24413 | 48 |
by (relation "measure (size o fst)") simp_all |
24333 | 49 |
|
50 |
constdefs |
|
51 |
bin_rec :: "'a => 'a => (int => bit => 'a => 'a) => int => 'a" |
|
52 |
"bin_rec f1 f2 f3 bin == bin_rec' (bin, f1, f2, f3)" |
|
53 |
||
24417 | 54 |
lemma bin_rec_Pls: "bin_rec f1 f2 f3 Numeral.Pls = f1" |
55 |
unfolding bin_rec_def by simp |
|
24333 | 56 |
|
24417 | 57 |
lemma bin_rec_Min: "bin_rec f1 f2 f3 Numeral.Min = f2" |
58 |
unfolding bin_rec_def by simp |
|
24333 | 59 |
|
60 |
lemma bin_rec_Bit: |
|
61 |
"f = bin_rec f1 f2 f3 ==> f3 Numeral.Pls bit.B0 f1 = f1 ==> |
|
62 |
f3 Numeral.Min bit.B1 f2 = f2 ==> f (w BIT b) = f3 w b (f w)" |
|
63 |
apply clarify |
|
64 |
apply (unfold bin_rec_def) |
|
65 |
apply (rule bin_rec'.simps [THEN trans]) |
|
66 |
apply auto |
|
67 |
done |
|
68 |
||
69 |
lemmas bin_rec_simps = refl [THEN bin_rec_Bit] bin_rec_Pls bin_rec_Min |
|
70 |
||
24413 | 71 |
subsection {* Sign of binary integers *} |
24364 | 72 |
|
73 |
consts |
|
74 |
bin_sign :: "int => int" |
|
75 |
||
76 |
defs |
|
77 |
bin_sign_def : "bin_sign == bin_rec Numeral.Pls Numeral.Min (%w b s. s)" |
|
78 |
||
24413 | 79 |
lemmas bin_rest_simps = |
80 |
bin_rest_Pls bin_rest_Min bin_rest_BIT |
|
24364 | 81 |
|
24413 | 82 |
lemmas bin_last_simps = |
83 |
bin_last_Pls bin_last_Min bin_last_BIT |
|
24333 | 84 |
|
85 |
lemma bin_sign_simps [simp]: |
|
86 |
"bin_sign Numeral.Pls = Numeral.Pls" |
|
87 |
"bin_sign Numeral.Min = Numeral.Min" |
|
88 |
"bin_sign (w BIT b) = bin_sign w" |
|
89 |
unfolding bin_sign_def by (auto simp: bin_rec_simps) |
|
90 |
||
91 |
lemma bin_r_l_extras [simp]: |
|
92 |
"bin_last 0 = bit.B0" |
|
93 |
"bin_last (- 1) = bit.B1" |
|
94 |
"bin_last -1 = bit.B1" |
|
95 |
"bin_last 1 = bit.B1" |
|
96 |
"bin_rest 1 = 0" |
|
97 |
"bin_rest 0 = 0" |
|
98 |
"bin_rest (- 1) = - 1" |
|
99 |
"bin_rest -1 = -1" |
|
100 |
apply (unfold number_of_Min) |
|
101 |
apply (unfold Pls_def [symmetric] Min_def [symmetric]) |
|
102 |
apply (unfold numeral_1_eq_1 [symmetric]) |
|
103 |
apply (auto simp: number_of_eq) |
|
104 |
done |
|
105 |
||
106 |
lemma bin_last_mod: |
|
107 |
"bin_last w = (if w mod 2 = 0 then bit.B0 else bit.B1)" |
|
24413 | 108 |
apply (subgoal_tac "bin_last w = |
109 |
(if number_of w mod 2 = (0::int) then bit.B0 else bit.B1)") |
|
110 |
apply (simp only: number_of_is_id) |
|
111 |
apply (induct w rule: int_bin_induct, simp_all) |
|
24333 | 112 |
done |
113 |
||
114 |
lemma bin_rest_div: |
|
115 |
"bin_rest w = w div 2" |
|
24413 | 116 |
apply (subgoal_tac "number_of (bin_rest w) = number_of w div (2::int)") |
117 |
apply (simp only: number_of_is_id) |
|
118 |
apply (induct w rule: int_bin_induct, simp_all) |
|
24333 | 119 |
done |
120 |
||
121 |
lemma Bit_div2 [simp]: "(w BIT b) div 2 = w" |
|
122 |
unfolding bin_rest_div [symmetric] by auto |
|
123 |
||
24409 | 124 |
subsection {* Testing bit positions *} |
125 |
||
126 |
consts |
|
127 |
bin_nth :: "int => nat => bool" |
|
128 |
||
129 |
primrec |
|
130 |
Z : "bin_nth w 0 = (bin_last w = bit.B1)" |
|
131 |
Suc : "bin_nth w (Suc n) = bin_nth (bin_rest w) n" |
|
132 |
||
24333 | 133 |
lemma bin_nth_lem [rule_format]: |
134 |
"ALL y. bin_nth x = bin_nth y --> x = y" |
|
135 |
apply (induct x rule: bin_induct) |
|
136 |
apply safe |
|
137 |
apply (erule rev_mp) |
|
138 |
apply (induct_tac y rule: bin_induct) |
|
139 |
apply safe |
|
140 |
apply (drule_tac x=0 in fun_cong, force) |
|
141 |
apply (erule notE, rule ext, |
|
142 |
drule_tac x="Suc x" in fun_cong, force) |
|
143 |
apply (drule_tac x=0 in fun_cong, force) |
|
144 |
apply (erule rev_mp) |
|
145 |
apply (induct_tac y rule: bin_induct) |
|
146 |
apply safe |
|
147 |
apply (drule_tac x=0 in fun_cong, force) |
|
148 |
apply (erule notE, rule ext, |
|
149 |
drule_tac x="Suc x" in fun_cong, force) |
|
150 |
apply (drule_tac x=0 in fun_cong, force) |
|
151 |
apply (case_tac y rule: bin_exhaust) |
|
152 |
apply clarify |
|
153 |
apply (erule allE) |
|
154 |
apply (erule impE) |
|
155 |
prefer 2 |
|
156 |
apply (erule BIT_eqI) |
|
157 |
apply (drule_tac x=0 in fun_cong, force) |
|
158 |
apply (rule ext) |
|
159 |
apply (drule_tac x="Suc ?x" in fun_cong, force) |
|
160 |
done |
|
161 |
||
162 |
lemma bin_nth_eq_iff: "(bin_nth x = bin_nth y) = (x = y)" |
|
163 |
by (auto elim: bin_nth_lem) |
|
164 |
||
165 |
lemmas bin_eqI = ext [THEN bin_nth_eq_iff [THEN iffD1], standard] |
|
166 |
||
167 |
lemma bin_nth_Pls [simp]: "~ bin_nth Numeral.Pls n" |
|
168 |
by (induct n) auto |
|
169 |
||
170 |
lemma bin_nth_Min [simp]: "bin_nth Numeral.Min n" |
|
171 |
by (induct n) auto |
|
172 |
||
173 |
lemma bin_nth_0_BIT: "bin_nth (w BIT b) 0 = (b = bit.B1)" |
|
174 |
by auto |
|
175 |
||
176 |
lemma bin_nth_Suc_BIT: "bin_nth (w BIT b) (Suc n) = bin_nth w n" |
|
177 |
by auto |
|
178 |
||
179 |
lemma bin_nth_minus [simp]: "0 < n ==> bin_nth (w BIT b) n = bin_nth w (n - 1)" |
|
180 |
by (cases n) auto |
|
181 |
||
182 |
lemmas bin_nth_0 = bin_nth.simps(1) |
|
183 |
lemmas bin_nth_Suc = bin_nth.simps(2) |
|
184 |
||
185 |
lemmas bin_nth_simps = |
|
186 |
bin_nth_0 bin_nth_Suc bin_nth_Pls bin_nth_Min bin_nth_minus |
|
187 |
||
24364 | 188 |
lemma bin_sign_rest [simp]: |
189 |
"bin_sign (bin_rest w) = (bin_sign w)" |
|
190 |
by (case_tac w rule: bin_exhaust) auto |
|
191 |
||
192 |
subsection {* Truncating binary integers *} |
|
193 |
||
194 |
consts |
|
195 |
bintrunc :: "nat => int => int" |
|
196 |
primrec |
|
197 |
Z : "bintrunc 0 bin = Numeral.Pls" |
|
198 |
Suc : "bintrunc (Suc n) bin = bintrunc n (bin_rest bin) BIT (bin_last bin)" |
|
199 |
||
200 |
consts |
|
201 |
sbintrunc :: "nat => int => int" |
|
202 |
primrec |
|
203 |
Z : "sbintrunc 0 bin = |
|
204 |
(case bin_last bin of bit.B1 => Numeral.Min | bit.B0 => Numeral.Pls)" |
|
205 |
Suc : "sbintrunc (Suc n) bin = sbintrunc n (bin_rest bin) BIT (bin_last bin)" |
|
206 |
||
24333 | 207 |
lemma sign_bintr: |
208 |
"!!w. bin_sign (bintrunc n w) = Numeral.Pls" |
|
209 |
by (induct n) auto |
|
210 |
||
211 |
lemma bintrunc_mod2p: |
|
212 |
"!!w. bintrunc n w = (w mod 2 ^ n :: int)" |
|
213 |
apply (induct n, clarsimp) |
|
214 |
apply (simp add: bin_last_mod bin_rest_div Bit_def zmod_zmult2_eq |
|
215 |
cong: number_of_False_cong) |
|
216 |
done |
|
217 |
||
218 |
lemma sbintrunc_mod2p: |
|
219 |
"!!w. sbintrunc n w = ((w + 2 ^ n) mod 2 ^ (Suc n) - 2 ^ n :: int)" |
|
220 |
apply (induct n) |
|
221 |
apply clarsimp |
|
222 |
apply (subst zmod_zadd_left_eq) |
|
223 |
apply (simp add: bin_last_mod) |
|
224 |
apply (simp add: number_of_eq) |
|
225 |
apply clarsimp |
|
226 |
apply (simp add: bin_last_mod bin_rest_div Bit_def |
|
227 |
cong: number_of_False_cong) |
|
228 |
apply (clarsimp simp: zmod_zmult_zmult1 [symmetric] |
|
229 |
zmod_zdiv_equality [THEN diff_eq_eq [THEN iffD2 [THEN sym]]]) |
|
230 |
apply (rule trans [symmetric, OF _ emep1]) |
|
231 |
apply auto |
|
232 |
apply (auto simp: even_def) |
|
233 |
done |
|
234 |
||
24413 | 235 |
text "Simplifications for (s)bintrunc" |
24333 | 236 |
|
237 |
lemma bin_sign_lem: |
|
238 |
"!!bin. (bin_sign (sbintrunc n bin) = Numeral.Min) = bin_nth bin n" |
|
239 |
apply (induct n) |
|
240 |
apply (case_tac bin rule: bin_exhaust, case_tac b, auto)+ |
|
241 |
done |
|
242 |
||
243 |
lemma nth_bintr: |
|
244 |
"!!w m. bin_nth (bintrunc m w) n = (n < m & bin_nth w n)" |
|
245 |
apply (induct n) |
|
246 |
apply (case_tac m, auto)[1] |
|
247 |
apply (case_tac m, auto)[1] |
|
248 |
done |
|
249 |
||
250 |
lemma nth_sbintr: |
|
251 |
"!!w m. bin_nth (sbintrunc m w) n = |
|
252 |
(if n < m then bin_nth w n else bin_nth w m)" |
|
253 |
apply (induct n) |
|
254 |
apply (case_tac m, simp_all split: bit.splits)[1] |
|
255 |
apply (case_tac m, simp_all split: bit.splits)[1] |
|
256 |
done |
|
257 |
||
258 |
lemma bin_nth_Bit: |
|
259 |
"bin_nth (w BIT b) n = (n = 0 & b = bit.B1 | (EX m. n = Suc m & bin_nth w m))" |
|
260 |
by (cases n) auto |
|
261 |
||
262 |
lemma bintrunc_bintrunc_l: |
|
263 |
"n <= m ==> (bintrunc m (bintrunc n w) = bintrunc n w)" |
|
264 |
by (rule bin_eqI) (auto simp add : nth_bintr) |
|
265 |
||
266 |
lemma sbintrunc_sbintrunc_l: |
|
267 |
"n <= m ==> (sbintrunc m (sbintrunc n w) = sbintrunc n w)" |
|
268 |
by (rule bin_eqI) (auto simp: nth_sbintr min_def) |
|
269 |
||
270 |
lemma bintrunc_bintrunc_ge: |
|
271 |
"n <= m ==> (bintrunc n (bintrunc m w) = bintrunc n w)" |
|
272 |
by (rule bin_eqI) (auto simp: nth_bintr) |
|
273 |
||
274 |
lemma bintrunc_bintrunc_min [simp]: |
|
275 |
"bintrunc m (bintrunc n w) = bintrunc (min m n) w" |
|
276 |
apply (unfold min_def) |
|
277 |
apply (rule bin_eqI) |
|
278 |
apply (auto simp: nth_bintr) |
|
279 |
done |
|
280 |
||
281 |
lemma sbintrunc_sbintrunc_min [simp]: |
|
282 |
"sbintrunc m (sbintrunc n w) = sbintrunc (min m n) w" |
|
283 |
apply (unfold min_def) |
|
284 |
apply (rule bin_eqI) |
|
285 |
apply (auto simp: nth_sbintr) |
|
286 |
done |
|
287 |
||
288 |
lemmas bintrunc_Pls = |
|
289 |
bintrunc.Suc [where bin="Numeral.Pls", simplified bin_last_simps bin_rest_simps] |
|
290 |
||
291 |
lemmas bintrunc_Min [simp] = |
|
292 |
bintrunc.Suc [where bin="Numeral.Min", simplified bin_last_simps bin_rest_simps] |
|
293 |
||
294 |
lemmas bintrunc_BIT [simp] = |
|
295 |
bintrunc.Suc [where bin="?w BIT ?b", simplified bin_last_simps bin_rest_simps] |
|
296 |
||
297 |
lemmas bintrunc_Sucs = bintrunc_Pls bintrunc_Min bintrunc_BIT |
|
298 |
||
299 |
lemmas sbintrunc_Suc_Pls = |
|
300 |
sbintrunc.Suc [where bin="Numeral.Pls", simplified bin_last_simps bin_rest_simps] |
|
301 |
||
302 |
lemmas sbintrunc_Suc_Min = |
|
303 |
sbintrunc.Suc [where bin="Numeral.Min", simplified bin_last_simps bin_rest_simps] |
|
304 |
||
305 |
lemmas sbintrunc_Suc_BIT [simp] = |
|
306 |
sbintrunc.Suc [where bin="?w BIT ?b", simplified bin_last_simps bin_rest_simps] |
|
307 |
||
308 |
lemmas sbintrunc_Sucs = sbintrunc_Suc_Pls sbintrunc_Suc_Min sbintrunc_Suc_BIT |
|
309 |
||
310 |
lemmas sbintrunc_Pls = |
|
311 |
sbintrunc.Z [where bin="Numeral.Pls", |
|
312 |
simplified bin_last_simps bin_rest_simps bit.simps] |
|
313 |
||
314 |
lemmas sbintrunc_Min = |
|
315 |
sbintrunc.Z [where bin="Numeral.Min", |
|
316 |
simplified bin_last_simps bin_rest_simps bit.simps] |
|
317 |
||
318 |
lemmas sbintrunc_0_BIT_B0 [simp] = |
|
319 |
sbintrunc.Z [where bin="?w BIT bit.B0", |
|
320 |
simplified bin_last_simps bin_rest_simps bit.simps] |
|
321 |
||
322 |
lemmas sbintrunc_0_BIT_B1 [simp] = |
|
323 |
sbintrunc.Z [where bin="?w BIT bit.B1", |
|
324 |
simplified bin_last_simps bin_rest_simps bit.simps] |
|
325 |
||
326 |
lemmas sbintrunc_0_simps = |
|
327 |
sbintrunc_Pls sbintrunc_Min sbintrunc_0_BIT_B0 sbintrunc_0_BIT_B1 |
|
328 |
||
329 |
lemmas bintrunc_simps = bintrunc.Z bintrunc_Sucs |
|
330 |
lemmas sbintrunc_simps = sbintrunc_0_simps sbintrunc_Sucs |
|
331 |
||
332 |
lemma bintrunc_minus: |
|
333 |
"0 < n ==> bintrunc (Suc (n - 1)) w = bintrunc n w" |
|
334 |
by auto |
|
335 |
||
336 |
lemma sbintrunc_minus: |
|
337 |
"0 < n ==> sbintrunc (Suc (n - 1)) w = sbintrunc n w" |
|
338 |
by auto |
|
339 |
||
340 |
lemmas bintrunc_minus_simps = |
|
341 |
bintrunc_Sucs [THEN [2] bintrunc_minus [symmetric, THEN trans], standard] |
|
342 |
lemmas sbintrunc_minus_simps = |
|
343 |
sbintrunc_Sucs [THEN [2] sbintrunc_minus [symmetric, THEN trans], standard] |
|
344 |
||
345 |
lemma bintrunc_n_Pls [simp]: |
|
346 |
"bintrunc n Numeral.Pls = Numeral.Pls" |
|
347 |
by (induct n) auto |
|
348 |
||
349 |
lemma sbintrunc_n_PM [simp]: |
|
350 |
"sbintrunc n Numeral.Pls = Numeral.Pls" |
|
351 |
"sbintrunc n Numeral.Min = Numeral.Min" |
|
352 |
by (induct n) auto |
|
353 |
||
354 |
lemmas thobini1 = arg_cong [where f = "%w. w BIT ?b"] |
|
355 |
||
356 |
lemmas bintrunc_BIT_I = trans [OF bintrunc_BIT thobini1] |
|
357 |
lemmas bintrunc_Min_I = trans [OF bintrunc_Min thobini1] |
|
358 |
||
359 |
lemmas bmsts = bintrunc_minus_simps [THEN thobini1 [THEN [2] trans], standard] |
|
360 |
lemmas bintrunc_Pls_minus_I = bmsts(1) |
|
361 |
lemmas bintrunc_Min_minus_I = bmsts(2) |
|
362 |
lemmas bintrunc_BIT_minus_I = bmsts(3) |
|
363 |
||
364 |
lemma bintrunc_0_Min: "bintrunc 0 Numeral.Min = Numeral.Pls" |
|
365 |
by auto |
|
366 |
lemma bintrunc_0_BIT: "bintrunc 0 (w BIT b) = Numeral.Pls" |
|
367 |
by auto |
|
368 |
||
369 |
lemma bintrunc_Suc_lem: |
|
370 |
"bintrunc (Suc n) x = y ==> m = Suc n ==> bintrunc m x = y" |
|
371 |
by auto |
|
372 |
||
373 |
lemmas bintrunc_Suc_Ialts = |
|
374 |
bintrunc_Min_I bintrunc_BIT_I [THEN bintrunc_Suc_lem, standard] |
|
375 |
||
376 |
lemmas sbintrunc_BIT_I = trans [OF sbintrunc_Suc_BIT thobini1] |
|
377 |
||
378 |
lemmas sbintrunc_Suc_Is = |
|
379 |
sbintrunc_Sucs [THEN thobini1 [THEN [2] trans], standard] |
|
380 |
||
381 |
lemmas sbintrunc_Suc_minus_Is = |
|
382 |
sbintrunc_minus_simps [THEN thobini1 [THEN [2] trans], standard] |
|
383 |
||
384 |
lemma sbintrunc_Suc_lem: |
|
385 |
"sbintrunc (Suc n) x = y ==> m = Suc n ==> sbintrunc m x = y" |
|
386 |
by auto |
|
387 |
||
388 |
lemmas sbintrunc_Suc_Ialts = |
|
389 |
sbintrunc_Suc_Is [THEN sbintrunc_Suc_lem, standard] |
|
390 |
||
391 |
lemma sbintrunc_bintrunc_lt: |
|
392 |
"m > n ==> sbintrunc n (bintrunc m w) = sbintrunc n w" |
|
393 |
by (rule bin_eqI) (auto simp: nth_sbintr nth_bintr) |
|
394 |
||
395 |
lemma bintrunc_sbintrunc_le: |
|
396 |
"m <= Suc n ==> bintrunc m (sbintrunc n w) = bintrunc m w" |
|
397 |
apply (rule bin_eqI) |
|
398 |
apply (auto simp: nth_sbintr nth_bintr) |
|
399 |
apply (subgoal_tac "x=n", safe, arith+)[1] |
|
400 |
apply (subgoal_tac "x=n", safe, arith+)[1] |
|
401 |
done |
|
402 |
||
403 |
lemmas bintrunc_sbintrunc [simp] = order_refl [THEN bintrunc_sbintrunc_le] |
|
404 |
lemmas sbintrunc_bintrunc [simp] = lessI [THEN sbintrunc_bintrunc_lt] |
|
405 |
lemmas bintrunc_bintrunc [simp] = order_refl [THEN bintrunc_bintrunc_l] |
|
406 |
lemmas sbintrunc_sbintrunc [simp] = order_refl [THEN sbintrunc_sbintrunc_l] |
|
407 |
||
408 |
lemma bintrunc_sbintrunc' [simp]: |
|
409 |
"0 < n \<Longrightarrow> bintrunc n (sbintrunc (n - 1) w) = bintrunc n w" |
|
410 |
by (cases n) (auto simp del: bintrunc.Suc) |
|
411 |
||
412 |
lemma sbintrunc_bintrunc' [simp]: |
|
413 |
"0 < n \<Longrightarrow> sbintrunc (n - 1) (bintrunc n w) = sbintrunc (n - 1) w" |
|
414 |
by (cases n) (auto simp del: bintrunc.Suc) |
|
415 |
||
416 |
lemma bin_sbin_eq_iff: |
|
417 |
"bintrunc (Suc n) x = bintrunc (Suc n) y <-> |
|
418 |
sbintrunc n x = sbintrunc n y" |
|
419 |
apply (rule iffI) |
|
420 |
apply (rule box_equals [OF _ sbintrunc_bintrunc sbintrunc_bintrunc]) |
|
421 |
apply simp |
|
422 |
apply (rule box_equals [OF _ bintrunc_sbintrunc bintrunc_sbintrunc]) |
|
423 |
apply simp |
|
424 |
done |
|
425 |
||
426 |
lemma bin_sbin_eq_iff': |
|
427 |
"0 < n \<Longrightarrow> bintrunc n x = bintrunc n y <-> |
|
428 |
sbintrunc (n - 1) x = sbintrunc (n - 1) y" |
|
429 |
by (cases n) (simp_all add: bin_sbin_eq_iff del: bintrunc.Suc) |
|
430 |
||
431 |
lemmas bintrunc_sbintruncS0 [simp] = bintrunc_sbintrunc' [unfolded One_nat_def] |
|
432 |
lemmas sbintrunc_bintruncS0 [simp] = sbintrunc_bintrunc' [unfolded One_nat_def] |
|
433 |
||
434 |
lemmas bintrunc_bintrunc_l' = le_add1 [THEN bintrunc_bintrunc_l] |
|
435 |
lemmas sbintrunc_sbintrunc_l' = le_add1 [THEN sbintrunc_sbintrunc_l] |
|
436 |
||
437 |
(* although bintrunc_minus_simps, if added to default simpset, |
|
438 |
tends to get applied where it's not wanted in developing the theories, |
|
439 |
we get a version for when the word length is given literally *) |
|
440 |
||
441 |
lemmas nat_non0_gr = |
|
442 |
trans [OF iszero_def [THEN Not_eq_iff [THEN iffD2]] neq0_conv, standard] |
|
443 |
||
444 |
lemmas bintrunc_pred_simps [simp] = |
|
445 |
bintrunc_minus_simps [of "number_of bin", simplified nobm1, standard] |
|
446 |
||
447 |
lemmas sbintrunc_pred_simps [simp] = |
|
448 |
sbintrunc_minus_simps [of "number_of bin", simplified nobm1, standard] |
|
449 |
||
450 |
lemma no_bintr_alt: |
|
451 |
"number_of (bintrunc n w) = w mod 2 ^ n" |
|
452 |
by (simp add: number_of_eq bintrunc_mod2p) |
|
453 |
||
454 |
lemma no_bintr_alt1: "bintrunc n = (%w. w mod 2 ^ n :: int)" |
|
455 |
by (rule ext) (rule bintrunc_mod2p) |
|
456 |
||
457 |
lemma range_bintrunc: "range (bintrunc n) = {i. 0 <= i & i < 2 ^ n}" |
|
458 |
apply (unfold no_bintr_alt1) |
|
459 |
apply (auto simp add: image_iff) |
|
460 |
apply (rule exI) |
|
461 |
apply (auto intro: int_mod_lem [THEN iffD1, symmetric]) |
|
462 |
done |
|
463 |
||
464 |
lemma no_bintr: |
|
465 |
"number_of (bintrunc n w) = (number_of w mod 2 ^ n :: int)" |
|
466 |
by (simp add : bintrunc_mod2p number_of_eq) |
|
467 |
||
468 |
lemma no_sbintr_alt2: |
|
469 |
"sbintrunc n = (%w. (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n :: int)" |
|
470 |
by (rule ext) (simp add : sbintrunc_mod2p) |
|
471 |
||
472 |
lemma no_sbintr: |
|
473 |
"number_of (sbintrunc n w) = |
|
474 |
((number_of w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n :: int)" |
|
475 |
by (simp add : no_sbintr_alt2 number_of_eq) |
|
476 |
||
477 |
lemma range_sbintrunc: |
|
478 |
"range (sbintrunc n) = {i. - (2 ^ n) <= i & i < 2 ^ n}" |
|
479 |
apply (unfold no_sbintr_alt2) |
|
480 |
apply (auto simp add: image_iff eq_diff_eq) |
|
481 |
apply (rule exI) |
|
482 |
apply (auto intro: int_mod_lem [THEN iffD1, symmetric]) |
|
483 |
done |
|
484 |
||
485 |
lemmas sb_inc_lem = int_mod_ge' |
|
486 |
[where n = "2 ^ (Suc ?k)" and b = "?a + 2 ^ ?k", |
|
487 |
simplified zless2p, OF _ TrueI] |
|
488 |
||
489 |
lemmas sb_inc_lem' = |
|
490 |
iffD1 [OF less_diff_eq, THEN sb_inc_lem, simplified OrderedGroup.diff_0] |
|
491 |
||
492 |
lemma sbintrunc_inc: |
|
493 |
"x < - (2 ^ n) ==> x + 2 ^ (Suc n) <= sbintrunc n x" |
|
494 |
unfolding no_sbintr_alt2 by (drule sb_inc_lem') simp |
|
495 |
||
496 |
lemmas sb_dec_lem = int_mod_le' |
|
497 |
[where n = "2 ^ (Suc ?k)" and b = "?a + 2 ^ ?k", |
|
498 |
simplified zless2p, OF _ TrueI, simplified] |
|
499 |
||
500 |
lemmas sb_dec_lem' = iffD1 [OF diff_le_eq', THEN sb_dec_lem, simplified] |
|
501 |
||
502 |
lemma sbintrunc_dec: |
|
503 |
"x >= (2 ^ n) ==> x - 2 ^ (Suc n) >= sbintrunc n x" |
|
504 |
unfolding no_sbintr_alt2 by (drule sb_dec_lem') simp |
|
505 |
||
506 |
lemmas zmod_uminus' = zmod_uminus [where b="?c"] |
|
507 |
lemmas zpower_zmod' = zpower_zmod [where m="?c" and y="?k"] |
|
508 |
||
509 |
lemmas brdmod1s' [symmetric] = |
|
510 |
zmod_zadd_left_eq zmod_zadd_right_eq |
|
511 |
zmod_zsub_left_eq zmod_zsub_right_eq |
|
512 |
zmod_zmult1_eq zmod_zmult1_eq_rev |
|
513 |
||
514 |
lemmas brdmods' [symmetric] = |
|
515 |
zpower_zmod' [symmetric] |
|
516 |
trans [OF zmod_zadd_left_eq zmod_zadd_right_eq] |
|
517 |
trans [OF zmod_zsub_left_eq zmod_zsub_right_eq] |
|
518 |
trans [OF zmod_zmult1_eq zmod_zmult1_eq_rev] |
|
519 |
zmod_uminus' [symmetric] |
|
520 |
zmod_zadd_left_eq [where b = "1"] |
|
521 |
zmod_zsub_left_eq [where b = "1"] |
|
522 |
||
523 |
lemmas bintr_arith1s = |
|
524 |
brdmod1s' [where c="2^?n", folded pred_def succ_def bintrunc_mod2p] |
|
525 |
lemmas bintr_ariths = |
|
526 |
brdmods' [where c="2^?n", folded pred_def succ_def bintrunc_mod2p] |
|
527 |
||
24364 | 528 |
lemmas m2pths [OF zless2p, standard] = pos_mod_sign pos_mod_bound |
529 |
||
24333 | 530 |
lemma bintr_ge0: "(0 :: int) <= number_of (bintrunc n w)" |
531 |
by (simp add : no_bintr m2pths) |
|
532 |
||
533 |
lemma bintr_lt2p: "number_of (bintrunc n w) < (2 ^ n :: int)" |
|
534 |
by (simp add : no_bintr m2pths) |
|
535 |
||
536 |
lemma bintr_Min: |
|
537 |
"number_of (bintrunc n Numeral.Min) = (2 ^ n :: int) - 1" |
|
538 |
by (simp add : no_bintr m1mod2k) |
|
539 |
||
540 |
lemma sbintr_ge: "(- (2 ^ n) :: int) <= number_of (sbintrunc n w)" |
|
541 |
by (simp add : no_sbintr m2pths) |
|
542 |
||
543 |
lemma sbintr_lt: "number_of (sbintrunc n w) < (2 ^ n :: int)" |
|
544 |
by (simp add : no_sbintr m2pths) |
|
545 |
||
546 |
lemma bintrunc_Suc: |
|
547 |
"bintrunc (Suc n) bin = bintrunc n (bin_rest bin) BIT bin_last bin" |
|
548 |
by (case_tac bin rule: bin_exhaust) auto |
|
549 |
||
550 |
lemma sign_Pls_ge_0: |
|
551 |
"(bin_sign bin = Numeral.Pls) = (number_of bin >= (0 :: int))" |
|
552 |
by (induct bin rule: bin_induct) auto |
|
553 |
||
554 |
lemma sign_Min_lt_0: |
|
555 |
"(bin_sign bin = Numeral.Min) = (number_of bin < (0 :: int))" |
|
556 |
by (induct bin rule: bin_induct) auto |
|
557 |
||
558 |
lemmas sign_Min_neg = trans [OF sign_Min_lt_0 neg_def [symmetric]] |
|
559 |
||
560 |
lemma bin_rest_trunc: |
|
561 |
"!!bin. (bin_rest (bintrunc n bin)) = bintrunc (n - 1) (bin_rest bin)" |
|
562 |
by (induct n) auto |
|
563 |
||
564 |
lemma bin_rest_power_trunc [rule_format] : |
|
565 |
"(bin_rest ^ k) (bintrunc n bin) = |
|
566 |
bintrunc (n - k) ((bin_rest ^ k) bin)" |
|
567 |
by (induct k) (auto simp: bin_rest_trunc) |
|
568 |
||
569 |
lemma bin_rest_trunc_i: |
|
570 |
"bintrunc n (bin_rest bin) = bin_rest (bintrunc (Suc n) bin)" |
|
571 |
by auto |
|
572 |
||
573 |
lemma bin_rest_strunc: |
|
574 |
"!!bin. bin_rest (sbintrunc (Suc n) bin) = sbintrunc n (bin_rest bin)" |
|
575 |
by (induct n) auto |
|
576 |
||
577 |
lemma bintrunc_rest [simp]: |
|
578 |
"!!bin. bintrunc n (bin_rest (bintrunc n bin)) = bin_rest (bintrunc n bin)" |
|
579 |
apply (induct n, simp) |
|
580 |
apply (case_tac bin rule: bin_exhaust) |
|
581 |
apply (auto simp: bintrunc_bintrunc_l) |
|
582 |
done |
|
583 |
||
584 |
lemma sbintrunc_rest [simp]: |
|
585 |
"!!bin. sbintrunc n (bin_rest (sbintrunc n bin)) = bin_rest (sbintrunc n bin)" |
|
586 |
apply (induct n, simp) |
|
587 |
apply (case_tac bin rule: bin_exhaust) |
|
588 |
apply (auto simp: bintrunc_bintrunc_l split: bit.splits) |
|
589 |
done |
|
590 |
||
591 |
lemma bintrunc_rest': |
|
592 |
"bintrunc n o bin_rest o bintrunc n = bin_rest o bintrunc n" |
|
593 |
by (rule ext) auto |
|
594 |
||
595 |
lemma sbintrunc_rest' : |
|
596 |
"sbintrunc n o bin_rest o sbintrunc n = bin_rest o sbintrunc n" |
|
597 |
by (rule ext) auto |
|
598 |
||
599 |
lemma rco_lem: |
|
600 |
"f o g o f = g o f ==> f o (g o f) ^ n = g ^ n o f" |
|
601 |
apply (rule ext) |
|
602 |
apply (induct_tac n) |
|
603 |
apply (simp_all (no_asm)) |
|
604 |
apply (drule fun_cong) |
|
605 |
apply (unfold o_def) |
|
606 |
apply (erule trans) |
|
607 |
apply simp |
|
608 |
done |
|
609 |
||
610 |
lemma rco_alt: "(f o g) ^ n o f = f o (g o f) ^ n" |
|
611 |
apply (rule ext) |
|
612 |
apply (induct n) |
|
613 |
apply (simp_all add: o_def) |
|
614 |
done |
|
615 |
||
616 |
lemmas rco_bintr = bintrunc_rest' |
|
617 |
[THEN rco_lem [THEN fun_cong], unfolded o_def] |
|
618 |
lemmas rco_sbintr = sbintrunc_rest' |
|
619 |
[THEN rco_lem [THEN fun_cong], unfolded o_def] |
|
620 |
||
24364 | 621 |
subsection {* Splitting and concatenation *} |
622 |
||
623 |
consts |
|
624 |
bin_split :: "nat => int => int * int" |
|
625 |
primrec |
|
626 |
Z : "bin_split 0 w = (w, Numeral.Pls)" |
|
627 |
Suc : "bin_split (Suc n) w = (let (w1, w2) = bin_split n (bin_rest w) |
|
628 |
in (w1, w2 BIT bin_last w))" |
|
629 |
||
630 |
consts |
|
631 |
bin_cat :: "int => nat => int => int" |
|
632 |
primrec |
|
633 |
Z : "bin_cat w 0 v = w" |
|
634 |
Suc : "bin_cat w (Suc n) v = bin_cat w n (bin_rest v) BIT bin_last v" |
|
635 |
||
636 |
subsection {* Miscellaneous lemmas *} |
|
637 |
||
638 |
lemmas funpow_minus_simp = |
|
639 |
trans [OF gen_minus [where f = "power f"] funpow_Suc, standard] |
|
640 |
||
641 |
lemmas funpow_pred_simp [simp] = |
|
642 |
funpow_minus_simp [of "number_of bin", simplified nobm1, standard] |
|
643 |
||
644 |
lemmas replicate_minus_simp = |
|
645 |
trans [OF gen_minus [where f = "%n. replicate n x"] replicate.replicate_Suc, |
|
646 |
standard] |
|
647 |
||
648 |
lemmas replicate_pred_simp [simp] = |
|
649 |
replicate_minus_simp [of "number_of bin", simplified nobm1, standard] |
|
650 |
||
651 |
lemmas power_Suc_no [simp] = power_Suc [of "number_of ?a"] |
|
652 |
||
653 |
lemmas power_minus_simp = |
|
654 |
trans [OF gen_minus [where f = "power f"] power_Suc, standard] |
|
655 |
||
656 |
lemmas power_pred_simp = |
|
657 |
power_minus_simp [of "number_of bin", simplified nobm1, standard] |
|
658 |
lemmas power_pred_simp_no [simp] = power_pred_simp [where f= "number_of ?f"] |
|
659 |
||
660 |
lemma list_exhaust_size_gt0: |
|
661 |
assumes y: "\<And>a list. y = a # list \<Longrightarrow> P" |
|
662 |
shows "0 < length y \<Longrightarrow> P" |
|
663 |
apply (cases y, simp) |
|
664 |
apply (rule y) |
|
665 |
apply fastsimp |
|
666 |
done |
|
667 |
||
668 |
lemma list_exhaust_size_eq0: |
|
669 |
assumes y: "y = [] \<Longrightarrow> P" |
|
670 |
shows "length y = 0 \<Longrightarrow> P" |
|
671 |
apply (cases y) |
|
672 |
apply (rule y, simp) |
|
673 |
apply simp |
|
674 |
done |
|
675 |
||
676 |
lemma size_Cons_lem_eq: |
|
677 |
"y = xa # list ==> size y = Suc k ==> size list = k" |
|
678 |
by auto |
|
679 |
||
680 |
lemma size_Cons_lem_eq_bin: |
|
681 |
"y = xa # list ==> size y = number_of (Numeral.succ k) ==> |
|
682 |
size list = number_of k" |
|
683 |
by (auto simp: pred_def succ_def split add : split_if_asm) |
|
684 |
||
24333 | 685 |
lemmas ls_splits = |
686 |
prod.split split_split prod.split_asm split_split_asm split_if_asm |
|
687 |
||
688 |
lemma not_B1_is_B0: "y \<noteq> bit.B1 \<Longrightarrow> y = bit.B0" |
|
689 |
by (cases y) auto |
|
690 |
||
691 |
lemma B1_ass_B0: |
|
692 |
assumes y: "y = bit.B0 \<Longrightarrow> y = bit.B1" |
|
693 |
shows "y = bit.B1" |
|
694 |
apply (rule classical) |
|
695 |
apply (drule not_B1_is_B0) |
|
696 |
apply (erule y) |
|
697 |
done |
|
698 |
||
699 |
-- "simplifications for specific word lengths" |
|
700 |
lemmas n2s_ths [THEN eq_reflection] = add_2_eq_Suc add_2_eq_Suc' |
|
701 |
||
702 |
lemmas s2n_ths = n2s_ths [symmetric] |
|
703 |
||
704 |
||
705 |
end |