author | Chengsong |
Thu, 01 Dec 2022 08:49:19 +0000 | |
changeset 629 | 96e009a446d5 |
parent 543 | b2bea5968b89 |
permissions | -rw-r--r-- |
365 | 1 |
|
393 | 2 |
theory SizeBound3 |
492 | 3 |
imports "ClosedFormsBounds" |
365 | 4 |
begin |
5 |
||
6 |
section \<open>Bit-Encodings\<close> |
|
7 |
||
8 |
datatype bit = Z | S |
|
9 |
||
10 |
fun code :: "val \<Rightarrow> bit list" |
|
11 |
where |
|
12 |
"code Void = []" |
|
13 |
| "code (Char c) = []" |
|
14 |
| "code (Left v) = Z # (code v)" |
|
15 |
| "code (Right v) = S # (code v)" |
|
16 |
| "code (Seq v1 v2) = (code v1) @ (code v2)" |
|
17 |
| "code (Stars []) = [S]" |
|
18 |
| "code (Stars (v # vs)) = (Z # code v) @ code (Stars vs)" |
|
19 |
||
20 |
||
21 |
fun |
|
22 |
Stars_add :: "val \<Rightarrow> val \<Rightarrow> val" |
|
23 |
where |
|
24 |
"Stars_add v (Stars vs) = Stars (v # vs)" |
|
25 |
||
26 |
function |
|
27 |
decode' :: "bit list \<Rightarrow> rexp \<Rightarrow> (val * bit list)" |
|
28 |
where |
|
29 |
"decode' ds ZERO = (Void, [])" |
|
30 |
| "decode' ds ONE = (Void, ds)" |
|
31 |
| "decode' ds (CH d) = (Char d, ds)" |
|
32 |
| "decode' [] (ALT r1 r2) = (Void, [])" |
|
33 |
| "decode' (Z # ds) (ALT r1 r2) = (let (v, ds') = decode' ds r1 in (Left v, ds'))" |
|
34 |
| "decode' (S # ds) (ALT r1 r2) = (let (v, ds') = decode' ds r2 in (Right v, ds'))" |
|
35 |
| "decode' ds (SEQ r1 r2) = (let (v1, ds') = decode' ds r1 in |
|
36 |
let (v2, ds'') = decode' ds' r2 in (Seq v1 v2, ds''))" |
|
37 |
| "decode' [] (STAR r) = (Void, [])" |
|
38 |
| "decode' (S # ds) (STAR r) = (Stars [], ds)" |
|
39 |
| "decode' (Z # ds) (STAR r) = (let (v, ds') = decode' ds r in |
|
40 |
let (vs, ds'') = decode' ds' (STAR r) |
|
41 |
in (Stars_add v vs, ds''))" |
|
42 |
by pat_completeness auto |
|
43 |
||
44 |
lemma decode'_smaller: |
|
45 |
assumes "decode'_dom (ds, r)" |
|
46 |
shows "length (snd (decode' ds r)) \<le> length ds" |
|
47 |
using assms |
|
48 |
apply(induct ds r) |
|
49 |
apply(auto simp add: decode'.psimps split: prod.split) |
|
50 |
using dual_order.trans apply blast |
|
51 |
by (meson dual_order.trans le_SucI) |
|
52 |
||
53 |
termination "decode'" |
|
54 |
apply(relation "inv_image (measure(%cs. size cs) <*lex*> measure(%s. size s)) (%(ds,r). (r,ds))") |
|
55 |
apply(auto dest!: decode'_smaller) |
|
56 |
by (metis less_Suc_eq_le snd_conv) |
|
57 |
||
58 |
definition |
|
59 |
decode :: "bit list \<Rightarrow> rexp \<Rightarrow> val option" |
|
60 |
where |
|
61 |
"decode ds r \<equiv> (let (v, ds') = decode' ds r |
|
62 |
in (if ds' = [] then Some v else None))" |
|
63 |
||
64 |
lemma decode'_code_Stars: |
|
65 |
assumes "\<forall>v\<in>set vs. \<Turnstile> v : r \<and> (\<forall>x. decode' (code v @ x) r = (v, x)) \<and> flat v \<noteq> []" |
|
66 |
shows "decode' (code (Stars vs) @ ds) (STAR r) = (Stars vs, ds)" |
|
67 |
using assms |
|
68 |
apply(induct vs) |
|
69 |
apply(auto) |
|
70 |
done |
|
71 |
||
72 |
lemma decode'_code: |
|
73 |
assumes "\<Turnstile> v : r" |
|
74 |
shows "decode' ((code v) @ ds) r = (v, ds)" |
|
75 |
using assms |
|
76 |
apply(induct v r arbitrary: ds) |
|
77 |
apply(auto) |
|
78 |
using decode'_code_Stars by blast |
|
79 |
||
80 |
lemma decode_code: |
|
81 |
assumes "\<Turnstile> v : r" |
|
82 |
shows "decode (code v) r = Some v" |
|
83 |
using assms unfolding decode_def |
|
84 |
by (smt append_Nil2 decode'_code old.prod.case) |
|
85 |
||
86 |
||
87 |
section {* Annotated Regular Expressions *} |
|
88 |
||
89 |
datatype arexp = |
|
90 |
AZERO |
|
91 |
| AONE "bit list" |
|
92 |
| ACHAR "bit list" char |
|
93 |
| ASEQ "bit list" arexp arexp |
|
94 |
| AALTs "bit list" "arexp list" |
|
95 |
| ASTAR "bit list" arexp |
|
96 |
||
97 |
abbreviation |
|
98 |
"AALT bs r1 r2 \<equiv> AALTs bs [r1, r2]" |
|
99 |
||
100 |
fun asize :: "arexp \<Rightarrow> nat" where |
|
101 |
"asize AZERO = 1" |
|
102 |
| "asize (AONE cs) = 1" |
|
103 |
| "asize (ACHAR cs c) = 1" |
|
104 |
| "asize (AALTs cs rs) = Suc (sum_list (map asize rs))" |
|
105 |
| "asize (ASEQ cs r1 r2) = Suc (asize r1 + asize r2)" |
|
106 |
| "asize (ASTAR cs r) = Suc (asize r)" |
|
107 |
||
108 |
fun |
|
109 |
erase :: "arexp \<Rightarrow> rexp" |
|
110 |
where |
|
111 |
"erase AZERO = ZERO" |
|
112 |
| "erase (AONE _) = ONE" |
|
113 |
| "erase (ACHAR _ c) = CH c" |
|
114 |
| "erase (AALTs _ []) = ZERO" |
|
115 |
| "erase (AALTs _ [r]) = (erase r)" |
|
116 |
| "erase (AALTs bs (r#rs)) = ALT (erase r) (erase (AALTs bs rs))" |
|
117 |
| "erase (ASEQ _ r1 r2) = SEQ (erase r1) (erase r2)" |
|
118 |
| "erase (ASTAR _ r) = STAR (erase r)" |
|
119 |
||
493 | 120 |
fun rerase :: "arexp \<Rightarrow> rrexp" |
121 |
where |
|
122 |
"rerase AZERO = RZERO" |
|
123 |
| "rerase (AONE _) = RONE" |
|
124 |
| "rerase (ACHAR _ c) = RCHAR c" |
|
125 |
| "rerase (AALTs bs rs) = RALTS (map rerase rs)" |
|
126 |
| "rerase (ASEQ _ r1 r2) = RSEQ (rerase r1) (rerase r2)" |
|
127 |
| "rerase (ASTAR _ r) = RSTAR (rerase r)" |
|
128 |
||
129 |
||
365 | 130 |
|
131 |
fun fuse :: "bit list \<Rightarrow> arexp \<Rightarrow> arexp" where |
|
132 |
"fuse bs AZERO = AZERO" |
|
133 |
| "fuse bs (AONE cs) = AONE (bs @ cs)" |
|
134 |
| "fuse bs (ACHAR cs c) = ACHAR (bs @ cs) c" |
|
135 |
| "fuse bs (AALTs cs rs) = AALTs (bs @ cs) rs" |
|
136 |
| "fuse bs (ASEQ cs r1 r2) = ASEQ (bs @ cs) r1 r2" |
|
137 |
| "fuse bs (ASTAR cs r) = ASTAR (bs @ cs) r" |
|
138 |
||
139 |
lemma fuse_append: |
|
140 |
shows "fuse (bs1 @ bs2) r = fuse bs1 (fuse bs2 r)" |
|
141 |
apply(induct r) |
|
142 |
apply(auto) |
|
143 |
done |
|
144 |
||
385 | 145 |
lemma fuse_Nil: |
146 |
shows "fuse [] r = r" |
|
147 |
by (induct r)(simp_all) |
|
148 |
||
393 | 149 |
(* |
385 | 150 |
lemma map_fuse_Nil: |
151 |
shows "map (fuse []) rs = rs" |
|
152 |
by (induct rs)(simp_all add: fuse_Nil) |
|
393 | 153 |
*) |
365 | 154 |
|
155 |
fun intern :: "rexp \<Rightarrow> arexp" where |
|
156 |
"intern ZERO = AZERO" |
|
157 |
| "intern ONE = AONE []" |
|
158 |
| "intern (CH c) = ACHAR [] c" |
|
159 |
| "intern (ALT r1 r2) = AALT [] (fuse [Z] (intern r1)) |
|
160 |
(fuse [S] (intern r2))" |
|
161 |
| "intern (SEQ r1 r2) = ASEQ [] (intern r1) (intern r2)" |
|
162 |
| "intern (STAR r) = ASTAR [] (intern r)" |
|
163 |
||
164 |
||
165 |
fun retrieve :: "arexp \<Rightarrow> val \<Rightarrow> bit list" where |
|
166 |
"retrieve (AONE bs) Void = bs" |
|
167 |
| "retrieve (ACHAR bs c) (Char d) = bs" |
|
168 |
| "retrieve (AALTs bs [r]) v = bs @ retrieve r v" |
|
169 |
| "retrieve (AALTs bs (r#rs)) (Left v) = bs @ retrieve r v" |
|
170 |
| "retrieve (AALTs bs (r#rs)) (Right v) = bs @ retrieve (AALTs [] rs) v" |
|
171 |
| "retrieve (ASEQ bs r1 r2) (Seq v1 v2) = bs @ retrieve r1 v1 @ retrieve r2 v2" |
|
172 |
| "retrieve (ASTAR bs r) (Stars []) = bs @ [S]" |
|
173 |
| "retrieve (ASTAR bs r) (Stars (v#vs)) = |
|
174 |
bs @ [Z] @ retrieve r v @ retrieve (ASTAR [] r) (Stars vs)" |
|
175 |
||
176 |
||
177 |
||
178 |
fun |
|
179 |
bnullable :: "arexp \<Rightarrow> bool" |
|
180 |
where |
|
181 |
"bnullable (AZERO) = False" |
|
182 |
| "bnullable (AONE bs) = True" |
|
183 |
| "bnullable (ACHAR bs c) = False" |
|
184 |
| "bnullable (AALTs bs rs) = (\<exists>r \<in> set rs. bnullable r)" |
|
185 |
| "bnullable (ASEQ bs r1 r2) = (bnullable r1 \<and> bnullable r2)" |
|
186 |
| "bnullable (ASTAR bs r) = True" |
|
187 |
||
393 | 188 |
abbreviation |
189 |
bnullables :: "arexp list \<Rightarrow> bool" |
|
190 |
where |
|
191 |
"bnullables rs \<equiv> (\<exists>r \<in> set rs. bnullable r)" |
|
192 |
||
365 | 193 |
fun |
393 | 194 |
bmkeps :: "arexp \<Rightarrow> bit list" and |
195 |
bmkepss :: "arexp list \<Rightarrow> bit list" |
|
365 | 196 |
where |
197 |
"bmkeps(AONE bs) = bs" |
|
198 |
| "bmkeps(ASEQ bs r1 r2) = bs @ (bmkeps r1) @ (bmkeps r2)" |
|
393 | 199 |
| "bmkeps(AALTs bs rs) = bs @ (bmkepss rs)" |
365 | 200 |
| "bmkeps(ASTAR bs r) = bs @ [S]" |
393 | 201 |
| "bmkepss [] = []" |
202 |
| "bmkepss (r # rs) = (if bnullable(r) then (bmkeps r) else (bmkepss rs))" |
|
203 |
||
204 |
lemma bmkepss1: |
|
205 |
assumes "\<not> bnullables rs1" |
|
206 |
shows "bmkepss (rs1 @ rs2) = bmkepss rs2" |
|
207 |
using assms |
|
208 |
by (induct rs1) (auto) |
|
209 |
||
210 |
lemma bmkepss2: |
|
211 |
assumes "bnullables rs1" |
|
212 |
shows "bmkepss (rs1 @ rs2) = bmkepss rs1" |
|
213 |
using assms |
|
214 |
by (induct rs1) (auto) |
|
365 | 215 |
|
216 |
||
217 |
fun |
|
218 |
bder :: "char \<Rightarrow> arexp \<Rightarrow> arexp" |
|
219 |
where |
|
220 |
"bder c (AZERO) = AZERO" |
|
221 |
| "bder c (AONE bs) = AZERO" |
|
222 |
| "bder c (ACHAR bs d) = (if c = d then AONE bs else AZERO)" |
|
223 |
| "bder c (AALTs bs rs) = AALTs bs (map (bder c) rs)" |
|
224 |
| "bder c (ASEQ bs r1 r2) = |
|
225 |
(if bnullable r1 |
|
226 |
then AALT bs (ASEQ [] (bder c r1) r2) (fuse (bmkeps r1) (bder c r2)) |
|
227 |
else ASEQ bs (bder c r1) r2)" |
|
228 |
| "bder c (ASTAR bs r) = ASEQ bs (fuse [Z] (bder c r)) (ASTAR [] r)" |
|
229 |
||
230 |
||
231 |
fun |
|
232 |
bders :: "arexp \<Rightarrow> string \<Rightarrow> arexp" |
|
233 |
where |
|
234 |
"bders r [] = r" |
|
235 |
| "bders r (c#s) = bders (bder c r) s" |
|
236 |
||
237 |
lemma bders_append: |
|
238 |
"bders r (s1 @ s2) = bders (bders r s1) s2" |
|
239 |
apply(induct s1 arbitrary: r s2) |
|
240 |
apply(simp_all) |
|
241 |
done |
|
242 |
||
243 |
lemma bnullable_correctness: |
|
244 |
shows "nullable (erase r) = bnullable r" |
|
245 |
apply(induct r rule: erase.induct) |
|
246 |
apply(simp_all) |
|
247 |
done |
|
248 |
||
493 | 249 |
lemma rbnullable_correctness: |
250 |
shows "rnullable (rerase r) = bnullable r" |
|
251 |
apply(induct r rule: rerase.induct) |
|
252 |
apply simp+ |
|
253 |
done |
|
254 |
||
255 |
||
365 | 256 |
lemma erase_fuse: |
257 |
shows "erase (fuse bs r) = erase r" |
|
258 |
apply(induct r rule: erase.induct) |
|
259 |
apply(simp_all) |
|
260 |
done |
|
261 |
||
493 | 262 |
lemma rerase_fuse: |
263 |
shows "rerase (fuse bs r) = rerase r" |
|
264 |
apply(induct r) |
|
265 |
apply simp+ |
|
266 |
done |
|
267 |
||
268 |
||
269 |
||
270 |
||
365 | 271 |
lemma erase_intern [simp]: |
272 |
shows "erase (intern r) = r" |
|
273 |
apply(induct r) |
|
274 |
apply(simp_all add: erase_fuse) |
|
275 |
done |
|
276 |
||
277 |
lemma erase_bder [simp]: |
|
278 |
shows "erase (bder a r) = der a (erase r)" |
|
279 |
apply(induct r rule: erase.induct) |
|
280 |
apply(simp_all add: erase_fuse bnullable_correctness) |
|
281 |
done |
|
282 |
||
493 | 283 |
|
284 |
||
285 |
||
286 |
||
365 | 287 |
lemma erase_bders [simp]: |
288 |
shows "erase (bders r s) = ders s (erase r)" |
|
289 |
apply(induct s arbitrary: r ) |
|
290 |
apply(simp_all) |
|
291 |
done |
|
292 |
||
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
293 |
lemma bnullable_fuse: |
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
294 |
shows "bnullable (fuse bs r) = bnullable r" |
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
295 |
apply(induct r arbitrary: bs) |
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
296 |
apply(auto) |
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
297 |
done |
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
298 |
|
365 | 299 |
lemma retrieve_encode_STARS: |
300 |
assumes "\<forall>v\<in>set vs. \<Turnstile> v : r \<and> code v = retrieve (intern r) v" |
|
301 |
shows "code (Stars vs) = retrieve (ASTAR [] (intern r)) (Stars vs)" |
|
302 |
using assms |
|
303 |
apply(induct vs) |
|
304 |
apply(simp_all) |
|
305 |
done |
|
306 |
||
307 |
||
308 |
lemma retrieve_fuse2: |
|
309 |
assumes "\<Turnstile> v : (erase r)" |
|
310 |
shows "retrieve (fuse bs r) v = bs @ retrieve r v" |
|
311 |
using assms |
|
312 |
apply(induct r arbitrary: v bs) |
|
313 |
apply(auto elim: Prf_elims)[4] |
|
314 |
defer |
|
315 |
using retrieve_encode_STARS |
|
316 |
apply(auto elim!: Prf_elims)[1] |
|
317 |
apply(case_tac vs) |
|
318 |
apply(simp) |
|
319 |
apply(simp) |
|
320 |
(* AALTs case *) |
|
321 |
apply(simp) |
|
322 |
apply(case_tac x2a) |
|
323 |
apply(simp) |
|
324 |
apply(auto elim!: Prf_elims)[1] |
|
325 |
apply(simp) |
|
326 |
apply(case_tac list) |
|
327 |
apply(simp) |
|
328 |
apply(auto) |
|
329 |
apply(auto elim!: Prf_elims)[1] |
|
330 |
done |
|
331 |
||
332 |
lemma retrieve_fuse: |
|
333 |
assumes "\<Turnstile> v : r" |
|
334 |
shows "retrieve (fuse bs (intern r)) v = bs @ retrieve (intern r) v" |
|
335 |
using assms |
|
336 |
by (simp_all add: retrieve_fuse2) |
|
337 |
||
338 |
||
339 |
lemma retrieve_code: |
|
340 |
assumes "\<Turnstile> v : r" |
|
341 |
shows "code v = retrieve (intern r) v" |
|
342 |
using assms |
|
343 |
apply(induct v r ) |
|
344 |
apply(simp_all add: retrieve_fuse retrieve_encode_STARS) |
|
345 |
done |
|
346 |
||
393 | 347 |
(* |
365 | 348 |
lemma bnullable_Hdbmkeps_Hd: |
349 |
assumes "bnullable a" |
|
350 |
shows "bmkeps (AALTs bs (a # rs)) = bs @ (bmkeps a)" |
|
393 | 351 |
using assms by simp |
352 |
*) |
|
365 | 353 |
|
354 |
||
355 |
lemma r2: |
|
356 |
assumes "x \<in> set rs" "bnullable x" |
|
357 |
shows "bnullable (AALTs bs rs)" |
|
358 |
using assms |
|
359 |
apply(induct rs) |
|
360 |
apply(auto) |
|
361 |
done |
|
362 |
||
363 |
lemma r3: |
|
364 |
assumes "\<not> bnullable r" |
|
393 | 365 |
"bnullables rs" |
365 | 366 |
shows "retrieve (AALTs bs rs) (mkeps (erase (AALTs bs rs))) = |
367 |
retrieve (AALTs bs (r # rs)) (mkeps (erase (AALTs bs (r # rs))))" |
|
368 |
using assms |
|
369 |
apply(induct rs arbitrary: r bs) |
|
370 |
apply(auto)[1] |
|
371 |
apply(auto) |
|
372 |
using bnullable_correctness apply blast |
|
373 |
apply(auto simp add: bnullable_correctness mkeps_nullable retrieve_fuse2) |
|
374 |
apply(subst retrieve_fuse2[symmetric]) |
|
375 |
apply (smt bnullable.simps(4) bnullable_correctness erase.simps(5) erase.simps(6) insert_iff list.exhaust list.set(2) mkeps.simps(3) mkeps_nullable) |
|
376 |
apply(simp) |
|
377 |
apply(case_tac "bnullable a") |
|
378 |
apply (smt append_Nil2 bnullable.simps(4) bnullable_correctness erase.simps(5) erase.simps(6) fuse.simps(4) insert_iff list.exhaust list.set(2) mkeps.simps(3) mkeps_nullable retrieve_fuse2) |
|
379 |
apply(drule_tac x="a" in meta_spec) |
|
380 |
apply(drule_tac x="bs" in meta_spec) |
|
381 |
apply(drule meta_mp) |
|
382 |
apply(simp) |
|
383 |
apply(drule meta_mp) |
|
384 |
apply(auto) |
|
385 |
apply(subst retrieve_fuse2[symmetric]) |
|
386 |
apply(case_tac rs) |
|
387 |
apply(simp) |
|
388 |
apply(auto)[1] |
|
389 |
apply (simp add: bnullable_correctness) |
|
393 | 390 |
|
365 | 391 |
apply (metis append_Nil2 bnullable_correctness erase_fuse fuse.simps(4) list.set_intros(1) mkeps.simps(3) mkeps_nullable nullable.simps(4) r2) |
392 |
apply (simp add: bnullable_correctness) |
|
393 |
apply (metis append_Nil2 bnullable_correctness erase.simps(6) erase_fuse fuse.simps(4) list.set_intros(2) mkeps.simps(3) mkeps_nullable r2) |
|
394 |
apply(simp) |
|
395 |
done |
|
396 |
||
397 |
lemma t: |
|
393 | 398 |
assumes "\<forall>r \<in> set rs. bnullable r \<longrightarrow> bmkeps r = retrieve r (mkeps (erase r))" |
399 |
"bnullables rs" |
|
400 |
shows "bs @ bmkepss rs = retrieve (AALTs bs rs) (mkeps (erase (AALTs bs rs)))" |
|
401 |
using assms |
|
365 | 402 |
apply(induct rs arbitrary: bs) |
403 |
apply(auto) |
|
393 | 404 |
apply (metis (no_types, opaque_lifting) bmkepss.cases bnullable_correctness erase.simps(5) erase.simps(6) mkeps.simps(3) retrieve.simps(3) retrieve.simps(4)) |
405 |
apply (metis r3) |
|
406 |
apply (metis (no_types, lifting) bmkepss.cases bnullable_correctness empty_iff erase.simps(6) list.set(1) mkeps.simps(3) retrieve.simps(4)) |
|
407 |
apply (metis r3) |
|
408 |
done |
|
493 | 409 |
|
410 |
||
365 | 411 |
lemma bmkeps_retrieve: |
393 | 412 |
assumes "bnullable r" |
365 | 413 |
shows "bmkeps r = retrieve r (mkeps (erase r))" |
414 |
using assms |
|
415 |
apply(induct r) |
|
393 | 416 |
apply(auto) |
417 |
using t by auto |
|
365 | 418 |
|
419 |
lemma bder_retrieve: |
|
420 |
assumes "\<Turnstile> v : der c (erase r)" |
|
421 |
shows "retrieve (bder c r) v = retrieve r (injval (erase r) c v)" |
|
393 | 422 |
using assms |
365 | 423 |
apply(induct r arbitrary: v rule: erase.induct) |
424 |
apply(simp) |
|
425 |
apply(erule Prf_elims) |
|
426 |
apply(simp) |
|
427 |
apply(erule Prf_elims) |
|
428 |
apply(simp) |
|
429 |
apply(case_tac "c = ca") |
|
430 |
apply(simp) |
|
431 |
apply(erule Prf_elims) |
|
432 |
apply(simp) |
|
433 |
apply(simp) |
|
434 |
apply(erule Prf_elims) |
|
435 |
apply(simp) |
|
436 |
apply(erule Prf_elims) |
|
437 |
apply(simp) |
|
438 |
apply(simp) |
|
439 |
apply(rename_tac "r\<^sub>1" "r\<^sub>2" rs v) |
|
440 |
apply(erule Prf_elims) |
|
441 |
apply(simp) |
|
442 |
apply(simp) |
|
443 |
apply(case_tac rs) |
|
444 |
apply(simp) |
|
445 |
apply(simp) |
|
446 |
apply (smt Prf_elims(3) injval.simps(2) injval.simps(3) retrieve.simps(4) retrieve.simps(5) same_append_eq) |
|
447 |
apply(simp) |
|
448 |
apply(case_tac "nullable (erase r1)") |
|
449 |
apply(simp) |
|
450 |
apply(erule Prf_elims) |
|
451 |
apply(subgoal_tac "bnullable r1") |
|
452 |
prefer 2 |
|
453 |
using bnullable_correctness apply blast |
|
454 |
apply(simp) |
|
455 |
apply(erule Prf_elims) |
|
456 |
apply(simp) |
|
457 |
apply(subgoal_tac "bnullable r1") |
|
458 |
prefer 2 |
|
459 |
using bnullable_correctness apply blast |
|
460 |
apply(simp) |
|
461 |
apply(simp add: retrieve_fuse2) |
|
462 |
apply(simp add: bmkeps_retrieve) |
|
463 |
apply(simp) |
|
464 |
apply(erule Prf_elims) |
|
465 |
apply(simp) |
|
466 |
using bnullable_correctness apply blast |
|
467 |
apply(rename_tac bs r v) |
|
468 |
apply(simp) |
|
469 |
apply(erule Prf_elims) |
|
470 |
apply(clarify) |
|
471 |
apply(erule Prf_elims) |
|
472 |
apply(clarify) |
|
473 |
apply(subst injval.simps) |
|
474 |
apply(simp del: retrieve.simps) |
|
475 |
apply(subst retrieve.simps) |
|
476 |
apply(subst retrieve.simps) |
|
477 |
apply(simp) |
|
478 |
apply(simp add: retrieve_fuse2) |
|
479 |
done |
|
480 |
||
481 |
||
482 |
||
483 |
lemma MAIN_decode: |
|
484 |
assumes "\<Turnstile> v : ders s r" |
|
485 |
shows "Some (flex r id s v) = decode (retrieve (bders (intern r) s) v) r" |
|
486 |
using assms |
|
487 |
proof (induct s arbitrary: v rule: rev_induct) |
|
488 |
case Nil |
|
489 |
have "\<Turnstile> v : ders [] r" by fact |
|
490 |
then have "\<Turnstile> v : r" by simp |
|
491 |
then have "Some v = decode (retrieve (intern r) v) r" |
|
492 |
using decode_code retrieve_code by auto |
|
493 |
then show "Some (flex r id [] v) = decode (retrieve (bders (intern r) []) v) r" |
|
494 |
by simp |
|
495 |
next |
|
496 |
case (snoc c s v) |
|
497 |
have IH: "\<And>v. \<Turnstile> v : ders s r \<Longrightarrow> |
|
498 |
Some (flex r id s v) = decode (retrieve (bders (intern r) s) v) r" by fact |
|
499 |
have asm: "\<Turnstile> v : ders (s @ [c]) r" by fact |
|
500 |
then have asm2: "\<Turnstile> injval (ders s r) c v : ders s r" |
|
501 |
by (simp add: Prf_injval ders_append) |
|
502 |
have "Some (flex r id (s @ [c]) v) = Some (flex r id s (injval (ders s r) c v))" |
|
503 |
by (simp add: flex_append) |
|
504 |
also have "... = decode (retrieve (bders (intern r) s) (injval (ders s r) c v)) r" |
|
505 |
using asm2 IH by simp |
|
506 |
also have "... = decode (retrieve (bder c (bders (intern r) s)) v) r" |
|
507 |
using asm by (simp_all add: bder_retrieve ders_append) |
|
508 |
finally show "Some (flex r id (s @ [c]) v) = |
|
509 |
decode (retrieve (bders (intern r) (s @ [c])) v) r" by (simp add: bders_append) |
|
510 |
qed |
|
511 |
||
512 |
definition blexer where |
|
513 |
"blexer r s \<equiv> if bnullable (bders (intern r) s) then |
|
514 |
decode (bmkeps (bders (intern r) s)) r else None" |
|
515 |
||
516 |
lemma blexer_correctness: |
|
517 |
shows "blexer r s = lexer r s" |
|
518 |
proof - |
|
519 |
{ define bds where "bds \<equiv> bders (intern r) s" |
|
520 |
define ds where "ds \<equiv> ders s r" |
|
521 |
assume asm: "nullable ds" |
|
522 |
have era: "erase bds = ds" |
|
523 |
unfolding ds_def bds_def by simp |
|
524 |
have mke: "\<Turnstile> mkeps ds : ds" |
|
525 |
using asm by (simp add: mkeps_nullable) |
|
526 |
have "decode (bmkeps bds) r = decode (retrieve bds (mkeps ds)) r" |
|
527 |
using bmkeps_retrieve |
|
393 | 528 |
using asm era |
529 |
using bnullable_correctness by force |
|
365 | 530 |
also have "... = Some (flex r id s (mkeps ds))" |
531 |
using mke by (simp_all add: MAIN_decode ds_def bds_def) |
|
532 |
finally have "decode (bmkeps bds) r = Some (flex r id s (mkeps ds))" |
|
533 |
unfolding bds_def ds_def . |
|
534 |
} |
|
535 |
then show "blexer r s = lexer r s" |
|
536 |
unfolding blexer_def lexer_flex |
|
393 | 537 |
by (auto simp add: bnullable_correctness[symmetric]) |
365 | 538 |
qed |
539 |
||
540 |
||
541 |
fun distinctBy :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b set \<Rightarrow> 'a list" |
|
542 |
where |
|
543 |
"distinctBy [] f acc = []" |
|
544 |
| "distinctBy (x#xs) f acc = |
|
545 |
(if (f x) \<in> acc then distinctBy xs f acc |
|
546 |
else x # (distinctBy xs f ({f x} \<union> acc)))" |
|
547 |
||
393 | 548 |
|
365 | 549 |
|
550 |
fun flts :: "arexp list \<Rightarrow> arexp list" |
|
551 |
where |
|
552 |
"flts [] = []" |
|
553 |
| "flts (AZERO # rs) = flts rs" |
|
554 |
| "flts ((AALTs bs rs1) # rs) = (map (fuse bs) rs1) @ flts rs" |
|
555 |
| "flts (r1 # rs) = r1 # flts rs" |
|
556 |
||
557 |
||
558 |
||
559 |
fun bsimp_ASEQ :: "bit list \<Rightarrow> arexp \<Rightarrow> arexp \<Rightarrow> arexp" |
|
560 |
where |
|
561 |
"bsimp_ASEQ _ AZERO _ = AZERO" |
|
562 |
| "bsimp_ASEQ _ _ AZERO = AZERO" |
|
563 |
| "bsimp_ASEQ bs1 (AONE bs2) r2 = fuse (bs1 @ bs2) r2" |
|
564 |
| "bsimp_ASEQ bs1 r1 r2 = ASEQ bs1 r1 r2" |
|
565 |
||
393 | 566 |
lemma bsimp_ASEQ0[simp]: |
567 |
shows "bsimp_ASEQ bs r1 AZERO = AZERO" |
|
568 |
by (case_tac r1)(simp_all) |
|
569 |
||
570 |
lemma bsimp_ASEQ1: |
|
571 |
assumes "r1 \<noteq> AZERO" "r2 \<noteq> AZERO" "\<nexists>bs. r1 = AONE bs" |
|
572 |
shows "bsimp_ASEQ bs r1 r2 = ASEQ bs r1 r2" |
|
573 |
using assms |
|
574 |
apply(induct bs r1 r2 rule: bsimp_ASEQ.induct) |
|
575 |
apply(auto) |
|
576 |
done |
|
577 |
||
578 |
lemma bsimp_ASEQ2[simp]: |
|
579 |
shows "bsimp_ASEQ bs1 (AONE bs2) r2 = fuse (bs1 @ bs2) r2" |
|
580 |
by (case_tac r2) (simp_all) |
|
581 |
||
365 | 582 |
|
583 |
fun bsimp_AALTs :: "bit list \<Rightarrow> arexp list \<Rightarrow> arexp" |
|
584 |
where |
|
585 |
"bsimp_AALTs _ [] = AZERO" |
|
586 |
| "bsimp_AALTs bs1 [r] = fuse bs1 r" |
|
587 |
| "bsimp_AALTs bs1 rs = AALTs bs1 rs" |
|
588 |
||
589 |
||
493 | 590 |
|
591 |
||
365 | 592 |
fun bsimp :: "arexp \<Rightarrow> arexp" |
593 |
where |
|
594 |
"bsimp (ASEQ bs1 r1 r2) = bsimp_ASEQ bs1 (bsimp r1) (bsimp r2)" |
|
493 | 595 |
| "bsimp (AALTs bs1 rs) = bsimp_AALTs bs1 (distinctBy (flts (map bsimp rs)) rerase {}) " |
365 | 596 |
| "bsimp r = r" |
597 |
||
598 |
||
599 |
fun |
|
600 |
bders_simp :: "arexp \<Rightarrow> string \<Rightarrow> arexp" |
|
601 |
where |
|
602 |
"bders_simp r [] = r" |
|
603 |
| "bders_simp r (c # s) = bders_simp (bsimp (bder c r)) s" |
|
604 |
||
605 |
definition blexer_simp where |
|
606 |
"blexer_simp r s \<equiv> if bnullable (bders_simp (intern r) s) then |
|
374 | 607 |
decode (bmkeps (bders_simp (intern r) s)) r else None" |
365 | 608 |
|
393 | 609 |
|
365 | 610 |
|
611 |
lemma bders_simp_append: |
|
612 |
shows "bders_simp r (s1 @ s2) = bders_simp (bders_simp r s1) s2" |
|
613 |
apply(induct s1 arbitrary: r s2) |
|
374 | 614 |
apply(simp_all) |
365 | 615 |
done |
616 |
||
617 |
lemma L_bsimp_ASEQ: |
|
393 | 618 |
"L (erase (ASEQ bs r1 r2)) = L (erase (bsimp_ASEQ bs r1 r2))" |
365 | 619 |
apply(induct bs r1 r2 rule: bsimp_ASEQ.induct) |
620 |
apply(simp_all) |
|
621 |
by (metis erase_fuse fuse.simps(4)) |
|
622 |
||
493 | 623 |
lemma RL_bsimp_ASEQ: |
624 |
"RL (rerase (ASEQ bs r1 r2)) = RL (rerase (bsimp_ASEQ bs r1 r2))" |
|
625 |
apply(induct bs r1 r2 rule: bsimp_ASEQ.induct) |
|
626 |
apply simp+ |
|
627 |
done |
|
628 |
||
365 | 629 |
lemma L_bsimp_AALTs: |
630 |
"L (erase (AALTs bs rs)) = L (erase (bsimp_AALTs bs rs))" |
|
631 |
apply(induct bs rs rule: bsimp_AALTs.induct) |
|
632 |
apply(simp_all add: erase_fuse) |
|
633 |
done |
|
634 |
||
493 | 635 |
lemma RL_bsimp_AALTs: |
636 |
"RL (rerase (AALTs bs rs)) = RL (rerase (bsimp_AALTs bs rs))" |
|
637 |
apply(induct bs rs rule: bsimp_AALTs.induct) |
|
638 |
apply(simp_all add: rerase_fuse) |
|
639 |
done |
|
640 |
||
641 |
lemma RL_rerase_AALTs: |
|
642 |
shows "RL (rerase (AALTs bs rs)) = \<Union> (RL ` rerase ` (set rs))" |
|
643 |
apply(induct rs) |
|
644 |
apply(simp) |
|
645 |
apply(simp) |
|
646 |
done |
|
647 |
||
365 | 648 |
lemma L_erase_AALTs: |
649 |
shows "L (erase (AALTs bs rs)) = \<Union> (L ` erase ` (set rs))" |
|
650 |
apply(induct rs) |
|
651 |
apply(simp) |
|
652 |
apply(simp) |
|
653 |
apply(case_tac rs) |
|
654 |
apply(simp) |
|
655 |
apply(simp) |
|
656 |
done |
|
657 |
||
493 | 658 |
|
659 |
||
365 | 660 |
lemma L_erase_flts: |
661 |
shows "\<Union> (L ` erase ` (set (flts rs))) = \<Union> (L ` erase ` (set rs))" |
|
662 |
apply(induct rs rule: flts.induct) |
|
393 | 663 |
apply(simp_all) |
365 | 664 |
apply(auto) |
665 |
using L_erase_AALTs erase_fuse apply auto[1] |
|
666 |
by (simp add: L_erase_AALTs erase_fuse) |
|
667 |
||
493 | 668 |
|
669 |
lemma RL_erase_flts: |
|
670 |
shows "\<Union> (RL ` rerase ` (set (flts rs))) = |
|
671 |
\<Union> (RL ` rerase ` (set rs))" |
|
672 |
apply(induct rs rule: flts.induct) |
|
673 |
apply simp+ |
|
674 |
apply auto |
|
675 |
apply (smt (verit) fuse.elims rerase.simps(2) rerase.simps(3) rerase.simps(4) rerase.simps(5) rerase.simps(6)) |
|
676 |
by (smt (verit, best) fuse.elims rerase.simps(2) rerase.simps(3) rerase.simps(4) rerase.simps(5) rerase.simps(6)) |
|
677 |
||
678 |
||
679 |
||
680 |
||
365 | 681 |
lemma L_erase_dB_acc: |
393 | 682 |
shows "(\<Union> (L ` acc) \<union> (\<Union> (L ` erase ` (set (distinctBy rs erase acc))))) |
683 |
= \<Union> (L ` acc) \<union> \<Union> (L ` erase ` (set rs))" |
|
365 | 684 |
apply(induction rs arbitrary: acc) |
393 | 685 |
apply simp_all |
365 | 686 |
by (smt (z3) SUP_absorb UN_insert sup_assoc sup_commute) |
687 |
||
493 | 688 |
lemma RL_rerase_dB_acc: |
689 |
shows "(\<Union> (RL ` acc) \<union> (\<Union> (RL ` rerase ` (set (distinctBy rs rerase acc))))) |
|
690 |
= \<Union> (RL ` acc) \<union> \<Union> (RL ` rerase ` (set rs))" |
|
691 |
apply(induction rs arbitrary: acc) |
|
692 |
apply simp_all |
|
693 |
by (smt (z3) SUP_absorb UN_insert sup_assoc sup_commute) |
|
694 |
||
695 |
||
696 |
||
393 | 697 |
|
365 | 698 |
lemma L_erase_dB: |
393 | 699 |
shows "(\<Union> (L ` erase ` (set (distinctBy rs erase {})))) = \<Union> (L ` erase ` (set rs))" |
365 | 700 |
by (metis L_erase_dB_acc Un_commute Union_image_empty) |
701 |
||
493 | 702 |
lemma RL_rerase_dB: |
703 |
shows "(\<Union> (RL ` rerase ` (set (distinctBy rs rerase {})))) = \<Union> (RL ` rerase ` (set rs))" |
|
704 |
by (metis RL_rerase_dB_acc Un_commute Union_image_empty) |
|
705 |
||
365 | 706 |
|
493 | 707 |
lemma RL_bsimp_rerase: |
708 |
shows "RL (rerase r) = RL (rerase (bsimp r))" |
|
709 |
apply(induct r) |
|
710 |
apply(simp) |
|
711 |
apply(simp) |
|
712 |
apply(simp) |
|
713 |
apply(auto simp add: Sequ_def)[1] |
|
714 |
apply(subst RL_bsimp_ASEQ[symmetric]) |
|
715 |
apply(auto simp add: Sequ_def)[1] |
|
716 |
apply(subst (asm) RL_bsimp_ASEQ[symmetric]) |
|
717 |
apply(auto simp add: Sequ_def)[1] |
|
718 |
apply(simp) |
|
719 |
apply(subst RL_bsimp_AALTs[symmetric]) |
|
720 |
defer |
|
721 |
apply(simp) |
|
722 |
using RL_erase_flts RL_rerase_dB by force |
|
723 |
||
724 |
lemma rder_correctness: |
|
725 |
shows "RL (rder c r) = Der c (RL r)" |
|
726 |
by (simp add: RL_rder) |
|
727 |
||
728 |
||
729 |
||
730 |
||
731 |
lemma asize_rsize: |
|
732 |
shows "rsize (rerase r) = asize r" |
|
733 |
apply(induct r) |
|
734 |
apply simp+ |
|
735 |
||
736 |
apply (metis (mono_tags, lifting) comp_apply map_eq_conv) |
|
737 |
by simp |
|
738 |
||
739 |
lemma rb_nullability: |
|
740 |
shows " rnullable (rerase a) = bnullable a" |
|
741 |
apply(induct a) |
|
742 |
apply simp+ |
|
743 |
done |
|
744 |
||
745 |
lemma fuse_anything_doesnt_matter: |
|
746 |
shows "rerase a = rerase (fuse bs a)" |
|
747 |
by (smt (verit) fuse.elims rerase.simps(2) rerase.simps(3) rerase.simps(4) rerase.simps(5) rerase.simps(6)) |
|
748 |
||
749 |
||
750 |
||
751 |
lemma rder_bder_rerase: |
|
752 |
shows "rder c (rerase r ) = rerase (bder c r)" |
|
753 |
apply (induct r) |
|
754 |
apply simp+ |
|
755 |
apply(case_tac "bnullable r1") |
|
756 |
apply simp |
|
757 |
apply (simp add: rb_nullability rerase_fuse) |
|
758 |
using rb_nullability apply presburger |
|
759 |
apply simp |
|
760 |
apply simp |
|
761 |
using rerase_fuse by presburger |
|
762 |
||
763 |
||
764 |
||
765 |
||
766 |
lemma RL_bder_preserved: |
|
767 |
shows "RL (rerase a1) = RL (rerase a2) \<Longrightarrow> RL (rerase (bder c a1 )) = RL (rerase (bder c a2))" |
|
768 |
by (metis rder_bder_rerase rder_correctness) |
|
769 |
||
770 |
||
771 |
lemma RL_bders_simp: |
|
772 |
shows "RL (rerase (bders_simp r s)) = RL (rerase (bders r s))" |
|
365 | 773 |
apply(induct s arbitrary: r rule: rev_induct) |
374 | 774 |
apply(simp) |
493 | 775 |
apply(simp add: bders_append) |
365 | 776 |
apply(simp add: bders_simp_append) |
493 | 777 |
apply(simp add: RL_bsimp_rerase[symmetric]) |
778 |
using RL_bder_preserved by blast |
|
779 |
||
365 | 780 |
|
781 |
||
393 | 782 |
lemma bmkeps_fuse: |
365 | 783 |
assumes "bnullable r" |
784 |
shows "bmkeps (fuse bs r) = bs @ bmkeps r" |
|
393 | 785 |
by (metis assms bmkeps_retrieve bnullable_correctness erase_fuse mkeps_nullable retrieve_fuse2) |
786 |
||
787 |
lemma bmkepss_fuse: |
|
788 |
assumes "bnullables rs" |
|
789 |
shows "bmkepss (map (fuse bs) rs) = bs @ bmkepss rs" |
|
790 |
using assms |
|
791 |
apply(induct rs arbitrary: bs) |
|
792 |
apply(auto simp add: bmkeps_fuse bnullable_fuse) |
|
793 |
done |
|
365 | 794 |
|
795 |
||
796 |
lemma b4: |
|
797 |
shows "bnullable (bders_simp r s) = bnullable (bders r s)" |
|
393 | 798 |
proof - |
493 | 799 |
have "RL (rerase (bders_simp r s)) = RL (rerase (bders r s))" |
800 |
by (simp add: RL_bders_simp) |
|
393 | 801 |
then show "bnullable (bders_simp r s) = bnullable (bders r s)" |
493 | 802 |
using RL_rnullable rb_nullability by blast |
393 | 803 |
qed |
365 | 804 |
|
805 |
||
806 |
lemma bder_fuse: |
|
807 |
shows "bder c (fuse bs a) = fuse bs (bder c a)" |
|
808 |
apply(induct a arbitrary: bs c) |
|
809 |
apply(simp_all) |
|
810 |
done |
|
811 |
||
385 | 812 |
|
813 |
||
814 |
||
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
815 |
inductive |
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
816 |
rrewrite:: "arexp \<Rightarrow> arexp \<Rightarrow> bool" ("_ \<leadsto> _" [99, 99] 99) |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
817 |
and |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
818 |
srewrite:: "arexp list \<Rightarrow> arexp list \<Rightarrow> bool" (" _ s\<leadsto> _" [100, 100] 100) |
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
819 |
where |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
820 |
bs1: "ASEQ bs AZERO r2 \<leadsto> AZERO" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
821 |
| bs2: "ASEQ bs r1 AZERO \<leadsto> AZERO" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
822 |
| bs3: "ASEQ bs1 (AONE bs2) r \<leadsto> fuse (bs1@bs2) r" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
823 |
| bs4: "r1 \<leadsto> r2 \<Longrightarrow> ASEQ bs r1 r3 \<leadsto> ASEQ bs r2 r3" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
824 |
| bs5: "r3 \<leadsto> r4 \<Longrightarrow> ASEQ bs r1 r3 \<leadsto> ASEQ bs r1 r4" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
825 |
| bs6: "AALTs bs [] \<leadsto> AZERO" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
826 |
| bs7: "AALTs bs [r] \<leadsto> fuse bs r" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
827 |
| bs10: "rs1 s\<leadsto> rs2 \<Longrightarrow> AALTs bs rs1 \<leadsto> AALTs bs rs2" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
828 |
| ss1: "[] s\<leadsto> []" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
829 |
| ss2: "rs1 s\<leadsto> rs2 \<Longrightarrow> (r # rs1) s\<leadsto> (r # rs2)" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
830 |
| ss3: "r1 \<leadsto> r2 \<Longrightarrow> (r1 # rs) s\<leadsto> (r2 # rs)" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
831 |
| ss4: "(AZERO # rs) s\<leadsto> rs" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
832 |
| ss5: "(AALTs bs1 rs1 # rsb) s\<leadsto> ((map (fuse bs1) rs1) @ rsb)" |
493 | 833 |
| ss6: "rerase a1 = rerase a2 \<Longrightarrow> (rsa@[a1]@rsb@[a2]@rsc) s\<leadsto> (rsa@[a1]@rsb@rsc)" |
365 | 834 |
|
393 | 835 |
|
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
836 |
inductive |
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
837 |
rrewrites:: "arexp \<Rightarrow> arexp \<Rightarrow> bool" ("_ \<leadsto>* _" [100, 100] 100) |
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
838 |
where |
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
839 |
rs1[intro, simp]:"r \<leadsto>* r" |
365 | 840 |
| rs2[intro]: "\<lbrakk>r1 \<leadsto>* r2; r2 \<leadsto> r3\<rbrakk> \<Longrightarrow> r1 \<leadsto>* r3" |
841 |
||
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
842 |
inductive |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
843 |
srewrites:: "arexp list \<Rightarrow> arexp list \<Rightarrow> bool" ("_ s\<leadsto>* _" [100, 100] 100) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
844 |
where |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
845 |
sss1[intro, simp]:"rs s\<leadsto>* rs" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
846 |
| sss2[intro]: "\<lbrakk>rs1 s\<leadsto> rs2; rs2 s\<leadsto>* rs3\<rbrakk> \<Longrightarrow> rs1 s\<leadsto>* rs3" |
365 | 847 |
|
848 |
||
849 |
lemma r_in_rstar : "r1 \<leadsto> r2 \<Longrightarrow> r1 \<leadsto>* r2" |
|
850 |
using rrewrites.intros(1) rrewrites.intros(2) by blast |
|
851 |
||
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
852 |
lemma rrewrites_trans[trans]: |
365 | 853 |
assumes a1: "r1 \<leadsto>* r2" and a2: "r2 \<leadsto>* r3" |
854 |
shows "r1 \<leadsto>* r3" |
|
855 |
using a2 a1 |
|
856 |
apply(induct r2 r3 arbitrary: r1 rule: rrewrites.induct) |
|
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
857 |
apply(auto) |
365 | 858 |
done |
859 |
||
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
860 |
lemma srewrites_trans[trans]: |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
861 |
assumes a1: "r1 s\<leadsto>* r2" and a2: "r2 s\<leadsto>* r3" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
862 |
shows "r1 s\<leadsto>* r3" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
863 |
using a1 a2 |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
864 |
apply(induct r1 r2 arbitrary: r3 rule: srewrites.induct) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
865 |
apply(auto) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
866 |
done |
365 | 867 |
|
868 |
||
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
869 |
|
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
870 |
lemma contextrewrites0: |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
871 |
"rs1 s\<leadsto>* rs2 \<Longrightarrow> AALTs bs rs1 \<leadsto>* AALTs bs rs2" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
872 |
apply(induct rs1 rs2 rule: srewrites.inducts) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
873 |
apply simp |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
874 |
using bs10 r_in_rstar rrewrites_trans by blast |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
875 |
|
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
876 |
lemma contextrewrites1: |
393 | 877 |
"r \<leadsto>* r' \<Longrightarrow> AALTs bs (r # rs) \<leadsto>* AALTs bs (r' # rs)" |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
878 |
apply(induct r r' rule: rrewrites.induct) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
879 |
apply simp |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
880 |
using bs10 ss3 by blast |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
881 |
|
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
882 |
lemma srewrite1: |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
883 |
shows "rs1 s\<leadsto> rs2 \<Longrightarrow> (rs @ rs1) s\<leadsto> (rs @ rs2)" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
884 |
apply(induct rs) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
885 |
apply(auto) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
886 |
using ss2 by auto |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
887 |
|
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
888 |
lemma srewrites1: |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
889 |
shows "rs1 s\<leadsto>* rs2 \<Longrightarrow> (rs @ rs1) s\<leadsto>* (rs @ rs2)" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
890 |
apply(induct rs1 rs2 rule: srewrites.induct) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
891 |
apply(auto) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
892 |
using srewrite1 by blast |
365 | 893 |
|
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
894 |
lemma srewrite2: |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
895 |
shows "r1 \<leadsto> r2 \<Longrightarrow> True" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
896 |
and "rs1 s\<leadsto> rs2 \<Longrightarrow> (rs1 @ rs) s\<leadsto>* (rs2 @ rs)" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
897 |
apply(induct rule: rrewrite_srewrite.inducts) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
898 |
apply(auto) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
899 |
apply (metis append_Cons append_Nil srewrites1) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
900 |
apply(meson srewrites.simps ss3) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
901 |
apply (meson srewrites.simps ss4) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
902 |
apply (meson srewrites.simps ss5) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
903 |
by (metis append_Cons append_Nil srewrites.simps ss6) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
904 |
|
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
905 |
|
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
906 |
lemma srewrites3: |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
907 |
shows "rs1 s\<leadsto>* rs2 \<Longrightarrow> (rs1 @ rs) s\<leadsto>* (rs2 @ rs)" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
908 |
apply(induct rs1 rs2 arbitrary: rs rule: srewrites.induct) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
909 |
apply(auto) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
910 |
by (meson srewrite2(2) srewrites_trans) |
365 | 911 |
|
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
912 |
(* |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
913 |
lemma srewrites4: |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
914 |
assumes "rs3 s\<leadsto>* rs4" "rs1 s\<leadsto>* rs2" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
915 |
shows "(rs1 @ rs3) s\<leadsto>* (rs2 @ rs4)" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
916 |
using assms |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
917 |
apply(induct rs3 rs4 arbitrary: rs1 rs2 rule: srewrites.induct) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
918 |
apply (simp add: srewrites3) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
919 |
using srewrite1 by blast |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
920 |
*) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
921 |
|
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
922 |
lemma srewrites6: |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
923 |
assumes "r1 \<leadsto>* r2" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
924 |
shows "[r1] s\<leadsto>* [r2]" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
925 |
using assms |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
926 |
apply(induct r1 r2 rule: rrewrites.induct) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
927 |
apply(auto) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
928 |
by (meson srewrites.simps srewrites_trans ss3) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
929 |
|
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
930 |
lemma srewrites7: |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
931 |
assumes "rs3 s\<leadsto>* rs4" "r1 \<leadsto>* r2" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
932 |
shows "(r1 # rs3) s\<leadsto>* (r2 # rs4)" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
933 |
using assms |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
934 |
by (smt (verit, best) append_Cons append_Nil srewrites1 srewrites3 srewrites6 srewrites_trans) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
935 |
|
393 | 936 |
lemma ss6_stronger_aux: |
493 | 937 |
shows "(rs1 @ rs2) s\<leadsto>* (rs1 @ distinctBy rs2 rerase (set (map rerase rs1)))" |
393 | 938 |
apply(induct rs2 arbitrary: rs1) |
939 |
apply(auto) |
|
940 |
apply (smt (verit, best) append.assoc append.right_neutral append_Cons append_Nil split_list srewrite2(2) srewrites_trans ss6) |
|
941 |
apply(drule_tac x="rs1 @ [a]" in meta_spec) |
|
942 |
apply(simp) |
|
943 |
done |
|
944 |
||
543 | 945 |
|
946 |
||
393 | 947 |
lemma ss6_stronger: |
493 | 948 |
shows "rs1 s\<leadsto>* distinctBy rs1 rerase {}" |
949 |
by (metis append_Nil list.set(1) list.simps(8) ss6_stronger_aux) |
|
393 | 950 |
|
951 |
||
952 |
lemma rewrite_preserves_fuse: |
|
953 |
shows "r2 \<leadsto> r3 \<Longrightarrow> fuse bs r2 \<leadsto> fuse bs r3" |
|
954 |
and "rs2 s\<leadsto> rs3 \<Longrightarrow> map (fuse bs) rs2 s\<leadsto>* map (fuse bs) rs3" |
|
955 |
proof(induct rule: rrewrite_srewrite.inducts) |
|
956 |
case (bs3 bs1 bs2 r) |
|
957 |
then show ?case |
|
958 |
by (metis fuse.simps(5) fuse_append rrewrite_srewrite.bs3) |
|
959 |
next |
|
960 |
case (bs7 bs r) |
|
961 |
then show ?case |
|
962 |
by (metis fuse.simps(4) fuse_append rrewrite_srewrite.bs7) |
|
963 |
next |
|
964 |
case (ss2 rs1 rs2 r) |
|
965 |
then show ?case |
|
966 |
using srewrites7 by force |
|
967 |
next |
|
968 |
case (ss3 r1 r2 rs) |
|
969 |
then show ?case by (simp add: r_in_rstar srewrites7) |
|
970 |
next |
|
971 |
case (ss5 bs1 rs1 rsb) |
|
972 |
then show ?case |
|
973 |
apply(simp) |
|
974 |
by (metis (mono_tags, lifting) comp_def fuse_append map_eq_conv rrewrite_srewrite.ss5 srewrites.simps) |
|
975 |
next |
|
976 |
case (ss6 a1 a2 rsa rsb rsc) |
|
977 |
then show ?case |
|
978 |
apply(simp only: map_append) |
|
493 | 979 |
by (smt (verit, ccfv_threshold) rerase_fuse list.simps(8) list.simps(9) rrewrite_srewrite.ss6 srewrites.simps) |
393 | 980 |
qed (auto intro: rrewrite_srewrite.intros) |
981 |
||
982 |
||
983 |
lemma rewrites_fuse: |
|
984 |
assumes "r1 \<leadsto>* r2" |
|
985 |
shows "fuse bs r1 \<leadsto>* fuse bs r2" |
|
986 |
using assms |
|
987 |
apply(induction r1 r2 arbitrary: bs rule: rrewrites.induct) |
|
988 |
apply(auto intro: rewrite_preserves_fuse rrewrites_trans) |
|
989 |
done |
|
365 | 990 |
|
991 |
||
374 | 992 |
lemma star_seq: |
993 |
assumes "r1 \<leadsto>* r2" |
|
994 |
shows "ASEQ bs r1 r3 \<leadsto>* ASEQ bs r2 r3" |
|
995 |
using assms |
|
996 |
apply(induct r1 r2 arbitrary: r3 rule: rrewrites.induct) |
|
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
997 |
apply(auto intro: rrewrite_srewrite.intros) |
374 | 998 |
done |
365 | 999 |
|
374 | 1000 |
lemma star_seq2: |
1001 |
assumes "r3 \<leadsto>* r4" |
|
1002 |
shows "ASEQ bs r1 r3 \<leadsto>* ASEQ bs r1 r4" |
|
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1003 |
using assms |
374 | 1004 |
apply(induct r3 r4 arbitrary: r1 rule: rrewrites.induct) |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1005 |
apply(auto intro: rrewrite_srewrite.intros) |
374 | 1006 |
done |
365 | 1007 |
|
374 | 1008 |
lemma continuous_rewrite: |
1009 |
assumes "r1 \<leadsto>* AZERO" |
|
1010 |
shows "ASEQ bs1 r1 r2 \<leadsto>* AZERO" |
|
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1011 |
using assms bs1 star_seq by blast |
365 | 1012 |
|
393 | 1013 |
(* |
1014 |
lemma continuous_rewrite2: |
|
1015 |
assumes "r1 \<leadsto>* AONE bs" |
|
1016 |
shows "ASEQ bs1 r1 r2 \<leadsto>* (fuse (bs1 @ bs) r2)" |
|
1017 |
using assms by (meson bs3 rrewrites.simps star_seq) |
|
1018 |
*) |
|
365 | 1019 |
|
374 | 1020 |
lemma bsimp_aalts_simpcases: |
1021 |
shows "AONE bs \<leadsto>* bsimp (AONE bs)" |
|
1022 |
and "AZERO \<leadsto>* bsimp AZERO" |
|
1023 |
and "ACHAR bs c \<leadsto>* bsimp (ACHAR bs c)" |
|
1024 |
by (simp_all) |
|
365 | 1025 |
|
393 | 1026 |
lemma bsimp_AALTs_rewrites: |
1027 |
shows "AALTs bs1 rs \<leadsto>* bsimp_AALTs bs1 rs" |
|
1028 |
by (smt (verit) bs6 bs7 bsimp_AALTs.elims rrewrites.simps) |
|
385 | 1029 |
|
1030 |
lemma trivialbsimp_srewrites: |
|
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1031 |
"\<lbrakk>\<And>x. x \<in> set rs \<Longrightarrow> x \<leadsto>* f x \<rbrakk> \<Longrightarrow> rs s\<leadsto>* (map f rs)" |
365 | 1032 |
apply(induction rs) |
1033 |
apply simp |
|
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1034 |
apply(simp) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1035 |
using srewrites7 by auto |
365 | 1036 |
|
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1037 |
|
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1038 |
|
393 | 1039 |
lemma fltsfrewrites: "rs s\<leadsto>* flts rs" |
1040 |
apply(induction rs rule: flts.induct) |
|
1041 |
apply(auto intro: rrewrite_srewrite.intros) |
|
1042 |
apply (meson srewrites.simps srewrites1 ss5) |
|
1043 |
using rs1 srewrites7 apply presburger |
|
1044 |
using srewrites7 apply force |
|
1045 |
apply (simp add: srewrites7) |
|
1046 |
by (simp add: srewrites7) |
|
365 | 1047 |
|
393 | 1048 |
lemma bnullable0: |
1049 |
shows "r1 \<leadsto> r2 \<Longrightarrow> bnullable r1 = bnullable r2" |
|
1050 |
and "rs1 s\<leadsto> rs2 \<Longrightarrow> bnullables rs1 = bnullables rs2" |
|
1051 |
apply(induct rule: rrewrite_srewrite.inducts) |
|
1052 |
apply(auto simp add: bnullable_fuse) |
|
493 | 1053 |
apply (meson UnCI bnullable_fuse imageI) |
1054 |
by (metis rb_nullability) |
|
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1055 |
|
373 | 1056 |
|
1057 |
lemma rewritesnullable: |
|
393 | 1058 |
assumes "r1 \<leadsto>* r2" |
1059 |
shows "bnullable r1 = bnullable r2" |
|
1060 |
using assms |
|
365 | 1061 |
apply(induction r1 r2 rule: rrewrites.induct) |
1062 |
apply simp |
|
393 | 1063 |
using bnullable0(1) by auto |
365 | 1064 |
|
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1065 |
lemma rewrite_bmkeps_aux: |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1066 |
shows "r1 \<leadsto> r2 \<Longrightarrow> (bnullable r1 \<and> bnullable r2 \<Longrightarrow> bmkeps r1 = bmkeps r2)" |
393 | 1067 |
and "rs1 s\<leadsto> rs2 \<Longrightarrow> (bnullables rs1 \<and> bnullables rs2 \<Longrightarrow> bmkepss rs1 = bmkepss rs2)" |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1068 |
proof (induct rule: rrewrite_srewrite.inducts) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1069 |
case (bs3 bs1 bs2 r) |
393 | 1070 |
then show ?case by (simp add: bmkeps_fuse) |
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1071 |
next |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1072 |
case (bs7 bs r) |
393 | 1073 |
then show ?case by (simp add: bmkeps_fuse) |
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1074 |
next |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1075 |
case (ss3 r1 r2 rs) |
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1076 |
then show ?case |
393 | 1077 |
by (metis bmkepss.simps(2) bnullable0(1)) |
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1078 |
next |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1079 |
case (ss5 bs1 rs1 rsb) |
393 | 1080 |
then show ?case |
1081 |
by (simp add: bmkepss1 bmkepss2 bmkepss_fuse bnullable_fuse) |
|
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1082 |
next |
393 | 1083 |
case (ss6 a1 a2 rsa rsb rsc) |
1084 |
then show ?case |
|
493 | 1085 |
by (smt (z3) append_is_Nil_conv bmkepss1 bmkepss2 in_set_conv_decomp_first list.set_intros(1) rb_nullability set_ConsD) |
393 | 1086 |
qed (auto) |
365 | 1087 |
|
373 | 1088 |
lemma rewrites_bmkeps: |
1089 |
assumes "r1 \<leadsto>* r2" "bnullable r1" |
|
1090 |
shows "bmkeps r1 = bmkeps r2" |
|
1091 |
using assms |
|
1092 |
proof(induction r1 r2 rule: rrewrites.induct) |
|
1093 |
case (rs1 r) |
|
1094 |
then show "bmkeps r = bmkeps r" by simp |
|
1095 |
next |
|
1096 |
case (rs2 r1 r2 r3) |
|
1097 |
then have IH: "bmkeps r1 = bmkeps r2" by simp |
|
1098 |
have a1: "bnullable r1" by fact |
|
1099 |
have a2: "r1 \<leadsto>* r2" by fact |
|
1100 |
have a3: "r2 \<leadsto> r3" by fact |
|
1101 |
have a4: "bnullable r2" using a1 a2 by (simp add: rewritesnullable) |
|
393 | 1102 |
then have "bmkeps r2 = bmkeps r3" |
1103 |
using a3 bnullable0(1) rewrite_bmkeps_aux(1) by blast |
|
373 | 1104 |
then show "bmkeps r1 = bmkeps r3" using IH by simp |
1105 |
qed |
|
365 | 1106 |
|
393 | 1107 |
|
1108 |
lemma rewrites_to_bsimp: |
|
1109 |
shows "r \<leadsto>* bsimp r" |
|
1110 |
proof (induction r rule: bsimp.induct) |
|
1111 |
case (1 bs1 r1 r2) |
|
1112 |
have IH1: "r1 \<leadsto>* bsimp r1" by fact |
|
1113 |
have IH2: "r2 \<leadsto>* bsimp r2" by fact |
|
1114 |
{ assume as: "bsimp r1 = AZERO \<or> bsimp r2 = AZERO" |
|
1115 |
with IH1 IH2 have "r1 \<leadsto>* AZERO \<or> r2 \<leadsto>* AZERO" by auto |
|
1116 |
then have "ASEQ bs1 r1 r2 \<leadsto>* AZERO" |
|
1117 |
by (metis bs2 continuous_rewrite rrewrites.simps star_seq2) |
|
1118 |
then have "ASEQ bs1 r1 r2 \<leadsto>* bsimp (ASEQ bs1 r1 r2)" using as by auto |
|
1119 |
} |
|
1120 |
moreover |
|
1121 |
{ assume "\<exists>bs. bsimp r1 = AONE bs" |
|
1122 |
then obtain bs where as: "bsimp r1 = AONE bs" by blast |
|
1123 |
with IH1 have "r1 \<leadsto>* AONE bs" by simp |
|
1124 |
then have "ASEQ bs1 r1 r2 \<leadsto>* fuse (bs1 @ bs) r2" using bs3 star_seq by blast |
|
1125 |
with IH2 have "ASEQ bs1 r1 r2 \<leadsto>* fuse (bs1 @ bs) (bsimp r2)" |
|
1126 |
using rewrites_fuse by (meson rrewrites_trans) |
|
1127 |
then have "ASEQ bs1 r1 r2 \<leadsto>* bsimp (ASEQ bs1 (AONE bs) r2)" by simp |
|
1128 |
then have "ASEQ bs1 r1 r2 \<leadsto>* bsimp (ASEQ bs1 r1 r2)" by (simp add: as) |
|
1129 |
} |
|
1130 |
moreover |
|
1131 |
{ assume as1: "bsimp r1 \<noteq> AZERO" "bsimp r2 \<noteq> AZERO" and as2: "(\<nexists>bs. bsimp r1 = AONE bs)" |
|
1132 |
then have "bsimp_ASEQ bs1 (bsimp r1) (bsimp r2) = ASEQ bs1 (bsimp r1) (bsimp r2)" |
|
1133 |
by (simp add: bsimp_ASEQ1) |
|
1134 |
then have "ASEQ bs1 r1 r2 \<leadsto>* bsimp_ASEQ bs1 (bsimp r1) (bsimp r2)" using as1 as2 IH1 IH2 |
|
1135 |
by (metis rrewrites_trans star_seq star_seq2) |
|
1136 |
then have "ASEQ bs1 r1 r2 \<leadsto>* bsimp (ASEQ bs1 r1 r2)" by simp |
|
1137 |
} |
|
1138 |
ultimately show "ASEQ bs1 r1 r2 \<leadsto>* bsimp (ASEQ bs1 r1 r2)" by blast |
|
1139 |
next |
|
1140 |
case (2 bs1 rs) |
|
1141 |
have IH: "\<And>x. x \<in> set rs \<Longrightarrow> x \<leadsto>* bsimp x" by fact |
|
1142 |
then have "rs s\<leadsto>* (map bsimp rs)" by (simp add: trivialbsimp_srewrites) |
|
1143 |
also have "... s\<leadsto>* flts (map bsimp rs)" by (simp add: fltsfrewrites) |
|
493 | 1144 |
also have "... s\<leadsto>* distinctBy (flts (map bsimp rs)) rerase {}" by (simp add: ss6_stronger) |
1145 |
finally have "AALTs bs1 rs \<leadsto>* AALTs bs1 (distinctBy (flts (map bsimp rs)) rerase {})" |
|
1146 |
using contextrewrites0 by auto |
|
1147 |
also have "... \<leadsto>* bsimp_AALTs bs1 (distinctBy (flts (map bsimp rs)) rerase {})" |
|
393 | 1148 |
by (simp add: bsimp_AALTs_rewrites) |
1149 |
finally show "AALTs bs1 rs \<leadsto>* bsimp (AALTs bs1 rs)" by simp |
|
1150 |
qed (simp_all) |
|
1151 |
||
1152 |
||
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1153 |
lemma to_zero_in_alt: |
393 | 1154 |
shows "AALT bs (ASEQ [] AZERO r) r2 \<leadsto> AALT bs AZERO r2" |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1155 |
by (simp add: bs1 bs10 ss3) |
365 | 1156 |
|
1157 |
||
1158 |
||
374 | 1159 |
lemma bder_fuse_list: |
1160 |
shows "map (bder c \<circ> fuse bs1) rs1 = map (fuse bs1 \<circ> bder c) rs1" |
|
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1161 |
apply(induction rs1) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1162 |
apply(simp_all add: bder_fuse) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1163 |
done |
365 | 1164 |
|
1165 |
||
393 | 1166 |
lemma rewrite_preserves_bder: |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1167 |
shows "r1 \<leadsto> r2 \<Longrightarrow> (bder c r1) \<leadsto>* (bder c r2)" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1168 |
and "rs1 s\<leadsto> rs2 \<Longrightarrow> map (bder c) rs1 s\<leadsto>* map (bder c) rs2" |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1169 |
proof(induction rule: rrewrite_srewrite.inducts) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1170 |
case (bs1 bs r2) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1171 |
then show ?case |
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1172 |
by (simp add: continuous_rewrite) |
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1173 |
next |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1174 |
case (bs2 bs r1) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1175 |
then show ?case |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1176 |
apply(auto) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1177 |
apply (meson bs6 contextrewrites0 rrewrite_srewrite.bs2 rs2 ss3 ss4 sss1 sss2) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1178 |
by (simp add: r_in_rstar rrewrite_srewrite.bs2) |
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1179 |
next |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1180 |
case (bs3 bs1 bs2 r) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1181 |
then show ?case |
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1182 |
apply(simp) |
393 | 1183 |
|
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1184 |
by (metis (no_types, lifting) bder_fuse bs10 bs7 fuse_append rrewrites.simps ss4 to_zero_in_alt) |
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1185 |
next |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1186 |
case (bs4 r1 r2 bs r3) |
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1187 |
have as: "r1 \<leadsto> r2" by fact |
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1188 |
have IH: "bder c r1 \<leadsto>* bder c r2" by fact |
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1189 |
from as IH show "bder c (ASEQ bs r1 r3) \<leadsto>* bder c (ASEQ bs r2 r3)" |
393 | 1190 |
by (metis bder.simps(5) bnullable0(1) contextrewrites1 rewrite_bmkeps_aux(1) star_seq) |
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1191 |
next |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1192 |
case (bs5 r3 r4 bs r1) |
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1193 |
have as: "r3 \<leadsto> r4" by fact |
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1194 |
have IH: "bder c r3 \<leadsto>* bder c r4" by fact |
385 | 1195 |
from as IH show "bder c (ASEQ bs r1 r3) \<leadsto>* bder c (ASEQ bs r1 r4)" |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1196 |
apply(simp) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1197 |
apply(auto) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1198 |
using contextrewrites0 r_in_rstar rewrites_fuse srewrites6 srewrites7 star_seq2 apply presburger |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1199 |
using star_seq2 by blast |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1200 |
next |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1201 |
case (bs6 bs) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1202 |
then show ?case |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1203 |
using rrewrite_srewrite.bs6 by force |
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1204 |
next |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1205 |
case (bs7 bs r) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1206 |
then show ?case |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1207 |
by (simp add: bder_fuse r_in_rstar rrewrite_srewrite.bs7) |
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1208 |
next |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1209 |
case (bs10 rs1 rs2 bs) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1210 |
then show ?case |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1211 |
using contextrewrites0 by force |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1212 |
next |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1213 |
case ss1 |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1214 |
then show ?case by simp |
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1215 |
next |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1216 |
case (ss2 rs1 rs2 r) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1217 |
then show ?case |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1218 |
by (simp add: srewrites7) |
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1219 |
next |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1220 |
case (ss3 r1 r2 rs) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1221 |
then show ?case |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1222 |
by (simp add: srewrites7) |
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1223 |
next |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1224 |
case (ss4 rs) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1225 |
then show ?case |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1226 |
using rrewrite_srewrite.ss4 by fastforce |
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1227 |
next |
392
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1228 |
case (ss5 bs1 rs1 rsb) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1229 |
then show ?case |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1230 |
apply(simp) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1231 |
using bder_fuse_list map_map rrewrite_srewrite.ss5 srewrites.simps by blast |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1232 |
next |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1233 |
case (ss6 a1 a2 bs rsa rsb) |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1234 |
then show ?case |
8194086c2a8a
simplified version
Christian Urban <christian.urban@kcl.ac.uk>
parents:
385
diff
changeset
|
1235 |
apply(simp only: map_append) |
493 | 1236 |
by (smt (verit, best) rder_bder_rerase list.simps(8) list.simps(9) local.ss6 rrewrite_srewrite.ss6 srewrites.simps) |
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1237 |
qed |
365 | 1238 |
|
393 | 1239 |
lemma rewrites_preserves_bder: |
373 | 1240 |
assumes "r1 \<leadsto>* r2" |
1241 |
shows "bder c r1 \<leadsto>* bder c r2" |
|
1242 |
using assms |
|
1243 |
apply(induction r1 r2 rule: rrewrites.induct) |
|
393 | 1244 |
apply(simp_all add: rewrite_preserves_bder rrewrites_trans) |
373 | 1245 |
done |
365 | 1246 |
|
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1247 |
|
373 | 1248 |
lemma central: |
1249 |
shows "bders r s \<leadsto>* bders_simp r s" |
|
1250 |
proof(induct s arbitrary: r rule: rev_induct) |
|
1251 |
case Nil |
|
1252 |
then show "bders r [] \<leadsto>* bders_simp r []" by simp |
|
1253 |
next |
|
1254 |
case (snoc x xs) |
|
1255 |
have IH: "\<And>r. bders r xs \<leadsto>* bders_simp r xs" by fact |
|
1256 |
have "bders r (xs @ [x]) = bders (bders r xs) [x]" by (simp add: bders_append) |
|
1257 |
also have "... \<leadsto>* bders (bders_simp r xs) [x]" using IH |
|
393 | 1258 |
by (simp add: rewrites_preserves_bder) |
373 | 1259 |
also have "... \<leadsto>* bders_simp (bders_simp r xs) [x]" using IH |
393 | 1260 |
by (simp add: rewrites_to_bsimp) |
373 | 1261 |
finally show "bders r (xs @ [x]) \<leadsto>* bders_simp r (xs @ [x])" |
1262 |
by (simp add: bders_simp_append) |
|
1263 |
qed |
|
365 | 1264 |
|
393 | 1265 |
lemma main_aux: |
373 | 1266 |
assumes "bnullable (bders r s)" |
1267 |
shows "bmkeps (bders r s) = bmkeps (bders_simp r s)" |
|
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1268 |
proof - |
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1269 |
have "bders r s \<leadsto>* bders_simp r s" by (rule central) |
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1270 |
then |
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1271 |
show "bmkeps (bders r s) = bmkeps (bders_simp r s)" using assms |
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1272 |
by (rule rewrites_bmkeps) |
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1273 |
qed |
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1274 |
|
365 | 1275 |
|
385 | 1276 |
|
1277 |
||
393 | 1278 |
theorem main_blexer_simp: |
373 | 1279 |
shows "blexer r s = blexer_simp r s" |
381
0c666a0c57d7
isarfied some proofs
Christian Urban <christian.urban@kcl.ac.uk>
parents:
379
diff
changeset
|
1280 |
unfolding blexer_def blexer_simp_def |
393 | 1281 |
using b4 main_aux by simp |
365 | 1282 |
|
1283 |
||
373 | 1284 |
theorem blexersimp_correctness: |
1285 |
shows "lexer r s = blexer_simp r s" |
|
393 | 1286 |
using blexer_correctness main_blexer_simp by simp |
365 | 1287 |
|
1288 |
||
492 | 1289 |
|
1290 |
lemma rerase_preimage: |
|
1291 |
shows "rerase r = RZERO \<Longrightarrow> r = AZERO" |
|
1292 |
apply(case_tac r) |
|
1293 |
apply simp+ |
|
1294 |
done |
|
1295 |
||
1296 |
lemma rerase_preimage2: |
|
1297 |
shows "rerase r = RONE \<Longrightarrow> \<exists>bs. r = AONE bs" |
|
1298 |
apply(case_tac r) |
|
1299 |
apply simp+ |
|
1300 |
done |
|
1301 |
||
1302 |
lemma rerase_preimage3: |
|
1303 |
shows "rerase r= RCHAR c \<Longrightarrow> \<exists>bs. r = ACHAR bs c" |
|
1304 |
apply(case_tac r) |
|
1305 |
apply simp+ |
|
1306 |
done |
|
1307 |
||
1308 |
lemma rerase_preimage4: |
|
1309 |
shows "rerase r = RSEQ r1 r2 \<Longrightarrow> \<exists>bs a1 a2. r = ASEQ bs a1 a2 \<and> rerase a1 = r1 \<and> rerase a2 = r2" |
|
1310 |
apply(case_tac r) |
|
1311 |
apply(simp)+ |
|
1312 |
done |
|
1313 |
||
1314 |
lemma rerase_preimage5: |
|
1315 |
shows "rerase r = RALTS rs \<Longrightarrow> \<exists>bs as. r = AALTs bs as \<and> map rerase as = rs" |
|
1316 |
apply(case_tac r) |
|
1317 |
apply(simp)+ |
|
1318 |
done |
|
1319 |
||
1320 |
lemma rerase_preimage6: |
|
1321 |
shows "rerase r = RSTAR r0 \<Longrightarrow> \<exists>bs a0. r = ASTAR bs a0 \<and> rerase a0 = r0 " |
|
1322 |
apply(case_tac r) |
|
1323 |
apply(simp)+ |
|
1324 |
done |
|
1325 |
||
1326 |
lemma map_rder_bder: |
|
1327 |
shows "\<lbrakk> \<And>xa a. \<lbrakk>xa \<in> set x; rerase a = xa\<rbrakk> \<Longrightarrow> rerase (bder c a) = rder c xa; |
|
1328 |
map rerase as = x\<rbrakk> \<Longrightarrow> |
|
1329 |
map rerase (map (bder c) as) = map (rder c) x" |
|
1330 |
apply(induct x arbitrary: as) |
|
1331 |
apply simp+ |
|
1332 |
by force |
|
378 | 1333 |
|
1334 |
||
492 | 1335 |
lemma der_rerase: |
1336 |
shows "rerase a = r \<Longrightarrow> rerase (bder c a) = rder c r" |
|
1337 |
apply(induct r arbitrary: a) |
|
1338 |
apply simp |
|
1339 |
using rerase_preimage apply fastforce |
|
1340 |
using rerase_preimage2 apply force |
|
1341 |
apply (metis bder.simps(3) rder.simps(3) rerase.simps(1) rerase.simps(2) rerase_preimage3) |
|
1342 |
apply(insert rerase_preimage4)[1] |
|
1343 |
apply(subgoal_tac " \<exists>bs a1 a2. a = ASEQ bs a1 a2 \<and> rerase a1 = r1 \<and> rerase a2 = r2") |
|
1344 |
prefer 2 |
|
1345 |
apply presburger |
|
1346 |
apply(erule exE)+ |
|
1347 |
apply(erule conjE)+ |
|
1348 |
apply(subgoal_tac " rerase (bder c a1) = rder c r1") |
|
1349 |
prefer 2 |
|
1350 |
apply blast |
|
1351 |
apply(subgoal_tac " rerase (bder c a2) = rder c r2") |
|
1352 |
prefer 2 |
|
1353 |
apply blast |
|
1354 |
apply(case_tac "rnullable r1") |
|
1355 |
apply(subgoal_tac "bnullable a1") |
|
1356 |
apply simp |
|
1357 |
using fuse_anything_doesnt_matter apply presburger |
|
1358 |
using rb_nullability apply blast |
|
1359 |
apply (metis bder.simps(5) rb_nullability rder.simps(5) rerase.simps(5)) |
|
1360 |
apply simp |
|
1361 |
apply(insert rerase_preimage5)[1] |
|
1362 |
apply(subgoal_tac "\<exists>bs as. a = AALTs bs as \<and> map rerase as = x") |
|
1363 |
prefer 2 |
|
1364 |
||
1365 |
apply blast |
|
1366 |
apply(erule exE)+ |
|
1367 |
apply(erule conjE)+ |
|
1368 |
apply(subgoal_tac "map rerase (map (bder c) as) = map (rder c) x") |
|
1369 |
using bder.simps(4) rerase.simps(4) apply presburger |
|
1370 |
using map_rder_bder apply blast |
|
1371 |
using fuse_anything_doesnt_matter rerase_preimage6 by force |
|
1372 |
||
1373 |
lemma der_rerase2: |
|
1374 |
shows "rerase (bder c a) = rder c (rerase a)" |
|
1375 |
using der_rerase by force |
|
1376 |
||
1377 |
lemma ders_rerase: |
|
1378 |
shows "rerase (bders a s) = rders (rerase a) s" |
|
1379 |
apply(induct s rule: rev_induct) |
|
1380 |
apply simp |
|
1381 |
by (simp add: bders_append der_rerase rders_append) |
|
1382 |
||
1383 |
lemma rerase_bsimp_ASEQ: |
|
1384 |
shows "rerase (bsimp_ASEQ x1 a1 a2) = rsimp_SEQ (rerase a1) (rerase a2)" |
|
1385 |
by (smt (verit, ccfv_SIG) SizeBound3.bsimp_ASEQ0 SizeBound3.bsimp_ASEQ2 basic_rsimp_SEQ_property1 basic_rsimp_SEQ_property2 basic_rsimp_SEQ_property3 bsimp_ASEQ.simps(1) bsimp_ASEQ1 fuse_anything_doesnt_matter rerase.simps(1) rerase.simps(2) rerase.simps(5) rerase_preimage rerase_preimage2 rsimp_SEQ.simps(1)) |
|
1386 |
||
1387 |
lemma rerase_bsimp_AALTs: |
|
1388 |
shows "rerase (bsimp_AALTs bs rs) = rsimp_ALTs (map rerase rs)" |
|
1389 |
apply(induct rs arbitrary: bs) |
|
1390 |
apply simp+ |
|
1391 |
by (smt (verit, ccfv_threshold) Cons_eq_map_conv bsimp_AALTs.elims fuse_anything_doesnt_matter list.discI list.inject list.simps(9) rerase.simps(4) rsimp_ALTs.elims) |
|
1392 |
||
1393 |
||
1394 |
fun anonalt :: "arexp \<Rightarrow> bool" |
|
1395 |
where |
|
1396 |
"anonalt (AALTs bs2 rs) = False" |
|
1397 |
| "anonalt r = True" |
|
1398 |
||
1399 |
||
1400 |
fun agood :: "arexp \<Rightarrow> bool" where |
|
1401 |
"agood AZERO = False" |
|
1402 |
| "agood (AONE cs) = True" |
|
1403 |
| "agood (ACHAR cs c) = True" |
|
1404 |
| "agood (AALTs cs []) = False" |
|
1405 |
| "agood (AALTs cs [r]) = False" |
|
493 | 1406 |
| "agood (AALTs cs (r1#r2#rs)) = (distinct (map rerase (r1 # r2 # rs)) \<and>(\<forall>r' \<in> set (r1#r2#rs). agood r' \<and> anonalt r'))" |
492 | 1407 |
| "agood (ASEQ _ AZERO _) = False" |
1408 |
| "agood (ASEQ _ (AONE _) _) = False" |
|
1409 |
| "agood (ASEQ _ _ AZERO) = False" |
|
1410 |
| "agood (ASEQ cs r1 r2) = (agood r1 \<and> agood r2)" |
|
1411 |
| "agood (ASTAR cs r) = True" |
|
1412 |
||
1413 |
||
1414 |
fun anonnested :: "arexp \<Rightarrow> bool" |
|
1415 |
where |
|
1416 |
"anonnested (AALTs bs2 []) = True" |
|
1417 |
| "anonnested (AALTs bs2 ((AALTs bs1 rs1) # rs2)) = False" |
|
1418 |
| "anonnested (AALTs bs2 (r # rs2)) = anonnested (AALTs bs2 rs2)" |
|
1419 |
| "anonnested r = True" |
|
365 | 1420 |
|
1421 |
||
492 | 1422 |
lemma k0: |
1423 |
shows "flts (r # rs1) = flts [r] @ flts rs1" |
|
1424 |
apply(induct r arbitrary: rs1) |
|
1425 |
apply(auto) |
|
1426 |
done |
|
1427 |
||
1428 |
lemma k00: |
|
1429 |
shows "flts (rs1 @ rs2) = flts rs1 @ flts rs2" |
|
1430 |
apply(induct rs1 arbitrary: rs2) |
|
1431 |
apply(auto) |
|
1432 |
by (metis append.assoc k0) |
|
1433 |
||
1434 |
||
1435 |
lemma bbbbs: |
|
1436 |
assumes "agood r" "r = AALTs bs1 rs" |
|
1437 |
shows "bsimp_AALTs bs (flts [r]) = AALTs bs (map (fuse bs1) rs)" |
|
1438 |
using assms |
|
1439 |
by (smt (verit, ccfv_SIG) Cons_eq_map_conv agood.simps(4) agood.simps(5) append.right_neutral bsimp_AALTs.elims flts.simps(1) flts.simps(3) list.map_disc_iff) |
|
1440 |
||
1441 |
lemma bbbbs1: |
|
1442 |
shows "anonalt r \<or> (\<exists>bs rs. r = AALTs bs rs)" |
|
1443 |
by (meson anonalt.elims(3)) |
|
1444 |
||
1445 |
||
1446 |
||
1447 |
lemma good_fuse: |
|
1448 |
shows "agood (fuse bs r) = agood r" |
|
1449 |
apply(induct r arbitrary: bs) |
|
1450 |
apply(auto) |
|
1451 |
apply(case_tac r1) |
|
1452 |
apply(simp_all) |
|
1453 |
apply(case_tac r2) |
|
1454 |
apply(simp_all) |
|
1455 |
apply(case_tac r2) |
|
1456 |
apply(simp_all) |
|
1457 |
apply(case_tac r2) |
|
1458 |
apply(simp_all) |
|
1459 |
apply(case_tac r2) |
|
1460 |
apply(simp_all) |
|
1461 |
apply(case_tac r1) |
|
1462 |
apply(simp_all) |
|
1463 |
apply(case_tac r2) |
|
1464 |
apply(simp_all) |
|
1465 |
apply(case_tac r2) |
|
1466 |
apply(simp_all) |
|
1467 |
apply(case_tac r2) |
|
1468 |
apply(simp_all) |
|
1469 |
apply(case_tac r2) |
|
1470 |
apply(simp_all) |
|
1471 |
apply(case_tac x2a) |
|
1472 |
apply(simp_all) |
|
1473 |
apply(case_tac list) |
|
1474 |
apply(simp_all) |
|
1475 |
apply(case_tac x2a) |
|
1476 |
apply(simp_all) |
|
1477 |
apply(case_tac list) |
|
1478 |
apply(simp_all) |
|
1479 |
done |
|
1480 |
||
1481 |
lemma good0: |
|
493 | 1482 |
assumes "rs \<noteq> Nil" "\<forall>r \<in> set rs. anonalt r" "distinct (map rerase rs)" |
492 | 1483 |
shows "agood (bsimp_AALTs bs rs) \<longleftrightarrow> (\<forall>r \<in> set rs. agood r)" |
1484 |
using assms |
|
1485 |
apply(induct bs rs rule: bsimp_AALTs.induct) |
|
1486 |
apply simp+ |
|
1487 |
apply (simp add: good_fuse) |
|
1488 |
apply(subgoal_tac "bsimp_AALTs bs1 (v # vb # vc) = AALTs bs1 (v # vb # vc)") |
|
1489 |
prefer 2 |
|
1490 |
||
1491 |
||
1492 |
using bsimp_AALTs.simps(3) apply presburger |
|
1493 |
apply(simp only:) |
|
493 | 1494 |
apply(subgoal_tac "agood (AALTs bs1 (v # vb # vc)) = ((\<forall>r \<in> (set (v # vb # vc)). agood r \<and> anonalt r) \<and> distinct (map rerase (v # vb # vc)))") |
492 | 1495 |
prefer 2 |
1496 |
using agood.simps(6) apply blast |
|
1497 |
apply(simp only:) |
|
493 | 1498 |
apply(subgoal_tac "((\<forall>r\<in>set (v # vb # vc). agood r \<and> anonalt r) \<and> distinct (map rerase (v # vb # vc))) = ((\<forall>r\<in>set (v # vb # vc). agood r) \<and> distinct (map rerase (v # vb # vc)))") |
492 | 1499 |
prefer 2 |
1500 |
apply blast |
|
1501 |
apply(simp only:) |
|
493 | 1502 |
apply(subgoal_tac "((\<forall>r \<in> set (v # vb # vc). agood r) \<and> distinct (map rerase (v # vb # vc))) = (\<forall> r \<in> set (v # vb # vc). agood r) ") |
492 | 1503 |
prefer 2 |
493 | 1504 |
apply(subgoal_tac "distinct (map rerase (v # vb # vc)) = True") |
1505 |
prefer 2 |
|
492 | 1506 |
apply blast |
1507 |
prefer 2 |
|
1508 |
apply force |
|
1509 |
apply simp |
|
1510 |
done |
|
1511 |
||
1512 |
||
1513 |
lemma nn11a: |
|
1514 |
assumes "anonalt r" |
|
1515 |
shows "anonalt (fuse bs r)" |
|
1516 |
using assms |
|
1517 |
apply(induct r) |
|
1518 |
apply(auto) |
|
1519 |
done |
|
1520 |
||
378 | 1521 |
|
1522 |
||
1523 |
||
492 | 1524 |
lemma flts0: |
1525 |
assumes "r \<noteq> AZERO" "anonalt r" |
|
1526 |
shows "flts [r] \<noteq> []" |
|
1527 |
using assms |
|
1528 |
apply(induct r) |
|
1529 |
apply(simp_all) |
|
1530 |
done |
|
1531 |
||
1532 |
lemma flts1: |
|
1533 |
assumes "agood r" |
|
1534 |
shows "flts [r] \<noteq> []" |
|
1535 |
using assms |
|
1536 |
apply(induct r) |
|
1537 |
apply(simp_all) |
|
1538 |
apply(case_tac x2a) |
|
1539 |
apply(simp) |
|
1540 |
apply(simp) |
|
1541 |
done |
|
1542 |
||
1543 |
lemma flts2: |
|
1544 |
assumes "agood r" |
|
1545 |
shows "\<forall>r' \<in> set (flts [r]). agood r' \<and> anonalt r'" |
|
1546 |
using assms |
|
1547 |
apply(induct r) |
|
1548 |
apply(simp) |
|
1549 |
apply(simp) |
|
1550 |
apply(simp) |
|
1551 |
prefer 2 |
|
1552 |
apply(simp) |
|
1553 |
apply(auto)[1] |
|
1554 |
apply (metis bsimp_AALTs.elims agood.simps(4) agood.simps(5) agood.simps(6) good_fuse) |
|
1555 |
apply (metis bsimp_AALTs.elims agood.simps(4) agood.simps(5) agood.simps(6) nn11a) |
|
1556 |
apply fastforce |
|
1557 |
apply(simp) |
|
1558 |
done |
|
1559 |
||
1560 |
||
1561 |
lemma flts3: |
|
1562 |
assumes "\<forall>r \<in> set rs. agood r \<or> r = AZERO" |
|
1563 |
shows "\<forall>r \<in> set (flts rs). agood r" |
|
1564 |
using assms |
|
1565 |
apply(induct rs arbitrary: rule: flts.induct) |
|
1566 |
apply(simp_all) |
|
1567 |
by (metis SizeBound3.flts2 UnE append.right_neutral flts.simps(1) flts.simps(3) list.set_map) |
|
1568 |
||
1569 |
||
1570 |
lemma flts3b: |
|
1571 |
assumes "\<exists>r\<in>set rs. agood r" |
|
1572 |
shows "flts rs \<noteq> []" |
|
1573 |
using assms |
|
1574 |
apply(induct rs arbitrary: rule: flts.induct) |
|
1575 |
apply(simp) |
|
1576 |
apply(simp) |
|
1577 |
apply(simp) |
|
1578 |
apply(auto) |
|
1579 |
done |
|
1580 |
||
1581 |
lemma k0a: |
|
1582 |
shows "flts [AALTs bs rs] = map (fuse bs) rs" |
|
1583 |
apply(simp) |
|
1584 |
done |
|
1585 |
||
1586 |
||
1587 |
lemma k0b: |
|
1588 |
assumes "anonalt r" "r \<noteq> AZERO" |
|
1589 |
shows "flts [r] = [r]" |
|
1590 |
using assms |
|
1591 |
apply(case_tac r) |
|
1592 |
apply(simp_all) |
|
1593 |
done |
|
1594 |
||
1595 |
lemma flts4_nogood: |
|
1596 |
shows "bsimp_AALTs bs0 rs = AZERO \<Longrightarrow> \<forall>r \<in> set rs. \<not> agood r" |
|
1597 |
by (metis SizeBound3.flts3b arexp.distinct(7) bsimp_AALTs.elims flts.simps(1) flts.simps(2) fuse_anything_doesnt_matter rerase.simps(1) rerase_preimage) |
|
1598 |
||
1599 |
lemma flts4_append: |
|
1600 |
shows "bsimp_AALTs bs0 (rs @ flts rsa) = AZERO \<Longrightarrow> bsimp_AALTs bs0 rs = AZERO" |
|
1601 |
by (metis append_is_Nil_conv arexp.simps(13) bsimp_AALTs.elims bsimp_AALTs.simps(1) butlast_append butlast_snoc) |
|
1602 |
||
1603 |
lemma flts4: |
|
1604 |
assumes "bsimp_AALTs bs (flts rs) = AZERO" |
|
1605 |
shows "\<forall>r \<in> set rs. \<not> agood r" |
|
1606 |
using assms |
|
1607 |
apply(induct rs arbitrary: bs rule: flts.induct) |
|
1608 |
apply(auto) |
|
1609 |
defer |
|
1610 |
apply (metis (no_types, lifting) Nil_is_append_conv append_self_conv2 bsimp_AALTs.elims butlast.simps(2) butlast_append flts3b anonalt.simps(1) anonalt.simps(2)) |
|
1611 |
apply (metis agood.simps(1) agood.simps(2) arexp.distinct(7) bsimp_AALTs.elims good_fuse list.distinct(1) list.inject) |
|
1612 |
apply (metis agood.simps(1) agood.simps(3) arexp.distinct(7) bsimp_AALTs.elims good_fuse list.distinct(1) list.inject) |
|
1613 |
apply (metis agood.simps(1) arexp.distinct(7) bsimp_AALTs.elims good_fuse list.discI list.inject) |
|
1614 |
apply (metis arexp.distinct(7) bsimp_AALTs.elims list.distinct(1) list.inject) |
|
1615 |
apply (metis agood.simps(1) agood.simps(33) arexp.distinct(7) bsimp_AALTs.elims good_fuse list.discI list.inject) |
|
1616 |
by (metis SizeBound3.bbbbs SizeBound3.k0a arexp.simps(13) flts4_append) |
|
1617 |
||
1618 |
||
1619 |
lemma flts_nil: |
|
1620 |
assumes "\<forall>y. asize y < Suc (sum_list (map asize rs)) \<longrightarrow> |
|
1621 |
agood (bsimp y) \<or> bsimp y = AZERO" |
|
1622 |
and "\<forall>r\<in>set rs. \<not> agood (bsimp r)" |
|
1623 |
shows "flts (map bsimp rs) = []" |
|
1624 |
using assms |
|
1625 |
apply(induct rs) |
|
1626 |
apply(simp) |
|
1627 |
apply(simp) |
|
1628 |
apply(subst k0) |
|
1629 |
apply(simp) |
|
1630 |
by force |
|
1631 |
||
1632 |
lemma flts_nil2: |
|
1633 |
assumes "\<forall>y. asize y < Suc (sum_list (map asize rs)) \<longrightarrow> |
|
1634 |
agood (bsimp y) \<or> bsimp y = AZERO" |
|
1635 |
and "bsimp_AALTs bs (flts (map bsimp rs)) = AZERO" |
|
1636 |
shows "flts (map bsimp rs) = []" |
|
1637 |
using assms |
|
1638 |
apply(induct rs arbitrary: bs) |
|
1639 |
apply(simp) |
|
1640 |
apply(simp) |
|
1641 |
apply(subst k0) |
|
1642 |
apply(simp) |
|
1643 |
apply(subst (asm) k0) |
|
1644 |
apply(auto) |
|
1645 |
apply (metis flts.simps(1) flts.simps(2) flts4 k0 less_add_Suc1 list.set_intros(1)) |
|
1646 |
by (metis flts.simps(2) flts4 k0 less_add_Suc1 list.set_intros(1)) |
|
1647 |
||
1648 |
||
1649 |
||
1650 |
lemma good_SEQ: |
|
1651 |
assumes "r1 \<noteq> AZERO" "r2 \<noteq> AZERO" "\<forall>bs. r1 \<noteq> AONE bs" |
|
1652 |
shows "agood (ASEQ bs r1 r2) = (agood r1 \<and> agood r2)" |
|
1653 |
using assms |
|
1654 |
apply(case_tac r1) |
|
1655 |
apply(simp_all) |
|
1656 |
apply(case_tac r2) |
|
1657 |
apply(simp_all) |
|
1658 |
apply(case_tac r2) |
|
1659 |
apply(simp_all) |
|
1660 |
apply(case_tac r2) |
|
1661 |
apply(simp_all) |
|
1662 |
apply(case_tac r2) |
|
1663 |
apply(simp_all) |
|
1664 |
done |
|
1665 |
||
1666 |
lemma asize0: |
|
1667 |
shows "0 < asize r" |
|
1668 |
apply(induct r) |
|
1669 |
apply(auto) |
|
1670 |
done |
|
1671 |
||
1672 |
lemma nn1qq: |
|
1673 |
assumes "anonnested (AALTs bs rs)" |
|
1674 |
shows "\<nexists>bs1 rs1. AALTs bs1 rs1 \<in> set rs" |
|
1675 |
using assms |
|
1676 |
apply(induct rs rule: flts.induct) |
|
1677 |
apply(auto) |
|
1678 |
done |
|
1679 |
||
1680 |
lemma nn1c: |
|
1681 |
assumes "\<forall>r \<in> set rs. anonnested r" |
|
1682 |
shows "\<forall>r \<in> set (flts rs). anonalt r" |
|
1683 |
using assms |
|
1684 |
apply(induct rs rule: flts.induct) |
|
1685 |
apply(auto) |
|
1686 |
apply(rule nn11a) |
|
1687 |
by (metis nn1qq anonalt.elims(3)) |
|
1688 |
||
1689 |
lemma n0: |
|
1690 |
shows "anonnested (AALTs bs rs) \<longleftrightarrow> (\<forall>r \<in> set rs. anonalt r)" |
|
1691 |
apply(induct rs arbitrary: bs) |
|
1692 |
apply(auto) |
|
1693 |
apply (metis SizeBound3.nn1qq anonalt.elims(3) list.set_intros(1)) |
|
1694 |
apply (metis SizeBound3.bbbbs1 SizeBound3.nn1qq list.set_intros(2)) |
|
1695 |
by (metis anonalt.elims(2) anonnested.simps(3) anonnested.simps(4) anonnested.simps(5) anonnested.simps(6) anonnested.simps(7)) |
|
1696 |
||
1697 |
||
1698 |
||
1699 |
lemma nn1bb: |
|
1700 |
assumes "\<forall>r \<in> set rs. anonalt r" |
|
1701 |
shows "anonnested (bsimp_AALTs bs rs)" |
|
1702 |
using assms |
|
1703 |
apply(induct bs rs rule: bsimp_AALTs.induct) |
|
1704 |
apply(auto) |
|
1705 |
apply (metis nn11a anonalt.simps(1) anonnested.elims(3)) |
|
1706 |
using n0 by auto |
|
1707 |
||
1708 |
lemma nn10: |
|
1709 |
assumes "anonnested (AALTs cs rs)" |
|
1710 |
shows "anonnested (AALTs (bs @ cs) rs)" |
|
1711 |
using assms |
|
1712 |
apply(induct rs arbitrary: cs bs) |
|
1713 |
apply(simp_all) |
|
1714 |
apply(case_tac a) |
|
1715 |
apply(simp_all) |
|
1716 |
done |
|
1717 |
||
1718 |
lemma nn1a: |
|
1719 |
assumes "anonnested r" |
|
1720 |
shows "anonnested (fuse bs r)" |
|
1721 |
using assms |
|
1722 |
apply(induct bs r arbitrary: rule: fuse.induct) |
|
1723 |
apply(simp_all add: nn10) |
|
1724 |
done |
|
1725 |
||
1726 |
lemma dB_keeps_property: |
|
493 | 1727 |
shows "(\<forall>r \<in> set rs. P r) \<longrightarrow> (\<forall>r \<in> set (distinctBy rs rerase rset). P r)" |
492 | 1728 |
apply(induct rs arbitrary: rset) |
1729 |
apply simp+ |
|
1730 |
done |
|
1731 |
||
1732 |
lemma dB_filters_out: |
|
493 | 1733 |
shows "rerase a \<in> rset \<Longrightarrow> a \<notin> set (distinctBy rs rerase (rset))" |
492 | 1734 |
apply(induct rs arbitrary: rset) |
1735 |
apply simp |
|
1736 |
apply(case_tac "a = aa") |
|
1737 |
apply simp+ |
|
1738 |
done |
|
1739 |
||
1740 |
lemma dB_will_be_distinct: |
|
493 | 1741 |
shows "distinct (distinctBy rs rerase (insert (rerase a) rset)) \<Longrightarrow> distinct (a # (distinctBy rs rerase (insert (rerase a) rset)))" |
492 | 1742 |
using dB_filters_out by force |
1743 |
||
1744 |
||
1745 |
lemma dB_does_the_job2: |
|
493 | 1746 |
shows "distinct (distinctBy rs rerase rset)" |
492 | 1747 |
apply(induct rs arbitrary: rset) |
1748 |
apply simp |
|
493 | 1749 |
apply(case_tac "rerase a \<in> rset") |
492 | 1750 |
apply simp |
493 | 1751 |
apply(drule_tac x = "insert (rerase a) rset" in meta_spec) |
1752 |
apply(subgoal_tac "distinctBy (a # rs) rerase rset = a # distinctBy rs rerase (insert (rerase a ) rset)") |
|
492 | 1753 |
apply(simp only:) |
1754 |
using dB_will_be_distinct apply presburger |
|
1755 |
by auto |
|
1756 |
||
1757 |
lemma dB_does_more_job: |
|
493 | 1758 |
shows "distinct (map rerase (distinctBy rs rerase rset))" |
492 | 1759 |
apply(induct rs arbitrary:rset) |
1760 |
apply simp |
|
493 | 1761 |
apply(case_tac "rerase a \<in> rset") |
492 | 1762 |
apply simp+ |
1763 |
using dB_filters_out by force |
|
1764 |
||
1765 |
lemma dB_mono2: |
|
493 | 1766 |
shows "rset \<subseteq> rset' \<Longrightarrow> distinctBy rs rerase rset = [] \<Longrightarrow> distinctBy rs rerase rset' = []" |
492 | 1767 |
apply(induct rs arbitrary: rset rset') |
1768 |
apply simp+ |
|
1769 |
by (meson in_mono list.discI) |
|
1770 |
||
1771 |
||
1772 |
lemma nn1b: |
|
1773 |
shows "anonnested (bsimp r)" |
|
1774 |
apply(induct r) |
|
1775 |
apply(simp_all) |
|
1776 |
apply(case_tac "bsimp r1 = AZERO") |
|
1777 |
apply(simp) |
|
1778 |
apply(case_tac "bsimp r2 = AZERO") |
|
1779 |
apply(simp) |
|
1780 |
apply(case_tac "\<exists>bs. bsimp r1 = AONE bs") |
|
1781 |
apply(auto)[1] |
|
1782 |
apply (simp add: nn1a) |
|
1783 |
apply(subst bsimp_ASEQ1) |
|
1784 |
apply(auto) |
|
1785 |
apply(rule nn1bb) |
|
1786 |
apply(auto) |
|
1787 |
apply(subgoal_tac "\<forall>r \<in> set (flts (map bsimp x2a)). anonalt r") |
|
1788 |
prefer 2 |
|
1789 |
apply (metis (mono_tags, lifting) SizeBound3.nn1c image_iff list.set_map) |
|
493 | 1790 |
apply(subgoal_tac "\<forall>r \<in> set (distinctBy (flts (map bsimp x2a)) rerase {}). anonalt r") |
1791 |
prefer 2 |
|
492 | 1792 |
using dB_keeps_property apply presburger |
1793 |
by blast |
|
1794 |
||
1795 |
lemma induct_smaller_elem_list: |
|
1796 |
shows "\<forall>r \<in> set list. asize r < Suc (sum_list (map asize list))" |
|
1797 |
apply(induct list) |
|
1798 |
apply simp+ |
|
1799 |
by fastforce |
|
1800 |
||
1801 |
lemma no0s_fltsbsimp: |
|
1802 |
shows "\<forall>r \<in> set (flts (map bsimp rs)). r \<noteq> AZERO" |
|
1803 |
oops |
|
1804 |
||
1805 |
lemma flts_all0s: |
|
1806 |
shows "\<forall>r \<in> set rs. r = AZERO \<Longrightarrow> flts rs = []" |
|
1807 |
apply(induct rs) |
|
1808 |
apply simp+ |
|
1809 |
done |
|
1810 |
||
1811 |
||
1812 |
||
1813 |
||
1814 |
||
1815 |
lemma good1: |
|
1816 |
shows "agood (bsimp a) \<or> bsimp a = AZERO" |
|
1817 |
apply(induct a taking: asize rule: measure_induct) |
|
1818 |
apply(case_tac x) |
|
1819 |
apply(simp) |
|
1820 |
apply(simp) |
|
1821 |
apply(simp) |
|
1822 |
prefer 3 |
|
1823 |
apply(simp) |
|
1824 |
prefer 2 |
|
1825 |
(* AALTs case *) |
|
1826 |
apply(simp only:) |
|
1827 |
apply(case_tac "x52") |
|
1828 |
apply(simp) |
|
1829 |
(* AALTs list at least one - case *) |
|
1830 |
apply(simp only: ) |
|
1831 |
apply(frule_tac x="a" in spec) |
|
1832 |
apply(drule mp) |
|
1833 |
apply(simp) |
|
1834 |
(* either first element is agood, or AZERO *) |
|
1835 |
apply(erule disjE) |
|
1836 |
prefer 2 |
|
1837 |
apply(simp) |
|
1838 |
(* in the AZERO case, the size is smaller *) |
|
1839 |
apply(drule_tac x="AALTs x51 list" in spec) |
|
1840 |
apply(drule mp) |
|
1841 |
apply(simp add: asize0) |
|
1842 |
apply(subst (asm) bsimp.simps) |
|
1843 |
apply(subst (asm) bsimp.simps) |
|
1844 |
apply(assumption) |
|
1845 |
(* in the agood case *) |
|
1846 |
apply(frule_tac x="AALTs x51 list" in spec) |
|
1847 |
apply(drule mp) |
|
1848 |
apply(simp add: asize0) |
|
1849 |
apply(erule disjE) |
|
1850 |
apply(rule disjI1) |
|
1851 |
apply(simp add: good0) |
|
1852 |
apply(subst good0) |
|
1853 |
using SizeBound3.flts3b distinctBy.elims apply force |
|
1854 |
apply(subgoal_tac " Ball (set ( (flts (bsimp a # map bsimp list)))) anonalt") |
|
1855 |
prefer 2 |
|
1856 |
apply (metis Cons_eq_map_conv SizeBound3.nn1b SizeBound3.nn1c ex_map_conv) |
|
1857 |
using dB_keeps_property apply presburger |
|
543 | 1858 |
|
493 | 1859 |
|
492 | 1860 |
using dB_does_more_job apply presburger |
1861 |
apply(subgoal_tac " Ball (set ( (flts (bsimp a # map bsimp list)) )) agood") |
|
1862 |
using dB_keeps_property apply presburger |
|
1863 |
apply(subgoal_tac "\<forall>r \<in> (set (bsimp a # map bsimp list)). (agood r) \<or> (r = AZERO)") |
|
1864 |
using SizeBound3.flts3 apply blast |
|
1865 |
||
1866 |
apply(subgoal_tac "\<forall>r \<in> set list. asize r < Suc (asize a + sum_list (map asize list))") |
|
1867 |
||
1868 |
apply simp |
|
1869 |
||
1870 |
apply(subgoal_tac "\<forall>r \<in> set list. asize r < Suc (sum_list (map asize list))") |
|
1871 |
||
1872 |
apply fastforce |
|
1873 |
using induct_smaller_elem_list apply blast |
|
1874 |
||
1875 |
||
1876 |
||
1877 |
||
1878 |
apply simp |
|
1879 |
apply(subgoal_tac "agood (bsimp a)") |
|
1880 |
prefer 2 |
|
1881 |
apply blast |
|
1882 |
apply(subgoal_tac "\<forall>r \<in> (set (map bsimp list)). r = AZERO \<or> agood r") |
|
1883 |
prefer 2 |
|
1884 |
apply (metis add_Suc_right image_iff induct_smaller_elem_list list.set_map trans_less_add2) |
|
1885 |
apply(subgoal_tac "\<not>(\<exists> r\<in>set (map bsimp list). agood r)") |
|
1886 |
prefer 2 |
|
1887 |
apply (metis SizeBound3.flts3 SizeBound3.flts3b distinctBy.elims empty_iff flts4_nogood list.set_intros(1)) |
|
1888 |
apply(subgoal_tac "\<forall>r \<in> set (map bsimp list). r = AZERO") |
|
1889 |
prefer 2 |
|
1890 |
apply blast |
|
1891 |
apply(subgoal_tac "flts (map bsimp list) = []") |
|
1892 |
prefer 2 |
|
1893 |
using flts_all0s apply presburger |
|
1894 |
apply (smt (verit, del_insts) SizeBound3.flts2 SizeBound3.good0 SizeBound3.k0 dB_does_more_job dB_keeps_property distinctBy.simps(1) flts.simps(1)) |
|
1895 |
apply(subgoal_tac "agood (bsimp x42) \<or> bsimp x42 = AZERO") |
|
1896 |
apply(subgoal_tac "agood (bsimp x43) \<or> bsimp x43 = AZERO") |
|
1897 |
apply(case_tac "bsimp x42 = AZERO") |
|
1898 |
apply simp |
|
1899 |
apply(case_tac "bsimp x43 = AZERO") |
|
1900 |
apply simp |
|
1901 |
apply(case_tac "\<exists>bs'. bsimp x42 = AONE bs'") |
|
1902 |
apply(erule exE) |
|
1903 |
apply simp |
|
1904 |
using good_fuse apply presburger |
|
1905 |
apply simp |
|
1906 |
apply(subgoal_tac "bsimp_ASEQ x41 (bsimp x42) (bsimp x43) = ASEQ x41 (bsimp x42) (bsimp x43)") |
|
1907 |
prefer 2 |
|
1908 |
using bsimp_ASEQ1 apply presburger |
|
1909 |
using SizeBound3.good_SEQ apply presburger |
|
1910 |
using asize.simps(5) less_add_Suc2 apply presburger |
|
1911 |
by simp |
|
1912 |
||
1913 |
||
1914 |
||
1915 |
||
1916 |
||
1917 |
||
1918 |
lemma good1a: |
|
493 | 1919 |
assumes "RL (rerase a) \<noteq> {}" |
492 | 1920 |
shows "agood (bsimp a)" |
1921 |
using good1 assms |
|
493 | 1922 |
using RL_bsimp_rerase by force |
492 | 1923 |
|
1924 |
||
1925 |
||
1926 |
lemma flts_append: |
|
1927 |
"flts (xs1 @ xs2) = flts xs1 @ flts xs2" |
|
1928 |
apply(induct xs1 arbitrary: xs2 rule: rev_induct) |
|
1929 |
apply(auto) |
|
1930 |
apply(case_tac xs) |
|
1931 |
apply(auto) |
|
1932 |
apply(case_tac x) |
|
1933 |
apply(auto) |
|
1934 |
apply(case_tac x) |
|
1935 |
apply(auto) |
|
1936 |
done |
|
1937 |
||
1938 |
lemma g1: |
|
1939 |
assumes "agood (bsimp_AALTs bs rs)" |
|
1940 |
shows "bsimp_AALTs bs rs = AALTs bs rs \<or> (\<exists>r. rs = [r] \<and> bsimp_AALTs bs [r] = fuse bs r)" |
|
1941 |
using assms |
|
1942 |
apply(induct rs arbitrary: bs) |
|
1943 |
apply(simp) |
|
1944 |
apply(case_tac rs) |
|
1945 |
apply(simp only:) |
|
1946 |
apply(simp) |
|
1947 |
apply(case_tac list) |
|
1948 |
apply(simp) |
|
1949 |
by simp |
|
1950 |
||
1951 |
lemma flts_0: |
|
1952 |
assumes "anonnested (AALTs bs rs)" |
|
1953 |
shows "\<forall>r \<in> set (flts rs). r \<noteq> AZERO" |
|
1954 |
using assms |
|
1955 |
apply(induct rs arbitrary: bs rule: flts.induct) |
|
1956 |
apply(simp) |
|
1957 |
apply(simp) |
|
1958 |
defer |
|
1959 |
apply(simp) |
|
1960 |
apply(simp) |
|
1961 |
apply(simp) |
|
1962 |
apply(simp) |
|
1963 |
apply(rule ballI) |
|
1964 |
apply(simp) |
|
1965 |
done |
|
1966 |
||
1967 |
lemma flts_0a: |
|
1968 |
assumes "anonnested (AALTs bs rs)" |
|
1969 |
shows "AZERO \<notin> set (flts rs)" |
|
1970 |
using assms |
|
1971 |
using flts_0 by blast |
|
1972 |
||
1973 |
lemma qqq1: |
|
1974 |
shows "AZERO \<notin> set (flts (map bsimp rs))" |
|
1975 |
by (metis ex_map_conv flts3 agood.simps(1) good1) |
|
1976 |
||
1977 |
||
1978 |
fun nonazero :: "arexp \<Rightarrow> bool" |
|
1979 |
where |
|
1980 |
"nonazero AZERO = False" |
|
1981 |
| "nonazero r = True" |
|
1982 |
||
1983 |
lemma flts_concat: |
|
1984 |
shows "flts rs = concat (map (\<lambda>r. flts [r]) rs)" |
|
1985 |
apply(induct rs) |
|
1986 |
apply(auto) |
|
1987 |
apply(subst k0) |
|
1988 |
apply(simp) |
|
1989 |
done |
|
1990 |
||
1991 |
lemma flts_single1: |
|
1992 |
assumes "anonalt r" "nonazero r" |
|
1993 |
shows "flts [r] = [r]" |
|
1994 |
using assms |
|
1995 |
apply(induct r) |
|
1996 |
apply(auto) |
|
1997 |
done |
|
1998 |
||
1999 |
lemma flts_qq: |
|
2000 |
assumes "\<forall>y. asize y < Suc (sum_list (map asize rs)) \<longrightarrow> agood y \<longrightarrow> bsimp y = y" |
|
2001 |
"\<forall>r'\<in>set rs. agood r' \<and> anonalt r'" |
|
2002 |
shows "flts (map bsimp rs) = rs" |
|
2003 |
using assms |
|
2004 |
apply(induct rs) |
|
2005 |
apply(simp) |
|
2006 |
apply(simp) |
|
2007 |
apply(subst k0) |
|
2008 |
apply(subgoal_tac "flts [bsimp a] = [a]") |
|
2009 |
prefer 2 |
|
2010 |
apply(drule_tac x="a" in spec) |
|
2011 |
apply(drule mp) |
|
2012 |
apply(simp) |
|
2013 |
apply(auto)[1] |
|
2014 |
using agood.simps(1) k0b apply blast |
|
2015 |
apply(auto)[1] |
|
2016 |
done |
|
2017 |
||
2018 |
||
2019 |
||
2020 |
||
2021 |
||
2022 |
||
2023 |
lemma nonalt_flts : shows |
|
2024 |
"\<forall>r \<in> set (flts (map bsimp rs)). r \<noteq> AZERO" |
|
2025 |
using SizeBound3.qqq1 by force |
|
493 | 2026 |
|
2027 |
||
2028 |
lemma rerase_map_bsimp: |
|
2029 |
assumes "\<And> r. r \<in> set x2a \<Longrightarrow> rerase (bsimp r) = (rsimp \<circ> rerase) r" |
|
2030 |
shows " map rerase (map bsimp x2a) = map (rsimp \<circ> rerase) x2a " |
|
2031 |
using assms |
|
2032 |
||
2033 |
apply(induct x2a) |
|
2034 |
apply simp |
|
2035 |
apply(subgoal_tac "a \<in> set (a # x2a)") |
|
2036 |
prefer 2 |
|
2037 |
apply (meson list.set_intros(1)) |
|
2038 |
apply(subgoal_tac "rerase (bsimp a ) = (rsimp \<circ> rerase) a") |
|
2039 |
apply simp |
|
2040 |
by presburger |
|
2041 |
||
2042 |
||
2043 |
||
2044 |
lemma rerase_flts: |
|
2045 |
shows "map rerase (flts rs) = rflts (map rerase rs)" |
|
2046 |
apply(induct rs) |
|
2047 |
apply simp+ |
|
2048 |
apply(case_tac a) |
|
2049 |
apply simp+ |
|
2050 |
apply(simp add: rerase_fuse) |
|
2051 |
apply simp |
|
2052 |
done |
|
2053 |
||
2054 |
lemma rerase_dB: |
|
2055 |
shows "map rerase (distinctBy rs rerase acc) = rdistinct (map rerase rs) acc" |
|
2056 |
apply(induct rs arbitrary: acc) |
|
2057 |
apply simp+ |
|
2058 |
done |
|
492 | 2059 |
|
2060 |
||
2061 |
||
2062 |
||
493 | 2063 |
lemma rerase_earlier_later_same: |
2064 |
assumes " \<And>r. r \<in> set x2a \<Longrightarrow> rerase (bsimp r) = rsimp (rerase r)" |
|
2065 |
shows " (map rerase (distinctBy (flts (map bsimp x2a)) rerase {})) = |
|
2066 |
(rdistinct (rflts (map (rsimp \<circ> rerase) x2a)) {})" |
|
2067 |
apply(subst rerase_dB) |
|
2068 |
apply(subst rerase_flts) |
|
2069 |
apply(subst rerase_map_bsimp) |
|
2070 |
apply auto |
|
2071 |
using assms |
|
2072 |
apply simp |
|
2073 |
done |
|
492 | 2074 |
|
2075 |
||
2076 |
||
2077 |
lemma simp_rerase: |
|
2078 |
shows "rerase (bsimp a) = rsimp (rerase a)" |
|
2079 |
apply(induct a) |
|
2080 |
apply simp+ |
|
2081 |
using rerase_bsimp_ASEQ apply presburger |
|
2082 |
apply simp |
|
2083 |
apply(subst rerase_bsimp_AALTs) |
|
493 | 2084 |
apply(subgoal_tac "map rerase (distinctBy (flts (map bsimp x2a)) rerase {}) = rdistinct (rflts (map (rsimp \<circ> rerase) x2a)) {}") |
2085 |
apply simp |
|
2086 |
using rerase_earlier_later_same apply presburger |
|
492 | 2087 |
apply simp |
493 | 2088 |
done |
2089 |
||
2090 |
||
492 | 2091 |
|
2092 |
lemma rders_simp_size: |
|
2093 |
shows " rders_simp (rerase r) s = rerase (bders_simp r s)" |
|
2094 |
apply(induct s rule: rev_induct) |
|
2095 |
apply simp |
|
2096 |
apply(subst rders_simp_append) |
|
2097 |
apply(subst bders_simp_append) |
|
2098 |
apply(subgoal_tac "rders_simp (rerase r ) xs = rerase (bders_simp r xs)") |
|
2099 |
apply(simp only:) |
|
2100 |
apply simp |
|
2101 |
apply (simp add: der_rerase2 simp_rerase) |
|
2102 |
by simp |
|
2103 |
||
2104 |
||
2105 |
||
2106 |
||
2107 |
corollary aders_simp_finiteness: |
|
2108 |
assumes "\<exists>N. \<forall>s. rsize (rders_simp (rerase r) s) \<le> N" |
|
2109 |
shows " \<exists>N. \<forall>s. asize (bders_simp r s) \<le> N" |
|
2110 |
using assms |
|
2111 |
apply(subgoal_tac "\<forall>s. asize (bders_simp r s) = rsize (rerase (bders_simp r s))") |
|
2112 |
apply(erule exE) |
|
2113 |
apply(rule_tac x = "N" in exI) |
|
2114 |
using rders_simp_size apply auto[1] |
|
2115 |
using asize_rsize by auto |
|
2116 |
||
2117 |
theorem annotated_size_bound: |
|
2118 |
shows "\<exists>N. \<forall>s. asize (bders_simp r s) \<le> N" |
|
2119 |
apply(insert aders_simp_finiteness) |
|
2120 |
by (simp add: rders_simp_bounded) |
|
2121 |
||
543 | 2122 |
unused_thms |
2123 |
||
2124 |
||
365 | 2125 |
end |