# HG changeset patch # User wu # Date 1292337091 0 # Node ID 90a57a533b0c51740bef95012fb486396fb1f235 # Parent b0cdf8f5a9af7439ec68ee9e0bd10b67e576824f Add new file for the new definition of the hard direction's simplification. Merging Operation is deleted All definitions are done. Proof still undone. diff -r b0cdf8f5a9af -r 90a57a533b0c Myhill.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Myhill.thy Tue Dec 14 14:31:31 2010 +0000 @@ -0,0 +1,1421 @@ +theory MyhillNerode + imports "Main" "List_Prefix" +begin + +text {* sequential composition of languages *} +definition + Seq :: "string set \ string set \ string set" ("_ ;; _" [100,100] 100) +where + "L1 ;; L2 = {s1 @ s2 | s1 s2. s1 \ L1 \ s2 \ L2}" + +inductive_set + Star :: "string set \ string set" ("_\" [101] 102) + for L :: "string set" +where + start[intro]: "[] \ L\" +| step[intro]: "\s1 \ L; s2 \ L\\ \ s1@s2 \ L\" + +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 + 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 + by (auto) (metis append_assoc)+ + finally show "X = X ;; A \ B" using eq by auto +next + assume "X = X ;; A \ B" + then have "B \ X" "X ;; A \ X" by auto + show "X = B ;; A\" sorry +qed + +datatype rexp = + NULL +| EMPTY +| CHAR char +| SEQ rexp rexp +| ALT rexp rexp +| STAR rexp + +consts L:: "'a \ string set" + +overloading L_rexp \ "L:: rexp \ string set" +begin + +fun + L_rexp :: "rexp \ string set" +where + "L_rexp (NULL) = {}" + | "L_rexp (EMPTY) = {[]}" + | "L_rexp (CHAR c) = {[c]}" + | "L_rexp (SEQ r1 r2) = (L_rexp r1) ;; (L_rexp r2)" + | "L_rexp (ALT r1 r2) = (L_rexp r1) \ (L_rexp r2)" + | "L_rexp (STAR r) = (L_rexp r)\" +end + +definition + folds :: "('a \ 'b \ 'b) \ 'b \ 'a set \ 'b" +where + "folds f z S \ SOME x. fold_graph f z S x" + +lemma folds_simp_null [simp]: + "finite rs \ x \ L (folds ALT NULL rs) \ (\r \ rs. x \ L r)" +apply (simp add: folds_def) +apply (rule someI2_ex) +apply (erule finite_imp_fold_graph) +apply (erule fold_graph.induct) +apply (auto) +done + +lemma [simp]: + shows "(x, y) \ {(x, y). P x y} \ P x y" +by simp + +definition + str_eq ("_ \_ _") +where + "x \Lang y \ (\z. x @ z \ Lang \ y @ z \ Lang)" + +definition + str_eq_rel ("\_") +where + "\Lang \ {(x, y). x \Lang y}" + + + +section {* finite \ regular *} + +definition + transitions :: "string set \ string set \ rexp set" ("_\\_") +where + "Y \\ X \ {CHAR c | c. Y ;; {[c]} \ X}" + +definition + transitions_rexp ("_ \\ _") +where + "Y \\ X \ folds ALT NULL (Y \\X)" + +definition + "init_rhs CS X \ if X = {[]} + then {({[]}, EMPTY)} + else if ([] \ X) + then insert ({[]}, EMPTY) {(Y, Y \\X) | Y. Y \ CS} + else {(Y, Y \\X) | Y. Y \ CS}" + +overloading L_rhs \ "L:: (string set \ rexp) set \ string set" +begin +fun L_rhs:: "(string set \ rexp) set \ string set" +where + "L_rhs rhs = \ {(Y;; L r) | Y r . (Y, r) \ rhs}" +end + +definition + "eqs CS \ (\X \ CS. {(X, init_rhs CS X)})" + +lemma [simp]: + shows "finite (Y \\ X)" +unfolding transitions_def +by auto + +lemma defined_by_str: + assumes "s \ X" + and "X \ UNIV // (\Lang)" + shows "X = (\Lang) `` {s}" +using assms +unfolding quotient_def Image_def +unfolding str_eq_rel_def str_eq_def +by auto + + + +(************ arden's lemma variation ********************) +definition + "rexp_of rhs X \ folds ALT NULL {r. (X, r) \ rhs}" + +definition + "arden_variate X rhs \ {(Y, SEQ r (STAR (rexp_of rhs X)))| Y r. (Y, r) \ rhs \ Y \ X}" + +(************* rhs/equations property **************) + +definition + "distinct_equas ES \ \ X rhs rhs'. (X, rhs) \ ES \ (X, rhs') \ ES \ rhs = rhs'" + +(*********** substitution of ES *************) + +text {* rhs_subst rhs X xrhs: substitude all occurence of X in rhs with xrhs *} +definition + "rhs_subst rhs X xrhs \ {(Y, r) | Y r. Y \ X \ (Y, r) \ rhs} \ + {(X, SEQ r\<^isub>1 r\<^isub>2 ) | r\<^isub>1 r\<^isub>2. (X, r\<^isub>1) \ xrhs \ (X, r\<^isub>2) \ rhs}" + +definition + "eqs_subst ES X xrhs \ {(Y, rhs_subst yrhs X xrhs) | Y yrhs. (Y, yrhs) \ ES}" + +text {* + Inv: Invairance of the equation-system, during the decrease of the equation-system, Inv holds. +*} + +definition + "ardenable ES \ \ X rhs. (X, rhs) \ ES \ ([] \ L (rexp_of rhs X)) \ X = L rhs" + +definition + "non_empty ES \ \ X rhs. (X, rhs) \ ES \ X \ {}" + +definition + "self_contained ES \ \ X xrhs. (X, xrhs) \ ES + \ (\ Y r.(Y, r) \ xrhs \ Y \ {[]} \ (\ yrhs. (Y, yrhs) \ ES))" + +definition + "Inv ES \ finite ES \ distinct_equas ES \ ardenable ES \ non_empty ES \ self_contained ES" + +lemma wf_iter [rule_format]: + fixes f + assumes step: "\ e. \P e; \ Q e\ \ (\ e'. P e' \ (f(e'), f(e)) \ less_than)" + shows pe: "P e \ (\ e'. P e' \ Q e')" +proof(induct e rule: wf_induct + [OF wf_inv_image[OF wf_less_than, where f = "f"]], clarify) + fix x + assume h [rule_format]: + "\y. (y, x) \ inv_image less_than f \ P y \ (\e'. P e' \ Q e')" + and px: "P x" + show "\e'. P e' \ Q e'" + proof(cases "Q x") + assume "Q x" with px show ?thesis by blast + next + assume nq: "\ Q x" + from step [OF px nq] + obtain e' where pe': "P e'" and ltf: "(f e', f x) \ less_than" by auto + show ?thesis + proof(rule h) + from ltf show "(e', x) \ inv_image less_than f" + by (simp add:inv_image_def) + next + from pe' show "P e'" . + qed + qed +qed + +text {* ******BEGIN: proving the initial equation-system satisfies Inv ****** *} + +lemma init_ES_satisfy_Inv: + assumes finite_CS: "finite (UNIV // (\Lang))" + and X_in_eq_cls: "X \ (UNIV // (\Lang))" + shows "Inv (eqs (UNIV // (\Lang)))" +proof - + have "finite (eqs (UNIV // (\Lang)))" using finite_CS + by (auto simp add:eqs_def) + moreover have "distinct_equas (eqs (UNIV // (\Lang)))" + by (auto simp:distinct_equas_def eqs_def) + moreover have "ardenable (eqs (UNIV // (\Lang)))" + proof- + have "\ X rhs. (X, rhs) \ (eqs (UNIV // (\Lang))) \ ([] \ L (rexp_of rhs X))" + proof + apply (auto simp:eqs_def rexp_of_def) + sorry + moreover have "\ X rhs. (X, rhs) \ (eqs (UNIV // (\Lang))) \ X = L rhs" + sorry + ultimately show ?thesis by (simp add:ardenable_def) + qed + moreover have "non_empty (eqs (UNIV // (\Lang)))" + by (auto simp:non_empty_def eqs_def quotient_def Image_def str_eq_rel_def str_eq_def) + moreover have "self_contained (eqs (UNIV // (\Lang)))" + by (auto simp:self_contained_def eqs_def init_rhs_def) + ultimately show ?thesis by (simp add:Inv_def) +qed + + +text {* ****** BEGIN: proving every equation-system's iteration step satisfies Inv ***** *} + +lemma iteration_step: + assumes Inv_ES: "Inv ES" + and X_in_ES: "\ xrhs. (X, xrhs) \ ES" + and not_T: "card ES > 1" + shows "(\ ES' xrhs'. Inv ES' \ (card ES', card ES) \ less_than \ (X, xrhs') \ ES')" +proof - + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + +lemma distinct_rhs_equations: + "(X, xrhs) \ equations (UNIV Quo Lang) \ distinct_rhs xrhs" +by (auto simp: equations_def equation_rhs_def distinct_rhs_def empty_rhs_def dest:no_two_cls_inters) + +lemma every_nonempty_eqclass_has_strings: + "\X \ (UNIV Quo Lang); X \ {[]}\ \ \ clist. clist \ X \ clist \ []" +by (auto simp:quot_def equiv_class_def equiv_str_def) + +lemma every_eqclass_is_derived_from_empty: + assumes not_empty: "X \ {[]}" + shows "X \ (UNIV Quo Lang) \ \ clist. {[]};{clist} \ X \ clist \ []" +using not_empty +apply (drule_tac every_nonempty_eqclass_has_strings, simp) +by (auto intro:exI[where x = clist] simp:lang_seq_def) + +lemma equiv_str_in_CS: + "\clist\Lang \ (UNIV Quo Lang)" +by (auto simp:quot_def) + +lemma has_str_imp_defined_by_str: + "\str \ X; X \ UNIV Quo Lang\ \ X = \str\Lang" +by (auto simp:quot_def equiv_class_def equiv_str_def) + +lemma every_eqclass_has_ascendent: + assumes has_str: "clist @ [c] \ X" + and in_CS: "X \ UNIV Quo Lang" + shows "\ Y. Y \ UNIV Quo Lang \ Y-c\X \ clist \ Y" (is "\ Y. ?P Y") +proof - + have "?P (\clist\Lang)" + proof - + have "\clist\Lang \ UNIV Quo Lang" + by (simp add:quot_def, rule_tac x = clist in exI, simp) + moreover have "\clist\Lang-c\X" + proof - + have "X = \(clist @ [c])\Lang" using has_str in_CS + by (auto intro!:has_str_imp_defined_by_str) + moreover have "\ sl. sl \ \clist\Lang \ sl @ [c] \ \(clist @ [c])\Lang" + by (auto simp:equiv_class_def equiv_str_def) + ultimately show ?thesis unfolding CT_def lang_seq_def + by auto + qed + moreover have "clist \ \clist\Lang" + by (auto simp:equiv_str_def equiv_class_def) + ultimately show "?P (\clist\Lang)" by simp + qed + thus ?thesis by blast +qed + +lemma finite_charset_rS: + "finite {CHAR c |c. Y-c\X}" +by (rule_tac A = UNIV and f = CHAR in finite_surj, auto) + +lemma l_eq_r_in_equations: + assumes X_in_equas: "(X, xrhs) \ equations (UNIV Quo Lang)" + shows "X = L xrhs" +proof (cases "X = {[]}") + case True + thus ?thesis using X_in_equas + by (simp add:equations_def equation_rhs_def lang_seq_def) +next + case False + show ?thesis + proof + show "X \ L xrhs" + proof + fix x + assume "(1)": "x \ X" + show "x \ L xrhs" + proof (cases "x = []") + assume empty: "x = []" + hence "x \ L (empty_rhs X)" using "(1)" + by (auto simp:empty_rhs_def lang_seq_def) + thus ?thesis using X_in_equas False empty "(1)" + unfolding equations_def equation_rhs_def by auto + next + assume not_empty: "x \ []" + hence "\ clist c. x = clist @ [c]" by (case_tac x rule:rev_cases, auto) + then obtain clist c where decom: "x = clist @ [c]" by blast + moreover have "\ Y. \Y \ UNIV Quo Lang; Y-c\X; clist \ Y\ + \ [c] \ L (folds ALT NULL {CHAR c |c. Y-c\X})" + proof - + fix Y + assume Y_is_eq_cl: "Y \ UNIV Quo Lang" + and Y_CT_X: "Y-c\X" + and clist_in_Y: "clist \ Y" + with finite_charset_rS + show "[c] \ L (folds ALT NULL {CHAR c |c. Y-c\X})" + by (auto simp :fold_alt_null_eqs) + qed + hence "\Xa. Xa \ UNIV Quo Lang \ clist @ [c] \ Xa ; L (folds ALT NULL {CHAR c |c. Xa-c\X})" + using X_in_equas False not_empty "(1)" decom + by (auto dest!:every_eqclass_has_ascendent simp:equations_def equation_rhs_def lang_seq_def) + then obtain Xa where + "Xa \ UNIV Quo Lang \ clist @ [c] \ Xa ; L (folds ALT NULL {CHAR c |c. Xa-c\X})" by blast + hence "x \ L {(S, folds ALT NULL {CHAR c |c. S-c\X}) |S. S \ UNIV Quo Lang}" + using X_in_equas "(1)" decom + by (auto simp add:equations_def equation_rhs_def intro!:exI[where x = Xa]) + thus "x \ L xrhs" using X_in_equas False not_empty unfolding equations_def equation_rhs_def + by auto + qed + qed + next + show "L xrhs \ X" + proof + fix x + assume "(2)": "x \ L xrhs" + have "(2_1)": "\ s1 s2 r Xa. \s1 \ Xa; s2 \ L (folds ALT NULL {CHAR c |c. Xa-c\X})\ \ s1 @ s2 \ X" + using finite_charset_rS + by (auto simp:CT_def lang_seq_def fold_alt_null_eqs) + have "(2_2)": "\ s1 s2 Xa r.\s1 \ Xa; s2 \ L r; (Xa, r) \ empty_rhs X\ \ s1 @ s2 \ X" + by (simp add:empty_rhs_def split:if_splits) + show "x \ X" using X_in_equas False "(2)" + by (auto intro:"(2_1)" "(2_2)" simp:equations_def equation_rhs_def lang_seq_def) + qed + qed +qed + + + +lemma no_EMPTY_equations: + "(X, xrhs) \ equations CS \ no_EMPTY_rhs xrhs" +apply (clarsimp simp add:equations_def equation_rhs_def) +apply (simp add:no_EMPTY_rhs_def empty_rhs_def, auto) +apply (subgoal_tac "finite {CHAR c |c. Xa-c\X}", drule_tac x = "[]" in fold_alt_null_eqs, clarsimp, rule finite_charset_rS)+ +done + +lemma init_ES_satisfy_ardenable: + "(X, xrhs) \ equations (UNIV Quo Lang) \ ardenable (X, xrhs)" + unfolding ardenable_def + by (auto intro:distinct_rhs_equations no_EMPTY_equations simp:l_eq_r_in_equations simp del:L_rhs.simps) + +lemma init_ES_satisfy_Inv: + assumes finite_CS: "finite (UNIV Quo Lang)" + and X_in_eq_cls: "X \ UNIV Quo Lang" + shows "Inv X (equations (UNIV Quo Lang))" +proof - + have "finite (equations (UNIV Quo Lang))" using finite_CS + by (auto simp:equations_def) + moreover have "\rhs. (X, rhs) \ equations (UNIV Quo Lang)" using X_in_eq_cls + by (simp add:equations_def) + moreover have "distinct_equas (equations (UNIV Quo Lang))" + by (auto simp:distinct_equas_def equations_def) + moreover have "\X xrhs. (X, xrhs) \ equations (UNIV Quo Lang) \ + rhs_eq_cls xrhs \ insert {[]} (left_eq_cls (equations (UNIV Quo Lang)))" + apply (simp add:left_eq_cls_def equations_def rhs_eq_cls_def equation_rhs_def) + by (auto simp:empty_rhs_def split:if_splits) + moreover have "\X xrhs. (X, xrhs) \ equations (UNIV Quo Lang) \ X \ {}" + by (clarsimp simp:equations_def empty_notin_CS intro:classical) + moreover have "\X xrhs. (X, xrhs) \ equations (UNIV Quo Lang) \ ardenable (X, xrhs)" + by (auto intro!:init_ES_satisfy_ardenable) + ultimately show ?thesis by (simp add:Inv_def) +qed + + +text {* *********** END: proving the initial equation-system satisfies Inv ******* *} + + + + + + + + + + + + + + + +text {* ****** BEGIN: proving every equation-system's iteration step satisfies Inv ***** *} + +lemma not_T_aux: "\\ TCon (insert a A); x = a\ + \ \y. x \ y \ y \ insert a A " +apply (case_tac "insert a A = {a}") +by (auto simp:TCon_def) + +lemma not_T_atleast_2[rule_format]: + "finite S \ \ x. x \ S \ (\ TCon S)\ (\ y. x \ y \ y \ S)" +apply (erule finite.induct, simp) +apply (clarify, case_tac "x = a") +by (erule not_T_aux, auto) + +lemma exist_another_equa: + "\\ TCon ES; finite ES; distinct_equas ES; (X, rhl) \ ES\ \ \ Y yrhl. (Y, yrhl) \ ES \ X \ Y" +apply (drule not_T_atleast_2, simp) +apply (clarsimp simp:distinct_equas_def) +apply (drule_tac x= X in spec, drule_tac x = rhl in spec, drule_tac x = b in spec) +by auto + +lemma Inv_mono_with_lambda: + assumes Inv_ES: "Inv X ES" + and X_noteq_Y: "X \ {[]}" + shows "Inv X (ES - {({[]}, yrhs)})" +proof - + have "finite (ES - {({[]}, yrhs)})" using Inv_ES + by (simp add:Inv_def) + moreover have "\rhs. (X, rhs) \ ES - {({[]}, yrhs)}" using Inv_ES X_noteq_Y + by (simp add:Inv_def) + moreover have "distinct_equas (ES - {({[]}, yrhs)})" using Inv_ES X_noteq_Y + apply (clarsimp simp:Inv_def distinct_equas_def) + by (drule_tac x = Xa in spec, simp) + moreover have "\X xrhs.(X, xrhs) \ ES - {({[]}, yrhs)} \ + ardenable (X, xrhs) \ X \ {}" using Inv_ES + by (clarify, simp add:Inv_def) + moreover + have "insert {[]} (left_eq_cls (ES - {({[]}, yrhs)})) = insert {[]} (left_eq_cls ES)" + by (auto simp:left_eq_cls_def) + hence "\X xrhs.(X, xrhs) \ ES - {({[]}, yrhs)} \ + rhs_eq_cls xrhs \ insert {[]} (left_eq_cls (ES - {({[]}, yrhs)}))" + using Inv_ES by (auto simp:Inv_def) + ultimately show ?thesis by (simp add:Inv_def) +qed + +lemma non_empty_card_prop: + "finite ES \ \e. e \ ES \ card ES - Suc 0 < card ES" +apply (erule finite.induct, simp) +apply (case_tac[!] "a \ A") +by (auto simp:insert_absorb) + +lemma ardenable_prop: + assumes not_lambda: "Y \ {[]}" + and ardable: "ardenable (Y, yrhs)" + shows "\ yrhs'. Y = L yrhs' \ distinct_rhs yrhs' \ rhs_eq_cls yrhs' = rhs_eq_cls yrhs - {Y}" (is "\ yrhs'. ?P yrhs'") +proof (cases "(\ reg. (Y, reg) \ yrhs)") + case True + thus ?thesis + proof + fix reg + assume self_contained: "(Y, reg) \ yrhs" + show ?thesis + proof - + have "?P (arden_variate Y reg yrhs)" + proof - + have "Y = L (arden_variate Y reg yrhs)" + using self_contained not_lambda ardable + by (rule_tac arden_variate_valid, simp_all add:ardenable_def) + moreover have "distinct_rhs (arden_variate Y reg yrhs)" + using ardable + by (auto simp:distinct_rhs_def arden_variate_def seq_rhs_r_def del_x_paired_def ardenable_def) + moreover have "rhs_eq_cls (arden_variate Y reg yrhs) = rhs_eq_cls yrhs - {Y}" + proof - + have "\ rhs r. rhs_eq_cls (seq_rhs_r rhs r) = rhs_eq_cls rhs" + apply (auto simp:rhs_eq_cls_def seq_rhs_r_def image_def) + by (rule_tac x = "SEQ ra r" in exI, rule_tac x = "(x, ra)" in bexI, simp+) + moreover have "\ rhs X. rhs_eq_cls (del_x_paired rhs X) = rhs_eq_cls rhs - {X}" + by (auto simp:rhs_eq_cls_def del_x_paired_def) + ultimately show ?thesis by (simp add:arden_variate_def) + qed + ultimately show ?thesis by simp + qed + thus ?thesis by (rule_tac x= "arden_variate Y reg yrhs" in exI, simp) + qed + qed +next + case False + hence "(2)": "rhs_eq_cls yrhs - {Y} = rhs_eq_cls yrhs" + by (auto simp:rhs_eq_cls_def) + show ?thesis + proof - + have "?P yrhs" using False ardable "(2)" + by (simp add:ardenable_def) + thus ?thesis by blast + qed +qed + +lemma equas_subst_f_del_no_other: + assumes self_contained: "(Y, rhs) \ ES" + shows "\ rhs'. (Y, rhs') \ (equas_subst_f X xrhs ` ES)" (is "\ rhs'. ?P rhs'") +proof - + have "\ rhs'. equas_subst_f X xrhs (Y, rhs) = (Y, rhs')" + by (auto simp:equas_subst_f_def) + then obtain rhs' where "equas_subst_f X xrhs (Y, rhs) = (Y, rhs')" by blast + hence "?P rhs'" unfolding image_def using self_contained + by (auto intro:bexI[where x = "(Y, rhs)"]) + thus ?thesis by blast +qed + +lemma del_x_paired_del_only_x: + "\X \ Y; (X, rhs) \ ES\ \ (X, rhs) \ del_x_paired ES Y" +by (auto simp:del_x_paired_def) + +lemma equas_subst_del_no_other: + "\(X, rhs) \ ES; X \ Y\ \ (\rhs. (X, rhs) \ equas_subst ES Y rhs')" +unfolding equas_subst_def +apply (drule_tac X = Y and xrhs = rhs' in equas_subst_f_del_no_other) +by (erule exE, drule del_x_paired_del_only_x, auto) + +lemma equas_subst_holds_distinct: + "distinct_equas ES \ distinct_equas (equas_subst ES Y rhs')" +apply (clarsimp simp add:equas_subst_def distinct_equas_def del_x_paired_def equas_subst_f_def) +by (auto split:if_splits) + +lemma del_x_paired_dels: + "(X, rhs) \ ES \ {Y. Y \ ES \ fst Y = X} \ ES \ {}" +by (auto) + +lemma del_x_paired_subset: + "(X, rhs) \ ES \ ES - {Y. Y \ ES \ fst Y = X} \ ES" +apply (drule del_x_paired_dels) +by auto + +lemma del_x_paired_card_less: + "\finite ES; (X, rhs) \ ES\ \ card (del_x_paired ES X) < card ES" +apply (simp add:del_x_paired_def) +apply (drule del_x_paired_subset) +by (auto intro:psubset_card_mono) + +lemma equas_subst_card_less: + "\finite ES; (Y, yrhs) \ ES\ \ card (equas_subst ES Y rhs') < card ES" +apply (simp add:equas_subst_def) +apply (frule_tac h = "equas_subst_f Y rhs'" in finite_imageI) +apply (drule_tac f = "equas_subst_f Y rhs'" in Finite_Set.card_image_le) +apply (drule_tac X = Y and xrhs = rhs' in equas_subst_f_del_no_other,erule exE) +by (drule del_x_paired_card_less, auto) + +lemma equas_subst_holds_distinct_rhs: + assumes dist': "distinct_rhs yrhs'" + and history: "\X xrhs. (X, xrhs) \ ES \ ardenable (X, xrhs)" + and X_in : "(X, xrhs) \ equas_subst ES Y yrhs'" + shows "distinct_rhs xrhs" +using X_in history +apply (clarsimp simp:equas_subst_def del_x_paired_def) +apply (drule_tac x = a in spec, drule_tac x = b in spec) +apply (simp add:ardenable_def equas_subst_f_def) +by (auto intro:rhs_subst_holds_distinct_rhs simp:dist' split:if_splits) + +lemma r_no_EMPTY_imp_seq_rhs_r_no_EMPTY: + "[] \ L r \ no_EMPTY_rhs (seq_rhs_r rhs r)" +by (auto simp:no_EMPTY_rhs_def seq_rhs_r_def lang_seq_def) + +lemma del_x_paired_holds_no_EMPTY: + "no_EMPTY_rhs yrhs \ no_EMPTY_rhs (del_x_paired yrhs Y)" +by (auto simp:no_EMPTY_rhs_def del_x_paired_def) + +lemma rhs_subst_holds_no_EMPTY: + "\no_EMPTY_rhs yrhs; (Y, r) \ yrhs; Y \ {[]}\ \ no_EMPTY_rhs (rhs_subst yrhs Y rhs' r)" +apply (auto simp:rhs_subst_def intro!:no_EMPTY_rhss_imp_merge_no_EMPTY r_no_EMPTY_imp_seq_rhs_r_no_EMPTY del_x_paired_holds_no_EMPTY) +by (auto simp:no_EMPTY_rhs_def) + +lemma equas_subst_holds_no_EMPTY: + assumes substor: "Y \ {[]}" + and history: "\X xrhs. (X, xrhs) \ ES \ ardenable (X, xrhs)" + and X_in:"(X, xrhs) \ equas_subst ES Y yrhs'" + shows "no_EMPTY_rhs xrhs" +proof- + from X_in have "\ (Z, zrhs) \ ES. (X, xrhs) = equas_subst_f Y yrhs' (Z, zrhs)" + by (auto simp add:equas_subst_def del_x_paired_def) + then obtain Z zrhs where Z_in: "(Z, zrhs) \ ES" + and X_Z: "(X, xrhs) = equas_subst_f Y yrhs' (Z, zrhs)" by blast + hence dist_zrhs: "distinct_rhs zrhs" using history + by (auto simp:ardenable_def) + show ?thesis + proof (cases "\ r. (Y, r) \ zrhs") + case True + then obtain r where Y_in_zrhs: "(Y, r) \ zrhs" .. + hence some: "(SOME r. (Y, r) \ zrhs) = r" using Z_in dist_zrhs + by (auto simp:distinct_rhs_def) + hence "no_EMPTY_rhs (rhs_subst zrhs Y yrhs' r)" + using substor Y_in_zrhs history Z_in + by (rule_tac rhs_subst_holds_no_EMPTY, auto simp:ardenable_def) + thus ?thesis using X_Z True some + by (simp add:equas_subst_def equas_subst_f_def) + next + case False + hence "(X, xrhs) = (Z, zrhs)" using Z_in X_Z + by (simp add:equas_subst_f_def) + thus ?thesis using history Z_in + by (auto simp:ardenable_def) + qed +qed + +lemma equas_subst_f_holds_left_eq_right: + assumes substor: "Y = L rhs'" + and history: "\X xrhs. (X, xrhs) \ ES \ distinct_rhs xrhs \ X = L xrhs" + and subst: "(X, xrhs) = equas_subst_f Y rhs' (Z, zrhs)" + and self_contained: "(Z, zrhs) \ ES" + shows "X = L xrhs" +proof (cases "\ r. (Y, r) \ zrhs") + case True + from True obtain r where "(1)":"(Y, r) \ zrhs" .. + show ?thesis + proof - + from history self_contained + have dist: "distinct_rhs zrhs" by auto + hence "(SOME r. (Y, r) \ zrhs) = r" using self_contained "(1)" + using distinct_rhs_def by (auto intro:some_equality) + moreover have "L zrhs = L (rhs_subst zrhs Y rhs' r)" using substor dist "(1)" self_contained + by (rule_tac rhs_subst_prop1, auto simp:distinct_equas_rhs_def) + ultimately show ?thesis using subst history self_contained + by (auto simp:equas_subst_f_def split:if_splits) + qed +next + case False + thus ?thesis using history subst self_contained + by (auto simp:equas_subst_f_def) +qed + +lemma equas_subst_holds_left_eq_right: + assumes history: "\X xrhs. (X, xrhs) \ ES \ ardenable (X, xrhs)" + and substor: "Y = L rhs'" + and X_in : "(X, xrhs) \ equas_subst ES Y yrhs'" + shows "\X xrhs. (X, xrhs) \ equas_subst ES Y rhs' \ X = L xrhs" +apply (clarsimp simp add:equas_subst_def del_x_paired_def) +using substor +apply (drule_tac equas_subst_f_holds_left_eq_right) +using history +by (auto simp:ardenable_def) + +lemma equas_subst_holds_ardenable: + assumes substor: "Y = L yrhs'" + and history: "\X xrhs. (X, xrhs) \ ES \ ardenable (X, xrhs)" + and X_in:"(X, xrhs) \ equas_subst ES Y yrhs'" + and dist': "distinct_rhs yrhs'" + and not_lambda: "Y \ {[]}" + shows "ardenable (X, xrhs)" +proof - + have "distinct_rhs xrhs" using history X_in dist' + by (auto dest:equas_subst_holds_distinct_rhs) + moreover have "no_EMPTY_rhs xrhs" using history X_in not_lambda + by (auto intro:equas_subst_holds_no_EMPTY) + moreover have "X = L xrhs" using history substor X_in + by (auto dest: equas_subst_holds_left_eq_right simp del:L_rhs.simps) + ultimately show ?thesis using ardenable_def by simp +qed + +lemma equas_subst_holds_cls_defined: + assumes X_in: "(X, xrhs) \ equas_subst ES Y yrhs'" + and Inv_ES: "Inv X' ES" + and subst: "(Y, yrhs) \ ES" + and cls_holds_but_Y: "rhs_eq_cls yrhs' = rhs_eq_cls yrhs - {Y}" + shows "rhs_eq_cls xrhs \ insert {[]} (left_eq_cls (equas_subst ES Y yrhs'))" +proof- + have tac: "\ A \ B; C \ D; E \ A \ B\ \ E \ B \ D" by auto + from X_in have "\ (Z, zrhs) \ ES. (X, xrhs) = equas_subst_f Y yrhs' (Z, zrhs)" + by (auto simp add:equas_subst_def del_x_paired_def) + then obtain Z zrhs where Z_in: "(Z, zrhs) \ ES" + and X_Z: "(X, xrhs) = equas_subst_f Y yrhs' (Z, zrhs)" by blast + hence "rhs_eq_cls zrhs \ insert {[]} (left_eq_cls ES)" using Inv_ES + by (auto simp:Inv_def) + moreover have "rhs_eq_cls yrhs' \ insert {[]} (left_eq_cls ES) - {Y}" + using Inv_ES subst cls_holds_but_Y + by (auto simp:Inv_def) + moreover have "rhs_eq_cls xrhs \ rhs_eq_cls zrhs \ rhs_eq_cls yrhs' - {Y}" + using X_Z cls_holds_but_Y + apply (clarsimp simp add:equas_subst_f_def rhs_subst_def split:if_splits) + by (auto simp:rhs_eq_cls_def merge_rhs_def del_x_paired_def seq_rhs_r_def) + moreover have "left_eq_cls (equas_subst ES Y yrhs') = left_eq_cls ES - {Y}" using subst + by (auto simp: left_eq_cls_def equas_subst_def del_x_paired_def equas_subst_f_def + dest: equas_subst_f_del_no_other + split: if_splits) + ultimately show ?thesis by blast +qed + +lemma iteration_step: + assumes Inv_ES: "Inv X ES" + and not_T: "\ TCon ES" + shows "(\ ES'. Inv X ES' \ (card ES', card ES) \ less_than)" +proof - + from Inv_ES not_T have another: "\Y yrhs. (Y, yrhs) \ ES \ X \ Y" unfolding Inv_def + by (clarify, rule_tac exist_another_equa[where X = X], auto) + then obtain Y yrhs where subst: "(Y, yrhs) \ ES" and not_X: " X \ Y" by blast + show ?thesis (is "\ ES'. ?P ES'") + proof (cases "Y = {[]}") + case True + --"in this situation, we pick a \"\\" equation, thus directly remove it from the equation-system" + have "?P (ES - {(Y, yrhs)})" + proof + show "Inv X (ES - {(Y, yrhs)})" using True not_X + by (simp add:Inv_ES Inv_mono_with_lambda) + next + show "(card (ES - {(Y, yrhs)}), card ES) \ less_than" using Inv_ES subst + by (auto elim:non_empty_card_prop[rule_format] simp:Inv_def) + qed + thus ?thesis by blast + next + case False + --"in this situation, we pick a equation and using ardenable to get a + rhs without itself in it, then use equas_subst to form a new equation-system" + hence "\ yrhs'. Y = L yrhs' \ distinct_rhs yrhs' \ rhs_eq_cls yrhs' = rhs_eq_cls yrhs - {Y}" + using subst Inv_ES + by (auto intro:ardenable_prop simp add:Inv_def simp del:L_rhs.simps) + then obtain yrhs' where Y'_l_eq_r: "Y = L yrhs'" + and dist_Y': "distinct_rhs yrhs'" + and cls_holds_but_Y: "rhs_eq_cls yrhs' = rhs_eq_cls yrhs - {Y}" by blast + hence "?P (equas_subst ES Y yrhs')" + proof - + have finite_del: "\ S x. finite S \ finite (del_x_paired S x)" + apply (rule_tac A = "del_x_paired S x" in finite_subset) + by (auto simp:del_x_paired_def) + have "finite (equas_subst ES Y yrhs')" using Inv_ES + by (auto intro!:finite_del simp:equas_subst_def Inv_def) + moreover have "\rhs. (X, rhs) \ equas_subst ES Y yrhs'" using Inv_ES not_X + by (auto intro:equas_subst_del_no_other simp:Inv_def) + moreover have "distinct_equas (equas_subst ES Y yrhs')" using Inv_ES + by (auto intro:equas_subst_holds_distinct simp:Inv_def) + moreover have "\X xrhs. (X, xrhs) \ equas_subst ES Y yrhs' \ ardenable (X, xrhs)" + using Inv_ES dist_Y' False Y'_l_eq_r + apply (clarsimp simp:Inv_def) + by (rule equas_subst_holds_ardenable, simp_all) + moreover have "\X xrhs. (X, xrhs) \ equas_subst ES Y yrhs' \ X \ {}" using Inv_ES + by (clarsimp simp:equas_subst_def Inv_def del_x_paired_def equas_subst_f_def split:if_splits, auto) + moreover have "\X xrhs. (X, xrhs) \ equas_subst ES Y yrhs' \ + rhs_eq_cls xrhs \ insert {[]} (left_eq_cls (equas_subst ES Y yrhs'))" + using Inv_ES subst cls_holds_but_Y + apply (rule_tac impI | rule_tac allI)+ + by (erule equas_subst_holds_cls_defined, auto) + moreover have "(card (equas_subst ES Y yrhs'), card ES) \ less_than"using Inv_ES subst + by (simp add:equas_subst_card_less Inv_def) + ultimately show "?P (equas_subst ES Y yrhs')" by (auto simp:Inv_def) + qed + thus ?thesis by blast + qed +qed + +text {* ***** END: proving every equation-system's iteration step satisfies Inv ************** *} + +lemma iteration_conc: + assumes history: "Inv X ES" + shows "\ ES'. Inv X ES' \ TCon ES'" (is "\ ES'. ?P ES'") +proof (cases "TCon ES") + case True + hence "?P ES" using history by simp + thus ?thesis by blast +next + case False + thus ?thesis using history iteration_step + by (rule_tac f = card in wf_iter, simp_all) +qed + +lemma eqset_imp_iff': "A = B \ \ x. x \ A \ x \ B" +apply (auto simp:mem_def) +done + +lemma set_cases2: + "\(A = {} \ R A); \ x. (A = {x}) \ R A; \ x y. \x \ y; x \ A; y \ A\ \ R A\ \ R A" +apply (case_tac "A = {}", simp) +by (case_tac "\ x. A = {x}", clarsimp, blast) + +lemma rhs_aux:"\distinct_rhs rhs; {Y. \r. (Y, r) \ rhs} = {X}\ \ (\r. rhs = {(X, r)})" +apply (rule_tac A = rhs in set_cases2, simp) +apply (drule_tac x = X in eqset_imp_iff, clarsimp) +apply (drule eqset_imp_iff',clarsimp) +apply (frule_tac x = a in spec, drule_tac x = aa in spec) +by (auto simp:distinct_rhs_def) + +lemma every_eqcl_has_reg: + assumes finite_CS: "finite (UNIV Quo Lang)" + and X_in_CS: "X \ (UNIV Quo Lang)" + shows "\ (reg::rexp). L reg = X" (is "\ r. ?E r") +proof- + have "\ES'. Inv X ES' \ TCon ES'" using finite_CS X_in_CS + by (auto intro:init_ES_satisfy_Inv iteration_conc) + then obtain ES' where Inv_ES': "Inv X ES'" + and TCon_ES': "TCon ES'" by blast + from Inv_ES' TCon_ES' + have "\ rhs. ES' = {(X, rhs)}" + apply (clarsimp simp:Inv_def TCon_def) + apply (rule_tac x = rhs in exI) + by (auto dest!:card_Suc_Diff1 simp:card_eq_0_iff) + then obtain rhs where ES'_single_equa: "ES' = {(X, rhs)}" .. + hence X_ardenable: "ardenable (X, rhs)" using Inv_ES' + by (simp add:Inv_def) + + from X_ardenable have X_l_eq_r: "X = L rhs" + by (simp add:ardenable_def) + hence rhs_not_empty: "rhs \ {}" using Inv_ES' ES'_single_equa + by (auto simp:Inv_def ardenable_def) + have rhs_eq_cls: "rhs_eq_cls rhs \ {X, {[]}}" + using Inv_ES' ES'_single_equa + by (auto simp:Inv_def ardenable_def left_eq_cls_def) + have X_not_empty: "X \ {}" using Inv_ES' ES'_single_equa + by (auto simp:Inv_def) + show ?thesis + proof (cases "X = {[]}") + case True + hence "?E EMPTY" by (simp) + thus ?thesis by blast + next + case False with X_ardenable + have "\ rhs'. X = L rhs' \ rhs_eq_cls rhs' = rhs_eq_cls rhs - {X} \ distinct_rhs rhs'" + by (drule_tac ardenable_prop, auto) + then obtain rhs' where X_eq_rhs': "X = L rhs'" + and rhs'_eq_cls: "rhs_eq_cls rhs' = rhs_eq_cls rhs - {X}" + and rhs'_dist : "distinct_rhs rhs'" by blast + have "rhs_eq_cls rhs' \ {{[]}}" using rhs_eq_cls False rhs'_eq_cls rhs_not_empty + by blast + hence "rhs_eq_cls rhs' = {{[]}}" using X_not_empty X_eq_rhs' + by (auto simp:rhs_eq_cls_def) + hence "\ r. rhs' = {({[]}, r)}" using rhs'_dist + by (auto intro:rhs_aux simp:rhs_eq_cls_def) + then obtain r where "rhs' = {({[]}, r)}" .. + hence "?E r" using X_eq_rhs' by (auto simp add:lang_seq_def) + thus ?thesis by blast + qed +qed + +text {* definition of a regular language *} + +abbreviation + reg :: "string set => bool" +where + "reg L1 \ (\r::rexp. L r = L1)" + +theorem myhill_nerode: + assumes finite_CS: "finite (UNIV Quo Lang)" + shows "\ (reg::rexp). Lang = L reg" (is "\ r. ?P r") +proof - + have has_r_each: "\C\{X \ UNIV Quo Lang. \x\X. x \ Lang}. \(r::rexp). C = L r" using finite_CS + by (auto dest:every_eqcl_has_reg) + have "\ (rS::rexp set). finite rS \ + (\ C \ {X \ UNIV Quo Lang. \x\X. x \ Lang}. \ r \ rS. C = L r) \ + (\ r \ rS. \ C \ {X \ UNIV Quo Lang. \x\X. x \ Lang}. C = L r)" + (is "\ rS. ?Q rS") + proof- + have "\ C. C \ {X \ UNIV Quo Lang. \x\X. x \ Lang} \ C = L (SOME (ra::rexp). C = L ra)" + using has_r_each + apply (erule_tac x = C in ballE, erule_tac exE) + by (rule_tac a = r in someI2, simp+) + hence "?Q ((\ C. SOME r. C = L r) ` {X \ UNIV Quo Lang. \x\X. x \ Lang})" using has_r_each + using finite_CS by auto + thus ?thesis by blast + qed + then obtain rS where finite_rS : "finite rS" + and has_r_each': "\ C \ {X \ UNIV Quo Lang. \x\X. x \ Lang}. \ r \ (rS::rexp set). C = L r" + and has_cl_each: "\ r \ (rS::rexp set). \ C \ {X \ UNIV Quo Lang. \x\X. x \ Lang}. C = L r" by blast + have "?P (folds ALT NULL rS)" + proof + show "Lang \ L (folds ALT NULL rS)" using finite_rS lang_eqs_union_of_eqcls[of Lang] has_r_each' + apply (clarsimp simp:fold_alt_null_eqs) by blast + next + show "L (folds ALT NULL rS) \ Lang" using finite_rS lang_eqs_union_of_eqcls[of Lang] has_cl_each + by (clarsimp simp:fold_alt_null_eqs) + qed + thus ?thesis by blast +qed + + +text {* tests by Christian *} + +(* 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 + 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 + qed +qed + +lemma eq_class_imp_eq_str: + "\x\lang = \y\lang \ x \lang\ y" +by (auto simp:equiv_class_def equiv_str_def) + +lemma finite_tag_image: + "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 +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 + +(* list_diff:: list substract, once different return tailer *) +fun list_diff :: "'a list \ 'a list \ 'a list" (infix "-" 51) +where + "list_diff [] xs = []" | + "list_diff (x#xs) [] = x#xs" | + "list_diff (x#xs) (y#ys) = (if x = y then list_diff xs ys else (x#xs))" + +lemma [simp]: "(x @ y) - x = y" +apply (induct x) +by (case_tac y, simp+) + +lemma [simp]: "x - x = []" +by (induct x, auto) + +lemma [simp]: "x = xa @ y \ x - xa = y " +by (induct x, auto) + +lemma [simp]: "x - [] = x" +by (induct x, auto) + +lemma [simp]: "xa \ x \ (x @ y) - xa = (x - xa) @ y" +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 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 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 + +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) +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 +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) +qed + +(****************** the STAR case ************************) + +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) + +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) +qed + +lemma star_prop[rule_format]: "x \ lang\ \ \ y. y \ lang\ \ x @ y \ lang\" +by (erule Star.induct, auto) + +lemma star_prop2: "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 postfix_prop: "y >>= (x @ y) \ x = []" +by (erule postfixE, induct x arbitrary:y, auto) + +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 +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) +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 + +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) + +end