|
1 (***************************************************************** |
|
2 |
|
3 Isabelle Tutorial |
|
4 ----------------- |
|
5 |
|
6 2st June 2009, Beijing |
|
7 |
|
8 *) |
|
9 |
|
10 theory Lec3 |
|
11 imports "Main" |
|
12 begin |
|
13 |
|
14 |
|
15 definition |
|
16 lang_seq :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ; _") |
|
17 where |
|
18 "L1 ; L2 = {s1@s2 | s1 s2. s1 \<in> L1 \<and> s2 \<in> L2}" |
|
19 |
|
20 fun |
|
21 lang_pow :: "string set \<Rightarrow> nat \<Rightarrow> string set" ("_ \<up> _") |
|
22 where |
|
23 "L \<up> 0 = {[]}" |
|
24 | "L \<up> (Suc i) = L ; (L \<up> i)" |
|
25 |
|
26 definition |
|
27 lang_star :: "string set \<Rightarrow> string set" ("_\<star>") |
|
28 where |
|
29 "L\<star> \<equiv> \<Union>i. (L \<up> i)" |
|
30 |
|
31 lemma lang_seq_cases: |
|
32 shows "(s \<in> L1 ; L2) = (\<exists>s1 s2. s = s1@s2 \<and> s1\<in>L1 \<and> s2\<in>L2)" |
|
33 by (simp add: lang_seq_def) |
|
34 |
|
35 lemma lang_seq_union: |
|
36 shows "(L1 \<union> L2);L3 = (L1;L3) \<union> (L2;L3)" |
|
37 and "L1;(L2 \<union> L3) = (L1;L2) \<union> (L1;L3)" |
|
38 unfolding lang_seq_def by auto |
|
39 |
|
40 lemma lang_seq_empty: |
|
41 shows "{[]} ; L = L" |
|
42 and "L ; {[]} = L" |
|
43 unfolding lang_seq_def by auto |
|
44 |
|
45 lemma lang_seq_assoc: |
|
46 shows "(L1 ; L2) ; L3 = L1 ; (L2 ; L3)" |
|
47 by (simp add: lang_seq_def Collect_def mem_def expand_fun_eq) |
|
48 (metis append_assoc) |
|
49 |
|
50 lemma silly: |
|
51 shows "[] \<in> L \<up> 0" |
|
52 by simp |
|
53 |
|
54 lemma lang_star_empty: |
|
55 shows "{[]} \<union> (L\<star>) = L\<star>" |
|
56 unfolding lang_star_def |
|
57 by (auto intro: silly) |
|
58 |
|
59 lemma lang_star_in_empty: |
|
60 shows "[] \<in> L\<star>" |
|
61 unfolding lang_star_def |
|
62 by (auto intro: silly) |
|
63 |
|
64 lemma lang_seq_subseteq: |
|
65 shows "L \<subseteq> (L'\<star>) ; L" |
|
66 and "L \<subseteq> L ; (L'\<star>)" |
|
67 proof - |
|
68 have "L = {[]} ; L" using lang_seq_empty by simp |
|
69 also have "\<dots> \<subseteq> ({[]} ; L) \<union> ((L'\<star>) ; L)" by auto |
|
70 also have "\<dots> = ({[]} \<union> (L'\<star>)) ; L" by (simp add: lang_seq_union[symmetric]) |
|
71 also have "\<dots> = (L'\<star>); L" using lang_star_empty by simp |
|
72 finally show "L \<subseteq> (L'\<star>); L" by simp |
|
73 next |
|
74 show "L \<subseteq> L ; (L'\<star>)" sorry |
|
75 qed |
|
76 |
|
77 lemma lang_star_subseteq: |
|
78 shows "L ; (L\<star>) \<subseteq> (L\<star>)" |
|
79 unfolding lang_star_def lang_seq_def |
|
80 apply(auto) |
|
81 apply(rule_tac x="Suc xa" in exI) |
|
82 apply(auto simp add: lang_seq_def) |
|
83 done |
|
84 |
|
85 (* regular expressions *) |
|
86 |
|
87 datatype rexp = |
|
88 EMPTY |
|
89 | CHAR char |
|
90 | SEQ rexp rexp |
|
91 | ALT rexp rexp |
|
92 | STAR rexp |
|
93 |
|
94 fun |
|
95 L :: "rexp \<Rightarrow> string set" |
|
96 where |
|
97 "L(EMPTY) = {[]}" |
|
98 | "L(CHAR c) = {[c]}" |
|
99 | "L(SEQ r1 r2) = (L r1) ; (L r2)" |
|
100 | "L(ALT r1 r2) = (L r1) \<union> (L r2)" |
|
101 | "L(STAR r) = (L r)\<star>" |
|
102 |
|
103 definition |
|
104 Ls :: "rexp set \<Rightarrow> string set" |
|
105 where |
|
106 "Ls R = (\<Union>r\<in>R. (L r))" |
|
107 |
|
108 lemma |
|
109 shows "Ls {} = {}" |
|
110 unfolding Ls_def by simp |
|
111 |
|
112 lemma Ls_union: |
|
113 "Ls (R1 \<union> R2) = (Ls R1) \<union> (Ls R2)" |
|
114 unfolding Ls_def by auto |
|
115 |
|
116 function |
|
117 dagger :: "rexp \<Rightarrow> char \<Rightarrow> rexp set" ("_ \<dagger> _") |
|
118 where |
|
119 r1: "(EMPTY) \<dagger> c = {}" |
|
120 | r2: "(CHAR c') \<dagger> c = (if c = c' then {EMPTY} else {})" |
|
121 | r3: "(ALT r1 r2) \<dagger> c = r1 \<dagger> c \<union> r2 \<dagger> c" |
|
122 | r4: "(SEQ EMPTY r2) \<dagger> c = r2 \<dagger> c" |
|
123 | r5: "(SEQ (CHAR c') r2) \<dagger> c = (if c= c' then {r2} else {})" |
|
124 | r6: "(SEQ (SEQ r11 r12) r2) \<dagger> c = (SEQ r11 (SEQ r12 r2)) \<dagger> c" |
|
125 | r7: "(SEQ (ALT r11 r12) r2) \<dagger> c = (SEQ r11 r2) \<dagger> c \<union> (SEQ r12 r2) \<dagger> c" |
|
126 | r8: "(SEQ (STAR r1) r2) \<dagger> c = |
|
127 r2 \<dagger> c \<union> {SEQ (SEQ r' (STAR r1)) r2 | r'. r' \<in> r1 \<dagger> c}" |
|
128 | r9: "(STAR r) \<dagger> c = {SEQ r' (STAR r) | r'. r' \<in> r \<dagger> c}" |
|
129 by (pat_completeness) (auto) |
|
130 |
|
131 termination |
|
132 dagger sorry |
|
133 |
|
134 definition |
|
135 OR :: "bool set \<Rightarrow> bool" |
|
136 where |
|
137 "OR S \<equiv> (\<exists>b\<in>S. b)" |
|
138 |
|
139 function |
|
140 matcher :: "rexp \<Rightarrow> string \<Rightarrow> bool" ("_ ! _") |
|
141 where |
|
142 s01: "EMPTY ! s = (s =[])" |
|
143 | s02: "CHAR c ! s = (s = [c])" |
|
144 | s03: "ALT r1 r2 ! s = (r1 ! s \<or> r2 ! s)" |
|
145 | s04: "STAR r ! [] = True" |
|
146 | s05: "STAR r ! c#s = (False \<or> OR {SEQ (r') (STAR r)!s | r'. r' \<in> r \<dagger> c})" |
|
147 | s06: "SEQ r1 r2 ! [] = (r1 ! [] \<and> r2 ! [])" |
|
148 | s07: "SEQ EMPTY r2 ! (c#s) = (r2 ! c#s)" |
|
149 | s08: "SEQ (CHAR c') r2 ! (c#s) = (if c'=c then r2 ! s else False)" |
|
150 | s09: "SEQ (SEQ r11 r12) r2 ! (c#s) = (SEQ r11 (SEQ r12 r2) ! c#s)" |
|
151 | s10: "SEQ (ALT r11 r12) r2 ! (c#s) = ((SEQ r11 r2) ! (c#s) \<or> (SEQ r12 r2) ! (c#s))" |
|
152 | s11: "SEQ (STAR r1) r2 ! (c#s) = |
|
153 (r2 ! (c#s) \<or> OR {SEQ r' (SEQ (STAR r1) r2) ! s | r'. r' \<in> r1 \<dagger> c})" |
|
154 by (pat_completeness) (auto) |
|
155 |
|
156 termination |
|
157 matcher sorry |
|
158 |
|
159 lemma "(CHAR a) ! [a]" by auto |
|
160 lemma "\<not>(CHAR a) ! [a,a]" by auto |
|
161 lemma "(STAR (CHAR a)) ! []" by auto |
|
162 lemma "(STAR (CHAR a)) ! [a,a]" by (auto simp add: OR_def) |
|
163 lemma "(SEQ (CHAR a) (SEQ (STAR (CHAR b)) (CHAR c))) ! [a,b,b,b,c]" |
|
164 by (auto simp add: OR_def) |
|
165 |
|
166 lemma holes: |
|
167 assumes a: "Ls (r \<dagger> c) = {s. c#s \<in> L r}" |
|
168 shows "Ls (r \<dagger> c) ; L (STAR r) = {s''. c#s'' \<in> L (STAR r)}" |
|
169 proof - |
|
170 have "Ls (r \<dagger> c) ; L (STAR r) = {s. c#s \<in> L r} ; L (STAR r)" by (simp add: a) |
|
171 also have "\<dots> = {s'. c#s' \<in> (L r ; L (STAR r))}" sorry |
|
172 also have "\<dots> = {s''. c#s'' \<in> L (STAR r)}" sorry |
|
173 finally show "Ls (r \<dagger> c) ; L (STAR r) = {s''. c#s'' \<in> L (STAR r)}" by simp |
|
174 qed |
|
175 |
|
176 lemma eq: |
|
177 shows "Ls (STAR r) \<dagger> c = (Ls (r \<dagger> c) ; L (STAR r))" |
|
178 proof |
|
179 show "Ls STAR r \<dagger> c \<subseteq> Ls r \<dagger> c ; L (STAR r)" |
|
180 by (auto simp add: lang_star_def lang_seq_def Ls_def) (blast) |
|
181 next |
|
182 show "Ls r \<dagger> c ; L (STAR r) \<subseteq> Ls STAR r \<dagger> c" |
|
183 apply(auto simp add: lang_star_def lang_seq_def Ls_def) |
|
184 apply(rule_tac x="SEQ xa (STAR r)" in exI) |
|
185 apply(simp add: lang_star_def lang_seq_def) |
|
186 apply(blast) |
|
187 done |
|
188 qed |
|
189 |
|
190 (* correctness of the matcher *) |
|
191 lemma dagger_holes: |
|
192 "Ls (r \<dagger> c) = {s. c#s \<in> L r}" |
|
193 proof (induct rule: dagger.induct) |
|
194 case (1 c) |
|
195 show "Ls (EMPTY \<dagger> c) = {s. c#s \<in> L EMPTY}" |
|
196 by (simp add: Ls_def) |
|
197 next |
|
198 case (2 c' c) |
|
199 show "Ls (CHAR c') \<dagger> c = {s. c#s \<in> L (CHAR c')}" |
|
200 proof (cases "c=c'") |
|
201 assume "c=c'" |
|
202 then show "Ls (CHAR c') \<dagger> c = {s. c#s \<in> L (CHAR c')}" |
|
203 by (simp add: Ls_def) |
|
204 next |
|
205 assume "c\<noteq>c'" |
|
206 then show "Ls (CHAR c') \<dagger> c = {s. c#s \<in> L (CHAR c')}" |
|
207 by (simp add: Ls_def) |
|
208 qed |
|
209 next |
|
210 case (3 r1 r2 c) |
|
211 have ih1: "Ls r1 \<dagger> c = {s. c#s \<in> L r1}" by fact |
|
212 have ih2: "Ls r2 \<dagger> c = {s. c#s \<in> L r2}" by fact |
|
213 show "Ls (ALT r1 r2) \<dagger> c = {s. c#s \<in> L (ALT r1 r2)}" |
|
214 by (auto simp add: Ls_union ih1 ih2) |
|
215 next |
|
216 case (4 r2 c) |
|
217 have ih: "Ls r2 \<dagger> c = {s. c#s \<in> L r2}" by fact |
|
218 show "Ls (SEQ EMPTY r2) \<dagger> c = {s. c#s \<in> L (SEQ EMPTY r2)}" |
|
219 by (simp add: ih lang_seq_empty) |
|
220 next |
|
221 case (5 c' r2 c) |
|
222 show "Ls (SEQ (CHAR c') r2) \<dagger> c = {s. c#s \<in> L (SEQ (CHAR c') r2)}" |
|
223 proof (cases "c=c'") |
|
224 assume "c=c'" |
|
225 then show "Ls (SEQ (CHAR c') r2) \<dagger> c = {s. c#s \<in> L (SEQ (CHAR c') r2)}" |
|
226 by (simp add: Ls_def lang_seq_def) |
|
227 next |
|
228 assume "c\<noteq>c'" |
|
229 then show "Ls (SEQ (CHAR c') r2) \<dagger> c = {s. c#s \<in> L (SEQ (CHAR c') r2)}" |
|
230 by (simp add: Ls_def lang_seq_def) |
|
231 qed |
|
232 next |
|
233 case (6 r11 r12 r2 c) |
|
234 have ih: "Ls (SEQ r11 (SEQ r12 r2)) \<dagger> c = {s. c#s \<in> L (SEQ r11 (SEQ r12 r2))}" by fact |
|
235 show "Ls (SEQ (SEQ r11 r12) r2) \<dagger> c = {s. c # s \<in> L (SEQ (SEQ r11 r12) r2)}" |
|
236 by (simp add: ih lang_seq_assoc) |
|
237 next |
|
238 case (7 r11 r12 r2 c) |
|
239 have ih1: "Ls (SEQ r11 r2) \<dagger> c = {s. c#s \<in> L (SEQ r11 r2)}" by fact |
|
240 have ih2: "Ls (SEQ r12 r2) \<dagger> c = {s. c#s \<in> L (SEQ r12 r2)}" by fact |
|
241 show "Ls (SEQ (ALT r11 r12) r2) \<dagger> c = {s. c#s \<in> L (SEQ (ALT r11 r12) r2)}" |
|
242 by (auto simp add: Ls_union ih1 ih2 lang_seq_union) |
|
243 next |
|
244 case (8 r1 r2 c) |
|
245 have ih1: "Ls r2 \<dagger> c = {s. c#s \<in> L r2}" by fact |
|
246 have ih2: "Ls r1 \<dagger> c = {s. c#s \<in> L r1}" by fact |
|
247 show "Ls (SEQ (STAR r1) r2) \<dagger> c = {s. c#s \<in> L (SEQ (STAR r1) r2)}" |
|
248 sorry |
|
249 next |
|
250 case (9 r c) |
|
251 have ih: "Ls r \<dagger> c = {s. c#s \<in> L r}" by fact |
|
252 show "Ls (STAR r) \<dagger> c = {s. c#s \<in> L (STAR r)}" |
|
253 by (simp only: eq holes[OF ih] del: r9) |
|
254 qed |
|
255 |
|
256 (* correctness of the matcher *) |
|
257 lemma macher_holes: |
|
258 shows "r ! s \<Longrightarrow> s \<in> L r" |
|
259 and "\<not> r ! s \<Longrightarrow> s \<notin> L r" |
|
260 proof (induct rule: matcher.induct) |
|
261 case (1 s) |
|
262 { case 1 |
|
263 have "EMPTY ! s" by fact |
|
264 then show "s \<in> L EMPTY" by simp |
|
265 next |
|
266 case 2 |
|
267 have "\<not> EMPTY ! s" by fact |
|
268 then show "s \<notin> L EMPTY" by simp |
|
269 } |
|
270 next |
|
271 case (2 c s) |
|
272 { case 1 |
|
273 have "CHAR c ! s" by fact |
|
274 then show "s \<in> L (CHAR c)" by simp |
|
275 next |
|
276 case 2 |
|
277 have "\<not> CHAR c ! s" by fact |
|
278 then show "s \<notin> L (CHAR c)" by simp |
|
279 } |
|
280 next |
|
281 case (3 r1 r2 s) |
|
282 have ih1: "r1 ! s \<Longrightarrow> s \<in> L r1" by fact |
|
283 have ih2: "\<not> r1 ! s \<Longrightarrow> s \<notin> L r1" by fact |
|
284 have ih3: "r2 ! s \<Longrightarrow> s \<in> L r2" by fact |
|
285 have ih4: "\<not> r2 ! s \<Longrightarrow> s \<notin> L r2" by fact |
|
286 { case 1 |
|
287 have "ALT r1 r2 ! s" by fact |
|
288 then show "s \<in> L (ALT r1 r2)" by (auto simp add: ih1 ih3) |
|
289 next |
|
290 case 2 |
|
291 have "\<not> ALT r1 r2 ! s" by fact |
|
292 then show "s \<notin> L (ALT r1 r2)" by (simp add: ih2 ih4) |
|
293 } |
|
294 next |
|
295 case (4 r) |
|
296 { case 1 |
|
297 have "STAR r ! []" by fact |
|
298 then show "[] \<in> L (STAR r)" by (simp add: lang_star_in_empty) |
|
299 next |
|
300 case 2 |
|
301 have "\<not> STAR r ! []" by fact |
|
302 then show "[] \<notin> L (STAR r)" by (simp) |
|
303 } |
|
304 next |
|
305 case (5 r c s) |
|
306 have ih1: "\<And>rx. SEQ rx (STAR r) ! s \<Longrightarrow> s \<in> L (SEQ rx (STAR r))" by fact |
|
307 have ih2: "\<And>rx. \<not>SEQ rx (STAR r) ! s \<Longrightarrow> s \<notin> L (SEQ rx (STAR r))" by fact |
|
308 { case 1 |
|
309 have as: "STAR r ! c#s" by fact |
|
310 then have "\<exists>r' \<in> r \<dagger> c. SEQ r' (STAR r) ! s" by (auto simp add: OR_def) |
|
311 then obtain r' where imp1: "r' \<in> r \<dagger> c" and imp2: "SEQ r' (STAR r) ! s" by blast |
|
312 from imp2 have "s \<in> L (SEQ r' (STAR r))" using ih1 by simp |
|
313 then have "s \<in> L r' ; L (STAR r)" by simp |
|
314 then have "c#s \<in> {[c]} ; (L r' ; L (STAR r))" by (simp add: lang_seq_def) |
|
315 also have "\<dots> \<subseteq> L r ; L (STAR r)" using imp1 |
|
316 apply(auto simp add: lang_seq_def) sorry |
|
317 also have "\<dots> \<subseteq> L (STAR r)" by (simp add: lang_star_subseteq) |
|
318 finally show "c#s \<in> L (STAR r)" by simp |
|
319 next |
|
320 case 2 |
|
321 have as: "\<not> STAR r ! c#s" by fact |
|
322 then have "\<forall>r'\<in> r \<dagger> c. \<not> (SEQ r' (STAR r) ! s)" |
|
323 by (auto simp add: OR_def) |
|
324 then have "\<forall>r'\<in> r \<dagger> c. s \<notin> L (SEQ r' (STAR r))" using ih2 by auto |
|
325 then obtain r' where "r'\<in> r \<dagger> c \<Longrightarrow> s \<notin> L (SEQ r' (STAR r))" by auto |
|
326 |
|
327 show "c#s \<notin> L (STAR r)" sorry |
|
328 } |
|
329 next |
|
330 case (6 r1 r2) |
|
331 have ih1: "r1 ! [] \<Longrightarrow> [] \<in> L r1" by fact |
|
332 have ih2: "\<not> r1 ! [] \<Longrightarrow> [] \<notin> L r1" by fact |
|
333 have ih3: "r2 ! [] \<Longrightarrow> [] \<in> L r2" by fact |
|
334 have ih4: "\<not> r2 ! [] \<Longrightarrow> [] \<notin> L r2" by fact |
|
335 { case 1 |
|
336 have as: "SEQ r1 r2 ! []" by fact |
|
337 then have "r1 ! [] \<and> r2 ! []" by (simp) |
|
338 then show "[] \<in> L (SEQ r1 r2)" using ih1 ih3 by (simp add: lang_seq_def) |
|
339 next |
|
340 case 2 |
|
341 have "\<not> SEQ r1 r2 ! []" by fact |
|
342 then have "(\<not> r1 ! []) \<or> (\<not> r2 ! [])" by (simp) |
|
343 then show "[] \<notin> L (SEQ r1 r2)" using ih2 ih4 |
|
344 by (auto simp add: lang_seq_def) |
|
345 } |
|
346 next |
|
347 case (7 r2 c s) |
|
348 have ih1: "r2 ! c#s \<Longrightarrow> c#s \<in> L r2" by fact |
|
349 have ih2: "\<not> r2 ! c#s \<Longrightarrow> c#s \<notin> L r2" by fact |
|
350 { case 1 |
|
351 have "SEQ EMPTY r2 ! c#s" by fact |
|
352 then show "c#s \<in> L (SEQ EMPTY r2)" |
|
353 using ih1 by (simp add: lang_seq_def) |
|
354 next |
|
355 case 2 |
|
356 have "\<not> SEQ EMPTY r2 ! c#s" by fact |
|
357 then show "c#s \<notin> L (SEQ EMPTY r2)" |
|
358 using ih2 by (simp add: lang_seq_def) |
|
359 } |
|
360 next |
|
361 case (8 c' r2 c s) |
|
362 have ih1: "\<lbrakk>c' = c; r2 ! s\<rbrakk> \<Longrightarrow> s \<in> L r2" by fact |
|
363 have ih2: "\<lbrakk>c' = c; \<not>r2 ! s\<rbrakk> \<Longrightarrow> s \<notin> L r2" by fact |
|
364 { case 1 |
|
365 have "SEQ (CHAR c') r2 ! c#s" by fact |
|
366 then show "c#s \<in> L (SEQ (CHAR c') r2)" |
|
367 using ih1 by (auto simp add: lang_seq_def split: if_splits) |
|
368 next |
|
369 case 2 |
|
370 have "\<not> SEQ (CHAR c') r2 ! c#s" by fact |
|
371 then show "c#s \<notin> L (SEQ (CHAR c') r2)" |
|
372 using ih2 by (auto simp add: lang_seq_def) |
|
373 } |
|
374 next |
|
375 case (9 r11 r12 r2 c s) |
|
376 have ih1: "SEQ r11 (SEQ r12 r2) ! c#s \<Longrightarrow> c#s \<in> L (SEQ r11 (SEQ r12 r2))" by fact |
|
377 have ih2: "\<not> SEQ r11 (SEQ r12 r2) ! c#s \<Longrightarrow> c#s \<notin> L (SEQ r11 (SEQ r12 r2))" by fact |
|
378 { case 1 |
|
379 have "SEQ (SEQ r11 r12) r2 ! c#s" by fact |
|
380 then show "c#s \<in> L (SEQ (SEQ r11 r12) r2)" |
|
381 using ih1 |
|
382 apply(auto simp add: lang_seq_def) |
|
383 apply(rule_tac x="s1@s1a" in exI) |
|
384 apply(rule_tac x="s2a" in exI) |
|
385 apply(simp) |
|
386 apply(blast) |
|
387 done |
|
388 next |
|
389 case 2 |
|
390 have "\<not> SEQ (SEQ r11 r12) r2 ! c#s" by fact |
|
391 then show "c#s \<notin> L (SEQ (SEQ r11 r12) r2)" |
|
392 using ih2 by (auto simp add: lang_seq_def) |
|
393 } |
|
394 next |
|
395 case (10 r11 r12 r2 c s) |
|
396 have ih1: "SEQ r11 r2 ! c#s \<Longrightarrow> c#s \<in> L (SEQ r11 r2)" by fact |
|
397 have ih2: "\<not> SEQ r11 r2 ! c#s \<Longrightarrow> c#s \<notin> L (SEQ r11 r2)" by fact |
|
398 have ih3: "SEQ r12 r2 ! c#s \<Longrightarrow> c#s \<in> L (SEQ r12 r2)" by fact |
|
399 have ih4: "\<not> SEQ r12 r2 ! c#s \<Longrightarrow> c#s \<notin> L (SEQ r12 r2)" by fact |
|
400 { case 1 |
|
401 have "SEQ (ALT r11 r12) r2 ! c#s" by fact |
|
402 then show "c#s \<in> L (SEQ (ALT r11 r12) r2)" |
|
403 using ih1 ih3 by (auto simp add: lang_seq_union) |
|
404 next |
|
405 case 2 |
|
406 have "\<not> SEQ (ALT r11 r12) r2 ! c#s" by fact |
|
407 then show " c#s \<notin> L (SEQ (ALT r11 r12) r2)" |
|
408 using ih2 ih4 by (simp add: lang_seq_union) |
|
409 } |
|
410 next |
|
411 case (11 r1 r2 c s) |
|
412 have ih1: "r2 ! c#s \<Longrightarrow> c#s \<in> L r2" by fact |
|
413 have ih2: "\<not>r2 ! c#s \<Longrightarrow> c#s \<notin> L r2" by fact |
|
414 { case 1 |
|
415 have "SEQ (STAR r1) r2 ! c#s" by fact |
|
416 then show "c#s \<in> L (SEQ (STAR r1) r2)" |
|
417 using ih1 sorry |
|
418 next |
|
419 case 2 |
|
420 have "\<not> SEQ (STAR r1) r2 ! c#s" by fact |
|
421 then show "c#s \<notin> L (SEQ (STAR r1) r2)" |
|
422 using ih2 sorry |
|
423 } |
|
424 qed |
|
425 |
|
426 |
|
427 end |
|
428 |
|
429 |
|
430 |
|
431 |
|
432 |
|
433 |
|
434 |
|
435 |
|
436 |
|
437 |