diff -r cef2893f353b -r c64241fa4dff Myhill.thy --- a/Myhill.thy Fri Dec 31 13:47:53 2010 +0000 +++ b/Myhill.thy Fri Jan 07 14:25:23 2011 +0000 @@ -3,8 +3,7 @@ begin text {* sequential composition of languages *} -definition - Seq :: "string set \ string set \ string set" ("_ ;; _" [100,100] 100) +definition Seq :: "string set \ string set \ string set" ("_ ;; _" [100,100] 100) where "L1 ;; L2 = {s1 @ s2 | s1 s2. s1 \ L1 \ s2 \ L2}" @@ -19,25 +18,22 @@ "(A \ B) ;; C = (A ;; C) \ (B ;; C)" by (auto simp:Seq_def) +lemma seq_intro: + "\x \ A; y \ B\ \ x @ y \ A ;; B " +by (auto simp:Seq_def) + lemma seq_assoc: "(A ;; B) ;; C = A ;; (B ;; C)" -unfolding Seq_def -apply(auto) -apply(metis) +apply(auto simp:Seq_def) +apply blast by (metis append_assoc) -lemma union_seq: - "\ {f x y ;; z| x y. P x y } = (\ {f x y|x y. P x y });; z" -apply (auto simp add:Seq_def) -apply metis -done - theorem ardens_revised: assumes nemp: "[] \ A" shows "(X = X ;; A \ B) \ (X = B ;; A\)" proof assume eq: "X = B ;; A\" - have "A\ = {[]} \ A\ ;; A" sorry + have "A\ = {[]} \ A\ ;; A" sorry then have "B ;; A\ = B ;; ({[]} \ A\ ;; A)" unfolding Seq_def by simp also have "\ = B \ B ;; (A\ ;; A)" unfolding Seq_def by auto also have "\ = B \ (B ;; A\) ;; A" unfolding Seq_def @@ -46,7 +42,7 @@ next assume "X = X ;; A \ B" then have "B \ X" "X ;; A \ X" by auto - show "X = B ;; A\" sorry + thus "X = B ;; A\" sorry qed datatype rexp = @@ -778,286 +774,71 @@ "UNIV // (\{[]}) \ {{[]}, UNIV - {[]}}" proof fix x - assume h: "x \ UNIV // \{[]}" - show "x \ {{[]}, UNIV - {[]}}" - - - have "\ s. s \{[]} [] \ s = []" - apply (auto simp add:str_eq_def) - apply blast - - hence "False" - apply (simp add:quotient_def) - - -lemma other_direction: - "Lang = L (r::rexp) \ finite (UNIV // (\Lang))" -proof (induct arbitrary:Lang rule:rexp.induct) - case NULL - have "UNIV // (\{}) \ {UNIV} " - by (auto simp:quotient_def str_eq_rel_def str_eq_def) - with prems show "?case" by (auto intro:finite_subset) -next - case EMPTY - have "UNIV // (\{[]}) \ {{[]}, UNIV - {[]}}" - sorry - with prems show ?case by (auto intro:finite_subset) -next - case (CHAR c) - have "UNIV // (\{[c]}) \ {{[]},{[c]}, UNIV - {[], [c]}}" - sorry - with prems show ?case by (auto intro:finite_subset) -next - case (SEQ r1 r2) - show ?case sorry -next - case (ALT r1 r1) - show ?case sorry -next - case (STAR r) - show ?case sorry -qed - - - - - - - - - - - - - - - - - - - - -apply (induct arbitrary:Lang rule:rexp.induct) -apply (simp add:QUOT_def equiv_class_def equiv_str_def) -by (simp_all add:quot_lambda quot_single quot_seq quot_alt quot_star) - -(* Alternative definition for Quo *) -definition - QUOT :: "string set \ (string set) set" -where - "QUOT Lang \ (\x. {\x\Lang})" - -lemma test: - "UNIV Quo Lang = QUOT Lang" -by (auto simp add: quot_def QUOT_def) - -lemma test1: - assumes finite_CS: "finite (QUOT Lang)" - shows "reg Lang" -using finite_CS -unfolding test[symmetric] -by (auto dest: myhill_nerode) - -lemma cons_one: "x @ y \ {z} \ x @ y = z" -by simp - -lemma quot_lambda: "QUOT {[]} = {{[]}, UNIV - {[]}}" -proof - show "QUOT {[]} \ {{[]}, UNIV - {[]}}" - proof - fix x - assume in_quot: "x \ QUOT {[]}" - show "x \ {{[]}, UNIV - {[]}}" - proof - - from in_quot - have "x = {[]} \ x = UNIV - {[]}" - unfolding QUOT_def equiv_class_def - proof - fix xa - assume in_univ: "xa \ UNIV" - and in_eqiv: "x \ {{y. xa \{[]}\ y}}" - show "x = {[]} \ x = UNIV - {[]}" - proof(cases "xa = []") - case True - hence "{y. xa \{[]}\ y} = {[]}" using in_eqiv - by (auto simp add:equiv_str_def) - thus ?thesis using in_eqiv by (rule_tac disjI1, simp) - next - case False - hence "{y. xa \{[]}\ y} = UNIV - {[]}" using in_eqiv - by (auto simp:equiv_str_def) - thus ?thesis using in_eqiv by (rule_tac disjI2, simp) - qed - qed - thus ?thesis by simp - qed - qed -next - show "{{[]}, UNIV - {[]}} \ QUOT {[]}" - proof - fix x - assume in_res: "x \ {{[]}, (UNIV::string set) - {[]}}" - show "x \ (QUOT {[]})" - proof - - have "x = {[]} \ x \ QUOT {[]}" - apply (simp add:QUOT_def equiv_class_def equiv_str_def) - by (rule_tac x = "[]" in exI, auto) - moreover have "x = UNIV - {[]} \ x \ QUOT {[]}" - apply (simp add:QUOT_def equiv_class_def equiv_str_def) - by (rule_tac x = "''a''" in exI, auto) - ultimately show ?thesis using in_res by blast - qed + assume "x \ UNIV // \{[]}" + then obtain y where h: "x = {z. (y, z) \ \{[]}}" unfolding quotient_def Image_def by blast + show "x \ {{[]}, UNIV - {[]}}" + proof (cases "y = []") + case True with h + have "x = {[]}" by (auto simp:str_eq_rel_def str_eq_def) + thus ?thesis by simp + next + case False with h + have "x = UNIV - {[]}" by (auto simp:str_eq_rel_def str_eq_def) + thus ?thesis by simp qed qed -lemma quot_single_aux: "\x \ []; x \ [c]\ \ x @ z \ [c]" -by (induct x, auto) - -lemma quot_single: "\ (c::char). QUOT {[c]} = {{[]}, {[c]}, UNIV - {[], [c]}}" -proof - - fix c::"char" - have exist_another: "\ a. a \ c" - apply (case_tac "c = CHR ''a''") - apply (rule_tac x = "CHR ''b''" in exI, simp) - by (rule_tac x = "CHR ''a''" in exI, simp) - show "QUOT {[c]} = {{[]}, {[c]}, UNIV - {[], [c]}}" - proof - show "QUOT {[c]} \ {{[]},{[c]}, UNIV - {[], [c]}}" - proof - fix x - assume in_quot: "x \ QUOT {[c]}" - show "x \ {{[]}, {[c]}, UNIV - {[],[c]}}" - proof - - from in_quot - have "x = {[]} \ x = {[c]} \ x = UNIV - {[],[c]}" - unfolding QUOT_def equiv_class_def - proof - fix xa - assume in_eqiv: "x \ {{y. xa \{[c]}\ y}}" - show "x = {[]} \ x = {[c]} \ x = UNIV - {[], [c]}" - proof- - have "xa = [] \ x = {[]}" using in_eqiv - by (auto simp add:equiv_str_def) - moreover have "xa = [c] \ x = {[c]}" - proof - - have "xa = [c] \ {y. xa \{[c]}\ y} = {[c]}" using in_eqiv - apply (simp add:equiv_str_def) - apply (rule set_ext, rule iffI, simp) - apply (drule_tac x = "[]" in spec, auto) - done - thus "xa = [c] \ x = {[c]}" using in_eqiv by simp - qed - moreover have "\xa \ []; xa \ [c]\ \ x = UNIV - {[],[c]}" - proof - - have "\xa \ []; xa \ [c]\ \ {y. xa \{[c]}\ y} = UNIV - {[],[c]}" - apply (clarsimp simp add:equiv_str_def) - apply (rule set_ext, rule iffI, simp) - apply (rule conjI) - apply (drule_tac x = "[c]" in spec, simp) - apply (drule_tac x = "[]" in spec, simp) - by (auto dest:quot_single_aux) - thus "\xa \ []; xa \ [c]\ \ x = UNIV - {[],[c]}" using in_eqiv by simp - qed - ultimately show ?thesis by blast - qed - qed - thus ?thesis by simp - qed - qed - next - show "{{[]}, {[c]}, UNIV - {[],[c]}} \ QUOT {[c]}" - proof - fix x - assume in_res: "x \ {{[]},{[c]}, (UNIV::string set) - {[],[c]}}" - show "x \ (QUOT {[c]})" - proof - - have "x = {[]} \ x \ QUOT {[c]}" - apply (simp add:QUOT_def equiv_class_def equiv_str_def) - by (rule_tac x = "[]" in exI, auto) - moreover have "x = {[c]} \ x \ QUOT {[c]}" - apply (simp add:QUOT_def equiv_class_def equiv_str_def) - apply (rule_tac x = "[c]" in exI, simp) - apply (rule set_ext, rule iffI, simp+) - by (drule_tac x = "[]" in spec, simp) - moreover have "x = UNIV - {[],[c]} \ x \ QUOT {[c]}" - using exist_another - apply (clarsimp simp add:QUOT_def equiv_class_def equiv_str_def) - apply (rule_tac x = "[a]" in exI, simp) - apply (rule set_ext, rule iffI, simp) - apply (clarsimp simp:quot_single_aux, simp) - apply (rule conjI) - apply (drule_tac x = "[c]" in spec, simp) - by (drule_tac x = "[]" in spec, simp) - ultimately show ?thesis using in_res by blast - qed - qed +lemma quot_char_subset: + "UNIV // (\{[c]}) \ {{[]},{[c]}, UNIV - {[], [c]}}" +proof + fix x + assume "x \ UNIV // \{[c]}" + then obtain y where h: "x = {z. (y, z) \ \{[c]}}" unfolding quotient_def Image_def by blast + show "x \ {{[]},{[c]}, UNIV - {[], [c]}}" + proof - + { assume "y = []" hence "x = {[]}" using h by (auto simp:str_eq_rel_def str_eq_def) + } moreover { + assume "y = [c]" hence "x = {[c]}" using h + by (auto dest!:spec[where x = "[]"] simp:str_eq_rel_def str_eq_def) + } moreover { + assume "y \ []" and "y \ [c]" + hence "\ z. (y @ z) \ [c]" by (case_tac y, auto) + moreover have "\ p. (p \ [] \ p \ [c]) = (\ q. p @ q \ [c])" by (case_tac p, auto) + ultimately have "x = UNIV - {[],[c]}" using h + by (auto simp add:str_eq_rel_def str_eq_def) + } ultimately show ?thesis by blast qed qed -lemma eq_class_imp_eq_str: - "\x\lang = \y\lang \ x \lang\ y" -by (auto simp:equiv_class_def equiv_str_def) +text {* *************** Some common lemmas for following ALT, SEQ & STAR cases ******************* *} -lemma finite_tag_image: +lemma finite_tag_imageI: "finite (range tag) \ finite (((op `) tag) ` S)" apply (rule_tac B = "Pow (tag ` UNIV)" in finite_subset) by (auto simp add:image_def Pow_def) -lemma str_inj_imps: - assumes str_inj: "\ m n. tag m = tag (n::string) \ m \lang\ n" - shows "inj_on ((op `) tag) (QUOT lang)" -proof (clarsimp simp add:inj_on_def QUOT_def) - fix x y - assume eq_tag: "tag ` \x\lang = tag ` \y\lang" - show "\x\lang = \y\lang" - proof - - have aux1:"\a b. a \ (\b\lang) \ (a \lang\ b)" - by (simp add:equiv_class_def equiv_str_def) - have aux2: "\ A B f. \f ` A = f ` B; A \ {}\ \ \ a b. a \ A \ b \ B \ f a = f b" - by auto - have aux3: "\ a l. \a\l \ {}" - by (auto simp:equiv_class_def equiv_str_def) - show ?thesis using eq_tag - apply (drule_tac aux2, simp add:aux3, clarsimp) - apply (drule_tac str_inj, (drule_tac aux1)+) - by (simp add:equiv_str_def equiv_class_def) - qed +lemma eq_class_equalI: + "\X \ UNIV // \lang; Y \ UNIV // \lang; x \ X; y \ Y; x \lang y\ \ X = Y" +by (auto simp:quotient_def str_eq_rel_def str_eq_def) + +lemma tag_image_injI: + assumes str_inj: "\ m n. tag m = tag (n::string) \ m \lang n" + shows "inj_on ((op `) tag) (UNIV // \lang)" +proof- + { fix X Y + assume X_in: "X \ UNIV // \lang" + and Y_in: "Y \ UNIV // \lang" + and tag_eq: "tag ` X = tag ` Y" + then obtain x y where "x \ X" and "y \ Y" and "tag x = tag y" + unfolding quotient_def Image_def str_eq_rel_def str_eq_def image_def + apply simp by blast + with X_in Y_in str_inj + have "X = Y" by (rule_tac eq_class_equalI, simp+) + } + thus ?thesis unfolding inj_on_def by auto qed -definition tag_str_ALT :: "string set \ string set \ string \ (string set \ string set)" -where - "tag_str_ALT L\<^isub>1 L\<^isub>2 x \ (\x\L\<^isub>1, \x\L\<^isub>2)" - -lemma tag_str_alt_range_finite: - assumes finite1: "finite (QUOT L\<^isub>1)" - and finite2: "finite (QUOT L\<^isub>2)" - shows "finite (range (tag_str_ALT L\<^isub>1 L\<^isub>2))" -proof - - have "{y. \x. y = (\x\L\<^isub>1, \x\L\<^isub>2)} \ (QUOT L\<^isub>1) \ (QUOT L\<^isub>2)" - by (auto simp:QUOT_def) - thus ?thesis using finite1 finite2 - by (auto simp: image_def tag_str_ALT_def UNION_def - intro: finite_subset[where B = "(QUOT L\<^isub>1) \ (QUOT L\<^isub>2)"]) -qed - -lemma tag_str_alt_inj: - "tag_str_ALT L\<^isub>1 L\<^isub>2 x = tag_str_ALT L\<^isub>1 L\<^isub>2 y \ x \(L\<^isub>1 \ L\<^isub>2)\ y" -apply (simp add:tag_str_ALT_def equiv_class_def equiv_str_def) -by blast - -lemma quot_alt: - assumes finite1: "finite (QUOT L\<^isub>1)" - and finite2: "finite (QUOT L\<^isub>2)" - shows "finite (QUOT (L\<^isub>1 \ L\<^isub>2))" -proof (rule_tac f = "(op `) (tag_str_ALT L\<^isub>1 L\<^isub>2)" in finite_imageD) - show "finite (op ` (tag_str_ALT L\<^isub>1 L\<^isub>2) ` QUOT (L\<^isub>1 \ L\<^isub>2))" - using finite_tag_image tag_str_alt_range_finite finite1 finite2 - by auto -next - show "inj_on (op ` (tag_str_ALT L\<^isub>1 L\<^isub>2)) (QUOT (L\<^isub>1 \ L\<^isub>2))" - apply (rule_tac str_inj_imps) - by (erule_tac tag_str_alt_inj) -qed +text {* **************** the SEQ case ************************ *} (* list_diff:: list substract, once different return tailer *) fun list_diff :: "'a list \ 'a list \ 'a list" (infix "-" 51) @@ -1079,273 +860,297 @@ lemma [simp]: "x - [] = x" by (induct x, auto) -lemma [simp]: "xa \ x \ (x @ y) - xa = (x - xa) @ y" +lemma [simp]: "(x - y = []) \ (x \ y)" +proof- + have "\xa. x = xa @ (x - y) \ xa \ y" + apply (rule list_diff.induct[of _ x y], simp+) + by (clarsimp, rule_tac x = "y # xa" in exI, simp+) + thus "(x - y = []) \ (x \ y)" by simp +qed + +lemma diff_prefix: + "\c \ a - b; b \ a\ \ b @ c \ a" by (auto elim:prefixE) -definition tag_str_SEQ:: "string set \ string set \ string \ (string set \ string set set)" -where - "tag_str_SEQ L\<^isub>1 L\<^isub>2 x \ if (\ xa \ x. xa \ L\<^isub>1) - then (\x\L\<^isub>1, {\(x - xa)\L\<^isub>2 | xa. xa \ x \ xa \ L\<^isub>1}) - else (\x\L\<^isub>1, {})" +lemma diff_diff_appd: + "\c < a - b; b < a\ \ (a - b) - c = a - (b @ c)" +apply (clarsimp simp:strict_prefix_def) +by (drule diff_prefix, auto elim:prefixE) -lemma tag_seq_eq_E: - "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y \ - ((\ xa \ x. xa \ L\<^isub>1) \ \x\L\<^isub>1 = \y\L\<^isub>1 \ - {\(x - xa)\L\<^isub>2 | xa. xa \ x \ xa \ L\<^isub>1} = {\(y - ya)\L\<^isub>2 | ya. ya \ y \ ya \ L\<^isub>1} ) \ - ((\ xa \ x. xa \ L\<^isub>1) \ \x\L\<^isub>1 = \y\L\<^isub>1)" -by (simp add:tag_str_SEQ_def split:if_splits, blast) +lemma app_eq_cases[rule_format]: + "\ x . x @ y = m @ n \ (x \ m \ m \ x)" +apply (induct y, simp) +apply (clarify, drule_tac x = "x @ [a]" in spec) +by (clarsimp, auto simp:prefix_def) + +lemma app_eq_dest: + "x @ y = m @ n \ (x \ m \ (m - x) @ n = y) \ (m \ x \ (x - m) @ y = n)" +by (frule_tac app_eq_cases, auto elim:prefixE) + +definition + "tag_str_SEQ L\<^isub>1 L\<^isub>2 x \ ((\L\<^isub>1) `` {x}, {(\L\<^isub>2) `` {x - xa}| xa. xa \ x \ xa \ L\<^isub>1})" lemma tag_str_seq_range_finite: - assumes finite1: "finite (QUOT L\<^isub>1)" - and finite2: "finite (QUOT L\<^isub>2)" - shows "finite (range (tag_str_SEQ L\<^isub>1 L\<^isub>2))" -proof - - have "(range (tag_str_SEQ L\<^isub>1 L\<^isub>2)) \ (QUOT L\<^isub>1) \ (Pow (QUOT L\<^isub>2))" - by (auto simp:image_def tag_str_SEQ_def QUOT_def) - thus ?thesis using finite1 finite2 - by (rule_tac B = "(QUOT L\<^isub>1) \ (Pow (QUOT L\<^isub>2))" in finite_subset, auto) -qed - -lemma app_in_seq_decom[rule_format]: - "\ x. x @ z \ L\<^isub>1 ; L\<^isub>2 \ (\ xa \ x. xa \ L\<^isub>1 \ (x - xa) @ z \ L\<^isub>2) \ - (\ za \ z. (x @ za) \ L\<^isub>1 \ (z - za) \ L\<^isub>2)" -apply (induct z) -apply (simp, rule allI, rule impI, rule disjI1) -apply (clarsimp simp add:lang_seq_def) -apply (rule_tac x = s1 in exI, simp) -apply (rule allI | rule impI)+ -apply (drule_tac x = "x @ [a]" in spec, simp) -apply (erule exE | erule conjE | erule disjE)+ -apply (rule disjI2, rule_tac x = "[a]" in exI, simp) -apply (rule disjI1, rule_tac x = xa in exI, simp) -apply (erule exE | erule conjE)+ -apply (rule disjI2, rule_tac x = "a # za" in exI, simp) -done + "\finite (UNIV // \L\<^isub>1); finite (UNIV // \L\<^isub>2)\ \ finite (range (tag_str_SEQ L\<^isub>1 L\<^isub>2))" +apply (rule_tac B = "(UNIV // \L\<^isub>1) \ (Pow (UNIV // \L\<^isub>2))" in finite_subset) +by (auto simp:tag_str_SEQ_def Image_def quotient_def split:if_splits) -lemma tag_str_seq_inj: - assumes tag_eq: "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y" - shows "(x::string) \(L\<^isub>1 ; L\<^isub>2)\ y" -proof - - have aux: "\ x y z. \tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y; x @ z \ L\<^isub>1 ; L\<^isub>2\ - \ y @ z \ L\<^isub>1 ; L\<^isub>2" - proof (drule app_in_seq_decom, erule disjE) - fix x y z - assume tag_eq: "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y" - and x_gets_l2: "\xa\x. xa \ L\<^isub>1 \ (x - xa) @ z \ L\<^isub>2" - from x_gets_l2 have "\ xa \ x. xa \ L\<^isub>1" by blast - hence xy_l2:"{\(x - xa)\L\<^isub>2 | xa. xa \ x \ xa \ L\<^isub>1} = {\(y - ya)\L\<^isub>2 | ya. ya \ y \ ya \ L\<^isub>1}" - using tag_eq tag_seq_eq_E by blast - from x_gets_l2 obtain xa where xa_le_x: "xa \ x" - and xa_in_l1: "xa \ L\<^isub>1" - and rest_in_l2: "(x - xa) @ z \ L\<^isub>2" by blast - hence "\ ya. \(x - xa)\L\<^isub>2 = \(y - ya)\L\<^isub>2 \ ya \ y \ ya \ L\<^isub>1" using xy_l2 by auto - then obtain ya where ya_le_x: "ya \ y" - and ya_in_l1: "ya \ L\<^isub>1" - and rest_eq: "\(x - xa)\L\<^isub>2 = \(y - ya)\L\<^isub>2" by blast - from rest_eq rest_in_l2 have "(y - ya) @ z \ L\<^isub>2" - by (auto simp:equiv_class_def equiv_str_def) - hence "ya @ ((y - ya) @ z) \ L\<^isub>1 ; L\<^isub>2" using ya_in_l1 - by (auto simp:lang_seq_def) - thus "y @ z \ L\<^isub>1 ; L\<^isub>2" using ya_le_x - by (erule_tac prefixE, simp) - next - fix x y z - assume tag_eq: "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y" - and x_gets_l1: "\za\z. x @ za \ L\<^isub>1 \ z - za \ L\<^isub>2" - from tag_eq tag_seq_eq_E have x_y_eq: "\x\L\<^isub>1 = \y\L\<^isub>1" by blast - from x_gets_l1 obtain za where za_le_z: "za \ z" - and x_za_in_l1: "(x @ za) \ L\<^isub>1" - and rest_in_l2: "z - za \ L\<^isub>2" by blast - from x_y_eq x_za_in_l1 have y_za_in_l1: "y @ za \ L\<^isub>1" - by (auto simp:equiv_class_def equiv_str_def) - hence "(y @ za) @ (z - za) \ L\<^isub>1 ; L\<^isub>2" using rest_in_l2 - apply (simp add:lang_seq_def) - by (rule_tac x = "y @ za" in exI, rule_tac x = "z - za" in exI, simp) - thus "y @ z \ L\<^isub>1 ; L\<^isub>2" using za_le_z - by (erule_tac prefixE, simp) - qed - show ?thesis using tag_eq - apply (simp add:equiv_str_def) - by (auto intro:aux) +lemma append_seq_elim: + assumes "x @ y \ L\<^isub>1 ;; L\<^isub>2" + shows "(\ xa \ x. xa \ L\<^isub>1 \ (x - xa) @ y \ L\<^isub>2) \ (\ ya \ y. (x @ ya) \ L\<^isub>1 \ (y - ya) \ L\<^isub>2)" +proof- + from assms obtain s\<^isub>1 s\<^isub>2 where "x @ y = s\<^isub>1 @ s\<^isub>2" and in_seq: "s\<^isub>1 \ L\<^isub>1 \ s\<^isub>2 \ L\<^isub>2" + by (auto simp:Seq_def) + hence "(x \ s\<^isub>1 \ (s\<^isub>1 - x) @ s\<^isub>2 = y) \ (s\<^isub>1 \ x \ (x - s\<^isub>1) @ y = s\<^isub>2)" + using app_eq_dest by auto + moreover have "\x \ s\<^isub>1; (s\<^isub>1 - x) @ s\<^isub>2 = y\ \ \ ya \ y. (x @ ya) \ L\<^isub>1 \ (y - ya) \ L\<^isub>2" using in_seq + by (rule_tac x = "s\<^isub>1 - x" in exI, auto elim:prefixE) + moreover have "\s\<^isub>1 \ x; (x - s\<^isub>1) @ y = s\<^isub>2\ \ \ xa \ x. xa \ L\<^isub>1 \ (x - xa) @ y \ L\<^isub>2" using in_seq + by (rule_tac x = s\<^isub>1 in exI, auto) + ultimately show ?thesis by blast qed -lemma quot_seq: - assumes finite1: "finite (QUOT L\<^isub>1)" - and finite2: "finite (QUOT L\<^isub>2)" - shows "finite (QUOT (L\<^isub>1;L\<^isub>2))" -proof (rule_tac f = "(op `) (tag_str_SEQ L\<^isub>1 L\<^isub>2)" in finite_imageD) - show "finite (op ` (tag_str_SEQ L\<^isub>1 L\<^isub>2) ` QUOT (L\<^isub>1 ; L\<^isub>2))" - using finite_tag_image tag_str_seq_range_finite finite1 finite2 - by auto +lemma tag_str_SEQ_injI: + "tag_str_SEQ L\<^isub>1 L\<^isub>2 m = tag_str_SEQ L\<^isub>1 L\<^isub>2 n \ m \(L\<^isub>1 ;; L\<^isub>2) n" +proof- + { fix x y z + assume xz_in_seq: "x @ z \ L\<^isub>1 ;; L\<^isub>2" + and tag_xy: "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y" + have"y @ z \ L\<^isub>1 ;; L\<^isub>2" + proof- + have "(\ xa \ x. xa \ L\<^isub>1 \ (x - xa) @ z \ L\<^isub>2) \ (\ za \ z. (x @ za) \ L\<^isub>1 \ (z - za) \ L\<^isub>2)" + using xz_in_seq append_seq_elim by simp + moreover { + fix xa + assume h1: "xa \ x" and h2: "xa \ L\<^isub>1" and h3: "(x - xa) @ z \ L\<^isub>2" + obtain ya where "ya \ y" and "ya \ L\<^isub>1" and "(y - ya) @ z \ L\<^isub>2" + proof - + have "\ ya. ya \ y \ ya \ L\<^isub>1 \ (x - xa) \L\<^isub>2 (y - ya)" + proof - + have "{\L\<^isub>2 `` {x - xa} |xa. xa \ x \ xa \ L\<^isub>1} = + {\L\<^isub>2 `` {y - xa} |xa. xa \ y \ xa \ L\<^isub>1}" (is "?Left = ?Right") + using h1 tag_xy by (auto simp:tag_str_SEQ_def) + moreover have "\L\<^isub>2 `` {x - xa} \ ?Left" using h1 h2 by auto + ultimately have "\L\<^isub>2 `` {x - xa} \ ?Right" by simp + thus ?thesis by (auto simp:Image_def str_eq_rel_def str_eq_def) + qed + with prems show ?thesis by (auto simp:str_eq_rel_def str_eq_def) + qed + hence "y @ z \ L\<^isub>1 ;; L\<^isub>2" by (erule_tac prefixE, auto simp:Seq_def) + } moreover { + fix za + assume h1: "za \ z" and h2: "(x @ za) \ L\<^isub>1" and h3: "z - za \ L\<^isub>2" + hence "y @ za \ L\<^isub>1" + proof- + have "\L\<^isub>1 `` {x} = \L\<^isub>1 `` {y}" using h1 tag_xy by (auto simp:tag_str_SEQ_def) + with h2 show ?thesis by (auto simp:Image_def str_eq_rel_def str_eq_def) + qed + with h1 h3 have "y @ z \ L\<^isub>1 ;; L\<^isub>2" by (drule_tac A = L\<^isub>1 in seq_intro, auto elim:prefixE) + } + ultimately show ?thesis by blast + qed + } thus "tag_str_SEQ L\<^isub>1 L\<^isub>2 m = tag_str_SEQ L\<^isub>1 L\<^isub>2 n \ m \(L\<^isub>1 ;; L\<^isub>2) n" + by (auto simp add: str_eq_def str_eq_rel_def) +qed + +lemma quot_seq_finiteI: + assumes finite1: "finite (UNIV // \(L\<^isub>1::string set))" + and finite2: "finite (UNIV // \L\<^isub>2)" + shows "finite (UNIV // \(L\<^isub>1 ;; L\<^isub>2))" +proof(rule_tac f = "(op `) (tag_str_SEQ L\<^isub>1 L\<^isub>2)" in finite_imageD) + show "finite (op ` (tag_str_SEQ L\<^isub>1 L\<^isub>2) ` UNIV // \L\<^isub>1 ;; L\<^isub>2)" using finite1 finite2 + by (auto intro:finite_tag_imageI tag_str_seq_range_finite) next - show "inj_on (op ` (tag_str_SEQ L\<^isub>1 L\<^isub>2)) (QUOT (L\<^isub>1 ; L\<^isub>2))" - apply (rule_tac str_inj_imps) - by (erule_tac tag_str_seq_inj) + show "inj_on (op ` (tag_str_SEQ L\<^isub>1 L\<^isub>2)) (UNIV // \L\<^isub>1 ;; L\<^isub>2)" + apply (rule tag_image_injI) + apply (rule tag_str_SEQ_injI) + by (auto intro:tag_image_injI tag_str_SEQ_injI simp:) qed -(****************** the STAR case ************************) +text {* **************** the ALT case ************************ *} + +definition + "tag_str_ALT L\<^isub>1 L\<^isub>2 (x::string) \ ((\L\<^isub>1) `` {x}, (\L\<^isub>2) `` {x})" -lemma app_eq_elim[rule_format]: - "\ a. \ b x y. a @ b = x @ y \ (\ aa ab. a = aa @ ab \ x = aa \ y = ab @ b) \ - (\ ba bb. b = ba @ bb \ x = a @ ba \ y = bb \ ba \ [])" - apply (induct_tac a rule:List.induct, simp) - apply (rule allI | rule impI)+ - by (case_tac x, auto) +lemma tag_str_alt_range_finite: + "\finite (UNIV // \L\<^isub>1); finite (UNIV // \L\<^isub>2)\ \ finite (range (tag_str_ALT L\<^isub>1 L\<^isub>2))" +apply (rule_tac B = "(UNIV // \L\<^isub>1) \ (UNIV // \L\<^isub>2)" in finite_subset) +by (auto simp:tag_str_ALT_def Image_def quotient_def) -definition tag_str_STAR:: "string set \ string \ string set set" -where - "tag_str_STAR L\<^isub>1 x \ if (x = []) - then {} - else {\x\<^isub>2\L\<^isub>1 | x\<^isub>1 x\<^isub>2. x = x\<^isub>1 @ x\<^isub>2 \ x\<^isub>1 \ L\<^isub>1\ \ x\<^isub>2 \ []}" - -lemma tag_str_star_range_finite: - assumes finite1: "finite (QUOT L\<^isub>1)" - shows "finite (range (tag_str_STAR L\<^isub>1))" -proof - - have "range (tag_str_STAR L\<^isub>1) \ Pow (QUOT L\<^isub>1)" - by (auto simp:image_def tag_str_STAR_def QUOT_def) - thus ?thesis using finite1 - by (rule_tac B = "Pow (QUOT L\<^isub>1)" in finite_subset, auto) +lemma quot_union_finiteI: + assumes finite1: "finite (UNIV // \(L\<^isub>1::string set))" + and finite2: "finite (UNIV // \L\<^isub>2)" + shows "finite (UNIV // \(L\<^isub>1 \ L\<^isub>2))" +proof(rule_tac f = "(op `) (tag_str_ALT L\<^isub>1 L\<^isub>2)" in finite_imageD) + show "finite (op ` (tag_str_ALT L\<^isub>1 L\<^isub>2) ` UNIV // \L\<^isub>1 \ L\<^isub>2)" using finite1 finite2 + by (auto intro:finite_tag_imageI tag_str_alt_range_finite) +next + show "inj_on (op ` (tag_str_ALT L\<^isub>1 L\<^isub>2)) (UNIV // \L\<^isub>1 \ L\<^isub>2)" + proof- + have "\m n. tag_str_ALT L\<^isub>1 L\<^isub>2 m = tag_str_ALT L\<^isub>1 L\<^isub>2 n \ m \(L\<^isub>1 \ L\<^isub>2) n" + unfolding tag_str_ALT_def str_eq_def Image_def str_eq_rel_def by auto + thus ?thesis by (auto intro:tag_image_injI) + qed qed -lemma star_prop[rule_format]: "x \ lang\ \ \ y. y \ lang\ \ x @ y \ lang\" +text {* **************** the Star case ****************** *} + +lemma finite_set_has_max: "\finite A; A \ {}\ \ (\ max \ A. \ a \ A. f a <= (f max :: nat))" +proof (induct rule:finite.induct) + case emptyI thus ?case by simp +next + case (insertI A a) + show ?case + proof (cases "A = {}") + case True thus ?thesis by (rule_tac x = a in bexI, auto) + next + case False + with prems obtain max where h1: "max \ A" and h2: "\a\A. f a \ f max" by blast + show ?thesis + proof (cases "f a \ f max") + assume "f a \ f max" + with h1 h2 show ?thesis by (rule_tac x = max in bexI, auto) + next + assume "\ (f a \ f max)" + thus ?thesis using h2 by (rule_tac x = a in bexI, auto) + qed + qed +qed + +lemma star_intro1[rule_format]: "x \ lang\ \ \ y. y \ lang\ \ x @ y \ lang\" by (erule Star.induct, auto) -lemma star_prop2: "y \ lang \ y \ lang\" +lemma star_intro2: "y \ lang \ y \ lang\" by (drule step[of y lang "[]"], auto simp:start) -lemma star_prop3[rule_format]: "x \ lang\ \ \y . y \ lang \ x @ y \ lang\" -by (erule Star.induct, auto intro:star_prop2) +lemma star_intro3[rule_format]: "x \ lang\ \ \y . y \ lang \ x @ y \ lang\" +by (erule Star.induct, auto intro:star_intro2) + +lemma star_decom: "\x \ lang\; x \ []\ \(\ a b. x = a @ b \ a \ [] \ a \ lang \ b \ lang\)" +by (induct x rule: Star.induct, simp, blast) -lemma postfix_prop: "y >>= (x @ y) \ x = []" -by (erule postfixE, induct x arbitrary:y, auto) +lemma finite_strict_prefix_set: "finite {xa. xa < (x::string)}" +apply (induct x rule:rev_induct, simp) +apply (subgoal_tac "{xa. xa < xs @ [x]} = {xa. xa < xs} \ {xs}") +by (auto simp:strict_prefix_def) + +definition + "tag_str_STAR L\<^isub>1 x \ {(\L\<^isub>1) `` {x - xa} | xa. xa < x \ xa \ L\<^isub>1\}" + +lemma tag_str_star_range_finite: + "finite (UNIV // \L\<^isub>1) \ finite (range (tag_str_STAR L\<^isub>1))" +apply (rule_tac B = "Pow (UNIV // \L\<^isub>1)" in finite_subset) +by (auto simp:tag_str_STAR_def Image_def quotient_def split:if_splits) -lemma inj_aux: - "\(m @ z) \ L\<^isub>1\; m \L\<^isub>1\ yb; xa @ m = x; xa \ L\<^isub>1\; m \ []; - \ xa xb. x = xa @ xb \ xa \ L\<^isub>1\ \ xb \ [] \ xb @ z \ L\<^isub>1\ \ xb >>= m\ - \ (yb @ z) \ L\<^isub>1\" -proof- - have "\s. s \ L\<^isub>1\ \ \ m z yb. (s = m @ z \ m \L\<^isub>1\ yb \ x = xa @ m \ xa \ L\<^isub>1\ \ m \ [] \ - (\ xa xb. x = xa @ xb \ xa \ L\<^isub>1\ \ xb \ [] \ xb @ z \ L\<^isub>1\ \ xb >>= m) \ (yb @ z) \ L\<^isub>1\)" - apply (erule Star.induct, simp) - apply (rule allI | rule impI | erule conjE)+ - apply (drule app_eq_elim) - apply (erule disjE | erule exE | erule conjE)+ - apply simp - apply (simp (no_asm) only:append_assoc[THEN sym]) - apply (rule step) - apply (simp add:equiv_str_def) - apply simp - - apply (erule exE | erule conjE)+ - apply (rotate_tac 3) - apply (frule_tac x = "xa @ s1" in spec) - apply (rotate_tac 12) - apply (drule_tac x = ba in spec) - apply (erule impE) - apply ( simp add:star_prop3) - apply (simp) - apply (drule postfix_prop) - apply simp - done - thus "\(m @ z) \ L\<^isub>1\; m \L\<^isub>1\ yb; xa @ m = x; xa \ L\<^isub>1\; m \ []; - \ xa xb. x = xa @ xb \ xa \ L\<^isub>1\ \ xb \ [] \ xb @ z \ L\<^isub>1\ \ xb >>= m\ - \ (yb @ z) \ L\<^isub>1\" by blast +lemma tag_str_STAR_injI: + "tag_str_STAR L\<^isub>1 m = tag_str_STAR L\<^isub>1 n \ m \(L\<^isub>1\) n" +proof- + { fix x y z + assume xz_in_star: "x @ z \ L\<^isub>1\" + and tag_xy: "tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 y" + have "y @ z \ L\<^isub>1\" + proof(cases "x = []") + case True + with tag_xy have "y = []" by (auto simp:tag_str_STAR_def strict_prefix_def) + thus ?thesis using xz_in_star True by simp + next + case False + obtain x_max where h1: "x_max < x" and h2: "x_max \ L\<^isub>1\" and h3: "(x - x_max) @ z \ L\<^isub>1\" + and h4:"\ xa < x. xa \ L\<^isub>1\ \ (x - xa) @ z \ L\<^isub>1\ \ length xa \ length x_max" + proof- + let ?S = "{xa. xa < x \ xa \ L\<^isub>1\ \ (x - xa) @ z \ L\<^isub>1\}" + have "finite ?S" + by (rule_tac B = "{xa. xa < x}" in finite_subset, auto simp:finite_strict_prefix_set) + moreover have "?S \ {}" using False xz_in_star + by (simp, rule_tac x = "[]" in exI, auto simp:strict_prefix_def) + ultimately have "\ max \ ?S. \ a \ ?S. length a \ length max" using finite_set_has_max by blast + with prems show ?thesis by blast + qed + obtain ya where h5: "ya < y" and h6: "ya \ L\<^isub>1\" and h7: "(x - x_max) \L\<^isub>1 (y - ya)" + proof- + from tag_xy have "{\L\<^isub>1 `` {x - xa} |xa. xa < x \ xa \ L\<^isub>1\} = + {\L\<^isub>1 `` {y - xa} |xa. xa < y \ xa \ L\<^isub>1\}" (is "?left = ?right") + by (auto simp:tag_str_STAR_def) + moreover have "\L\<^isub>1 `` {x - x_max} \ ?left" using h1 h2 by auto + ultimately have "\L\<^isub>1 `` {x - x_max} \ ?right" by simp + with prems show ?thesis apply (simp add:Image_def str_eq_rel_def str_eq_def) by blast + qed + have "(y - ya) @ z \ L\<^isub>1\" + proof- + from h3 h1 obtain a b where a_in: "a \ L\<^isub>1" and a_neq: "a \ []" and b_in: "b \ L\<^isub>1\" + and ab_max: "(x - x_max) @ z = a @ b" + by (drule_tac star_decom, auto simp:strict_prefix_def elim:prefixE) + have "(x - x_max) \ a \ (a - (x - x_max)) @ b = z" + proof - + have "((x - x_max) \ a \ (a - (x - x_max)) @ b = z) \ (a < (x - x_max) \ ((x - x_max) - a) @ z = b)" + using app_eq_dest[OF ab_max] by (auto simp:strict_prefix_def) + moreover { + assume np: "a < (x - x_max)" and b_eqs: " ((x - x_max) - a) @ z = b" + have "False" + proof - + let ?x_max' = "x_max @ a" + have "?x_max' < x" using np h1 by (clarsimp simp:strict_prefix_def diff_prefix) + moreover have "?x_max' \ L\<^isub>1\" using a_in h2 by (simp add:star_intro3) + moreover have "(x - ?x_max') @ z \ L\<^isub>1\" using b_eqs b_in np h1 by (simp add:diff_diff_appd) + moreover have "\ (length ?x_max' \ length x_max)" using a_neq by simp + ultimately show ?thesis using h4 by blast + qed + } ultimately show ?thesis by blast + qed + then obtain za where z_decom: "z = za @ b" and x_za: "(x - x_max) @ za \ L\<^isub>1" + using a_in by (auto elim:prefixE) + from x_za h7 have "(y - ya) @ za \ L\<^isub>1" by (auto simp:str_eq_def) + with z_decom b_in show ?thesis by (auto dest!:step[of "(y - ya) @ za"]) + qed + with h5 h6 show ?thesis by (drule_tac star_intro1, auto simp:strict_prefix_def elim:prefixE) + qed + } thus "tag_str_STAR L\<^isub>1 m = tag_str_STAR L\<^isub>1 n \ m \(L\<^isub>1\) n" + by (auto simp add:str_eq_def str_eq_rel_def) qed - -lemma min_postfix_exists[rule_format]: - "finite A \ A \ {} \ (\ a \ A. \ b \ A. ((b >>= a) \ (a >>= b))) - \ (\ min. (min \ A \ (\ a \ A. a >>= min)))" -apply (erule finite.induct) -apply simp -apply simp -apply (case_tac "A = {}") -apply simp -apply clarsimp -apply (case_tac "a >>= min") -apply (rule_tac x = min in exI, simp) -apply (rule_tac x = a in exI, simp) -apply clarify -apply (rotate_tac 5) -apply (erule_tac x = aa in ballE) defer apply simp -apply (erule_tac ys = min in postfix_trans) -apply (erule_tac x = min in ballE) -by simp+ - -lemma tag_str_star_inj: - "tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 (y::string) \ x \L\<^isub>1\\ y" -proof - - have aux: "\ x y z. \tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 y; x @ z \ L\<^isub>1\\ \ y @ z \ L\<^isub>1\" - proof- - fix x y z - assume tag_eq: "tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 y" - and x_z: "x @ z \ L\<^isub>1\" - show "y @ z \ L\<^isub>1\" - proof (cases "x = []") - case True - with tag_eq have "y = []" by (simp add:tag_str_STAR_def split:if_splits, blast) - thus ?thesis using x_z True by simp - next - case False - hence not_empty: "{xb. \ xa. x = xa @ xb \ xa \ L\<^isub>1\ \ xb \ [] \ xb @ z \ L\<^isub>1\} \ {}" using x_z - by (simp, rule_tac x = x in exI, rule_tac x = "[]" in exI, simp add:start) - have finite_set: "finite {xb. \ xa. x = xa @ xb \ xa \ L\<^isub>1\ \ xb \ [] \ xb @ z \ L\<^isub>1\}" - apply (rule_tac B = "{xb. \ xa. x = xa @ xb}" in finite_subset) - apply auto - apply (induct x, simp) - apply (subgoal_tac "{xb. \xa. a # x = xa @ xb} = insert (a # x) {xb. \xa. x = xa @ xb}") - apply auto - by (case_tac xaa, simp+) - have comparable: "\ a \ {xb. \ xa. x = xa @ xb \ xa \ L\<^isub>1\ \ xb \ [] \ xb @ z \ L\<^isub>1\}. - \ b \ {xb. \ xa. x = xa @ xb \ xa \ L\<^isub>1\ \ xb \ [] \ xb @ z \ L\<^isub>1\}. - ((b >>= a) \ (a >>= b))" - by (auto simp:postfix_def, drule app_eq_elim, blast) - hence "\ min. min \ {xb. \ xa. x = xa @ xb \ xa \ L\<^isub>1\ \ xb \ [] \ xb @ z \ L\<^isub>1\} - \ (\ xa xb. x = xa @ xb \ xa \ L\<^isub>1\ \ xb \ [] \ xb @ z \ L\<^isub>1\ \ xb >>= min)" - using finite_set not_empty comparable - apply (drule_tac min_postfix_exists, simp) - by (erule exE, rule_tac x = min in exI, auto) - then obtain min xa where x_decom: "x = xa @ min \ xa \ L\<^isub>1\" - and min_not_empty: "min \ []" - and min_z_in_star: "min @ z \ L\<^isub>1\" - and is_min: "\ xa xb. x = xa @ xb \ xa \ L\<^isub>1\ \ xb \ [] \ xb @ z \ L\<^isub>1\ \ xb >>= min" by blast - from x_decom min_not_empty have "\min\L\<^isub>1 \ tag_str_STAR L\<^isub>1 x" by (auto simp:tag_str_STAR_def) - hence "\ yb. \yb\L\<^isub>1 \ tag_str_STAR L\<^isub>1 y \ \min\L\<^isub>1 = \yb\L\<^isub>1" using tag_eq by auto - hence "\ ya yb. y = ya @ yb \ ya \ L\<^isub>1\ \ min \L\<^isub>1\ yb \ yb \ [] " - by (simp add:tag_str_STAR_def equiv_class_def equiv_str_def split:if_splits, blast) - then obtain ya yb where y_decom: "y = ya @ yb" - and ya_in_star: "ya \ L\<^isub>1\" - and yb_not_empty: "yb \ []" - and min_yb_eq: "min \L\<^isub>1\ yb" by blast - from min_z_in_star min_yb_eq min_not_empty is_min x_decom - have "yb @ z \ L\<^isub>1\" - by (rule_tac x = x and xa = xa in inj_aux, simp+) - thus ?thesis using ya_in_star y_decom - by (auto dest:star_prop) - qed - qed - show "tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 (y::string) \ x \L\<^isub>1\\ y" - by (auto intro:aux simp:equiv_str_def) +lemma quot_star_finiteI: + assumes finite: "finite (UNIV // \(L\<^isub>1::string set))" + shows "finite (UNIV // \(L\<^isub>1\))" +proof(rule_tac f = "(op `) (tag_str_STAR L\<^isub>1)" in finite_imageD) + show "finite (op ` (tag_str_STAR L\<^isub>1) ` UNIV // \L\<^isub>1\)" using finite + by (auto intro:finite_tag_imageI tag_str_star_range_finite) +next + show "inj_on (op ` (tag_str_STAR L\<^isub>1)) (UNIV // \L\<^isub>1\)" + by (auto intro:tag_image_injI tag_str_STAR_injI) qed -lemma quot_star: - assumes finite1: "finite (QUOT L\<^isub>1)" - shows "finite (QUOT (L\<^isub>1\))" -proof (rule_tac f = "(op `) (tag_str_STAR L\<^isub>1)" in finite_imageD) - show "finite (op ` (tag_str_STAR L\<^isub>1) ` QUOT (L\<^isub>1\))" - using finite_tag_image tag_str_star_range_finite finite1 - by auto -next - show "inj_on (op ` (tag_str_STAR L\<^isub>1)) (QUOT (L\<^isub>1\))" - apply (rule_tac str_inj_imps) - by (erule_tac tag_str_star_inj) -qed +text {* **************** the Other Direction ************ *} lemma other_direction: - "Lang = L (r::rexp) \ finite (QUOT Lang)" -apply (induct arbitrary:Lang rule:rexp.induct) -apply (simp add:QUOT_def equiv_class_def equiv_str_def) -by (simp_all add:quot_lambda quot_single quot_seq quot_alt quot_star) + "Lang = L (r::rexp) \ finite (UNIV // (\Lang))" +proof (induct arbitrary:Lang rule:rexp.induct) + case NULL + have "UNIV // (\{}) \ {UNIV} " + by (auto simp:quotient_def str_eq_rel_def str_eq_def) + with prems show "?case" by (auto intro:finite_subset) +next + case EMPTY + have "UNIV // (\{[]}) \ {{[]}, UNIV - {[]}}" by (rule quot_empty_subset) + with prems show ?case by (auto intro:finite_subset) +next + case (CHAR c) + have "UNIV // (\{[c]}) \ {{[]},{[c]}, UNIV - {[], [c]}}" by (rule quot_char_subset) + with prems show ?case by (auto intro:finite_subset) +next + case (SEQ r\<^isub>1 r\<^isub>2) + have "\finite (UNIV // \(L r\<^isub>1)); finite (UNIV // \(L r\<^isub>2))\ \ finite (UNIV // \(L r\<^isub>1 ;; L r\<^isub>2))" + by (erule quot_seq_finiteI, simp) + with prems show ?case by simp +next + case (ALT r\<^isub>1 r\<^isub>2) + have "\finite (UNIV // \(L r\<^isub>1)); finite (UNIV // \(L r\<^isub>2))\ \ finite (UNIV // \(L r\<^isub>1 \ L r\<^isub>2))" + by (erule quot_union_finiteI, simp) + with prems show ?case by simp +next + case (STAR r) + have "finite (UNIV // \(L r)) \ finite (UNIV // \((L r)\))" + by (erule quot_star_finiteI) + with prems show ?case by simp +qed -end +end \ No newline at end of file