author | Christian Urban <urbanc@in.tum.de> |
Mon, 11 Feb 2019 14:36:23 +0000 | |
changeset 308 | 496a37d816e9 |
parent 307 | ee1caac29bb2 |
child 313 | 3b8e3a156200 |
permissions | -rw-r--r-- |
148
702ed601349b
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1 |
|
702ed601349b
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2 |
theory Sulzmann |
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
3 |
imports "Lexer" |
148
702ed601349b
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4 |
begin |
702ed601349b
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5 |
|
154
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
6 |
section {* Bit-Encodings *} |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
7 |
|
289 | 8 |
datatype bit = Z | S |
9 |
||
154
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
10 |
fun |
289 | 11 |
code :: "val \<Rightarrow> bit list" |
154
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
12 |
where |
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
13 |
"code Void = []" |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
14 |
| "code (Char c) = []" |
289 | 15 |
| "code (Left v) = Z # (code v)" |
16 |
| "code (Right v) = S # (code v)" |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
17 |
| "code (Seq v1 v2) = (code v1) @ (code v2)" |
289 | 18 |
| "code (Stars []) = [S]" |
19 |
| "code (Stars (v # vs)) = (Z # code v) @ code (Stars vs)" |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
20 |
|
154
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
21 |
|
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
22 |
fun |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
23 |
Stars_add :: "val \<Rightarrow> val \<Rightarrow> val" |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
24 |
where |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
25 |
"Stars_add v (Stars vs) = Stars (v # vs)" |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
26 |
|
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
27 |
function |
289 | 28 |
decode' :: "bit list \<Rightarrow> rexp \<Rightarrow> (val * bit list)" |
154
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
29 |
where |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
30 |
"decode' ds ZERO = (Void, [])" |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
31 |
| "decode' ds ONE = (Void, ds)" |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
32 |
| "decode' ds (CHAR d) = (Char d, ds)" |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
33 |
| "decode' [] (ALT r1 r2) = (Void, [])" |
289 | 34 |
| "decode' (Z # ds) (ALT r1 r2) = (let (v, ds') = decode' ds r1 in (Left v, ds'))" |
35 |
| "decode' (S # ds) (ALT r1 r2) = (let (v, ds') = decode' ds r2 in (Right v, ds'))" |
|
154
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
36 |
| "decode' ds (SEQ r1 r2) = (let (v1, ds') = decode' ds r1 in |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
37 |
let (v2, ds'') = decode' ds' r2 in (Seq v1 v2, ds''))" |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
38 |
| "decode' [] (STAR r) = (Void, [])" |
289 | 39 |
| "decode' (S # ds) (STAR r) = (Stars [], ds)" |
40 |
| "decode' (Z # ds) (STAR r) = (let (v, ds') = decode' ds r in |
|
154
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
41 |
let (vs, ds'') = decode' ds' (STAR r) |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
42 |
in (Stars_add v vs, ds''))" |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
43 |
by pat_completeness auto |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
44 |
|
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
45 |
lemma decode'_smaller: |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
46 |
assumes "decode'_dom (ds, r)" |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
47 |
shows "length (snd (decode' ds r)) \<le> length ds" |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
48 |
using assms |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
49 |
apply(induct ds r) |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
50 |
apply(auto simp add: decode'.psimps split: prod.split) |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
51 |
using dual_order.trans apply blast |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
52 |
by (meson dual_order.trans le_SucI) |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
53 |
|
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
54 |
termination "decode'" |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
55 |
apply(relation "inv_image (measure(%cs. size cs) <*lex*> measure(%s. size s)) (%(ds,r). (r,ds))") |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
56 |
apply(auto dest!: decode'_smaller) |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
57 |
by (metis less_Suc_eq_le snd_conv) |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
58 |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
59 |
definition |
289 | 60 |
decode :: "bit list \<Rightarrow> rexp \<Rightarrow> val option" |
154
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
61 |
where |
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
62 |
"decode ds r \<equiv> (let (v, ds') = decode' ds r |
154
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
63 |
in (if ds' = [] then Some v else None))" |
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
64 |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
65 |
lemma decode'_code_Stars: |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
66 |
assumes "\<forall>v\<in>set vs. \<Turnstile> v : r \<and> (\<forall>x. decode' (code v @ x) r = (v, x)) \<and> flat v \<noteq> []" |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
67 |
shows "decode' (code (Stars vs) @ ds) (STAR r) = (Stars vs, ds)" |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
68 |
using assms |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
69 |
apply(induct vs) |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
70 |
apply(auto) |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
71 |
done |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
72 |
|
154
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
73 |
lemma decode'_code: |
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
74 |
assumes "\<Turnstile> v : r" |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
75 |
shows "decode' ((code v) @ ds) r = (v, ds)" |
154
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
76 |
using assms |
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
77 |
apply(induct v r arbitrary: ds) |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
78 |
apply(auto) |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
79 |
using decode'_code_Stars by blast |
154
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
80 |
|
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
81 |
lemma decode_code: |
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
82 |
assumes "\<Turnstile> v : r" |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
83 |
shows "decode (code v) r = Some v" |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
84 |
using assms unfolding decode_def |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
85 |
by (smt append_Nil2 decode'_code old.prod.case) |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
86 |
|
154
2de3cf684ba0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
148
diff
changeset
|
87 |
|
159
940530087f30
updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
88 |
datatype arexp = |
940530087f30
updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
89 |
AZERO |
289 | 90 |
| AONE "bit list" |
91 |
| ACHAR "bit list" char |
|
92 |
| ASEQ "bit list" arexp arexp |
|
93 |
| AALT "bit list" arexp arexp |
|
94 |
| ASTAR "bit list" arexp |
|
159
940530087f30
updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
95 |
|
289 | 96 |
fun fuse :: "bit list \<Rightarrow> arexp \<Rightarrow> arexp" where |
159
940530087f30
updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
97 |
"fuse bs AZERO = AZERO" |
940530087f30
updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
98 |
| "fuse bs (AONE cs) = AONE (bs @ cs)" |
940530087f30
updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
99 |
| "fuse bs (ACHAR cs c) = ACHAR (bs @ cs) c" |
940530087f30
updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
100 |
| "fuse bs (AALT cs r1 r2) = AALT (bs @ cs) r1 r2" |
940530087f30
updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
101 |
| "fuse bs (ASEQ cs r1 r2) = ASEQ (bs @ cs) r1 r2" |
940530087f30
updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
102 |
| "fuse bs (ASTAR cs r) = ASTAR (bs @ cs) r" |
940530087f30
updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
103 |
|
289 | 104 |
fun intern :: "rexp \<Rightarrow> arexp" where |
105 |
"intern ZERO = AZERO" |
|
106 |
| "intern ONE = AONE []" |
|
107 |
| "intern (CHAR c) = ACHAR [] c" |
|
108 |
| "intern (ALT r1 r2) = AALT [] (fuse [Z] (intern r1)) |
|
295 | 109 |
(fuse [S] (intern r2))" |
289 | 110 |
| "intern (SEQ r1 r2) = ASEQ [] (intern r1) (intern r2)" |
111 |
| "intern (STAR r) = ASTAR [] (intern r)" |
|
159
940530087f30
updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
112 |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
113 |
|
289 | 114 |
fun retrieve :: "arexp \<Rightarrow> val \<Rightarrow> bit list" where |
159
940530087f30
updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
115 |
"retrieve (AONE bs) Void = bs" |
940530087f30
updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
116 |
| "retrieve (ACHAR bs c) (Char d) = bs" |
940530087f30
updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
117 |
| "retrieve (AALT bs r1 r2) (Left v) = bs @ retrieve r1 v" |
940530087f30
updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
118 |
| "retrieve (AALT bs r1 r2) (Right v) = bs @ retrieve r2 v" |
940530087f30
updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
119 |
| "retrieve (ASEQ bs r1 r2) (Seq v1 v2) = bs @ retrieve r1 v1 @ retrieve r2 v2" |
289 | 120 |
| "retrieve (ASTAR bs r) (Stars []) = bs @ [S]" |
159
940530087f30
updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
121 |
| "retrieve (ASTAR bs r) (Stars (v#vs)) = |
289 | 122 |
bs @ [Z] @ retrieve r v @ retrieve (ASTAR [] r) (Stars vs)" |
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
123 |
|
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
124 |
fun |
289 | 125 |
erase :: "arexp \<Rightarrow> rexp" |
126 |
where |
|
127 |
"erase AZERO = ZERO" |
|
128 |
| "erase (AONE _) = ONE" |
|
129 |
| "erase (ACHAR _ c) = CHAR c" |
|
130 |
| "erase (AALT _ r1 r2) = ALT (erase r1) (erase r2)" |
|
131 |
| "erase (ASEQ _ r1 r2) = SEQ (erase r1) (erase r2)" |
|
132 |
| "erase (ASTAR _ r) = STAR (erase r)" |
|
133 |
||
134 |
fun |
|
135 |
bnullable :: "arexp \<Rightarrow> bool" |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
136 |
where |
289 | 137 |
"bnullable (AZERO) = False" |
138 |
| "bnullable (AONE bs) = True" |
|
139 |
| "bnullable (ACHAR bs c) = False" |
|
140 |
| "bnullable (AALT bs r1 r2) = (bnullable r1 \<or> bnullable r2)" |
|
141 |
| "bnullable (ASEQ bs r1 r2) = (bnullable r1 \<and> bnullable r2)" |
|
142 |
| "bnullable (ASTAR bs r) = True" |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
143 |
|
289 | 144 |
fun |
145 |
bmkeps :: "arexp \<Rightarrow> bit list" |
|
146 |
where |
|
147 |
"bmkeps(AONE bs) = bs" |
|
148 |
| "bmkeps(ASEQ bs r1 r2) = bs @ (bmkeps r1) @ (bmkeps r2)" |
|
149 |
| "bmkeps(AALT bs r1 r2) = (if bnullable(r1) then bs @ (bmkeps r1) else bs @ (bmkeps r2))" |
|
150 |
| "bmkeps(ASTAR bs r) = bs @ [S]" |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
151 |
|
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
152 |
|
159
940530087f30
updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
153 |
fun |
289 | 154 |
bder :: "char \<Rightarrow> arexp \<Rightarrow> arexp" |
159
940530087f30
updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
155 |
where |
289 | 156 |
"bder c (AZERO) = AZERO" |
157 |
| "bder c (AONE bs) = AZERO" |
|
158 |
| "bder c (ACHAR bs d) = (if c = d then AONE bs else AZERO)" |
|
159 |
| "bder c (AALT bs r1 r2) = AALT bs (bder c r1) (bder c r2)" |
|
160 |
| "bder c (ASEQ bs r1 r2) = |
|
161 |
(if bnullable r1 |
|
162 |
then AALT bs (ASEQ [] (bder c r1) r2) (fuse (bmkeps r1) (bder c r2)) |
|
163 |
else ASEQ bs (bder c r1) r2)" |
|
164 |
| "bder c (ASTAR bs r) = ASEQ bs (fuse [Z] (bder c r)) (ASTAR [] r)" |
|
159
940530087f30
updated programs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
165 |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
166 |
|
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
167 |
fun |
289 | 168 |
bders :: "arexp \<Rightarrow> string \<Rightarrow> arexp" |
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
169 |
where |
289 | 170 |
"bders r [] = r" |
171 |
| "bders r (c#s) = bders (bder c r) s" |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
172 |
|
289 | 173 |
lemma bders_append: |
174 |
"bders r (s1 @ s2) = bders (bders r s1) s2" |
|
287 | 175 |
apply(induct s1 arbitrary: r s2) |
176 |
apply(simp_all) |
|
177 |
done |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
178 |
|
289 | 179 |
lemma bnullable_correctness: |
180 |
shows "nullable (erase r) = bnullable r" |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
181 |
apply(induct r) |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
182 |
apply(simp_all) |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
183 |
done |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
184 |
|
289 | 185 |
lemma erase_fuse: |
186 |
shows "erase (fuse bs r) = erase r" |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
187 |
apply(induct r) |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
188 |
apply(simp_all) |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
189 |
done |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
190 |
|
289 | 191 |
lemma erase_intern[simp]: |
192 |
shows "erase (intern r) = r" |
|
287 | 193 |
apply(induct r) |
289 | 194 |
apply(simp_all add: erase_fuse) |
287 | 195 |
done |
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
196 |
|
289 | 197 |
lemma erase_bder[simp]: |
198 |
shows "erase (bder a r) = der a (erase r)" |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
199 |
apply(induct r) |
289 | 200 |
apply(simp_all add: erase_fuse bnullable_correctness) |
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
201 |
done |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
202 |
|
289 | 203 |
lemma erase_bders[simp]: |
204 |
shows "erase (bders r s) = ders s (erase r)" |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
205 |
apply(induct s arbitrary: r ) |
289 | 206 |
apply(simp_all) |
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
207 |
done |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
208 |
|
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
209 |
lemma retrieve_encode_STARS: |
289 | 210 |
assumes "\<forall>v\<in>set vs. \<Turnstile> v : r \<and> code v = retrieve (intern r) v" |
211 |
shows "code (Stars vs) = retrieve (ASTAR [] (intern r)) (Stars vs)" |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
212 |
using assms |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
213 |
apply(induct vs) |
289 | 214 |
apply(simp_all) |
215 |
done |
|
216 |
||
217 |
lemma retrieve_fuse2: |
|
218 |
assumes "\<Turnstile> v : (erase r)" |
|
219 |
shows "retrieve (fuse bs r) v = bs @ retrieve r v" |
|
220 |
using assms |
|
221 |
apply(induct r arbitrary: v bs) |
|
222 |
using retrieve_encode_STARS |
|
223 |
apply(auto elim!: Prf_elims) |
|
224 |
apply(case_tac vs) |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
225 |
apply(simp) |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
226 |
apply(simp) |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
227 |
done |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
228 |
|
289 | 229 |
lemma retrieve_fuse: |
230 |
assumes "\<Turnstile> v : r" |
|
231 |
shows "retrieve (fuse bs (intern r)) v = bs @ retrieve (intern r) v" |
|
232 |
using assms |
|
233 |
by (simp_all add: retrieve_fuse2) |
|
234 |
||
235 |
||
236 |
lemma retrieve_code: |
|
237 |
assumes "\<Turnstile> v : r" |
|
238 |
shows "code v = retrieve (intern r) v" |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
239 |
using assms |
289 | 240 |
apply(induct v r) |
241 |
apply(simp_all add: retrieve_fuse retrieve_encode_STARS) |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
242 |
done |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
243 |
|
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
244 |
|
289 | 245 |
lemma bmkeps_retrieve: |
246 |
assumes "nullable (erase r)" |
|
247 |
shows "bmkeps r = retrieve r (mkeps (erase r))" |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
248 |
using assms |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
249 |
apply(induct r) |
289 | 250 |
apply(auto simp add: bnullable_correctness) |
251 |
done |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
252 |
|
289 | 253 |
lemma bder_retrieve: |
254 |
assumes "\<Turnstile> v : der c (erase r)" |
|
255 |
shows "retrieve (bder c r) v = retrieve r (injval (erase r) c v)" |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
256 |
using assms |
289 | 257 |
apply(induct r arbitrary: v) |
258 |
apply(auto elim!: Prf_elims simp add: retrieve_fuse2 bnullable_correctness bmkeps_retrieve) |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
259 |
done |
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
260 |
|
289 | 261 |
lemma MAIN_decode: |
262 |
assumes "\<Turnstile> v : ders s r" |
|
263 |
shows "Some (flex r id s v) = decode (retrieve (bders (intern r) s) v) r" |
|
264 |
using assms |
|
265 |
proof (induct s arbitrary: v rule: rev_induct) |
|
266 |
case Nil |
|
267 |
have "\<Turnstile> v : ders [] r" by fact |
|
268 |
then have "\<Turnstile> v : r" by simp |
|
269 |
then have "Some v = decode (retrieve (intern r) v) r" |
|
270 |
using decode_code retrieve_code by auto |
|
271 |
then show "Some (flex r id [] v) = decode (retrieve (bders (intern r) []) v) r" |
|
272 |
by simp |
|
273 |
next |
|
274 |
case (snoc c s v) |
|
275 |
have IH: "\<And>v. \<Turnstile> v : ders s r \<Longrightarrow> |
|
276 |
Some (flex r id s v) = decode (retrieve (bders (intern r) s) v) r" by fact |
|
277 |
have asm: "\<Turnstile> v : ders (s @ [c]) r" by fact |
|
278 |
then have asm2: "\<Turnstile> injval (ders s r) c v : ders s r" |
|
279 |
by(simp add: Prf_injval ders_append) |
|
280 |
have "Some (flex r id (s @ [c]) v) = Some (flex r id s (injval (ders s r) c v))" |
|
281 |
by (simp add: flex_append) |
|
282 |
also have "... = decode (retrieve (bders (intern r) s) (injval (ders s r) c v)) r" |
|
283 |
using asm2 IH by simp |
|
284 |
also have "... = decode (retrieve (bder c (bders (intern r) s)) v) r" |
|
285 |
using asm by(simp_all add: bder_retrieve ders_append) |
|
286 |
finally show "Some (flex r id (s @ [c]) v) = |
|
287 |
decode (retrieve (bders (intern r) (s @ [c])) v) r" by (simp add: bders_append) |
|
288 |
qed |
|
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
289 |
|
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
290 |
|
289 | 291 |
definition blexer where |
292 |
"blexer r s \<equiv> if bnullable (bders (intern r) s) then |
|
293 |
decode (bmkeps (bders (intern r) s)) r else None" |
|
294 |
||
295 |
lemma blexer_correctness: |
|
296 |
shows "blexer r s = lexer r s" |
|
297 |
proof - |
|
298 |
{ define bds where "bds \<equiv> bders (intern r) s" |
|
299 |
define ds where "ds \<equiv> ders s r" |
|
300 |
assume asm: "nullable ds" |
|
301 |
have era: "erase bds = ds" |
|
302 |
unfolding ds_def bds_def by simp |
|
303 |
have mke: "\<Turnstile> mkeps ds : ds" |
|
304 |
using asm by (simp add: mkeps_nullable) |
|
305 |
have "decode (bmkeps bds) r = decode (retrieve bds (mkeps ds)) r" |
|
306 |
using bmkeps_retrieve |
|
307 |
using asm era by (simp add: bmkeps_retrieve) |
|
308 |
also have "... = Some (flex r id s (mkeps ds))" |
|
309 |
using mke by (simp_all add: MAIN_decode ds_def bds_def) |
|
310 |
finally have "decode (bmkeps bds) r = Some (flex r id s (mkeps ds))" |
|
311 |
unfolding bds_def ds_def . |
|
312 |
} |
|
313 |
then show "blexer r s = lexer r s" |
|
314 |
unfolding blexer_def lexer_flex |
|
293 | 315 |
apply(subst bnullable_correctness[symmetric]) |
316 |
apply(simp) |
|
317 |
done |
|
289 | 318 |
qed |
286
804fbb227568
added proof for bitcoded algorithm
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
319 |
|
295 | 320 |
|
321 |
||
148
702ed601349b
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
322 |
end |