author | Chengsong |
Mon, 10 Jul 2023 14:32:48 +0100 | |
changeset 663 | 0d1e68268d0f |
parent 397 | e1b74d618f1b |
permissions | -rw-r--r-- |
273 | 1 |
|
2 |
theory SpecExt |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
3 |
imports Main "HOL-Library.Sublist" (*"~~/src/HOL/Library/Sublist"*) |
273 | 4 |
begin |
5 |
||
6 |
section {* Sequential Composition of Languages *} |
|
7 |
||
8 |
definition |
|
9 |
Sequ :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ;; _" [100,100] 100) |
|
10 |
where |
|
11 |
"A ;; B = {s1 @ s2 | s1 s2. s1 \<in> A \<and> s2 \<in> B}" |
|
12 |
||
13 |
text {* Two Simple Properties about Sequential Composition *} |
|
14 |
||
15 |
lemma Sequ_empty_string [simp]: |
|
16 |
shows "A ;; {[]} = A" |
|
17 |
and "{[]} ;; A = A" |
|
18 |
by (simp_all add: Sequ_def) |
|
19 |
||
20 |
lemma Sequ_empty [simp]: |
|
21 |
shows "A ;; {} = {}" |
|
22 |
and "{} ;; A = {}" |
|
23 |
by (simp_all add: Sequ_def) |
|
24 |
||
25 |
lemma Sequ_assoc: |
|
26 |
shows "(A ;; B) ;; C = A ;; (B ;; C)" |
|
27 |
apply(auto simp add: Sequ_def) |
|
28 |
apply blast |
|
29 |
by (metis append_assoc) |
|
30 |
||
31 |
lemma Sequ_Union_in: |
|
32 |
shows "(A ;; (\<Union>x\<in> B. C x)) = (\<Union>x\<in> B. A ;; C x)" |
|
33 |
by (auto simp add: Sequ_def) |
|
34 |
||
35 |
section {* Semantic Derivative (Left Quotient) of Languages *} |
|
36 |
||
37 |
definition |
|
38 |
Der :: "char \<Rightarrow> string set \<Rightarrow> string set" |
|
39 |
where |
|
40 |
"Der c A \<equiv> {s. c # s \<in> A}" |
|
41 |
||
42 |
definition |
|
43 |
Ders :: "string \<Rightarrow> string set \<Rightarrow> string set" |
|
44 |
where |
|
45 |
"Ders s A \<equiv> {s'. s @ s' \<in> A}" |
|
46 |
||
47 |
lemma Der_null [simp]: |
|
48 |
shows "Der c {} = {}" |
|
49 |
unfolding Der_def |
|
50 |
by auto |
|
51 |
||
52 |
lemma Der_empty [simp]: |
|
53 |
shows "Der c {[]} = {}" |
|
54 |
unfolding Der_def |
|
55 |
by auto |
|
56 |
||
57 |
lemma Der_char [simp]: |
|
58 |
shows "Der c {[d]} = (if c = d then {[]} else {})" |
|
59 |
unfolding Der_def |
|
60 |
by auto |
|
61 |
||
62 |
lemma Der_union [simp]: |
|
63 |
shows "Der c (A \<union> B) = Der c A \<union> Der c B" |
|
64 |
unfolding Der_def |
|
65 |
by auto |
|
66 |
||
67 |
lemma Der_UNION [simp]: |
|
68 |
shows "Der c (\<Union>x\<in>A. B x) = (\<Union>x\<in>A. Der c (B x))" |
|
69 |
by (auto simp add: Der_def) |
|
70 |
||
71 |
lemma Der_Sequ [simp]: |
|
72 |
shows "Der c (A ;; B) = (Der c A) ;; B \<union> (if [] \<in> A then Der c B else {})" |
|
73 |
unfolding Der_def Sequ_def |
|
280 | 74 |
by (auto simp add: Cons_eq_append_conv) |
273 | 75 |
|
76 |
||
77 |
section {* Kleene Star for Languages *} |
|
78 |
||
79 |
inductive_set |
|
80 |
Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102) |
|
81 |
for A :: "string set" |
|
82 |
where |
|
83 |
start[intro]: "[] \<in> A\<star>" |
|
84 |
| step[intro]: "\<lbrakk>s1 \<in> A; s2 \<in> A\<star>\<rbrakk> \<Longrightarrow> s1 @ s2 \<in> A\<star>" |
|
85 |
||
86 |
(* Arden's lemma *) |
|
87 |
||
88 |
lemma Star_cases: |
|
89 |
shows "A\<star> = {[]} \<union> A ;; A\<star>" |
|
90 |
unfolding Sequ_def |
|
91 |
by (auto) (metis Star.simps) |
|
92 |
||
93 |
lemma Star_decomp: |
|
94 |
assumes "c # x \<in> A\<star>" |
|
95 |
shows "\<exists>s1 s2. x = s1 @ s2 \<and> c # s1 \<in> A \<and> s2 \<in> A\<star>" |
|
96 |
using assms |
|
97 |
by (induct x\<equiv>"c # x" rule: Star.induct) |
|
98 |
(auto simp add: append_eq_Cons_conv) |
|
99 |
||
100 |
lemma Star_Der_Sequ: |
|
101 |
shows "Der c (A\<star>) \<subseteq> (Der c A) ;; A\<star>" |
|
102 |
unfolding Der_def Sequ_def |
|
103 |
by(auto simp add: Star_decomp) |
|
104 |
||
105 |
||
106 |
lemma Der_star [simp]: |
|
107 |
shows "Der c (A\<star>) = (Der c A) ;; A\<star>" |
|
108 |
proof - |
|
109 |
have "Der c (A\<star>) = Der c ({[]} \<union> A ;; A\<star>)" |
|
110 |
by (simp only: Star_cases[symmetric]) |
|
111 |
also have "... = Der c (A ;; A\<star>)" |
|
112 |
by (simp only: Der_union Der_empty) (simp) |
|
113 |
also have "... = (Der c A) ;; A\<star> \<union> (if [] \<in> A then Der c (A\<star>) else {})" |
|
114 |
by simp |
|
115 |
also have "... = (Der c A) ;; A\<star>" |
|
116 |
using Star_Der_Sequ by auto |
|
117 |
finally show "Der c (A\<star>) = (Der c A) ;; A\<star>" . |
|
118 |
qed |
|
119 |
||
120 |
section {* Power operation for Sets *} |
|
121 |
||
122 |
fun |
|
123 |
Pow :: "string set \<Rightarrow> nat \<Rightarrow> string set" ("_ \<up> _" [101, 102] 101) |
|
124 |
where |
|
125 |
"A \<up> 0 = {[]}" |
|
126 |
| "A \<up> (Suc n) = A ;; (A \<up> n)" |
|
127 |
||
128 |
lemma Pow_empty [simp]: |
|
129 |
shows "[] \<in> A \<up> n \<longleftrightarrow> (n = 0 \<or> [] \<in> A)" |
|
130 |
by(induct n) (auto simp add: Sequ_def) |
|
131 |
||
132 |
lemma Pow_Suc_rev: |
|
133 |
"A \<up> (Suc n) = (A \<up> n) ;; A" |
|
134 |
apply(induct n arbitrary: A) |
|
135 |
apply(simp_all) |
|
136 |
by (metis Sequ_assoc) |
|
137 |
||
138 |
||
139 |
lemma Pow_decomp: |
|
140 |
assumes "c # x \<in> A \<up> n" |
|
141 |
shows "\<exists>s1 s2. x = s1 @ s2 \<and> c # s1 \<in> A \<and> s2 \<in> A \<up> (n - 1)" |
|
142 |
using assms |
|
143 |
apply(induct n) |
|
144 |
apply(auto simp add: Cons_eq_append_conv Sequ_def) |
|
145 |
apply(case_tac n) |
|
146 |
apply(auto simp add: Sequ_def) |
|
147 |
apply(blast) |
|
148 |
done |
|
149 |
||
150 |
lemma Star_Pow: |
|
151 |
assumes "s \<in> A\<star>" |
|
152 |
shows "\<exists>n. s \<in> A \<up> n" |
|
153 |
using assms |
|
154 |
apply(induct) |
|
155 |
apply(auto) |
|
156 |
apply(rule_tac x="Suc n" in exI) |
|
157 |
apply(auto simp add: Sequ_def) |
|
158 |
done |
|
159 |
||
160 |
lemma Pow_Star: |
|
161 |
assumes "s \<in> A \<up> n" |
|
162 |
shows "s \<in> A\<star>" |
|
163 |
using assms |
|
164 |
apply(induct n arbitrary: s) |
|
165 |
apply(auto simp add: Sequ_def) |
|
359 | 166 |
done |
167 |
||
273 | 168 |
lemma Der_Pow_0: |
169 |
shows "Der c (A \<up> 0) = {}" |
|
170 |
by(simp add: Der_def) |
|
171 |
||
172 |
lemma Der_Pow_Suc: |
|
173 |
shows "Der c (A \<up> (Suc n)) = (Der c A) ;; (A \<up> n)" |
|
174 |
unfolding Der_def Sequ_def |
|
175 |
apply(auto simp add: Cons_eq_append_conv Sequ_def dest!: Pow_decomp) |
|
176 |
apply(case_tac n) |
|
177 |
apply(force simp add: Sequ_def)+ |
|
178 |
done |
|
179 |
||
180 |
lemma Der_Pow [simp]: |
|
181 |
shows "Der c (A \<up> n) = (if n = 0 then {} else (Der c A) ;; (A \<up> (n - 1)))" |
|
182 |
apply(case_tac n) |
|
183 |
apply(simp_all del: Pow.simps add: Der_Pow_0 Der_Pow_Suc) |
|
184 |
done |
|
185 |
||
186 |
lemma Der_Pow_Sequ [simp]: |
|
187 |
shows "Der c (A ;; A \<up> n) = (Der c A) ;; (A \<up> n)" |
|
188 |
by (simp only: Pow.simps[symmetric] Der_Pow) (simp) |
|
189 |
||
190 |
||
191 |
lemma Pow_Sequ_Un: |
|
192 |
assumes "0 < x" |
|
193 |
shows "(\<Union>n \<in> {..x}. (A \<up> n)) = ({[]} \<union> (\<Union>n \<in> {..x - Suc 0}. A ;; (A \<up> n)))" |
|
194 |
using assms |
|
195 |
apply(auto simp add: Sequ_def) |
|
196 |
apply(smt Pow.elims Sequ_def Suc_le_mono Suc_pred atMost_iff empty_iff insert_iff mem_Collect_eq) |
|
197 |
apply(rule_tac x="Suc xa" in bexI) |
|
198 |
apply(auto simp add: Sequ_def) |
|
199 |
done |
|
200 |
||
201 |
lemma Pow_Sequ_Un2: |
|
202 |
assumes "0 < x" |
|
203 |
shows "(\<Union>n \<in> {x..}. (A \<up> n)) = (\<Union>n \<in> {x - Suc 0..}. A ;; (A \<up> n))" |
|
204 |
using assms |
|
205 |
apply(auto simp add: Sequ_def) |
|
206 |
apply(case_tac n) |
|
207 |
apply(auto simp add: Sequ_def) |
|
208 |
apply fastforce |
|
209 |
apply(case_tac x) |
|
210 |
apply(auto) |
|
211 |
apply(rule_tac x="Suc xa" in bexI) |
|
212 |
apply(auto simp add: Sequ_def) |
|
213 |
done |
|
214 |
||
215 |
section {* Regular Expressions *} |
|
216 |
||
217 |
datatype rexp = |
|
218 |
ZERO |
|
219 |
| ONE |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
220 |
| CH char |
273 | 221 |
| SEQ rexp rexp |
222 |
| ALT rexp rexp |
|
223 |
| STAR rexp |
|
224 |
| UPNTIMES rexp nat |
|
225 |
| NTIMES rexp nat |
|
226 |
| FROMNTIMES rexp nat |
|
227 |
| NMTIMES rexp nat nat |
|
359 | 228 |
| NOT rexp |
273 | 229 |
|
230 |
section {* Semantics of Regular Expressions *} |
|
231 |
||
232 |
fun |
|
233 |
L :: "rexp \<Rightarrow> string set" |
|
234 |
where |
|
235 |
"L (ZERO) = {}" |
|
236 |
| "L (ONE) = {[]}" |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
237 |
| "L (CH c) = {[c]}" |
273 | 238 |
| "L (SEQ r1 r2) = (L r1) ;; (L r2)" |
239 |
| "L (ALT r1 r2) = (L r1) \<union> (L r2)" |
|
240 |
| "L (STAR r) = (L r)\<star>" |
|
280 | 241 |
| "L (UPNTIMES r n) = (\<Union>i\<in>{..n} . (L r) \<up> i)" |
273 | 242 |
| "L (NTIMES r n) = (L r) \<up> n" |
280 | 243 |
| "L (FROMNTIMES r n) = (\<Union>i\<in>{n..} . (L r) \<up> i)" |
273 | 244 |
| "L (NMTIMES r n m) = (\<Union>i\<in>{n..m} . (L r) \<up> i)" |
359 | 245 |
| "L (NOT r) = ((UNIV:: string set) - L r)" |
273 | 246 |
|
247 |
section {* Nullable, Derivatives *} |
|
248 |
||
249 |
fun |
|
250 |
nullable :: "rexp \<Rightarrow> bool" |
|
251 |
where |
|
252 |
"nullable (ZERO) = False" |
|
253 |
| "nullable (ONE) = True" |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
254 |
| "nullable (CH c) = False" |
273 | 255 |
| "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)" |
256 |
| "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)" |
|
257 |
| "nullable (STAR r) = True" |
|
258 |
| "nullable (UPNTIMES r n) = True" |
|
259 |
| "nullable (NTIMES r n) = (if n = 0 then True else nullable r)" |
|
260 |
| "nullable (FROMNTIMES r n) = (if n = 0 then True else nullable r)" |
|
261 |
| "nullable (NMTIMES r n m) = (if m < n then False else (if n = 0 then True else nullable r))" |
|
359 | 262 |
| "nullable (NOT r) = (\<not> nullable r)" |
273 | 263 |
|
264 |
fun |
|
265 |
der :: "char \<Rightarrow> rexp \<Rightarrow> rexp" |
|
266 |
where |
|
267 |
"der c (ZERO) = ZERO" |
|
268 |
| "der c (ONE) = ZERO" |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
269 |
| "der c (CH d) = (if c = d then ONE else ZERO)" |
273 | 270 |
| "der c (ALT r1 r2) = ALT (der c r1) (der c r2)" |
271 |
| "der c (SEQ r1 r2) = |
|
272 |
(if nullable r1 |
|
273 |
then ALT (SEQ (der c r1) r2) (der c r2) |
|
274 |
else SEQ (der c r1) r2)" |
|
275 |
| "der c (STAR r) = SEQ (der c r) (STAR r)" |
|
276 |
| "der c (UPNTIMES r n) = (if n = 0 then ZERO else SEQ (der c r) (UPNTIMES r (n - 1)))" |
|
277 |
| "der c (NTIMES r n) = (if n = 0 then ZERO else SEQ (der c r) (NTIMES r (n - 1)))" |
|
294 | 278 |
| "der c (FROMNTIMES r n) = |
277 | 279 |
(if n = 0 |
280 |
then SEQ (der c r) (STAR r) |
|
281 |
else SEQ (der c r) (FROMNTIMES r (n - 1)))" |
|
273 | 282 |
| "der c (NMTIMES r n m) = |
283 |
(if m < n then ZERO |
|
284 |
else (if n = 0 then (if m = 0 then ZERO else |
|
285 |
SEQ (der c r) (UPNTIMES r (m - 1))) else |
|
286 |
SEQ (der c r) (NMTIMES r (n - 1) (m - 1))))" |
|
359 | 287 |
| "der c (NOT r) = NOT (der c r)" |
273 | 288 |
|
294 | 289 |
lemma |
290 |
"L(der c (UPNTIMES r m)) = |
|
291 |
L(if (m = 0) then ZERO else ALT ONE (SEQ(der c r) (UPNTIMES r (m - 1))))" |
|
292 |
apply(case_tac m) |
|
293 |
apply(simp) |
|
294 |
apply(simp del: der.simps) |
|
295 |
apply(simp only: der.simps) |
|
296 |
apply(simp add: Sequ_def) |
|
297 |
apply(auto) |
|
298 |
defer |
|
299 |
apply blast |
|
300 |
oops |
|
301 |
||
359 | 302 |
|
303 |
||
304 |
lemma |
|
305 |
assumes "der c r = ONE \<or> der c r = ZERO" |
|
306 |
shows "L (der c (NOT r)) \<noteq> L(if (der c r = ZERO) then ONE else |
|
307 |
if (der c r = ONE) then ZERO |
|
308 |
else NOT(der c r))" |
|
309 |
using assms |
|
310 |
apply(simp) |
|
311 |
apply(auto) |
|
312 |
done |
|
313 |
||
314 |
lemma |
|
315 |
"L (der c (NOT r)) = L(if (der c r = ZERO) then ONE else |
|
316 |
if (der c r = ONE) then ZERO |
|
317 |
else NOT(der c r))" |
|
318 |
apply(simp) |
|
319 |
apply(auto) |
|
320 |
oops |
|
321 |
||
294 | 322 |
lemma pow_add: |
323 |
assumes "s1 \<in> A \<up> n" "s2 \<in> A \<up> m" |
|
324 |
shows "s1 @ s2 \<in> A \<up> (n + m)" |
|
325 |
using assms |
|
326 |
apply(induct n arbitrary: m s1 s2) |
|
327 |
apply(auto simp add: Sequ_def) |
|
328 |
by blast |
|
329 |
||
330 |
lemma pow_add2: |
|
331 |
assumes "x \<in> A \<up> (m + n)" |
|
332 |
shows "x \<in> A \<up> m ;; A \<up> n" |
|
333 |
using assms |
|
334 |
apply(induct m arbitrary: n x) |
|
335 |
apply(auto simp add: Sequ_def) |
|
336 |
by (metis append.assoc) |
|
337 |
||
338 |
||
273 | 339 |
fun |
340 |
ders :: "string \<Rightarrow> rexp \<Rightarrow> rexp" |
|
341 |
where |
|
342 |
"ders [] r = r" |
|
343 |
| "ders (c # s) r = ders s (der c r)" |
|
344 |
||
345 |
||
346 |
lemma nullable_correctness: |
|
347 |
shows "nullable r \<longleftrightarrow> [] \<in> (L r)" |
|
348 |
by(induct r) (auto simp add: Sequ_def) |
|
349 |
||
350 |
||
351 |
lemma der_correctness: |
|
352 |
shows "L (der c r) = Der c (L r)" |
|
353 |
apply(induct r) |
|
354 |
apply(simp add: nullable_correctness del: Der_UNION) |
|
355 |
apply(simp add: nullable_correctness del: Der_UNION) |
|
356 |
apply(simp add: nullable_correctness del: Der_UNION) |
|
357 |
apply(simp add: nullable_correctness del: Der_UNION) |
|
358 |
apply(simp add: nullable_correctness del: Der_UNION) |
|
359 |
apply(simp add: nullable_correctness del: Der_UNION) |
|
359 | 360 |
prefer 2 |
361 |
apply(simp only: der.simps) |
|
362 |
apply(case_tac "x2 = 0") |
|
363 |
apply(simp) |
|
364 |
apply(simp del: Der_Sequ L.simps) |
|
365 |
apply(subst L.simps) |
|
366 |
apply(subst (2) L.simps) |
|
367 |
thm Der_UNION |
|
368 |
||
273 | 369 |
apply(simp add: nullable_correctness del: Der_UNION) |
370 |
apply(simp add: nullable_correctness del: Der_UNION) |
|
371 |
apply(rule impI) |
|
372 |
apply(subst Sequ_Union_in) |
|
373 |
apply(subst Der_Pow_Sequ[symmetric]) |
|
374 |
apply(subst Pow.simps[symmetric]) |
|
375 |
apply(subst Der_UNION[symmetric]) |
|
376 |
apply(subst Pow_Sequ_Un) |
|
377 |
apply(simp) |
|
378 |
apply(simp only: Der_union Der_empty) |
|
276 | 379 |
apply(simp) |
380 |
(* FROMNTIMES *) |
|
381 |
apply(simp add: nullable_correctness del: Der_UNION) |
|
277 | 382 |
apply(rule conjI) |
383 |
prefer 2 |
|
273 | 384 |
apply(subst Sequ_Union_in) |
385 |
apply(subst Der_Pow_Sequ[symmetric]) |
|
386 |
apply(subst Pow.simps[symmetric]) |
|
387 |
apply(case_tac x2) |
|
388 |
prefer 2 |
|
389 |
apply(subst Pow_Sequ_Un2) |
|
390 |
apply(simp) |
|
391 |
apply(simp) |
|
277 | 392 |
apply(auto simp add: Sequ_def Der_def)[1] |
393 |
apply(auto simp add: Sequ_def split: if_splits)[1] |
|
394 |
using Star_Pow apply fastforce |
|
395 |
using Pow_Star apply blast |
|
276 | 396 |
(* NMTIMES *) |
273 | 397 |
apply(simp add: nullable_correctness del: Der_UNION) |
398 |
apply(rule impI) |
|
399 |
apply(rule conjI) |
|
400 |
apply(rule impI) |
|
401 |
apply(subst Sequ_Union_in) |
|
402 |
apply(subst Der_Pow_Sequ[symmetric]) |
|
403 |
apply(subst Pow.simps[symmetric]) |
|
404 |
apply(subst Der_UNION[symmetric]) |
|
405 |
apply(case_tac x3a) |
|
406 |
apply(simp) |
|
407 |
apply(clarify) |
|
408 |
apply(auto simp add: Sequ_def Der_def Cons_eq_append_conv)[1] |
|
409 |
apply(rule_tac x="Suc xa" in bexI) |
|
410 |
apply(auto simp add: Sequ_def)[2] |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
411 |
apply (metis append_Cons) |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
412 |
apply(rule_tac x="xa - 1" in bexI) |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
413 |
apply(auto simp add: Sequ_def)[2] |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
414 |
apply (metis One_nat_def Pow_decomp) |
273 | 415 |
apply(rule impI)+ |
416 |
apply(subst Sequ_Union_in) |
|
417 |
apply(subst Der_Pow_Sequ[symmetric]) |
|
418 |
apply(subst Pow.simps[symmetric]) |
|
419 |
apply(subst Der_UNION[symmetric]) |
|
420 |
apply(case_tac x2) |
|
421 |
apply(simp) |
|
422 |
apply(simp del: Pow.simps) |
|
423 |
apply(auto simp add: Sequ_def Der_def) |
|
424 |
apply (metis One_nat_def Suc_le_D Suc_le_mono atLeastAtMost_iff diff_Suc_1 not_le) |
|
425 |
by fastforce |
|
426 |
||
427 |
||
428 |
||
429 |
lemma ders_correctness: |
|
430 |
shows "L (ders s r) = Ders s (L r)" |
|
431 |
by (induct s arbitrary: r) |
|
432 |
(simp_all add: Ders_def der_correctness Der_def) |
|
433 |
||
434 |
||
435 |
section {* Values *} |
|
436 |
||
437 |
datatype val = |
|
438 |
Void |
|
439 |
| Char char |
|
440 |
| Seq val val |
|
441 |
| Right val |
|
442 |
| Left val |
|
443 |
| Stars "val list" |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
444 |
| Nt val |
273 | 445 |
|
446 |
section {* The string behind a value *} |
|
447 |
||
448 |
fun |
|
449 |
flat :: "val \<Rightarrow> string" |
|
450 |
where |
|
451 |
"flat (Void) = []" |
|
452 |
| "flat (Char c) = [c]" |
|
453 |
| "flat (Left v) = flat v" |
|
454 |
| "flat (Right v) = flat v" |
|
455 |
| "flat (Seq v1 v2) = (flat v1) @ (flat v2)" |
|
456 |
| "flat (Stars []) = []" |
|
457 |
| "flat (Stars (v#vs)) = (flat v) @ (flat (Stars vs))" |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
458 |
| "flat (Nt v) = flat v" |
273 | 459 |
|
460 |
abbreviation |
|
461 |
"flats vs \<equiv> concat (map flat vs)" |
|
462 |
||
463 |
lemma flat_Stars [simp]: |
|
464 |
"flat (Stars vs) = flats vs" |
|
465 |
by (induct vs) (auto) |
|
466 |
||
467 |
lemma Star_concat: |
|
468 |
assumes "\<forall>s \<in> set ss. s \<in> A" |
|
469 |
shows "concat ss \<in> A\<star>" |
|
470 |
using assms by (induct ss) (auto) |
|
471 |
||
472 |
lemma Star_cstring: |
|
473 |
assumes "s \<in> A\<star>" |
|
474 |
shows "\<exists>ss. concat ss = s \<and> (\<forall>s \<in> set ss. s \<in> A \<and> s \<noteq> [])" |
|
475 |
using assms |
|
476 |
apply(induct rule: Star.induct) |
|
477 |
apply(auto)[1] |
|
478 |
apply(rule_tac x="[]" in exI) |
|
479 |
apply(simp) |
|
480 |
apply(erule exE) |
|
481 |
apply(clarify) |
|
482 |
apply(case_tac "s1 = []") |
|
483 |
apply(rule_tac x="ss" in exI) |
|
484 |
apply(simp) |
|
485 |
apply(rule_tac x="s1#ss" in exI) |
|
486 |
apply(simp) |
|
487 |
done |
|
488 |
||
489 |
lemma Aux: |
|
490 |
assumes "\<forall>s\<in>set ss. s = []" |
|
491 |
shows "concat ss = []" |
|
492 |
using assms |
|
493 |
by (induct ss) (auto) |
|
494 |
||
495 |
lemma Pow_cstring_nonempty: |
|
496 |
assumes "s \<in> A \<up> n" |
|
497 |
shows "\<exists>ss. concat ss = s \<and> length ss \<le> n \<and> (\<forall>s \<in> set ss. s \<in> A \<and> s \<noteq> [])" |
|
498 |
using assms |
|
499 |
apply(induct n arbitrary: s) |
|
500 |
apply(auto) |
|
501 |
apply(simp add: Sequ_def) |
|
502 |
apply(erule exE)+ |
|
503 |
apply(clarify) |
|
504 |
apply(drule_tac x="s2" in meta_spec) |
|
505 |
apply(simp) |
|
506 |
apply(clarify) |
|
507 |
apply(case_tac "s1 = []") |
|
508 |
apply(simp) |
|
509 |
apply(rule_tac x="ss" in exI) |
|
510 |
apply(simp) |
|
511 |
apply(rule_tac x="s1 # ss" in exI) |
|
512 |
apply(simp) |
|
513 |
done |
|
514 |
||
515 |
lemma Pow_cstring: |
|
516 |
assumes "s \<in> A \<up> n" |
|
517 |
shows "\<exists>ss1 ss2. concat (ss1 @ ss2) = s \<and> length (ss1 @ ss2) = n \<and> |
|
518 |
(\<forall>s \<in> set ss1. s \<in> A \<and> s \<noteq> []) \<and> (\<forall>s \<in> set ss2. s \<in> A \<and> s = [])" |
|
519 |
using assms |
|
520 |
apply(induct n arbitrary: s) |
|
521 |
apply(auto)[1] |
|
522 |
apply(simp only: Pow_Suc_rev) |
|
523 |
apply(simp add: Sequ_def) |
|
524 |
apply(erule exE)+ |
|
525 |
apply(clarify) |
|
526 |
apply(drule_tac x="s1" in meta_spec) |
|
527 |
apply(simp) |
|
528 |
apply(erule exE)+ |
|
529 |
apply(clarify) |
|
530 |
apply(case_tac "s2 = []") |
|
531 |
apply(simp) |
|
532 |
apply(rule_tac x="ss1" in exI) |
|
533 |
apply(rule_tac x="s2#ss2" in exI) |
|
534 |
apply(simp) |
|
535 |
apply(rule_tac x="ss1 @ [s2]" in exI) |
|
536 |
apply(rule_tac x="ss2" in exI) |
|
537 |
apply(simp) |
|
538 |
apply(subst Aux) |
|
539 |
apply(auto)[1] |
|
540 |
apply(subst Aux) |
|
541 |
apply(auto)[1] |
|
542 |
apply(simp) |
|
543 |
done |
|
544 |
||
545 |
||
546 |
section {* Lexical Values *} |
|
547 |
||
548 |
||
549 |
||
550 |
inductive |
|
551 |
Prf :: "val \<Rightarrow> rexp \<Rightarrow> bool" ("\<Turnstile> _ : _" [100, 100] 100) |
|
552 |
where |
|
553 |
"\<lbrakk>\<Turnstile> v1 : r1; \<Turnstile> v2 : r2\<rbrakk> \<Longrightarrow> \<Turnstile> Seq v1 v2 : SEQ r1 r2" |
|
554 |
| "\<Turnstile> v1 : r1 \<Longrightarrow> \<Turnstile> Left v1 : ALT r1 r2" |
|
555 |
| "\<Turnstile> v2 : r2 \<Longrightarrow> \<Turnstile> Right v2 : ALT r1 r2" |
|
556 |
| "\<Turnstile> Void : ONE" |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
557 |
| "\<Turnstile> Char c : CH c" |
273 | 558 |
| "\<lbrakk>\<forall>v \<in> set vs. \<Turnstile> v : r \<and> flat v \<noteq> []\<rbrakk> \<Longrightarrow> \<Turnstile> Stars vs : STAR r" |
559 |
| "\<lbrakk>\<forall>v \<in> set vs. \<Turnstile> v : r \<and> flat v \<noteq> []; length vs \<le> n\<rbrakk> \<Longrightarrow> \<Turnstile> Stars vs : UPNTIMES r n" |
|
560 |
| "\<lbrakk>\<forall>v \<in> set vs1. \<Turnstile> v : r \<and> flat v \<noteq> []; |
|
561 |
\<forall>v \<in> set vs2. \<Turnstile> v : r \<and> flat v = []; |
|
562 |
length (vs1 @ vs2) = n\<rbrakk> \<Longrightarrow> \<Turnstile> Stars (vs1 @ vs2) : NTIMES r n" |
|
563 |
| "\<lbrakk>\<forall>v \<in> set vs1. \<Turnstile> v : r \<and> flat v \<noteq> []; |
|
564 |
\<forall>v \<in> set vs2. \<Turnstile> v : r \<and> flat v = []; |
|
275 | 565 |
length (vs1 @ vs2) = n\<rbrakk> \<Longrightarrow> \<Turnstile> Stars (vs1 @ vs2) : FROMNTIMES r n" |
566 |
| "\<lbrakk>\<forall>v \<in> set vs. \<Turnstile> v : r \<and> flat v \<noteq> []; length vs > n\<rbrakk> \<Longrightarrow> \<Turnstile> Stars vs : FROMNTIMES r n" |
|
273 | 567 |
| "\<lbrakk>\<forall>v \<in> set vs1. \<Turnstile> v : r \<and> flat v \<noteq> []; |
568 |
\<forall>v \<in> set vs2. \<Turnstile> v : r \<and> flat v = []; |
|
275 | 569 |
length (vs1 @ vs2) = n; length (vs1 @ vs2) \<le> m\<rbrakk> \<Longrightarrow> \<Turnstile> Stars (vs1 @ vs2) : NMTIMES r n m" |
570 |
| "\<lbrakk>\<forall>v \<in> set vs. \<Turnstile> v : r \<and> flat v \<noteq> []; |
|
571 |
length vs > n; length vs \<le> m\<rbrakk> \<Longrightarrow> \<Turnstile> Stars vs : NMTIMES r n m" |
|
273 | 572 |
|
293 | 573 |
|
574 |
||
575 |
||
576 |
||
273 | 577 |
inductive_cases Prf_elims: |
578 |
"\<Turnstile> v : ZERO" |
|
579 |
"\<Turnstile> v : SEQ r1 r2" |
|
580 |
"\<Turnstile> v : ALT r1 r2" |
|
581 |
"\<Turnstile> v : ONE" |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
582 |
"\<Turnstile> v : CH c" |
273 | 583 |
"\<Turnstile> vs : STAR r" |
584 |
"\<Turnstile> vs : UPNTIMES r n" |
|
585 |
"\<Turnstile> vs : NTIMES r n" |
|
586 |
"\<Turnstile> vs : FROMNTIMES r n" |
|
587 |
"\<Turnstile> vs : NMTIMES r n m" |
|
588 |
||
589 |
lemma Prf_Stars_appendE: |
|
590 |
assumes "\<Turnstile> Stars (vs1 @ vs2) : STAR r" |
|
591 |
shows "\<Turnstile> Stars vs1 : STAR r \<and> \<Turnstile> Stars vs2 : STAR r" |
|
592 |
using assms |
|
593 |
by (auto intro: Prf.intros elim!: Prf_elims) |
|
594 |
||
274 | 595 |
|
596 |
||
273 | 597 |
lemma flats_empty: |
598 |
assumes "(\<forall>v\<in>set vs. flat v = [])" |
|
599 |
shows "flats vs = []" |
|
600 |
using assms |
|
601 |
by(induct vs) (simp_all) |
|
602 |
||
603 |
lemma Star_cval: |
|
604 |
assumes "\<forall>s\<in>set ss. \<exists>v. s = flat v \<and> \<Turnstile> v : r" |
|
605 |
shows "\<exists>vs. flats vs = concat ss \<and> (\<forall>v\<in>set vs. \<Turnstile> v : r \<and> flat v \<noteq> [])" |
|
606 |
using assms |
|
607 |
apply(induct ss) |
|
608 |
apply(auto) |
|
609 |
apply(rule_tac x="[]" in exI) |
|
610 |
apply(simp) |
|
611 |
apply(case_tac "flat v = []") |
|
612 |
apply(rule_tac x="vs" in exI) |
|
613 |
apply(simp) |
|
614 |
apply(rule_tac x="v#vs" in exI) |
|
615 |
apply(simp) |
|
616 |
done |
|
617 |
||
618 |
||
619 |
lemma flats_cval: |
|
620 |
assumes "\<forall>s\<in>set ss. \<exists>v. s = flat v \<and> \<Turnstile> v : r" |
|
621 |
shows "\<exists>vs1 vs2. flats (vs1 @ vs2) = concat ss \<and> length (vs1 @ vs2) = length ss \<and> |
|
622 |
(\<forall>v\<in>set vs1. \<Turnstile> v : r \<and> flat v \<noteq> []) \<and> |
|
623 |
(\<forall>v\<in>set vs2. \<Turnstile> v : r \<and> flat v = [])" |
|
624 |
using assms |
|
625 |
apply(induct ss rule: rev_induct) |
|
626 |
apply(rule_tac x="[]" in exI)+ |
|
627 |
apply(simp) |
|
628 |
apply(simp) |
|
629 |
apply(clarify) |
|
630 |
apply(case_tac "flat v = []") |
|
631 |
apply(rule_tac x="vs1" in exI) |
|
632 |
apply(rule_tac x="v#vs2" in exI) |
|
633 |
apply(simp) |
|
634 |
apply(rule_tac x="vs1 @ [v]" in exI) |
|
635 |
apply(rule_tac x="vs2" in exI) |
|
636 |
apply(simp) |
|
637 |
apply(subst (asm) (2) flats_empty) |
|
638 |
apply(simp) |
|
639 |
apply(simp) |
|
640 |
done |
|
641 |
||
642 |
lemma flats_cval_nonempty: |
|
643 |
assumes "\<forall>s\<in>set ss. \<exists>v. s = flat v \<and> \<Turnstile> v : r" |
|
644 |
shows "\<exists>vs. flats vs = concat ss \<and> length vs \<le> length ss \<and> |
|
645 |
(\<forall>v\<in>set vs. \<Turnstile> v : r \<and> flat v \<noteq> [])" |
|
646 |
using assms |
|
647 |
apply(induct ss) |
|
648 |
apply(rule_tac x="[]" in exI) |
|
649 |
apply(simp) |
|
650 |
apply(simp) |
|
651 |
apply(clarify) |
|
652 |
apply(case_tac "flat v = []") |
|
653 |
apply(rule_tac x="vs" in exI) |
|
654 |
apply(simp) |
|
655 |
apply(rule_tac x="v # vs" in exI) |
|
656 |
apply(simp) |
|
657 |
done |
|
658 |
||
659 |
lemma Pow_flats: |
|
660 |
assumes "\<forall>v \<in> set vs. flat v \<in> A" |
|
661 |
shows "flats vs \<in> A \<up> length vs" |
|
662 |
using assms |
|
663 |
by(induct vs)(auto simp add: Sequ_def) |
|
664 |
||
665 |
lemma Pow_flats_appends: |
|
666 |
assumes "\<forall>v \<in> set vs1. flat v \<in> A" "\<forall>v \<in> set vs2. flat v \<in> A" |
|
667 |
shows "flats vs1 @ flats vs2 \<in> A \<up> (length vs1 + length vs2)" |
|
668 |
using assms |
|
669 |
apply(induct vs1) |
|
670 |
apply(auto simp add: Sequ_def Pow_flats) |
|
671 |
done |
|
672 |
||
673 |
lemma L_flat_Prf1: |
|
674 |
assumes "\<Turnstile> v : r" |
|
675 |
shows "flat v \<in> L r" |
|
676 |
using assms |
|
677 |
apply(induct) |
|
678 |
apply(auto simp add: Sequ_def Star_concat Pow_flats) |
|
679 |
apply(meson Pow_flats atMost_iff) |
|
680 |
using Pow_flats_appends apply blast |
|
275 | 681 |
using Pow_flats_appends apply blast |
682 |
apply (meson Pow_flats atLeast_iff less_imp_le) |
|
683 |
apply(rule_tac x="length vs1 + length vs2" in bexI) |
|
273 | 684 |
apply(meson Pow_flats_appends atLeastAtMost_iff) |
275 | 685 |
apply(simp) |
686 |
apply(meson Pow_flats atLeastAtMost_iff less_or_eq_imp_le) |
|
273 | 687 |
done |
688 |
||
689 |
lemma L_flat_Prf2: |
|
690 |
assumes "s \<in> L r" |
|
691 |
shows "\<exists>v. \<Turnstile> v : r \<and> flat v = s" |
|
692 |
using assms |
|
693 |
proof(induct r arbitrary: s) |
|
694 |
case (STAR r s) |
|
695 |
have IH: "\<And>s. s \<in> L r \<Longrightarrow> \<exists>v. \<Turnstile> v : r \<and> flat v = s" by fact |
|
696 |
have "s \<in> L (STAR r)" by fact |
|
697 |
then obtain ss where "concat ss = s" "\<forall>s \<in> set ss. s \<in> L r \<and> s \<noteq> []" |
|
698 |
using Star_cstring by auto |
|
699 |
then obtain vs where "flats vs = s" "\<forall>v\<in>set vs. \<Turnstile> v : r \<and> flat v \<noteq> []" |
|
700 |
using IH Star_cval by metis |
|
701 |
then show "\<exists>v. \<Turnstile> v : STAR r \<and> flat v = s" |
|
702 |
using Prf.intros(6) flat_Stars by blast |
|
703 |
next |
|
704 |
case (SEQ r1 r2 s) |
|
705 |
then show "\<exists>v. \<Turnstile> v : SEQ r1 r2 \<and> flat v = s" |
|
706 |
unfolding Sequ_def L.simps by (fastforce intro: Prf.intros) |
|
707 |
next |
|
708 |
case (ALT r1 r2 s) |
|
709 |
then show "\<exists>v. \<Turnstile> v : ALT r1 r2 \<and> flat v = s" |
|
710 |
unfolding L.simps by (fastforce intro: Prf.intros) |
|
711 |
next |
|
712 |
case (NTIMES r n) |
|
713 |
have IH: "\<And>s. s \<in> L r \<Longrightarrow> \<exists>v. \<Turnstile> v : r \<and> flat v = s" by fact |
|
714 |
have "s \<in> L (NTIMES r n)" by fact |
|
715 |
then obtain ss1 ss2 where "concat (ss1 @ ss2) = s" "length (ss1 @ ss2) = n" |
|
716 |
"\<forall>s \<in> set ss1. s \<in> L r \<and> s \<noteq> []" "\<forall>s \<in> set ss2. s \<in> L r \<and> s = []" |
|
717 |
using Pow_cstring by force |
|
718 |
then obtain vs1 vs2 where "flats (vs1 @ vs2) = s" "length (vs1 @ vs2) = n" |
|
719 |
"\<forall>v\<in>set vs1. \<Turnstile> v : r \<and> flat v \<noteq> []" "\<forall>v\<in>set vs2. \<Turnstile> v : r \<and> flat v = []" |
|
720 |
using IH flats_cval |
|
721 |
apply - |
|
722 |
apply(drule_tac x="ss1 @ ss2" in meta_spec) |
|
723 |
apply(drule_tac x="r" in meta_spec) |
|
724 |
apply(drule meta_mp) |
|
725 |
apply(simp) |
|
726 |
apply (metis Un_iff) |
|
727 |
apply(clarify) |
|
728 |
apply(drule_tac x="vs1" in meta_spec) |
|
729 |
apply(drule_tac x="vs2" in meta_spec) |
|
730 |
apply(simp) |
|
731 |
done |
|
732 |
then show "\<exists>v. \<Turnstile> v : NTIMES r n \<and> flat v = s" |
|
733 |
using Prf.intros(8) flat_Stars by blast |
|
734 |
next |
|
735 |
case (FROMNTIMES r n) |
|
736 |
have IH: "\<And>s. s \<in> L r \<Longrightarrow> \<exists>v. \<Turnstile> v : r \<and> flat v = s" by fact |
|
737 |
have "s \<in> L (FROMNTIMES r n)" by fact |
|
276 | 738 |
then obtain ss1 ss2 k where "concat (ss1 @ ss2) = s" "length (ss1 @ ss2) = k" "n \<le> k" |
273 | 739 |
"\<forall>s \<in> set ss1. s \<in> L r \<and> s \<noteq> []" "\<forall>s \<in> set ss2. s \<in> L r \<and> s = []" |
275 | 740 |
using Pow_cstring by force |
276 | 741 |
then obtain vs1 vs2 where "flats (vs1 @ vs2) = s" "length (vs1 @ vs2) = k" "n \<le> k" |
273 | 742 |
"\<forall>v\<in>set vs1. \<Turnstile> v : r \<and> flat v \<noteq> []" "\<forall>v\<in>set vs2. \<Turnstile> v : r \<and> flat v = []" |
743 |
using IH flats_cval |
|
744 |
apply - |
|
745 |
apply(drule_tac x="ss1 @ ss2" in meta_spec) |
|
746 |
apply(drule_tac x="r" in meta_spec) |
|
747 |
apply(drule meta_mp) |
|
748 |
apply(simp) |
|
749 |
apply (metis Un_iff) |
|
750 |
apply(clarify) |
|
751 |
apply(drule_tac x="vs1" in meta_spec) |
|
752 |
apply(drule_tac x="vs2" in meta_spec) |
|
753 |
apply(simp) |
|
754 |
done |
|
755 |
then show "\<exists>v. \<Turnstile> v : FROMNTIMES r n \<and> flat v = s" |
|
275 | 756 |
apply(case_tac "length vs1 \<le> n") |
757 |
apply(rule_tac x="Stars (vs1 @ take (n - length vs1) vs2)" in exI) |
|
758 |
apply(simp) |
|
759 |
apply(subgoal_tac "flats (take (n - length vs1) vs2) = []") |
|
760 |
prefer 2 |
|
761 |
apply (meson flats_empty in_set_takeD) |
|
762 |
apply(clarify) |
|
763 |
apply(rule conjI) |
|
764 |
apply(rule Prf.intros) |
|
765 |
apply(simp) |
|
766 |
apply (meson in_set_takeD) |
|
767 |
apply(simp) |
|
768 |
apply(simp) |
|
769 |
apply (simp add: flats_empty) |
|
770 |
apply(rule_tac x="Stars vs1" in exI) |
|
771 |
apply(simp) |
|
772 |
apply(rule conjI) |
|
773 |
apply(rule Prf.intros(10)) |
|
774 |
apply(auto) |
|
775 |
done |
|
273 | 776 |
next |
777 |
case (NMTIMES r n m) |
|
778 |
have IH: "\<And>s. s \<in> L r \<Longrightarrow> \<exists>v. \<Turnstile> v : r \<and> flat v = s" by fact |
|
779 |
have "s \<in> L (NMTIMES r n m)" by fact |
|
780 |
then obtain ss1 ss2 k where "concat (ss1 @ ss2) = s" "length (ss1 @ ss2) = k" "n \<le> k" "k \<le> m" |
|
781 |
"\<forall>s \<in> set ss1. s \<in> L r \<and> s \<noteq> []" "\<forall>s \<in> set ss2. s \<in> L r \<and> s = []" |
|
782 |
using Pow_cstring by (auto, blast) |
|
783 |
then obtain vs1 vs2 where "flats (vs1 @ vs2) = s" "length (vs1 @ vs2) = k" "n \<le> k" "k \<le> m" |
|
784 |
"\<forall>v\<in>set vs1. \<Turnstile> v : r \<and> flat v \<noteq> []" "\<forall>v\<in>set vs2. \<Turnstile> v : r \<and> flat v = []" |
|
785 |
using IH flats_cval |
|
786 |
apply - |
|
787 |
apply(drule_tac x="ss1 @ ss2" in meta_spec) |
|
788 |
apply(drule_tac x="r" in meta_spec) |
|
789 |
apply(drule meta_mp) |
|
790 |
apply(simp) |
|
791 |
apply (metis Un_iff) |
|
792 |
apply(clarify) |
|
793 |
apply(drule_tac x="vs1" in meta_spec) |
|
794 |
apply(drule_tac x="vs2" in meta_spec) |
|
795 |
apply(simp) |
|
796 |
done |
|
797 |
then show "\<exists>v. \<Turnstile> v : NMTIMES r n m \<and> flat v = s" |
|
276 | 798 |
apply(case_tac "length vs1 \<le> n") |
799 |
apply(rule_tac x="Stars (vs1 @ take (n - length vs1) vs2)" in exI) |
|
273 | 800 |
apply(simp) |
276 | 801 |
apply(subgoal_tac "flats (take (n - length vs1) vs2) = []") |
802 |
prefer 2 |
|
803 |
apply (meson flats_empty in_set_takeD) |
|
804 |
apply(clarify) |
|
805 |
apply(rule conjI) |
|
806 |
apply(rule Prf.intros) |
|
807 |
apply(simp) |
|
808 |
apply (meson in_set_takeD) |
|
809 |
apply(simp) |
|
810 |
apply(simp) |
|
811 |
apply (simp add: flats_empty) |
|
812 |
apply(rule_tac x="Stars vs1" in exI) |
|
813 |
apply(simp) |
|
814 |
apply(rule conjI) |
|
815 |
apply(rule Prf.intros) |
|
816 |
apply(auto) |
|
817 |
done |
|
273 | 818 |
next |
819 |
case (UPNTIMES r n s) |
|
820 |
have IH: "\<And>s. s \<in> L r \<Longrightarrow> \<exists>v. \<Turnstile> v : r \<and> flat v = s" by fact |
|
821 |
have "s \<in> L (UPNTIMES r n)" by fact |
|
822 |
then obtain ss where "concat ss = s" "\<forall>s \<in> set ss. s \<in> L r \<and> s \<noteq> []" "length ss \<le> n" |
|
823 |
using Pow_cstring_nonempty by force |
|
824 |
then obtain vs where "flats vs = s" "\<forall>v\<in>set vs. \<Turnstile> v : r \<and> flat v \<noteq> []" "length vs \<le> n" |
|
825 |
using IH flats_cval_nonempty by (smt order.trans) |
|
826 |
then show "\<exists>v. \<Turnstile> v : UPNTIMES r n \<and> flat v = s" |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
827 |
using Prf.intros(7) flat_Stars by blast |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
828 |
next |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
829 |
case (NOT r) |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
830 |
then show ?case sorry |
273 | 831 |
qed (auto intro: Prf.intros) |
832 |
||
833 |
||
834 |
lemma L_flat_Prf: |
|
835 |
shows "L(r) = {flat v | v. \<Turnstile> v : r}" |
|
836 |
using L_flat_Prf1 L_flat_Prf2 by blast |
|
837 |
||
293 | 838 |
thm Prf.intros |
839 |
thm Prf.cases |
|
273 | 840 |
|
293 | 841 |
lemma |
842 |
assumes "\<Turnstile> v : (STAR r)" |
|
843 |
shows "\<Turnstile> v : (FROMNTIMES r 0)" |
|
844 |
using assms |
|
845 |
apply(erule_tac Prf.cases) |
|
846 |
apply(simp_all) |
|
847 |
apply(case_tac vs) |
|
848 |
apply(auto) |
|
849 |
apply(subst append_Nil[symmetric]) |
|
850 |
apply(rule Prf.intros) |
|
851 |
apply(auto) |
|
852 |
apply(simp add: Prf.intros) |
|
853 |
done |
|
854 |
||
855 |
lemma |
|
856 |
assumes "\<Turnstile> v : (FROMNTIMES r 0)" |
|
857 |
shows "\<Turnstile> v : (STAR r)" |
|
858 |
using assms |
|
859 |
apply(erule_tac Prf.cases) |
|
860 |
apply(simp_all) |
|
861 |
apply(rule Prf.intros) |
|
862 |
apply(simp) |
|
863 |
apply(rule Prf.intros) |
|
864 |
apply(simp) |
|
865 |
done |
|
273 | 866 |
|
867 |
section {* Sets of Lexical Values *} |
|
868 |
||
869 |
text {* |
|
870 |
Shows that lexical values are finite for a given regex and string. |
|
871 |
*} |
|
872 |
||
873 |
definition |
|
874 |
LV :: "rexp \<Rightarrow> string \<Rightarrow> val set" |
|
875 |
where "LV r s \<equiv> {v. \<Turnstile> v : r \<and> flat v = s}" |
|
876 |
||
877 |
lemma LV_simps: |
|
878 |
shows "LV ZERO s = {}" |
|
879 |
and "LV ONE s = (if s = [] then {Void} else {})" |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
880 |
and "LV (CH c) s = (if s = [c] then {Char c} else {})" |
273 | 881 |
and "LV (ALT r1 r2) s = Left ` LV r1 s \<union> Right ` LV r2 s" |
882 |
unfolding LV_def |
|
274 | 883 |
apply(auto intro: Prf.intros elim: Prf.cases) |
884 |
done |
|
273 | 885 |
|
886 |
abbreviation |
|
275 | 887 |
"Prefixes s \<equiv> {s'. prefix s' s}" |
273 | 888 |
|
889 |
abbreviation |
|
275 | 890 |
"Suffixes s \<equiv> {s'. suffix s' s}" |
273 | 891 |
|
892 |
abbreviation |
|
275 | 893 |
"SSuffixes s \<equiv> {s'. strict_suffix s' s}" |
273 | 894 |
|
895 |
lemma Suffixes_cons [simp]: |
|
896 |
shows "Suffixes (c # s) = Suffixes s \<union> {c # s}" |
|
275 | 897 |
by (auto simp add: suffix_def Cons_eq_append_conv) |
273 | 898 |
|
899 |
||
900 |
lemma finite_Suffixes: |
|
901 |
shows "finite (Suffixes s)" |
|
902 |
by (induct s) (simp_all) |
|
903 |
||
904 |
lemma finite_SSuffixes: |
|
905 |
shows "finite (SSuffixes s)" |
|
906 |
proof - |
|
907 |
have "SSuffixes s \<subseteq> Suffixes s" |
|
275 | 908 |
unfolding suffix_def strict_suffix_def by auto |
273 | 909 |
then show "finite (SSuffixes s)" |
910 |
using finite_Suffixes finite_subset by blast |
|
911 |
qed |
|
912 |
||
913 |
lemma finite_Prefixes: |
|
914 |
shows "finite (Prefixes s)" |
|
915 |
proof - |
|
916 |
have "finite (Suffixes (rev s))" |
|
917 |
by (rule finite_Suffixes) |
|
918 |
then have "finite (rev ` Suffixes (rev s))" by simp |
|
919 |
moreover |
|
920 |
have "rev ` (Suffixes (rev s)) = Prefixes s" |
|
275 | 921 |
unfolding suffix_def prefix_def image_def |
273 | 922 |
by (auto)(metis rev_append rev_rev_ident)+ |
923 |
ultimately show "finite (Prefixes s)" by simp |
|
924 |
qed |
|
925 |
||
276 | 926 |
definition |
927 |
"Stars_Cons V Vs \<equiv> {Stars (v # vs) | v vs. v \<in> V \<and> Stars vs \<in> Vs}" |
|
928 |
||
929 |
definition |
|
930 |
"Stars_Append Vs1 Vs2 \<equiv> {Stars (vs1 @ vs2) | vs1 vs2. Stars vs1 \<in> Vs1 \<and> Stars vs2 \<in> Vs2}" |
|
931 |
||
932 |
fun Stars_Pow :: "val set \<Rightarrow> nat \<Rightarrow> val set" |
|
933 |
where |
|
934 |
"Stars_Pow Vs 0 = {Stars []}" |
|
935 |
| "Stars_Pow Vs (Suc n) = Stars_Cons Vs (Stars_Pow Vs n)" |
|
936 |
||
937 |
lemma finite_Stars_Cons: |
|
938 |
assumes "finite V" "finite Vs" |
|
939 |
shows "finite (Stars_Cons V Vs)" |
|
940 |
using assms |
|
941 |
proof - |
|
942 |
from assms(2) have "finite (Stars -` Vs)" |
|
943 |
by(simp add: finite_vimageI inj_on_def) |
|
944 |
with assms(1) have "finite (V \<times> (Stars -` Vs))" |
|
945 |
by(simp) |
|
946 |
then have "finite ((\<lambda>(v, vs). Stars (v # vs)) ` (V \<times> (Stars -` Vs)))" |
|
947 |
by simp |
|
948 |
moreover have "Stars_Cons V Vs = (\<lambda>(v, vs). Stars (v # vs)) ` (V \<times> (Stars -` Vs))" |
|
949 |
unfolding Stars_Cons_def by auto |
|
950 |
ultimately show "finite (Stars_Cons V Vs)" |
|
951 |
by simp |
|
952 |
qed |
|
953 |
||
954 |
lemma finite_Stars_Append: |
|
955 |
assumes "finite Vs1" "finite Vs2" |
|
956 |
shows "finite (Stars_Append Vs1 Vs2)" |
|
957 |
using assms |
|
958 |
proof - |
|
959 |
define UVs1 where "UVs1 \<equiv> Stars -` Vs1" |
|
960 |
define UVs2 where "UVs2 \<equiv> Stars -` Vs2" |
|
961 |
from assms have "finite UVs1" "finite UVs2" |
|
962 |
unfolding UVs1_def UVs2_def |
|
963 |
by(simp_all add: finite_vimageI inj_on_def) |
|
964 |
then have "finite ((\<lambda>(vs1, vs2). Stars (vs1 @ vs2)) ` (UVs1 \<times> UVs2))" |
|
965 |
by simp |
|
966 |
moreover |
|
967 |
have "Stars_Append Vs1 Vs2 = (\<lambda>(vs1, vs2). Stars (vs1 @ vs2)) ` (UVs1 \<times> UVs2)" |
|
968 |
unfolding Stars_Append_def UVs1_def UVs2_def by auto |
|
969 |
ultimately show "finite (Stars_Append Vs1 Vs2)" |
|
970 |
by simp |
|
971 |
qed |
|
972 |
||
973 |
lemma finite_Stars_Pow: |
|
974 |
assumes "finite Vs" |
|
975 |
shows "finite (Stars_Pow Vs n)" |
|
976 |
by (induct n) (simp_all add: finite_Stars_Cons assms) |
|
977 |
||
273 | 978 |
lemma LV_STAR_finite: |
979 |
assumes "\<forall>s. finite (LV r s)" |
|
980 |
shows "finite (LV (STAR r) s)" |
|
981 |
proof(induct s rule: length_induct) |
|
982 |
fix s::"char list" |
|
983 |
assume "\<forall>s'. length s' < length s \<longrightarrow> finite (LV (STAR r) s')" |
|
984 |
then have IH: "\<forall>s' \<in> SSuffixes s. finite (LV (STAR r) s')" |
|
279 | 985 |
apply(auto simp add: strict_suffix_def suffix_def) |
986 |
by force |
|
275 | 987 |
define f where "f \<equiv> \<lambda>(v, vs). Stars (v # vs)" |
988 |
define S1 where "S1 \<equiv> \<Union>s' \<in> Prefixes s. LV r s'" |
|
276 | 989 |
define S2 where "S2 \<equiv> \<Union>s2 \<in> SSuffixes s. LV (STAR r) s2" |
273 | 990 |
have "finite S1" using assms |
991 |
unfolding S1_def by (simp_all add: finite_Prefixes) |
|
992 |
moreover |
|
993 |
with IH have "finite S2" unfolding S2_def |
|
276 | 994 |
by (auto simp add: finite_SSuffixes) |
273 | 995 |
ultimately |
276 | 996 |
have "finite ({Stars []} \<union> Stars_Cons S1 S2)" |
997 |
by (simp add: finite_Stars_Cons) |
|
273 | 998 |
moreover |
276 | 999 |
have "LV (STAR r) s \<subseteq> {Stars []} \<union> (Stars_Cons S1 S2)" |
1000 |
unfolding S1_def S2_def f_def LV_def Stars_Cons_def |
|
1001 |
unfolding prefix_def strict_suffix_def |
|
1002 |
unfolding image_def |
|
275 | 1003 |
apply(auto) |
1004 |
apply(case_tac x) |
|
273 | 1005 |
apply(auto elim: Prf_elims) |
1006 |
apply(erule Prf_elims) |
|
1007 |
apply(auto) |
|
1008 |
apply(case_tac vs) |
|
275 | 1009 |
apply(auto intro: Prf.intros) |
1010 |
apply(rule exI) |
|
1011 |
apply(rule conjI) |
|
276 | 1012 |
apply(rule_tac x="flats list" in exI) |
279 | 1013 |
apply(rule conjI) |
1014 |
apply(simp add: suffix_def) |
|
275 | 1015 |
apply(blast) |
276 | 1016 |
using Prf.intros(6) flat_Stars by blast |
273 | 1017 |
ultimately |
1018 |
show "finite (LV (STAR r) s)" by (simp add: finite_subset) |
|
1019 |
qed |
|
1020 |
||
1021 |
lemma LV_UPNTIMES_STAR: |
|
1022 |
"LV (UPNTIMES r n) s \<subseteq> LV (STAR r) s" |
|
1023 |
by(auto simp add: LV_def intro: Prf.intros elim: Prf_elims) |
|
1024 |
||
274 | 1025 |
lemma LV_NTIMES_3: |
1026 |
shows "LV (NTIMES r (Suc n)) [] = (\<lambda>(v,vs). Stars (v#vs)) ` (LV r [] \<times> (Stars -` (LV (NTIMES r n) [])))" |
|
1027 |
unfolding LV_def |
|
1028 |
apply(auto elim!: Prf_elims simp add: image_def) |
|
1029 |
apply(case_tac vs1) |
|
1030 |
apply(auto) |
|
1031 |
apply(case_tac vs2) |
|
1032 |
apply(auto) |
|
1033 |
apply(subst append.simps(1)[symmetric]) |
|
1034 |
apply(rule Prf.intros) |
|
1035 |
apply(auto) |
|
1036 |
apply(subst append.simps(1)[symmetric]) |
|
1037 |
apply(rule Prf.intros) |
|
1038 |
apply(auto) |
|
276 | 1039 |
done |
275 | 1040 |
|
276 | 1041 |
lemma LV_FROMNTIMES_3: |
1042 |
shows "LV (FROMNTIMES r (Suc n)) [] = |
|
1043 |
(\<lambda>(v,vs). Stars (v#vs)) ` (LV r [] \<times> (Stars -` (LV (FROMNTIMES r n) [])))" |
|
1044 |
unfolding LV_def |
|
1045 |
apply(auto elim!: Prf_elims simp add: image_def) |
|
1046 |
apply(case_tac vs1) |
|
1047 |
apply(auto) |
|
1048 |
apply(case_tac vs2) |
|
1049 |
apply(auto) |
|
1050 |
apply(subst append.simps(1)[symmetric]) |
|
1051 |
apply(rule Prf.intros) |
|
1052 |
apply(auto) |
|
1053 |
apply (metis le_imp_less_Suc length_greater_0_conv less_antisym list.exhaust list.set_intros(1) not_less_eq zero_le) |
|
1054 |
prefer 2 |
|
1055 |
using nth_mem apply blast |
|
1056 |
apply(case_tac vs1) |
|
1057 |
apply (smt Groups.add_ac(2) Prf.intros(9) add.right_neutral add_Suc_right append.simps(1) insert_iff length_append list.set(2) list.size(3) list.size(4)) |
|
275 | 1058 |
apply(auto) |
276 | 1059 |
done |
1060 |
||
1061 |
lemma LV_NTIMES_4: |
|
1062 |
"LV (NTIMES r n) [] = Stars_Pow (LV r []) n" |
|
1063 |
apply(induct n) |
|
1064 |
apply(simp add: LV_def) |
|
1065 |
apply(auto elim!: Prf_elims simp add: image_def)[1] |
|
1066 |
apply(subst append.simps[symmetric]) |
|
1067 |
apply(rule Prf.intros) |
|
1068 |
apply(simp_all) |
|
1069 |
apply(simp add: LV_NTIMES_3 image_def Stars_Cons_def) |
|
1070 |
apply blast |
|
1071 |
done |
|
274 | 1072 |
|
276 | 1073 |
lemma LV_NTIMES_5: |
1074 |
"LV (NTIMES r n) s \<subseteq> Stars_Append (LV (STAR r) s) (\<Union>i\<le>n. LV (NTIMES r i) [])" |
|
1075 |
apply(auto simp add: LV_def) |
|
1076 |
apply(auto elim!: Prf_elims) |
|
1077 |
apply(auto simp add: Stars_Append_def) |
|
1078 |
apply(rule_tac x="vs1" in exI) |
|
1079 |
apply(rule_tac x="vs2" in exI) |
|
1080 |
apply(auto) |
|
1081 |
using Prf.intros(6) apply(auto) |
|
1082 |
apply(rule_tac x="length vs2" in bexI) |
|
1083 |
thm Prf.intros |
|
1084 |
apply(subst append.simps(1)[symmetric]) |
|
1085 |
apply(rule Prf.intros) |
|
1086 |
apply(auto)[1] |
|
1087 |
apply(auto)[1] |
|
1088 |
apply(simp) |
|
1089 |
apply(simp) |
|
1090 |
done |
|
1091 |
||
1092 |
lemma ttty: |
|
1093 |
"LV (FROMNTIMES r n) [] = Stars_Pow (LV r []) n" |
|
1094 |
apply(induct n) |
|
1095 |
apply(simp add: LV_def) |
|
1096 |
apply(auto elim: Prf_elims simp add: image_def)[1] |
|
1097 |
prefer 2 |
|
1098 |
apply(subst append.simps[symmetric]) |
|
1099 |
apply(rule Prf.intros) |
|
1100 |
apply(simp_all) |
|
1101 |
apply(erule Prf_elims) |
|
1102 |
apply(case_tac vs1) |
|
1103 |
apply(simp) |
|
1104 |
apply(simp) |
|
1105 |
apply(case_tac x) |
|
1106 |
apply(simp_all) |
|
1107 |
apply(simp add: LV_FROMNTIMES_3 image_def Stars_Cons_def) |
|
1108 |
apply blast |
|
1109 |
done |
|
1110 |
||
1111 |
lemma LV_FROMNTIMES_5: |
|
1112 |
"LV (FROMNTIMES r n) s \<subseteq> Stars_Append (LV (STAR r) s) (\<Union>i\<le>n. LV (FROMNTIMES r i) [])" |
|
1113 |
apply(auto simp add: LV_def) |
|
1114 |
apply(auto elim!: Prf_elims) |
|
1115 |
apply(auto simp add: Stars_Append_def) |
|
1116 |
apply(rule_tac x="vs1" in exI) |
|
1117 |
apply(rule_tac x="vs2" in exI) |
|
1118 |
apply(auto) |
|
1119 |
using Prf.intros(6) apply(auto) |
|
1120 |
apply(rule_tac x="length vs2" in bexI) |
|
1121 |
thm Prf.intros |
|
1122 |
apply(subst append.simps(1)[symmetric]) |
|
1123 |
apply(rule Prf.intros) |
|
1124 |
apply(auto)[1] |
|
1125 |
apply(auto)[1] |
|
1126 |
apply(simp) |
|
1127 |
apply(simp) |
|
1128 |
apply(rule_tac x="vs" in exI) |
|
1129 |
apply(rule_tac x="[]" in exI) |
|
1130 |
apply(auto) |
|
1131 |
by (metis Prf.intros(9) append_Nil atMost_iff empty_iff le_imp_less_Suc less_antisym list.set(1) nth_mem zero_le) |
|
1132 |
||
1133 |
lemma LV_FROMNTIMES_6: |
|
1134 |
assumes "\<forall>s. finite (LV r s)" |
|
1135 |
shows "finite (LV (FROMNTIMES r n) s)" |
|
275 | 1136 |
apply(rule finite_subset) |
276 | 1137 |
apply(rule LV_FROMNTIMES_5) |
1138 |
apply(rule finite_Stars_Append) |
|
1139 |
apply(rule LV_STAR_finite) |
|
1140 |
apply(rule assms) |
|
1141 |
apply(rule finite_UN_I) |
|
1142 |
apply(auto) |
|
1143 |
by (simp add: assms finite_Stars_Pow ttty) |
|
275 | 1144 |
|
276 | 1145 |
lemma LV_NMTIMES_5: |
1146 |
"LV (NMTIMES r n m) s \<subseteq> Stars_Append (LV (STAR r) s) (\<Union>i\<le>n. LV (FROMNTIMES r i) [])" |
|
1147 |
apply(auto simp add: LV_def) |
|
1148 |
apply(auto elim!: Prf_elims) |
|
1149 |
apply(auto simp add: Stars_Append_def) |
|
1150 |
apply(rule_tac x="vs1" in exI) |
|
1151 |
apply(rule_tac x="vs2" in exI) |
|
1152 |
apply(auto) |
|
1153 |
using Prf.intros(6) apply(auto) |
|
1154 |
apply(rule_tac x="length vs2" in bexI) |
|
1155 |
thm Prf.intros |
|
1156 |
apply(subst append.simps(1)[symmetric]) |
|
1157 |
apply(rule Prf.intros) |
|
1158 |
apply(auto)[1] |
|
1159 |
apply(auto)[1] |
|
1160 |
apply(simp) |
|
1161 |
apply(simp) |
|
1162 |
apply(rule_tac x="vs" in exI) |
|
1163 |
apply(rule_tac x="[]" in exI) |
|
1164 |
apply(auto) |
|
1165 |
by (metis Prf.intros(9) append_Nil atMost_iff empty_iff le_imp_less_Suc less_antisym list.set(1) nth_mem zero_le) |
|
274 | 1166 |
|
276 | 1167 |
lemma LV_NMTIMES_6: |
1168 |
assumes "\<forall>s. finite (LV r s)" |
|
1169 |
shows "finite (LV (NMTIMES r n m) s)" |
|
1170 |
apply(rule finite_subset) |
|
1171 |
apply(rule LV_NMTIMES_5) |
|
1172 |
apply(rule finite_Stars_Append) |
|
1173 |
apply(rule LV_STAR_finite) |
|
1174 |
apply(rule assms) |
|
1175 |
apply(rule finite_UN_I) |
|
1176 |
apply(auto) |
|
1177 |
by (simp add: assms finite_Stars_Pow ttty) |
|
1178 |
||
1179 |
||
273 | 1180 |
lemma LV_finite: |
1181 |
shows "finite (LV r s)" |
|
1182 |
proof(induct r arbitrary: s) |
|
1183 |
case (ZERO s) |
|
1184 |
show "finite (LV ZERO s)" by (simp add: LV_simps) |
|
1185 |
next |
|
1186 |
case (ONE s) |
|
1187 |
show "finite (LV ONE s)" by (simp add: LV_simps) |
|
1188 |
next |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
1189 |
case (CH c s) |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
1190 |
show "finite (LV (CH c) s)" by (simp add: LV_simps) |
273 | 1191 |
next |
1192 |
case (ALT r1 r2 s) |
|
1193 |
then show "finite (LV (ALT r1 r2) s)" by (simp add: LV_simps) |
|
1194 |
next |
|
1195 |
case (SEQ r1 r2 s) |
|
275 | 1196 |
define f where "f \<equiv> \<lambda>(v1, v2). Seq v1 v2" |
1197 |
define S1 where "S1 \<equiv> \<Union>s' \<in> Prefixes s. LV r1 s'" |
|
1198 |
define S2 where "S2 \<equiv> \<Union>s' \<in> Suffixes s. LV r2 s'" |
|
273 | 1199 |
have IHs: "\<And>s. finite (LV r1 s)" "\<And>s. finite (LV r2 s)" by fact+ |
1200 |
then have "finite S1" "finite S2" unfolding S1_def S2_def |
|
1201 |
by (simp_all add: finite_Prefixes finite_Suffixes) |
|
1202 |
moreover |
|
1203 |
have "LV (SEQ r1 r2) s \<subseteq> f ` (S1 \<times> S2)" |
|
1204 |
unfolding f_def S1_def S2_def |
|
275 | 1205 |
unfolding LV_def image_def prefix_def suffix_def |
1206 |
apply (auto elim!: Prf_elims) |
|
1207 |
by (metis (mono_tags, lifting) mem_Collect_eq) |
|
273 | 1208 |
ultimately |
1209 |
show "finite (LV (SEQ r1 r2) s)" |
|
1210 |
by (simp add: finite_subset) |
|
1211 |
next |
|
1212 |
case (STAR r s) |
|
1213 |
then show "finite (LV (STAR r) s)" by (simp add: LV_STAR_finite) |
|
1214 |
next |
|
274 | 1215 |
case (UPNTIMES r n s) |
273 | 1216 |
have "\<And>s. finite (LV r s)" by fact |
274 | 1217 |
then show "finite (LV (UPNTIMES r n) s)" |
1218 |
by (meson LV_STAR_finite LV_UPNTIMES_STAR rev_finite_subset) |
|
1219 |
next |
|
1220 |
case (FROMNTIMES r n s) |
|
1221 |
have "\<And>s. finite (LV r s)" by fact |
|
1222 |
then show "finite (LV (FROMNTIMES r n) s)" |
|
276 | 1223 |
by (simp add: LV_FROMNTIMES_6) |
1224 |
next |
|
1225 |
case (NTIMES r n s) |
|
1226 |
have "\<And>s. finite (LV r s)" by fact |
|
1227 |
then show "finite (LV (NTIMES r n) s)" |
|
1228 |
by (metis (no_types, lifting) LV_NTIMES_4 LV_NTIMES_5 LV_STAR_finite finite_Stars_Append finite_Stars_Pow finite_UN_I finite_atMost finite_subset) |
|
1229 |
next |
|
1230 |
case (NMTIMES r n m s) |
|
1231 |
have "\<And>s. finite (LV r s)" by fact |
|
1232 |
then show "finite (LV (NMTIMES r n m) s)" |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
1233 |
by (simp add: LV_NMTIMES_6) |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
1234 |
next |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
1235 |
case (NOT r s) |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
1236 |
then show ?case sorry |
273 | 1237 |
qed |
1238 |
||
1239 |
||
1240 |
||
1241 |
section {* Our POSIX Definition *} |
|
1242 |
||
1243 |
inductive |
|
1244 |
Posix :: "string \<Rightarrow> rexp \<Rightarrow> val \<Rightarrow> bool" ("_ \<in> _ \<rightarrow> _" [100, 100, 100] 100) |
|
1245 |
where |
|
1246 |
Posix_ONE: "[] \<in> ONE \<rightarrow> Void" |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
1247 |
| Posix_CHAR: "[c] \<in> (CH c) \<rightarrow> (Char c)" |
273 | 1248 |
| Posix_ALT1: "s \<in> r1 \<rightarrow> v \<Longrightarrow> s \<in> (ALT r1 r2) \<rightarrow> (Left v)" |
1249 |
| Posix_ALT2: "\<lbrakk>s \<in> r2 \<rightarrow> v; s \<notin> L(r1)\<rbrakk> \<Longrightarrow> s \<in> (ALT r1 r2) \<rightarrow> (Right v)" |
|
1250 |
| Posix_SEQ: "\<lbrakk>s1 \<in> r1 \<rightarrow> v1; s2 \<in> r2 \<rightarrow> v2; |
|
1251 |
\<not>(\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (s1 @ s\<^sub>3) \<in> L r1 \<and> s\<^sub>4 \<in> L r2)\<rbrakk> \<Longrightarrow> |
|
1252 |
(s1 @ s2) \<in> (SEQ r1 r2) \<rightarrow> (Seq v1 v2)" |
|
1253 |
| Posix_STAR1: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> STAR r \<rightarrow> Stars vs; flat v \<noteq> []; |
|
1254 |
\<not>(\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (s1 @ s\<^sub>3) \<in> L r \<and> s\<^sub>4 \<in> L (STAR r))\<rbrakk> |
|
1255 |
\<Longrightarrow> (s1 @ s2) \<in> STAR r \<rightarrow> Stars (v # vs)" |
|
1256 |
| Posix_STAR2: "[] \<in> STAR r \<rightarrow> Stars []" |
|
276 | 1257 |
| Posix_NTIMES1: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> NTIMES r (n - 1) \<rightarrow> Stars vs; flat v \<noteq> []; 0 < n; |
1258 |
\<not>(\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (s1 @ s\<^sub>3) \<in> L r \<and> s\<^sub>4 \<in> L (NTIMES r (n - 1)))\<rbrakk> |
|
1259 |
\<Longrightarrow> (s1 @ s2) \<in> NTIMES r n \<rightarrow> Stars (v # vs)" |
|
1260 |
| Posix_NTIMES2: "\<lbrakk>\<forall>v \<in> set vs. [] \<in> r \<rightarrow> v; length vs = n\<rbrakk> |
|
1261 |
\<Longrightarrow> [] \<in> NTIMES r n \<rightarrow> Stars vs" |
|
1262 |
| Posix_UPNTIMES1: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> UPNTIMES r (n - 1) \<rightarrow> Stars vs; flat v \<noteq> []; 0 < n; |
|
1263 |
\<not>(\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (s1 @ s\<^sub>3) \<in> L r \<and> s\<^sub>4 \<in> L (UPNTIMES r (n - 1)))\<rbrakk> |
|
1264 |
\<Longrightarrow> (s1 @ s2) \<in> UPNTIMES r n \<rightarrow> Stars (v # vs)" |
|
1265 |
| Posix_UPNTIMES2: "[] \<in> UPNTIMES r n \<rightarrow> Stars []" |
|
1266 |
| Posix_FROMNTIMES2: "\<lbrakk>\<forall>v \<in> set vs. [] \<in> r \<rightarrow> v; length vs = n\<rbrakk> |
|
1267 |
\<Longrightarrow> [] \<in> FROMNTIMES r n \<rightarrow> Stars vs" |
|
1268 |
| Posix_FROMNTIMES1: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> FROMNTIMES r (n - 1) \<rightarrow> Stars vs; flat v \<noteq> []; 0 < n; |
|
1269 |
\<not>(\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (s1 @ s\<^sub>3) \<in> L r \<and> s\<^sub>4 \<in> L (FROMNTIMES r (n - 1)))\<rbrakk> |
|
1270 |
\<Longrightarrow> (s1 @ s2) \<in> FROMNTIMES r n \<rightarrow> Stars (v # vs)" |
|
277 | 1271 |
| Posix_FROMNTIMES3: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> STAR r \<rightarrow> Stars vs; flat v \<noteq> []; |
1272 |
\<not>(\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (s1 @ s\<^sub>3) \<in> L r \<and> s\<^sub>4 \<in> L (STAR r))\<rbrakk> |
|
1273 |
\<Longrightarrow> (s1 @ s2) \<in> FROMNTIMES r 0 \<rightarrow> Stars (v # vs)" |
|
276 | 1274 |
| Posix_NMTIMES2: "\<lbrakk>\<forall>v \<in> set vs. [] \<in> r \<rightarrow> v; length vs = n; n \<le> m\<rbrakk> |
1275 |
\<Longrightarrow> [] \<in> NMTIMES r n m \<rightarrow> Stars vs" |
|
278 | 1276 |
| Posix_NMTIMES1: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> NMTIMES r (n - 1) (m - 1) \<rightarrow> Stars vs; flat v \<noteq> []; 0 < n; n \<le> m; |
1277 |
\<not>(\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (s1 @ s\<^sub>3) \<in> L r \<and> s\<^sub>4 \<in> L (NMTIMES r (n - 1) (m - 1)))\<rbrakk> |
|
1278 |
\<Longrightarrow> (s1 @ s2) \<in> NMTIMES r n m \<rightarrow> Stars (v # vs)" |
|
1279 |
| Posix_NMTIMES3: "\<lbrakk>s1 \<in> r \<rightarrow> v; s2 \<in> UPNTIMES r (m - 1) \<rightarrow> Stars vs; flat v \<noteq> []; 0 < m; |
|
1280 |
\<not>(\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> (s1 @ s\<^sub>3) \<in> L r \<and> s\<^sub>4 \<in> L (UPNTIMES r (m - 1)))\<rbrakk> |
|
1281 |
\<Longrightarrow> (s1 @ s2) \<in> NMTIMES r 0 m \<rightarrow> Stars (v # vs)" |
|
276 | 1282 |
|
273 | 1283 |
inductive_cases Posix_elims: |
1284 |
"s \<in> ZERO \<rightarrow> v" |
|
1285 |
"s \<in> ONE \<rightarrow> v" |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
1286 |
"s \<in> CH c \<rightarrow> v" |
273 | 1287 |
"s \<in> ALT r1 r2 \<rightarrow> v" |
1288 |
"s \<in> SEQ r1 r2 \<rightarrow> v" |
|
1289 |
"s \<in> STAR r \<rightarrow> v" |
|
276 | 1290 |
"s \<in> NTIMES r n \<rightarrow> v" |
1291 |
"s \<in> UPNTIMES r n \<rightarrow> v" |
|
1292 |
"s \<in> FROMNTIMES r n \<rightarrow> v" |
|
1293 |
"s \<in> NMTIMES r n m \<rightarrow> v" |
|
1294 |
||
273 | 1295 |
lemma Posix1: |
1296 |
assumes "s \<in> r \<rightarrow> v" |
|
1297 |
shows "s \<in> L r" "flat v = s" |
|
1298 |
using assms |
|
276 | 1299 |
apply(induct s r v rule: Posix.induct) |
1300 |
apply(auto simp add: Sequ_def)[18] |
|
1301 |
apply(case_tac n) |
|
1302 |
apply(simp) |
|
1303 |
apply(simp add: Sequ_def) |
|
1304 |
apply(auto)[1] |
|
1305 |
apply(simp) |
|
1306 |
apply(clarify) |
|
1307 |
apply(rule_tac x="Suc x" in bexI) |
|
1308 |
apply(simp add: Sequ_def) |
|
1309 |
apply(auto)[5] |
|
1310 |
using nth_mem nullable.simps(9) nullable_correctness apply auto[1] |
|
1311 |
apply simp |
|
1312 |
apply(simp) |
|
1313 |
apply(clarify) |
|
1314 |
apply(rule_tac x="Suc x" in bexI) |
|
1315 |
apply(simp add: Sequ_def) |
|
1316 |
apply(auto)[3] |
|
277 | 1317 |
defer |
276 | 1318 |
apply(simp) |
1319 |
apply fastforce |
|
1320 |
apply(simp) |
|
1321 |
apply(simp) |
|
1322 |
apply(clarify) |
|
1323 |
apply(rule_tac x="Suc x" in bexI) |
|
1324 |
apply(auto simp add: Sequ_def)[2] |
|
1325 |
apply(simp) |
|
278 | 1326 |
apply(simp) |
1327 |
apply(clarify) |
|
1328 |
apply(rule_tac x="Suc x" in bexI) |
|
1329 |
apply(auto simp add: Sequ_def)[2] |
|
1330 |
apply(simp) |
|
1331 |
apply(simp add: Star.step Star_Pow) |
|
1332 |
done |
|
1333 |
||
273 | 1334 |
text {* |
1335 |
Our Posix definition determines a unique value. |
|
1336 |
*} |
|
276 | 1337 |
|
1338 |
lemma List_eq_zipI: |
|
1339 |
assumes "\<forall>(v1, v2) \<in> set (zip vs1 vs2). v1 = v2" |
|
1340 |
and "length vs1 = length vs2" |
|
1341 |
shows "vs1 = vs2" |
|
1342 |
using assms |
|
1343 |
apply(induct vs1 arbitrary: vs2) |
|
1344 |
apply(case_tac vs2) |
|
1345 |
apply(simp) |
|
1346 |
apply(simp) |
|
1347 |
apply(case_tac vs2) |
|
1348 |
apply(simp) |
|
1349 |
apply(simp) |
|
1350 |
done |
|
273 | 1351 |
|
1352 |
lemma Posix_determ: |
|
1353 |
assumes "s \<in> r \<rightarrow> v1" "s \<in> r \<rightarrow> v2" |
|
1354 |
shows "v1 = v2" |
|
1355 |
using assms |
|
1356 |
proof (induct s r v1 arbitrary: v2 rule: Posix.induct) |
|
1357 |
case (Posix_ONE v2) |
|
1358 |
have "[] \<in> ONE \<rightarrow> v2" by fact |
|
1359 |
then show "Void = v2" by cases auto |
|
1360 |
next |
|
1361 |
case (Posix_CHAR c v2) |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
1362 |
have "[c] \<in> CH c \<rightarrow> v2" by fact |
273 | 1363 |
then show "Char c = v2" by cases auto |
1364 |
next |
|
1365 |
case (Posix_ALT1 s r1 v r2 v2) |
|
1366 |
have "s \<in> ALT r1 r2 \<rightarrow> v2" by fact |
|
1367 |
moreover |
|
1368 |
have "s \<in> r1 \<rightarrow> v" by fact |
|
1369 |
then have "s \<in> L r1" by (simp add: Posix1) |
|
1370 |
ultimately obtain v' where eq: "v2 = Left v'" "s \<in> r1 \<rightarrow> v'" by cases auto |
|
1371 |
moreover |
|
1372 |
have IH: "\<And>v2. s \<in> r1 \<rightarrow> v2 \<Longrightarrow> v = v2" by fact |
|
1373 |
ultimately have "v = v'" by simp |
|
1374 |
then show "Left v = v2" using eq by simp |
|
1375 |
next |
|
1376 |
case (Posix_ALT2 s r2 v r1 v2) |
|
1377 |
have "s \<in> ALT r1 r2 \<rightarrow> v2" by fact |
|
1378 |
moreover |
|
1379 |
have "s \<notin> L r1" by fact |
|
1380 |
ultimately obtain v' where eq: "v2 = Right v'" "s \<in> r2 \<rightarrow> v'" |
|
1381 |
by cases (auto simp add: Posix1) |
|
1382 |
moreover |
|
1383 |
have IH: "\<And>v2. s \<in> r2 \<rightarrow> v2 \<Longrightarrow> v = v2" by fact |
|
1384 |
ultimately have "v = v'" by simp |
|
1385 |
then show "Right v = v2" using eq by simp |
|
1386 |
next |
|
1387 |
case (Posix_SEQ s1 r1 v1 s2 r2 v2 v') |
|
1388 |
have "(s1 @ s2) \<in> SEQ r1 r2 \<rightarrow> v'" |
|
1389 |
"s1 \<in> r1 \<rightarrow> v1" "s2 \<in> r2 \<rightarrow> v2" |
|
1390 |
"\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r1 \<and> s\<^sub>4 \<in> L r2)" by fact+ |
|
1391 |
then obtain v1' v2' where "v' = Seq v1' v2'" "s1 \<in> r1 \<rightarrow> v1'" "s2 \<in> r2 \<rightarrow> v2'" |
|
1392 |
apply(cases) apply (auto simp add: append_eq_append_conv2) |
|
1393 |
using Posix1(1) by fastforce+ |
|
1394 |
moreover |
|
1395 |
have IHs: "\<And>v1'. s1 \<in> r1 \<rightarrow> v1' \<Longrightarrow> v1 = v1'" |
|
1396 |
"\<And>v2'. s2 \<in> r2 \<rightarrow> v2' \<Longrightarrow> v2 = v2'" by fact+ |
|
1397 |
ultimately show "Seq v1 v2 = v'" by simp |
|
1398 |
next |
|
1399 |
case (Posix_STAR1 s1 r v s2 vs v2) |
|
1400 |
have "(s1 @ s2) \<in> STAR r \<rightarrow> v2" |
|
1401 |
"s1 \<in> r \<rightarrow> v" "s2 \<in> STAR r \<rightarrow> Stars vs" "flat v \<noteq> []" |
|
1402 |
"\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (STAR r))" by fact+ |
|
1403 |
then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" "s2 \<in> (STAR r) \<rightarrow> (Stars vs')" |
|
1404 |
apply(cases) apply (auto simp add: append_eq_append_conv2) |
|
1405 |
using Posix1(1) apply fastforce |
|
1406 |
apply (metis Posix1(1) Posix_STAR1.hyps(6) append_Nil append_Nil2) |
|
1407 |
using Posix1(2) by blast |
|
1408 |
moreover |
|
1409 |
have IHs: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2" |
|
1410 |
"\<And>v2. s2 \<in> STAR r \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+ |
|
1411 |
ultimately show "Stars (v # vs) = v2" by auto |
|
1412 |
next |
|
1413 |
case (Posix_STAR2 r v2) |
|
1414 |
have "[] \<in> STAR r \<rightarrow> v2" by fact |
|
1415 |
then show "Stars [] = v2" by cases (auto simp add: Posix1) |
|
276 | 1416 |
next |
1417 |
case (Posix_NTIMES2 vs r n v2) |
|
1418 |
then show "Stars vs = v2" |
|
1419 |
apply(erule_tac Posix_elims) |
|
1420 |
apply(auto) |
|
1421 |
apply (simp add: Posix1(2)) |
|
1422 |
apply(rule List_eq_zipI) |
|
1423 |
apply(auto) |
|
1424 |
by (meson in_set_zipE) |
|
1425 |
next |
|
1426 |
case (Posix_NTIMES1 s1 r v s2 n vs v2) |
|
1427 |
have "(s1 @ s2) \<in> NTIMES r n \<rightarrow> v2" |
|
1428 |
"s1 \<in> r \<rightarrow> v" "s2 \<in> NTIMES r (n - 1) \<rightarrow> Stars vs" "flat v \<noteq> []" |
|
1429 |
"\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (NTIMES r (n - 1 )))" by fact+ |
|
1430 |
then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" "s2 \<in> (NTIMES r (n - 1)) \<rightarrow> (Stars vs')" |
|
1431 |
apply(cases) apply (auto simp add: append_eq_append_conv2) |
|
1432 |
using Posix1(1) apply fastforce |
|
1433 |
apply (metis One_nat_def Posix1(1) Posix_NTIMES1.hyps(7) append.right_neutral append_self_conv2) |
|
1434 |
using Posix1(2) by blast |
|
1435 |
moreover |
|
1436 |
have IHs: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2" |
|
1437 |
"\<And>v2. s2 \<in> NTIMES r (n - 1) \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+ |
|
1438 |
ultimately show "Stars (v # vs) = v2" by auto |
|
1439 |
next |
|
1440 |
case (Posix_UPNTIMES1 s1 r v s2 n vs v2) |
|
1441 |
have "(s1 @ s2) \<in> UPNTIMES r n \<rightarrow> v2" |
|
1442 |
"s1 \<in> r \<rightarrow> v" "s2 \<in> UPNTIMES r (n - 1) \<rightarrow> Stars vs" "flat v \<noteq> []" |
|
1443 |
"\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (UPNTIMES r (n - 1 )))" by fact+ |
|
1444 |
then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" "s2 \<in> (UPNTIMES r (n - 1)) \<rightarrow> (Stars vs')" |
|
1445 |
apply(cases) apply (auto simp add: append_eq_append_conv2) |
|
1446 |
using Posix1(1) apply fastforce |
|
1447 |
apply (metis One_nat_def Posix1(1) Posix_UPNTIMES1.hyps(7) append.right_neutral append_self_conv2) |
|
1448 |
using Posix1(2) by blast |
|
1449 |
moreover |
|
1450 |
have IHs: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2" |
|
1451 |
"\<And>v2. s2 \<in> UPNTIMES r (n - 1) \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+ |
|
1452 |
ultimately show "Stars (v # vs) = v2" by auto |
|
1453 |
next |
|
1454 |
case (Posix_UPNTIMES2 r n v2) |
|
1455 |
then show "Stars [] = v2" |
|
1456 |
apply(erule_tac Posix_elims) |
|
1457 |
apply(auto) |
|
1458 |
by (simp add: Posix1(2)) |
|
1459 |
next |
|
1460 |
case (Posix_FROMNTIMES1 s1 r v s2 n vs v2) |
|
1461 |
have "(s1 @ s2) \<in> FROMNTIMES r n \<rightarrow> v2" |
|
277 | 1462 |
"s1 \<in> r \<rightarrow> v" "s2 \<in> FROMNTIMES r (n - 1) \<rightarrow> Stars vs" "flat v \<noteq> []" "0 < n" |
276 | 1463 |
"\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (FROMNTIMES r (n - 1 )))" by fact+ |
1464 |
then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" "s2 \<in> (FROMNTIMES r (n - 1)) \<rightarrow> (Stars vs')" |
|
1465 |
apply(cases) apply (auto simp add: append_eq_append_conv2) |
|
1466 |
using Posix1(1) Posix1(2) apply blast |
|
1467 |
apply(case_tac n) |
|
1468 |
apply(simp) |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
1469 |
apply(simp) |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
1470 |
apply (smt (verit, ccfv_threshold) L.simps(9) Posix1(1) UN_E append_eq_append_conv2) |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
1471 |
by (metis One_nat_def Posix1(1) Posix_FROMNTIMES1.hyps(7) append_Nil2 append_self_conv2) |
276 | 1472 |
moreover |
1473 |
have IHs: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2" |
|
1474 |
"\<And>v2. s2 \<in> FROMNTIMES r (n - 1) \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+ |
|
1475 |
ultimately show "Stars (v # vs) = v2" by auto |
|
1476 |
next |
|
1477 |
case (Posix_FROMNTIMES2 vs r n v2) |
|
1478 |
then show "Stars vs = v2" |
|
1479 |
apply(erule_tac Posix_elims) |
|
1480 |
apply(auto) |
|
1481 |
apply(rule List_eq_zipI) |
|
1482 |
apply(auto) |
|
277 | 1483 |
apply(meson in_set_zipE) |
1484 |
apply (simp add: Posix1(2)) |
|
1485 |
using Posix1(2) by blast |
|
276 | 1486 |
next |
277 | 1487 |
case (Posix_FROMNTIMES3 s1 r v s2 vs v2) |
1488 |
have "(s1 @ s2) \<in> FROMNTIMES r 0 \<rightarrow> v2" |
|
1489 |
"s1 \<in> r \<rightarrow> v" "s2 \<in> STAR r \<rightarrow> Stars vs" "flat v \<noteq> []" |
|
1490 |
"\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (STAR r))" by fact+ |
|
1491 |
then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" "s2 \<in> (STAR r) \<rightarrow> (Stars vs')" |
|
1492 |
apply(cases) apply (auto simp add: append_eq_append_conv2) |
|
1493 |
using Posix1(2) apply fastforce |
|
1494 |
using Posix1(1) apply fastforce |
|
1495 |
by (metis Posix1(1) Posix_FROMNTIMES3.hyps(6) append.right_neutral append_Nil) |
|
1496 |
moreover |
|
1497 |
have IHs: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2" |
|
1498 |
"\<And>v2. s2 \<in> STAR r \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+ |
|
1499 |
ultimately show "Stars (v # vs) = v2" by auto |
|
1500 |
next |
|
276 | 1501 |
case (Posix_NMTIMES1 s1 r v s2 n m vs v2) |
278 | 1502 |
have "(s1 @ s2) \<in> NMTIMES r n m \<rightarrow> v2" |
1503 |
"s1 \<in> r \<rightarrow> v" "s2 \<in> NMTIMES r (n - 1) (m - 1) \<rightarrow> Stars vs" "flat v \<noteq> []" |
|
1504 |
"0 < n" "n \<le> m" |
|
1505 |
"\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (NMTIMES r (n - 1) (m - 1)))" by fact+ |
|
1506 |
then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" |
|
1507 |
"s2 \<in> (NMTIMES r (n - 1) (m - 1)) \<rightarrow> (Stars vs')" |
|
1508 |
apply(cases) apply (auto simp add: append_eq_append_conv2) |
|
1509 |
using Posix1(1) Posix1(2) apply blast |
|
1510 |
apply(case_tac n) |
|
1511 |
apply(simp) |
|
1512 |
apply(simp) |
|
1513 |
apply(case_tac m) |
|
1514 |
apply(simp) |
|
1515 |
apply(simp) |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
1516 |
apply (metis L.simps(10) Posix1(1) UN_E append.right_neutral append_Nil) |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
359
diff
changeset
|
1517 |
by (metis One_nat_def Posix1(1) Posix_NMTIMES1.hyps(8) append_Nil self_append_conv) |
278 | 1518 |
moreover |
1519 |
have IHs: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2" |
|
1520 |
"\<And>v2. s2 \<in> NMTIMES r (n - 1) (m - 1) \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+ |
|
1521 |
ultimately show "Stars (v # vs) = v2" by auto |
|
276 | 1522 |
next |
1523 |
case (Posix_NMTIMES2 vs r n m v2) |
|
1524 |
then show "Stars vs = v2" |
|
278 | 1525 |
apply(erule_tac Posix_elims) |
1526 |
apply(simp) |
|
1527 |
apply(rule List_eq_zipI) |
|
1528 |
apply(auto) |
|
1529 |
apply (meson in_set_zipE) |
|
1530 |
apply (simp add: Posix1(2)) |
|
1531 |
apply(erule_tac Posix_elims) |
|
1532 |
apply(auto) |
|
1533 |
apply (simp add: Posix1(2))+ |
|
1534 |
done |
|
1535 |
next |
|
1536 |
case (Posix_NMTIMES3 s1 r v s2 m vs v2) |
|
1537 |
have "(s1 @ s2) \<in> NMTIMES r 0 m \<rightarrow> v2" |
|
1538 |
"s1 \<in> r \<rightarrow> v" "s2 \<in> UPNTIMES r (m - 1) \<rightarrow> Stars vs" "flat v \<noteq> []" "0 < m" |
|
1539 |
"\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (UPNTIMES r (m - 1 )))" by fact+ |
|
1540 |
then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \<in> r \<rightarrow> v'" "s2 \<in> (UPNTIMES r (m - 1)) \<rightarrow> (Stars vs')" |
|
1541 |
apply(cases) apply (auto simp add: append_eq_append_conv2) |
|
1542 |
using Posix1(2) apply blast |
|
1543 |
apply (smt L.simps(7) Posix1(1) UN_E append_eq_append_conv2) |
|
1544 |
by (metis One_nat_def Posix1(1) Posix_NMTIMES3.hyps(7) append.right_neutral append_Nil) |
|
1545 |
moreover |
|
1546 |
have IHs: "\<And>v2. s1 \<in> r \<rightarrow> v2 \<Longrightarrow> v = v2" |
|
1547 |
"\<And>v2. s2 \<in> UPNTIMES r (m - 1) \<rightarrow> v2 \<Longrightarrow> Stars vs = v2" by fact+ |
|
1548 |
ultimately show "Stars (v # vs) = v2" by auto |
|
273 | 1549 |
qed |
1550 |
||
1551 |
||
1552 |
text {* |
|
1553 |
Our POSIX value is a lexical value. |
|
1554 |
*} |
|
1555 |
||
1556 |
lemma Posix_LV: |
|
1557 |
assumes "s \<in> r \<rightarrow> v" |
|
1558 |
shows "v \<in> LV r s" |
|
1559 |
using assms unfolding LV_def |
|
1560 |
apply(induct rule: Posix.induct) |
|
276 | 1561 |
apply(auto simp add: intro!: Prf.intros elim!: Prf_elims)[7] |
1562 |
defer |
|
1563 |
defer |
|
1564 |
apply(auto simp add: intro!: Prf.intros elim!: Prf_elims)[2] |
|
1565 |
apply (metis (mono_tags, lifting) Prf.intros(9) append_Nil empty_iff flat_Stars flats_empty list.set(1) mem_Collect_eq) |
|
277 | 1566 |
apply(simp) |
1567 |
apply(case_tac n) |
|
1568 |
apply(simp) |
|
1569 |
apply(simp) |
|
1570 |
apply(erule Prf_elims) |
|
1571 |
apply(simp) |
|
276 | 1572 |
apply(subst append.simps(2)[symmetric]) |
277 | 1573 |
apply(rule Prf.intros) |
1574 |
apply(simp) |
|
1575 |
apply(simp) |
|
1576 |
apply(simp) |
|
1577 |
apply(simp) |
|
1578 |
apply(rule Prf.intros) |
|
1579 |
apply(simp) |
|
1580 |
apply(simp) |
|
1581 |
apply(simp) |
|
1582 |
apply(erule Prf_elims) |
|
1583 |
apply(simp) |
|
1584 |
apply(rule Prf.intros) |
|
1585 |
apply(simp) |
|
278 | 1586 |
apply(simp) |
1587 |
(* NTIMES *) |
|
1588 |
prefer 4 |
|
1589 |
apply(simp) |
|
1590 |
apply(case_tac n) |
|
277 | 1591 |
apply(simp) |
278 | 1592 |
apply(simp) |
1593 |
apply(clarify) |
|
1594 |
apply(rotate_tac 5) |
|
1595 |
apply(erule Prf_elims) |
|
1596 |
apply(simp) |
|
1597 |
apply(subst append.simps(2)[symmetric]) |
|
1598 |
apply(rule Prf.intros) |
|
1599 |
apply(simp) |
|
1600 |
apply(simp) |
|
1601 |
apply(simp) |
|
1602 |
prefer 4 |
|
1603 |
apply(simp) |
|
1604 |
apply (metis Prf.intros(8) length_removeAll_less less_irrefl_nat removeAll.simps(1) self_append_conv2) |
|
277 | 1605 |
(* NMTIMES *) |
278 | 1606 |
apply(simp) |
1607 |
apply (metis Prf.intros(11) append_Nil empty_iff list.set(1)) |
|
1608 |
apply(simp) |
|
1609 |
apply(rotate_tac 6) |
|
1610 |
apply(erule Prf_elims) |
|
1611 |
apply(simp) |
|
1612 |
apply(subst append.simps(2)[symmetric]) |
|
1613 |
apply(rule Prf.intros) |
|
1614 |
apply(simp) |
|
1615 |
apply(simp) |
|
1616 |
apply(simp) |
|
1617 |
apply(simp) |
|
1618 |
apply(rule Prf.intros) |
|
1619 |
apply(simp) |
|
1620 |
apply(simp) |
|
1621 |
apply(simp) |
|
1622 |
apply(simp) |
|
1623 |
apply(rotate_tac 6) |
|
1624 |
apply(erule Prf_elims) |
|
1625 |
apply(simp) |
|
1626 |
apply(rule Prf.intros) |
|
1627 |
apply(simp) |
|
1628 |
apply(simp) |
|
1629 |
apply(simp) |
|
1630 |
done |
|
276 | 1631 |
|
273 | 1632 |
end |