1 theory My |
1 theory My |
2 imports Main |
2 imports Main Infinite_Set |
3 begin |
3 begin |
4 |
4 |
5 |
5 |
6 definition |
6 definition |
7 lang_seq :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ; _" [100,100] 100) |
7 Seq :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ;; _" [100,100] 100) |
8 where |
8 where |
9 "L1 ; L2 = {s1 @ s2 | s1 s2. s1 \<in> L1 \<and> s2 \<in> L2}" |
9 "L1 ;; L2 = {s1 @ s2 | s1 s2. s1 \<in> L1 \<and> s2 \<in> L2}" |
10 |
10 |
11 inductive_set |
11 inductive_set |
12 Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102) |
12 Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102) |
13 for L :: "string set" |
13 for L :: "string set" |
14 where |
14 where |
15 start[intro]: "[] \<in> L\<star>" |
15 start[intro]: "[] \<in> L\<star>" |
16 | step[intro]: "\<lbrakk>s1 \<in> L; s2 \<in> L\<star>\<rbrakk> \<Longrightarrow> s1@s2 \<in> L\<star>" |
16 | step[intro]: "\<lbrakk>s1 \<in> L; s2 \<in> L\<star>\<rbrakk> \<Longrightarrow> s1@s2 \<in> L\<star>" |
17 |
17 |
|
18 lemma lang_star_cases: |
|
19 shows "L\<star> = {[]} \<union> L ;; L\<star>" |
|
20 unfolding Seq_def |
|
21 by (auto) (metis Star.simps) |
|
22 |
|
23 lemma lang_star_cases2: |
|
24 shows "L ;; L\<star> = L\<star> ;; L" |
|
25 sorry |
|
26 |
|
27 |
|
28 theorem ardens_revised: |
|
29 assumes nemp: "[] \<notin> A" |
|
30 shows "(X = X ;; A \<union> B) \<longleftrightarrow> (X = B ;; A\<star>)" |
|
31 proof |
|
32 assume eq: "X = B ;; A\<star>" |
|
33 have "A\<star> = {[]} \<union> A\<star> ;; A" sorry |
|
34 then have "B ;; A\<star> = B ;; ({[]} \<union> A\<star> ;; A)" unfolding Seq_def by simp |
|
35 also have "\<dots> = B \<union> B ;; (A\<star> ;; A)" unfolding Seq_def by auto |
|
36 also have "\<dots> = B \<union> (B ;; A\<star>) ;; A" unfolding Seq_def |
|
37 by (auto) (metis append_assoc)+ |
|
38 finally show "X = X ;; A \<union> B" using eq by auto |
|
39 next |
|
40 assume "X = X ;; A \<union> B" |
|
41 then have "B \<subseteq> X" "X ;; A \<subseteq> X" by auto |
|
42 show "X = B ;; A\<star>" sorry |
|
43 qed |
18 |
44 |
19 datatype rexp = |
45 datatype rexp = |
20 NULL |
46 NULL |
21 | EMPTY |
47 | EMPTY |
22 | CHAR char |
48 | CHAR char |
23 | SEQ rexp rexp |
49 | SEQ rexp rexp |
24 | ALT rexp rexp |
50 | ALT rexp rexp |
25 | STAR rexp |
51 | STAR rexp |
26 |
52 |
27 fun |
53 fun |
28 L_rexp :: "rexp \<Rightarrow> string set" |
54 Sem :: "rexp \<Rightarrow> string set" ("\<lparr>_\<rparr>" [0] 1000) |
29 where |
55 where |
30 "L_rexp (NULL) = {}" |
56 "\<lparr>NULL\<rparr> = {}" |
31 | "L_rexp (EMPTY) = {[]}" |
57 | "\<lparr>EMPTY\<rparr> = {[]}" |
32 | "L_rexp (CHAR c) = {[c]}" |
58 | "\<lparr>CHAR c\<rparr> = {[c]}" |
33 | "L_rexp (SEQ r1 r2) = (L_rexp r1) ; (L_rexp r2)" |
59 | "\<lparr>SEQ r1 r2\<rparr> = \<lparr>r1\<rparr> ;; \<lparr>r2\<rparr>" |
34 | "L_rexp (ALT r1 r2) = (L_rexp r1) \<union> (L_rexp r2)" |
60 | "\<lparr>ALT r1 r2\<rparr> = \<lparr>r1\<rparr> \<union> \<lparr>r2\<rparr>" |
35 | "L_rexp (STAR r) = (L_rexp r)\<star>" |
61 | "\<lparr>STAR r\<rparr> = \<lparr>r\<rparr>\<star>" |
36 |
62 |
37 definition |
63 definition |
38 folds :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b" |
64 folds :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b" |
39 where |
65 where |
40 "folds f z S \<equiv> SOME x. fold_graph f z S x" |
66 "folds f z S \<equiv> SOME x. fold_graph f z S x" |
41 |
67 |
42 lemma folds_simp_null [simp]: |
68 lemma folds_simp_null [simp]: |
43 "finite rs \<Longrightarrow> x \<in> L_rexp (folds ALT NULL rs) = (\<exists>r \<in> rs. x \<in> L_rexp r)" |
69 "finite rs \<Longrightarrow> x \<in> \<lparr>folds ALT NULL rs\<rparr> \<longleftrightarrow> (\<exists>r \<in> rs. x \<in> \<lparr>r\<rparr>)" |
44 apply (simp add: folds_def) |
70 apply (simp add: folds_def) |
45 apply (rule someI2_ex) |
71 apply (rule someI2_ex) |
46 apply (erule finite_imp_fold_graph) |
72 apply (erule finite_imp_fold_graph) |
47 apply (erule fold_graph.induct) |
73 apply (erule fold_graph.induct) |
48 apply (auto) |
74 apply (auto) |
49 done |
75 done |
50 |
76 |
51 lemma folds_simp_empty [simp]: |
77 lemma folds_simp_empty [simp]: |
52 "finite rs \<Longrightarrow> x \<in> L_rexp (folds ALT EMPTY rs) = ((\<exists>r \<in> rs. x \<in> L_rexp r) \<or> x = [])" |
78 "finite rs \<Longrightarrow> x \<in> \<lparr>folds ALT EMPTY rs\<rparr> \<longleftrightarrow> (\<exists>r \<in> rs. x \<in> \<lparr>r\<rparr>) \<or> x = []" |
53 apply (simp add: folds_def) |
79 apply (simp add: folds_def) |
54 apply (rule someI2_ex) |
80 apply (rule someI2_ex) |
55 apply (erule finite_imp_fold_graph) |
81 apply (erule finite_imp_fold_graph) |
56 apply (erule fold_graph.induct) |
82 apply (erule fold_graph.induct) |
57 apply (auto) |
83 apply (auto) |
58 done |
84 done |
59 |
85 |
60 lemma [simp]: |
86 lemma [simp]: |
61 "(x, y) \<in> {(x, y). P x y} \<longleftrightarrow> P x y" |
87 shows "(x, y) \<in> {(x, y). P x y} \<longleftrightarrow> P x y" |
62 by simp |
88 by simp |
63 |
89 |
64 definition |
90 definition |
65 str_eq ("_ \<approx>_ _") |
91 str_eq ("_ \<approx>_ _") |
66 where |
92 where |
130 qed |
156 qed |
131 |
157 |
132 lemma finite_regular_aux: |
158 lemma finite_regular_aux: |
133 fixes Lang :: "string set" |
159 fixes Lang :: "string set" |
134 assumes "finite (UNIV // (\<approx>Lang))" |
160 assumes "finite (UNIV // (\<approx>Lang))" |
135 shows "\<exists>rs. Lang = L_rexp (folds ALT NULL rs)" |
161 shows "\<exists>rs. Lang = \<lparr>folds ALT NULL rs\<rparr>" |
136 apply(subst lang_is_union_of_finals) |
162 apply(subst lang_is_union_of_finals) |
137 using assms |
163 using assms |
138 apply - |
164 apply - |
139 apply(drule finite_final) |
165 apply(drule finite_final) |
140 apply(drule_tac f="L_rexp" in finite_f_one_to_one) |
166 apply(drule_tac f="Sem" in finite_f_one_to_one) |
141 apply(clarify) |
167 apply(clarify) |
142 apply(drule final_rexp[OF assms]) |
168 apply(drule final_rexp[OF assms]) |
143 apply(auto)[1] |
169 apply(auto)[1] |
144 apply(clarify) |
170 apply(clarify) |
145 apply(rule_tac x="rs" in exI) |
171 apply(rule_tac x="rs" in exI) |
237 section {* finite \<Rightarrow> regular *} |
263 section {* finite \<Rightarrow> regular *} |
238 |
264 |
239 definition |
265 definition |
240 transitions :: "string set \<Rightarrow> string set \<Rightarrow> rexp set" ("_\<Turnstile>\<Rightarrow>_") |
266 transitions :: "string set \<Rightarrow> string set \<Rightarrow> rexp set" ("_\<Turnstile>\<Rightarrow>_") |
241 where |
267 where |
242 "Y \<Turnstile>\<Rightarrow> X \<equiv> {CHAR c | c. Y ; {[c]} \<subseteq> X}" |
268 "Y \<Turnstile>\<Rightarrow> X \<equiv> {CHAR c | c. Y ;; {[c]} \<subseteq> X}" |
243 |
269 |
244 definition |
270 definition |
245 transitions_rexp ("_ \<turnstile>\<rightarrow> _") |
271 transitions_rexp ("_ \<turnstile>\<rightarrow> _") |
246 where |
272 where |
247 "Y \<turnstile>\<rightarrow> X \<equiv> if [] \<in> X then folds ALT EMPTY (Y \<Turnstile>\<Rightarrow>X) else folds ALT NULL (Y \<Turnstile>\<Rightarrow>X)" |
273 "Y \<turnstile>\<rightarrow> X \<equiv> if [] \<in> X then folds ALT EMPTY (Y \<Turnstile>\<Rightarrow>X) else folds ALT NULL (Y \<Turnstile>\<Rightarrow>X)" |
248 |
274 |
249 definition |
275 definition |
250 "rhs CS X \<equiv> if X = {[]} then {({[]}, EMPTY)} else {(Y, Y \<turnstile>\<rightarrow>X) | Y. Y \<in> CS}" |
276 "rhs CS X \<equiv> if X = {[]} then {({[]}, EMPTY)} else {(Y, Y \<turnstile>\<rightarrow>X) | Y. Y \<in> CS}" |
251 |
277 |
252 definition |
278 definition |
253 "rhs_sem CS X \<equiv> \<Union> {(Y; L_rexp r) | Y r . (Y, r) \<in> rhs CS X}" |
279 "rhs_sem CS X \<equiv> \<Union> {(Y;; \<lparr>r\<rparr>) | Y r . (Y, r) \<in> rhs CS X}" |
254 |
280 |
255 definition |
281 definition |
256 "eqs CS \<equiv> (\<Union>X \<in> CS. {(X, rhs CS X)})" |
282 "eqs CS \<equiv> (\<Union>X \<in> CS. {(X, rhs CS X)})" |
257 |
283 |
258 definition |
284 definition |
274 by auto |
300 by auto |
275 |
301 |
276 lemma every_eqclass_has_transition: |
302 lemma every_eqclass_has_transition: |
277 assumes has_str: "s @ [c] \<in> X" |
303 assumes has_str: "s @ [c] \<in> X" |
278 and in_CS: "X \<in> UNIV // (\<approx>Lang)" |
304 and in_CS: "X \<in> UNIV // (\<approx>Lang)" |
279 obtains Y where "Y \<in> UNIV // (\<approx>Lang)" and "Y ; {[c]} \<subseteq> X" and "s \<in> Y" |
305 obtains Y where "Y \<in> UNIV // (\<approx>Lang)" and "Y ;; {[c]} \<subseteq> X" and "s \<in> Y" |
280 proof - |
306 proof - |
281 def Y \<equiv> "(\<approx>Lang) `` {s}" |
307 def Y \<equiv> "(\<approx>Lang) `` {s}" |
282 have "Y \<in> UNIV // (\<approx>Lang)" |
308 have "Y \<in> UNIV // (\<approx>Lang)" |
283 unfolding Y_def quotient_def by auto |
309 unfolding Y_def quotient_def by auto |
284 moreover |
310 moreover |
285 have "X = (\<approx>Lang) `` {s @ [c]}" |
311 have "X = (\<approx>Lang) `` {s @ [c]}" |
286 using has_str in_CS defined_by_str by blast |
312 using has_str in_CS defined_by_str by blast |
287 then have "Y ; {[c]} \<subseteq> X" |
313 then have "Y ;; {[c]} \<subseteq> X" |
288 unfolding Y_def Image_def lang_seq_def |
314 unfolding Y_def Image_def Seq_def |
289 unfolding str_eq_rel_def |
315 unfolding str_eq_rel_def |
290 by (auto) (simp add: str_eq_def) |
316 by (auto) (simp add: str_eq_def) |
291 moreover |
317 moreover |
292 have "s \<in> Y" unfolding Y_def |
318 have "s \<in> Y" unfolding Y_def |
293 unfolding Image_def str_eq_rel_def str_eq_def by simp |
319 unfolding Image_def str_eq_rel_def str_eq_def by simp |
294 moreover |
320 (*moreover |
295 have "True" by simp (* FIXME *) |
321 have "True" by simp FIXME *) |
296 note that |
322 ultimately show thesis by (blast intro: that) |
297 ultimately show thesis by blast |
|
298 qed |
323 qed |
299 |
324 |
300 lemma test: |
325 lemma test: |
301 assumes "[] \<in> X" |
326 assumes "[] \<in> X" |
302 shows "[] \<in> L_rexp (Y \<turnstile>\<rightarrow> X)" |
327 shows "[] \<in> \<lparr>Y \<turnstile>\<rightarrow> X\<rparr>" |
303 using assms |
328 using assms |
304 by (simp add: transitions_rexp_def) |
329 by (simp add: transitions_rexp_def) |
305 |
330 |
306 lemma rhs_sem: |
331 lemma rhs_sem: |
307 assumes "X \<in> (UNIV // (\<approx>Lang))" |
332 assumes "X \<in> (UNIV // (\<approx>Lang))" |
308 shows "X \<subseteq> rhs_sem (UNIV // (\<approx>Lang)) X" |
333 shows "X \<subseteq> rhs_sem (UNIV // (\<approx>Lang)) X" |
309 apply(case_tac "X = {[]}") |
334 apply(case_tac "X = {[]}") |
310 apply(simp) |
335 apply(simp) |
311 apply(simp add: rhs_sem_def rhs_def lang_seq_def) |
336 apply(simp add: rhs_sem_def rhs_def Seq_def) |
312 apply(rule subsetI) |
337 apply(rule subsetI) |
313 apply(case_tac "x = []") |
338 apply(case_tac "x = []") |
314 apply(simp add: rhs_sem_def rhs_def) |
339 apply(simp add: rhs_sem_def rhs_def) |
315 apply(rule_tac x = "X" in exI) |
340 apply(rule_tac x = "X" in exI) |
316 apply(simp) |
341 apply(simp) |
317 apply(rule_tac x = "X" in exI) |
342 apply(rule_tac x = "X" in exI) |
318 apply(simp add: assms) |
343 apply(simp add: assms) |
319 apply(simp add: transitions_rexp_def) |
344 apply(simp add: transitions_rexp_def) |
320 oops |
345 oops |
|
346 |
|
347 |
|
348 (* |
|
349 fun |
|
350 power :: "string \<Rightarrow> nat \<Rightarrow> string" (infixr "\<Up>" 100) |
|
351 where |
|
352 "s \<Up> 0 = s" |
|
353 | "s \<Up> (Suc n) = s @ (s \<Up> n)" |
|
354 |
|
355 definition |
|
356 "Lone = {(''0'' \<Up> n) @ (''1'' \<Up> n) | n. True }" |
|
357 |
|
358 lemma |
|
359 "infinite (UNIV // (\<approx>Lone))" |
|
360 unfolding infinite_iff_countable_subset |
|
361 apply(rule_tac x="\<lambda>n. {(''0'' \<Up> n) @ (''1'' \<Up> i) | i. i \<in> {..n} }" in exI) |
|
362 apply(auto) |
|
363 prefer 2 |
|
364 unfolding Lone_def |
|
365 unfolding quotient_def |
|
366 unfolding Image_def |
|
367 apply(simp) |
|
368 unfolding str_eq_rel_def |
|
369 unfolding str_eq_def |
|
370 apply(auto) |
|
371 apply(rule_tac x="''0'' \<Up> n" in exI) |
|
372 apply(auto) |
|
373 unfolding infinite_nat_iff_unbounded |
|
374 unfolding Lone_def |
|
375 *) |
|
376 |
|
377 |
|
378 |
|
379 text {* Derivatives *} |
|
380 |
|
381 definition |
|
382 DERS :: "string \<Rightarrow> string set \<Rightarrow> string set" |
|
383 where |
|
384 "DERS s L \<equiv> {s'. s @ s' \<in> L}" |
|
385 |
|
386 lemma |
|
387 shows "x \<approx>L y \<longleftrightarrow> DERS x L = DERS y L" |
|
388 unfolding DERS_def str_eq_def |
|
389 by auto |