theory MyhillNerode imports "Main" "List_Prefix"begintext {* sequential composition of languages *}definition lang_seq :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ; _" [100,100] 100)where "L1 ; L2 = {s1 @ s2 | s1 s2. s1 \<in> L1 \<and> s2 \<in> L2}"lemma lang_seq_empty: shows "{[]} ; L = L" and "L ; {[]} = L"unfolding lang_seq_def by autolemma lang_seq_null: shows "{} ; L = {}" and "L ; {} = {}"unfolding lang_seq_def by autolemma lang_seq_append: assumes a: "s1 \<in> L1" and b: "s2 \<in> L2" shows "s1@s2 \<in> L1 ; L2"unfolding lang_seq_defusing a b by auto lemma lang_seq_union: shows "(L1 \<union> L2); L3 = (L1; L3) \<union> (L2; L3)" and "L1; (L2 \<union> L3) = (L1; L2) \<union> (L1; L3)"unfolding lang_seq_def by autolemma lang_seq_assoc: shows "(L1 ; L2) ; L3 = L1 ; (L2 ; L3)"unfolding lang_seq_defapply(auto)apply(metis)by (metis append_assoc)section {* Kleene star for languages defined as least fixed point *}inductive_set Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102) for L :: "string set"where start[intro]: "[] \<in> L\<star>"| step[intro]: "\<lbrakk>s1 \<in> L; s2 \<in> L\<star>\<rbrakk> \<Longrightarrow> s1@s2 \<in> L\<star>"lemma lang_star_empty: shows "{}\<star> = {[]}"by (auto elim: Star.cases)lemma lang_star_cases: shows "L\<star> = {[]} \<union> L ; L\<star>"proof { fix x have "x \<in> L\<star> \<Longrightarrow> x \<in> {[]} \<union> L ; L\<star>" unfolding lang_seq_def by (induct rule: Star.induct) (auto) } then show "L\<star> \<subseteq> {[]} \<union> L ; L\<star>" by autonext show "{[]} \<union> L ; L\<star> \<subseteq> L\<star>" unfolding lang_seq_def by autoqedlemma lang_star_cases': shows "L\<star> = {[]} \<union> L\<star> ; L"proof { fix x have "x \<in> L\<star> \<Longrightarrow> x \<in> {[]} \<union> L\<star> ; L" unfolding lang_seq_def apply (induct rule: Star.induct) apply simp apply simp apply (erule disjE) apply (auto)[1] apply (erule exE | erule conjE)+ apply (rule disjI2) apply (rule_tac x = "s1 @ s1a" in exI) by auto } then show "L\<star> \<subseteq> {[]} \<union> L\<star> ; L" by autonext show "{[]} \<union> L\<star> ; L \<subseteq> L\<star>" unfolding lang_seq_def apply auto apply (erule Star.induct) apply auto apply (drule step[of _ _ "[]"]) by (auto intro:start)qedlemma lang_star_simple: shows "L \<subseteq> L\<star>"by (subst lang_star_cases) (auto simp only: lang_seq_def)lemma lang_star_prop0_aux: "s2 \<in> L\<star> \<Longrightarrow> \<forall> s1. s1 \<in> L \<longrightarrow> (\<exists> s3 s4. s3 \<in> L\<star> \<and> s4 \<in> L \<and> s1 @ s2 = s3 @ s4)" apply (erule Star.induct)apply (clarify, rule_tac x = "[]" in exI, rule_tac x = s1 in exI, simp add:start)apply (clarify, drule_tac x = s1 in spec)apply (drule mp, simp, clarify)apply (rule_tac x = "s1a @ s3" in exI, rule_tac x = s4 in exI)by autolemma lang_star_prop0: "\<lbrakk>s1 \<in> L; s2 \<in> L\<star>\<rbrakk> \<Longrightarrow> \<exists> s3 s4. s3 \<in> L\<star> \<and> s4 \<in> L \<and> s1 @ s2 = s3 @ s4" by (auto dest:lang_star_prop0_aux)lemma lang_star_prop1: assumes asm: "L1; L2 \<subseteq> L2" shows "L1\<star>; L2 \<subseteq> L2"proof - { fix s1 s2 assume minor: "s2 \<in> L2" assume major: "s1 \<in> L1\<star>" then have "s1@s2 \<in> L2" proof(induct rule: Star.induct) case start show "[]@s2 \<in> L2" using minor by simp next case (step s1 s1') have "s1 \<in> L1" by fact moreover have "s1'@s2 \<in> L2" by fact ultimately have "s1@(s1'@s2) \<in> L1; L2" by (auto simp add: lang_seq_def) with asm have "s1@(s1'@s2) \<in> L2" by auto then show "(s1@s1')@s2 \<in> L2" by simp qed } then show "L1\<star>; L2 \<subseteq> L2" by (auto simp add: lang_seq_def)qedlemma lang_star_prop2_aux: "s2 \<in> L2\<star> \<Longrightarrow> \<forall> s1. s1 \<in> L1 \<and> L1 ; L2 \<subseteq> L1 \<longrightarrow> s1 @ s2 \<in> L1"apply (erule Star.induct, simp)apply (clarify, drule_tac x = "s1a @ s1" in spec)by (auto simp:lang_seq_def)lemma lang_star_prop2: "L1; L2 \<subseteq> L1 \<Longrightarrow> L1 ; L2\<star> \<subseteq> L1"by (auto dest!:lang_star_prop2_aux simp:lang_seq_def)lemma lang_star_seq_subseteq: shows "L ; L\<star> \<subseteq> L\<star>"using lang_star_cases by blastlemma lang_star_double: shows "L\<star>; L\<star> = L\<star>"proof show "L\<star> ; L\<star> \<subseteq> L\<star>" using lang_star_prop1 lang_star_seq_subseteq by blastnext have "L\<star> \<subseteq> L\<star> \<union> L\<star>; (L; L\<star>)" by auto also have "\<dots> = L\<star>;{[]} \<union> L\<star>; (L; L\<star>)" by (simp add: lang_seq_empty) also have "\<dots> = L\<star>; ({[]} \<union> L; L\<star>)" by (simp only: lang_seq_union) also have "\<dots> = L\<star>; L\<star>" using lang_star_cases by simp finally show "L\<star> \<subseteq> L\<star> ; L\<star>" by simpqedlemma lang_star_seq_subseteq': shows "L\<star>; L \<subseteq> L\<star>"proof - have "L \<subseteq> L\<star>" by (rule lang_star_simple) then have "L\<star>; L \<subseteq> L\<star>; L\<star>" by (auto simp add: lang_seq_def) then show "L\<star>; L \<subseteq> L\<star>" using lang_star_double by blastqedlemma shows "L\<star> \<subseteq> L\<star>\<star>"by (rule lang_star_simple)section {* regular expressions *}datatype rexp = NULL| EMPTY| CHAR char| SEQ rexp rexp| ALT rexp rexp| STAR rexpconsts L:: "'a \<Rightarrow> string set"overloading L_rexp \<equiv> "L:: rexp \<Rightarrow> string set"beginfun L_rexp :: "rexp \<Rightarrow> 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) \<union> (L_rexp r2)" | "L_rexp (STAR r) = (L_rexp r)\<star>"endtext{* ************ now is the codes writen by chunhan ************************************* *}section {* Arden's Lemma revised *}lemma arden_aux1: assumes a: "X \<subseteq> X ; A \<union> B" and b: "[] \<notin> A" shows "x \<in> X \<Longrightarrow> x \<in> B ; A\<star>"apply (induct x taking:length rule:measure_induct)apply (subgoal_tac "x \<in> X ; A \<union> B")deferusing aapply (auto)[1]apply simpapply (erule disjE)deferapply (auto simp add:lang_seq_def) [1]apply (subgoal_tac "\<exists> x1 x2. x = x1 @ x2 \<and> x1 \<in> X \<and> x2 \<in> A")deferapply (auto simp add:lang_seq_def) [1]apply (erule exE | erule conjE)+apply simpapply (drule_tac x = x1 in spec)apply (simp)using bapply -apply (auto)[1]apply (subgoal_tac "x1 @ x2 \<in> (B ; A\<star>) ; A")deferapply (auto simp add:lang_seq_def)[1]by (metis Un_absorb1 lang_seq_assoc lang_seq_union(2) lang_star_double lang_star_simple mem_def sup1CI)theorem ardens_revised: assumes nemp: "[] \<notin> A" shows "(X = X ; A \<union> B) \<longleftrightarrow> (X = B ; A\<star>)"apply(rule iffI)deferapply(simp)apply(subst lang_star_cases')apply(subst lang_seq_union)apply(simp add: lang_seq_empty)apply(simp add: lang_seq_assoc)apply(auto)[1]proof - assume "X = X ; A \<union> B" then have as1: "X ; A \<union> B \<subseteq> X" and as2: "X \<subseteq> X ; A \<union> B" by simp_all from as1 have a: "X ; A \<subseteq> X" and b: "B \<subseteq> X" by simp_all from b have "B; A\<star> \<subseteq> X ; A\<star>" by (auto simp add: lang_seq_def) moreover from a have "X ; A\<star> \<subseteq> X" by (rule lang_star_prop2) ultimately have f1: "B ; A\<star> \<subseteq> X" by simp from as2 nemp have f2: "X \<subseteq> B; A\<star>" using arden_aux1 by auto from f1 f2 show "X = B; A\<star>" by autoqedsection {* equiv class' definition *}definition equiv_str :: "string \<Rightarrow> string set \<Rightarrow> string \<Rightarrow> bool" ("_ \<equiv>_\<equiv> _" [100, 100, 100] 100)where "x \<equiv>Lang\<equiv> y \<longleftrightarrow> (\<forall>z. x @ z \<in> Lang \<longleftrightarrow> y @ z \<in> Lang)"definition equiv_class :: "string \<Rightarrow> (string set) \<Rightarrow> string set" ("\<lbrakk>_\<rbrakk>_" [100, 100] 100)where "\<lbrakk>x\<rbrakk>Lang \<equiv> {y. x \<equiv>Lang\<equiv> y}"text {* Chunhan modifies Q to Quo *}definition quot :: "string set \<Rightarrow> string set \<Rightarrow> (string set) set" ("_ Quo _" [100, 100] 100)where "L1 Quo L2 \<equiv> { \<lbrakk>x\<rbrakk>L2 | x. x \<in> L1}" lemma lang_eqs_union_of_eqcls: "Lang = \<Union> {X. X \<in> (UNIV Quo Lang) \<and> (\<forall> x \<in> X. x \<in> Lang)}"proof show "Lang \<subseteq> \<Union>{X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}" proof fix x assume "x \<in> Lang" thus "x \<in> \<Union>{X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}" proof (simp add:quot_def) assume "(1)": "x \<in> Lang" show "\<exists>xa. (\<exists>x. xa = \<lbrakk>x\<rbrakk>Lang) \<and> (\<forall>x\<in>xa. x \<in> Lang) \<and> x \<in> xa" (is "\<exists>xa.?P xa") proof - have "?P (\<lbrakk>x\<rbrakk>Lang)" using "(1)" by (auto simp:equiv_class_def equiv_str_def dest: spec[where x = "[]"]) thus ?thesis by blast qed qed qed next show "\<Union>{X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang} \<subseteq> Lang" by autoqedlemma empty_notin_CS: "{} \<notin> UNIV Quo Lang"apply (clarsimp simp:quot_def equiv_class_def)by (rule_tac x = x in exI, auto simp:equiv_str_def)lemma no_two_cls_inters: "\<lbrakk>X \<in> UNIV Quo Lang; Y \<in> UNIV Quo Lang; X \<noteq> Y\<rbrakk> \<Longrightarrow> X \<inter> Y = {}"by (auto simp:quot_def equiv_class_def equiv_str_def)text {* equiv_class transition *}definition CT :: "string set \<Rightarrow> char \<Rightarrow> string set \<Rightarrow> bool" ("_-_\<rightarrow>_" [99,99]99)where "X-c\<rightarrow>Y \<equiv> ((X;{[c]}) \<subseteq> Y)"types t_equa_rhs = "(string set \<times> rexp) set"types t_equa = "(string set \<times> t_equa_rhs)"types t_equas = "t_equa set"text {* "empty_rhs" generates "\<lambda>" for init-state, just like "\<lambda>" for final states in Brzozowski method. But if the init-state is "{[]}" ("\<lambda>" itself) then empty set is returned, see definition of "equation_rhs" *} definition empty_rhs :: "string set \<Rightarrow> t_equa_rhs"where "empty_rhs X \<equiv> if ([] \<in> X) then {({[]}, EMPTY)} else {}"definition folds :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b"where "folds f z S \<equiv> SOME x. fold_graph f z S x"definition equation_rhs :: "(string set) set \<Rightarrow> string set \<Rightarrow> t_equa_rhs"where "equation_rhs CS X \<equiv> if (X = {[]}) then {({[]}, EMPTY)} else {(S, folds ALT NULL {CHAR c| c. S-c\<rightarrow>X} )| S. S \<in> CS} \<union> empty_rhs X"definition equations :: "(string set) set \<Rightarrow> t_equas"where "equations CS \<equiv> {(X, equation_rhs CS X) | X. X \<in> CS}"overloading L_rhs \<equiv> "L:: t_equa_rhs \<Rightarrow> string set"beginfun L_rhs:: "t_equa_rhs \<Rightarrow> string set"where "L_rhs rhs = {x. \<exists> X r. (X, r) \<in> rhs \<and> x \<in> X;(L r)}"enddefinition distinct_rhs :: "t_equa_rhs \<Rightarrow> bool"where "distinct_rhs rhs \<equiv> \<forall> X reg\<^isub>1 reg\<^isub>2. (X, reg\<^isub>1) \<in> rhs \<and> (X, reg\<^isub>2) \<in> rhs \<longrightarrow> reg\<^isub>1 = reg\<^isub>2"definition distinct_equas_rhs :: "t_equas \<Rightarrow> bool"where "distinct_equas_rhs equas \<equiv> \<forall> X rhs. (X, rhs) \<in> equas \<longrightarrow> distinct_rhs rhs"definition distinct_equas :: "t_equas \<Rightarrow> bool"where "distinct_equas equas \<equiv> \<forall> X rhs rhs'. (X, rhs) \<in> equas \<and> (X, rhs') \<in> equas \<longrightarrow> rhs = rhs'"definition seq_rhs_r :: "t_equa_rhs \<Rightarrow> rexp \<Rightarrow> t_equa_rhs"where "seq_rhs_r rhs r \<equiv> (\<lambda>(X, reg). (X, SEQ reg r)) ` rhs"definition del_x_paired :: "('a \<times> 'b) set \<Rightarrow> 'a \<Rightarrow> ('a \<times> 'b) set"where "del_x_paired S x \<equiv> S - {X. X \<in> S \<and> fst X = x}"definition arden_variate :: "string set \<Rightarrow> rexp \<Rightarrow> t_equa_rhs \<Rightarrow> t_equa_rhs"where "arden_variate X r rhs \<equiv> seq_rhs_r (del_x_paired rhs X) (STAR r)"definition no_EMPTY_rhs :: "t_equa_rhs \<Rightarrow> bool"where "no_EMPTY_rhs rhs \<equiv> \<forall> X r. (X, r) \<in> rhs \<and> X \<noteq> {[]} \<longrightarrow> [] \<notin> L r"definition no_EMPTY_equas :: "t_equas \<Rightarrow> bool"where "no_EMPTY_equas equas \<equiv> \<forall> X rhs. (X, rhs) \<in> equas \<longrightarrow> no_EMPTY_rhs rhs"lemma fold_alt_null_eqs: "finite rS \<Longrightarrow> x \<in> L (folds ALT NULL rS) = (\<exists> r \<in> rS. x \<in> L r)"apply (simp add:folds_def)apply (rule someI2_ex)apply (erule finite_imp_fold_graph)apply (erule fold_graph.induct)by auto (*??? how do this be in Isar ?? *)lemma seq_rhs_r_prop1: "L (seq_rhs_r rhs r) = (L rhs);(L r)"apply (auto simp:seq_rhs_r_def image_def lang_seq_def)apply (rule_tac x = "s1 @ s1a" in exI, rule_tac x = "s2a" in exI, simp)apply (rule_tac x = a in exI, rule_tac x = b in exI, simp)apply (rule_tac x = s1 in exI, rule_tac x = s1a in exI, simp)apply (rule_tac x = X in exI, rule_tac x = "SEQ ra r" in exI, simp)apply (rule conjI)apply (rule_tac x = "(X, ra)" in bexI, simp+) apply (rule_tac x = s1a in exI, rule_tac x = "s2a @ s2" in exI, simp)apply (simp add:lang_seq_def)by (rule_tac x = s2a in exI, rule_tac x = s2 in exI, simp)lemma del_x_paired_prop1: "\<lbrakk>distinct_rhs rhs; (X, r) \<in> rhs\<rbrakk> \<Longrightarrow> X ; L r \<union> L (del_x_paired rhs X) = L rhs" apply (simp add:del_x_paired_def) apply (simp add: distinct_rhs_def) apply(auto simp add: lang_seq_def) apply(metis) donelemma arden_variate_prop: assumes "(X, rx) \<in> rhs" shows "(\<forall> Y. Y \<noteq> X \<longrightarrow> (\<exists> r. (Y, r) \<in> rhs) = (\<exists> r. (Y, r) \<in> (arden_variate X rx rhs)))"proof (rule allI, rule impI) fix Y assume "(1)": "Y \<noteq> X" show "(\<exists>r. (Y, r) \<in> rhs) = (\<exists>r. (Y, r) \<in> arden_variate X rx rhs)" proof assume "(1_1)": "\<exists>r. (Y, r) \<in> rhs" show "\<exists>r. (Y, r) \<in> arden_variate X rx rhs" (is "\<exists>r. ?P r") proof - from "(1_1)" obtain r where "(Y, r) \<in> rhs" .. hence "?P (SEQ r (STAR rx))" proof (simp add:arden_variate_def image_def) have "(Y, r) \<in> rhs \<Longrightarrow> (Y, r) \<in> del_x_paired rhs X" by (auto simp:del_x_paired_def "(1)") thus "(Y, r) \<in> rhs \<Longrightarrow> (Y, SEQ r (STAR rx)) \<in> seq_rhs_r (del_x_paired rhs X) (STAR rx)" by (auto simp:seq_rhs_r_def) qed thus ?thesis by blast qed next assume "(2_1)": "\<exists>r. (Y, r) \<in> arden_variate X rx rhs" thus "\<exists>r. (Y, r) \<in> rhs" (is "\<exists> r. ?P r") by (auto simp:arden_variate_def del_x_paired_def seq_rhs_r_def image_def) qedqedtext {* arden_variate_valid: proves variation from "X = X;r + Y;ry + \<dots>" to "X = Y;(SEQ ry (STAR r)) + \<dots>" holds the law of "language of left equiv language of right" *}lemma arden_variate_valid: assumes X_not_empty: "X \<noteq> {[]}" and l_eq_r: "X = L rhs" and dist: "distinct_rhs rhs" and no_empty: "no_EMPTY_rhs rhs" and self_contained: "(X, r) \<in> rhs" shows "X = L (arden_variate X r rhs)" proof - have "[] \<notin> L r" using no_empty X_not_empty self_contained by (auto simp:no_EMPTY_rhs_def) hence ardens: "X = X;(L r) \<union> (L (del_x_paired rhs X)) \<longleftrightarrow> X = (L (del_x_paired rhs X)) ; (L r)\<star>" by (rule ardens_revised) have del_x: "X = X ; L r \<union> L (del_x_paired rhs X) \<longleftrightarrow> X = L rhs" using dist l_eq_r self_contained by (auto dest!:del_x_paired_prop1) show ?thesis proof show "X \<subseteq> L (arden_variate X r rhs)" proof fix x assume "(1_1)": "x \<in> X" with l_eq_r ardens del_x show "x \<in> L (arden_variate X r rhs)" by (simp add:arden_variate_def seq_rhs_r_prop1 del:L_rhs.simps) qed next show "L (arden_variate X r rhs) \<subseteq> X" proof fix x assume "(2_1)": "x \<in> L (arden_variate X r rhs)" with ardens del_x l_eq_r show "x \<in> X" by (simp add:arden_variate_def seq_rhs_r_prop1 del:L_rhs.simps) qed qedqedtext {* merge_rhs {(x1, r1), (x2, r2}, (x4, r4), \<dots>} {(x1, r1'), (x3, r3'), \<dots>} = {(x1, ALT r1 r1'}, (x2, r2), (x3, r3'), (x4, r4), \<dots>} *} definition merge_rhs :: "t_equa_rhs \<Rightarrow> t_equa_rhs \<Rightarrow> t_equa_rhs"where "merge_rhs rhs rhs' \<equiv> {(X, r). (\<exists> r1 r2. ((X,r1) \<in> rhs \<and> (X, r2) \<in> rhs') \<and> r = ALT r1 r2) \<or> (\<exists> r1. (X, r1) \<in> rhs \<and> (\<not> (\<exists> r2. (X, r2) \<in> rhs')) \<and> r = r1) \<or> (\<exists> r2. (X, r2) \<in> rhs' \<and> (\<not> (\<exists> r1. (X, r1) \<in> rhs)) \<and> r = r2) }" text {* rhs_subst rhs X=xrhs r: substitude all occurence of X in rhs((X,r) \<in> rhs) with xrhs *}definition rhs_subst :: "t_equa_rhs \<Rightarrow> string set \<Rightarrow> t_equa_rhs \<Rightarrow> rexp \<Rightarrow> t_equa_rhs"where "rhs_subst rhs X xrhs r \<equiv> merge_rhs (del_x_paired rhs X) (seq_rhs_r xrhs r)"definition equas_subst_f :: "string set \<Rightarrow> t_equa_rhs \<Rightarrow> t_equa \<Rightarrow> t_equa"where "equas_subst_f X xrhs equa \<equiv> let (Y, rhs) = equa in if (\<exists> r. (X, r) \<in> rhs) then (Y, rhs_subst rhs X xrhs (SOME r. (X, r) \<in> rhs)) else equa"definition equas_subst :: "t_equas \<Rightarrow> string set \<Rightarrow> t_equa_rhs \<Rightarrow> t_equas"where "equas_subst ES X xrhs \<equiv> del_x_paired (equas_subst_f X xrhs ` ES) X"lemma lang_seq_prop1: "x \<in> X ; L r \<Longrightarrow> x \<in> X ; (L r \<union> L r')"by (auto simp:lang_seq_def)lemma lang_seq_prop1': "x \<in> X; L r \<Longrightarrow> x \<in> X ; (L r' \<union> L r)"by (auto simp:lang_seq_def)lemma lang_seq_prop2: "x \<in> X; (L r \<union> L r') \<Longrightarrow> x \<in> X;L r \<or> x \<in> X;L r'"by (auto simp:lang_seq_def)lemma merge_rhs_prop1: shows "L (merge_rhs rhs rhs') = L rhs \<union> L rhs' "apply (auto simp add:merge_rhs_def dest!:lang_seq_prop2 intro:lang_seq_prop1)apply (rule_tac x = X in exI, rule_tac x = r1 in exI, simp)apply (case_tac "\<exists> r2. (X, r2) \<in> rhs'")apply (clarify, rule_tac x = X in exI, rule_tac x = "ALT r r2" in exI, simp add:lang_seq_prop1)apply (rule_tac x = X in exI, rule_tac x = r in exI, simp)apply (case_tac "\<exists> r1. (X, r1) \<in> rhs")apply (clarify, rule_tac x = X in exI, rule_tac x = "ALT r1 r" in exI, simp add:lang_seq_prop1')apply (rule_tac x = X in exI, rule_tac x = r in exI, simp)donelemma no_EMPTY_rhss_imp_merge_no_EMPTY: "\<lbrakk>no_EMPTY_rhs rhs; no_EMPTY_rhs rhs'\<rbrakk> \<Longrightarrow> no_EMPTY_rhs (merge_rhs rhs rhs')"apply (simp add:no_EMPTY_rhs_def merge_rhs_def)apply (clarify, (erule conjE | erule exE | erule disjE)+)by autolemma distinct_rhs_prop: "\<lbrakk>distinct_rhs rhs; (X, r1) \<in> rhs; (X, r2) \<in> rhs\<rbrakk> \<Longrightarrow> r1 = r2"by (auto simp:distinct_rhs_def)lemma merge_rhs_prop2: assumes dist_rhs: "distinct_rhs rhs" and dist_rhs':"distinct_rhs rhs'" shows "distinct_rhs (merge_rhs rhs rhs')"apply (auto simp:merge_rhs_def distinct_rhs_def)using dist_rhsapply (drule distinct_rhs_prop, simp+)using dist_rhs'apply (drule distinct_rhs_prop, simp+)using dist_rhsapply (drule distinct_rhs_prop, simp+)using dist_rhs'apply (drule distinct_rhs_prop, simp+)donelemma seq_rhs_r_holds_distinct: "distinct_rhs rhs \<Longrightarrow> distinct_rhs (seq_rhs_r rhs r)"by (auto simp:distinct_rhs_def seq_rhs_r_def)lemma seq_rhs_r_prop0: assumes l_eq_r: "X = L xrhs" shows "L (seq_rhs_r xrhs r) = X ; L r "using l_eq_rby (auto simp only:seq_rhs_r_prop1)lemma rhs_subst_prop1: assumes l_eq_r: "X = L xrhs" and dist: "distinct_rhs rhs" shows "(X, r) \<in> rhs \<Longrightarrow> L rhs = L (rhs_subst rhs X xrhs r)"apply (simp add:rhs_subst_def merge_rhs_prop1 del:L_rhs.simps)using l_eq_r apply (drule_tac r = r in seq_rhs_r_prop0, simp del:L_rhs.simps)using distby (auto dest!:del_x_paired_prop1 simp del:L_rhs.simps)lemma del_x_paired_holds_distinct_rhs: "distinct_rhs rhs \<Longrightarrow> distinct_rhs (del_x_paired rhs x)"by (auto simp:distinct_rhs_def del_x_paired_def)lemma rhs_subst_holds_distinct_rhs: "\<lbrakk>distinct_rhs rhs; distinct_rhs xrhs\<rbrakk> \<Longrightarrow> distinct_rhs (rhs_subst rhs X xrhs r)"apply (drule_tac r = r and rhs = xrhs in seq_rhs_r_holds_distinct)apply (drule_tac x = X in del_x_paired_holds_distinct_rhs)by (auto dest:merge_rhs_prop2[where rhs = "del_x_paired rhs X"] simp:rhs_subst_def)section {* myhill-nerode theorem *}definition left_eq_cls :: "t_equas \<Rightarrow> (string set) set"where "left_eq_cls ES \<equiv> {X. \<exists> rhs. (X, rhs) \<in> ES} "definition right_eq_cls :: "t_equas \<Rightarrow> (string set) set"where "right_eq_cls ES \<equiv> {Y. \<exists> X rhs r. (X, rhs) \<in> ES \<and> (Y, r) \<in> rhs }"definition rhs_eq_cls :: "t_equa_rhs \<Rightarrow> (string set) set"where "rhs_eq_cls rhs \<equiv> {Y. \<exists> r. (Y, r) \<in> rhs}"definition ardenable :: "t_equa \<Rightarrow> bool"where "ardenable equa \<equiv> let (X, rhs) = equa in distinct_rhs rhs \<and> no_EMPTY_rhs rhs \<and> X = L rhs"text {* Inv: Invairance of the equation-system, during the decrease of the equation-system, Inv holds.*}definition Inv :: "string set \<Rightarrow> t_equas \<Rightarrow> bool"where "Inv X ES \<equiv> finite ES \<and> (\<exists> rhs. (X, rhs) \<in> ES) \<and> distinct_equas ES \<and> (\<forall> X xrhs. (X, xrhs) \<in> ES \<longrightarrow> ardenable (X, xrhs) \<and> X \<noteq> {} \<and> rhs_eq_cls xrhs \<subseteq> insert {[]} (left_eq_cls ES))"text {* TCon: Termination Condition of the equation-system decreasion.*}definition TCon:: "'a set \<Rightarrow> bool"where "TCon ES \<equiv> card ES = 1"text {* The following is a iteration principle, and is the main framework for the proof: 1: We can form the initial equation-system using "equations" defined above, and prove it has invariance Inv by lemma "init_ES_satisfy_Inv"; 2: We can decrease the number of the equation-system using ardens_lemma_revised and substitution ("equas_subst", defined above), and prove it holds the property "step" in "wf_iter" by lemma "iteration_step" and finally using property Inv and TCon to prove the myhill-nerode theorem*}lemma wf_iter [rule_format]: fixes f assumes step: "\<And> e. \<lbrakk>P e; \<not> Q e\<rbrakk> \<Longrightarrow> (\<exists> e'. P e' \<and> (f(e'), f(e)) \<in> less_than)" shows pe: "P e \<longrightarrow> (\<exists> e'. P e' \<and> 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]: "\<forall>y. (y, x) \<in> inv_image less_than f \<longrightarrow> P y \<longrightarrow> (\<exists>e'. P e' \<and> Q e')" and px: "P x" show "\<exists>e'. P e' \<and> Q e'" proof(cases "Q x") assume "Q x" with px show ?thesis by blast next assume nq: "\<not> Q x" from step [OF px nq] obtain e' where pe': "P e'" and ltf: "(f e', f x) \<in> less_than" by auto show ?thesis proof(rule h) from ltf show "(e', x) \<in> inv_image less_than f" by (simp add:inv_image_def) next from pe' show "P e'" . qed qedqedtext {* ******BEGIN: proving the initial equation-system satisfies Inv ****** *}lemma distinct_rhs_equations: "(X, xrhs) \<in> equations (UNIV Quo Lang) \<Longrightarrow> 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: "\<lbrakk>X \<in> (UNIV Quo Lang); X \<noteq> {[]}\<rbrakk> \<Longrightarrow> \<exists> clist. clist \<in> X \<and> clist \<noteq> []"by (auto simp:quot_def equiv_class_def equiv_str_def)lemma every_eqclass_is_derived_from_empty: assumes not_empty: "X \<noteq> {[]}" shows "X \<in> (UNIV Quo Lang) \<Longrightarrow> \<exists> clist. {[]};{clist} \<subseteq> X \<and> clist \<noteq> []"using not_emptyapply (drule_tac every_nonempty_eqclass_has_strings, simp)by (auto intro:exI[where x = clist] simp:lang_seq_def)lemma equiv_str_in_CS: "\<lbrakk>clist\<rbrakk>Lang \<in> (UNIV Quo Lang)"by (auto simp:quot_def)lemma has_str_imp_defined_by_str: "\<lbrakk>str \<in> X; X \<in> UNIV Quo Lang\<rbrakk> \<Longrightarrow> X = \<lbrakk>str\<rbrakk>Lang"by (auto simp:quot_def equiv_class_def equiv_str_def)lemma every_eqclass_has_ascendent: assumes has_str: "clist @ [c] \<in> X" and in_CS: "X \<in> UNIV Quo Lang" shows "\<exists> Y. Y \<in> UNIV Quo Lang \<and> Y-c\<rightarrow>X \<and> clist \<in> Y" (is "\<exists> Y. ?P Y")proof - have "?P (\<lbrakk>clist\<rbrakk>Lang)" proof - have "\<lbrakk>clist\<rbrakk>Lang \<in> UNIV Quo Lang" by (simp add:quot_def, rule_tac x = clist in exI, simp) moreover have "\<lbrakk>clist\<rbrakk>Lang-c\<rightarrow>X" proof - have "X = \<lbrakk>(clist @ [c])\<rbrakk>Lang" using has_str in_CS by (auto intro!:has_str_imp_defined_by_str) moreover have "\<forall> sl. sl \<in> \<lbrakk>clist\<rbrakk>Lang \<longrightarrow> sl @ [c] \<in> \<lbrakk>(clist @ [c])\<rbrakk>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 \<in> \<lbrakk>clist\<rbrakk>Lang" by (auto simp:equiv_str_def equiv_class_def) ultimately show "?P (\<lbrakk>clist\<rbrakk>Lang)" by simp qed thus ?thesis by blastqedlemma finite_charset_rS: "finite {CHAR c |c. Y-c\<rightarrow>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) \<in> 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 \<subseteq> L xrhs" proof fix x assume "(1)": "x \<in> X" show "x \<in> L xrhs" proof (cases "x = []") assume empty: "x = []" hence "x \<in> 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 \<noteq> []" hence "\<exists> 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 "\<And> Y. \<lbrakk>Y \<in> UNIV Quo Lang; Y-c\<rightarrow>X; clist \<in> Y\<rbrakk> \<Longrightarrow> [c] \<in> L (folds ALT NULL {CHAR c |c. Y-c\<rightarrow>X})" proof - fix Y assume Y_is_eq_cl: "Y \<in> UNIV Quo Lang" and Y_CT_X: "Y-c\<rightarrow>X" and clist_in_Y: "clist \<in> Y" with finite_charset_rS show "[c] \<in> L (folds ALT NULL {CHAR c |c. Y-c\<rightarrow>X})" by (auto simp :fold_alt_null_eqs) qed hence "\<exists>Xa. Xa \<in> UNIV Quo Lang \<and> clist @ [c] \<in> Xa ; L (folds ALT NULL {CHAR c |c. Xa-c\<rightarrow>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 \<in> UNIV Quo Lang \<and> clist @ [c] \<in> Xa ; L (folds ALT NULL {CHAR c |c. Xa-c\<rightarrow>X})" by blast hence "x \<in> L {(S, folds ALT NULL {CHAR c |c. S-c\<rightarrow>X}) |S. S \<in> 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 \<in> L xrhs" using X_in_equas False not_empty unfolding equations_def equation_rhs_def by auto qed qed next show "L xrhs \<subseteq> X" proof fix x assume "(2)": "x \<in> L xrhs" have "(2_1)": "\<And> s1 s2 r Xa. \<lbrakk>s1 \<in> Xa; s2 \<in> L (folds ALT NULL {CHAR c |c. Xa-c\<rightarrow>X})\<rbrakk> \<Longrightarrow> s1 @ s2 \<in> X" using finite_charset_rS by (auto simp:CT_def lang_seq_def fold_alt_null_eqs) have "(2_2)": "\<And> s1 s2 Xa r.\<lbrakk>s1 \<in> Xa; s2 \<in> L r; (Xa, r) \<in> empty_rhs X\<rbrakk> \<Longrightarrow> s1 @ s2 \<in> X" by (simp add:empty_rhs_def split:if_splits) show "x \<in> X" using X_in_equas False "(2)" by (auto intro:"(2_1)" "(2_2)" simp:equations_def equation_rhs_def lang_seq_def) qed qedqedlemma no_EMPTY_equations: "(X, xrhs) \<in> equations CS \<Longrightarrow> 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\<rightarrow>X}", drule_tac x = "[]" in fold_alt_null_eqs, clarsimp, rule finite_charset_rS)+donelemma init_ES_satisfy_ardenable: "(X, xrhs) \<in> equations (UNIV Quo Lang) \<Longrightarrow> 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 \<in> 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 "\<exists>rhs. (X, rhs) \<in> 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 "\<forall>X xrhs. (X, xrhs) \<in> equations (UNIV Quo Lang) \<longrightarrow> rhs_eq_cls xrhs \<subseteq> 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 "\<forall>X xrhs. (X, xrhs) \<in> equations (UNIV Quo Lang) \<longrightarrow> X \<noteq> {}" by (clarsimp simp:equations_def empty_notin_CS intro:classical) moreover have "\<forall>X xrhs. (X, xrhs) \<in> equations (UNIV Quo Lang) \<longrightarrow> ardenable (X, xrhs)" by (auto intro!:init_ES_satisfy_ardenable) ultimately show ?thesis by (simp add:Inv_def)qedtext {* *********** END: proving the initial equation-system satisfies Inv ******* *}text {* ****** BEGIN: proving every equation-system's iteration step satisfies Inv ***** *}lemma not_T_aux: "\<lbrakk>\<not> TCon (insert a A); x = a\<rbrakk> \<Longrightarrow> \<exists>y. x \<noteq> y \<and> y \<in> insert a A "apply (case_tac "insert a A = {a}")by (auto simp:TCon_def)lemma not_T_atleast_2[rule_format]: "finite S \<Longrightarrow> \<forall> x. x \<in> S \<and> (\<not> TCon S)\<longrightarrow> (\<exists> y. x \<noteq> y \<and> y \<in> S)"apply (erule finite.induct, simp)apply (clarify, case_tac "x = a")by (erule not_T_aux, auto)lemma exist_another_equa: "\<lbrakk>\<not> TCon ES; finite ES; distinct_equas ES; (X, rhl) \<in> ES\<rbrakk> \<Longrightarrow> \<exists> Y yrhl. (Y, yrhl) \<in> ES \<and> X \<noteq> 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 autolemma Inv_mono_with_lambda: assumes Inv_ES: "Inv X ES" and X_noteq_Y: "X \<noteq> {[]}" shows "Inv X (ES - {({[]}, yrhs)})"proof - have "finite (ES - {({[]}, yrhs)})" using Inv_ES by (simp add:Inv_def) moreover have "\<exists>rhs. (X, rhs) \<in> 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 "\<forall>X xrhs.(X, xrhs) \<in> ES - {({[]}, yrhs)} \<longrightarrow> ardenable (X, xrhs) \<and> X \<noteq> {}" 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 "\<forall>X xrhs.(X, xrhs) \<in> ES - {({[]}, yrhs)} \<longrightarrow> rhs_eq_cls xrhs \<subseteq> insert {[]} (left_eq_cls (ES - {({[]}, yrhs)}))" using Inv_ES by (auto simp:Inv_def) ultimately show ?thesis by (simp add:Inv_def)qedlemma non_empty_card_prop: "finite ES \<Longrightarrow> \<forall>e. e \<in> ES \<longrightarrow> card ES - Suc 0 < card ES"apply (erule finite.induct, simp)apply (case_tac[!] "a \<in> A")by (auto simp:insert_absorb)lemma ardenable_prop: assumes not_lambda: "Y \<noteq> {[]}" and ardable: "ardenable (Y, yrhs)" shows "\<exists> yrhs'. Y = L yrhs' \<and> distinct_rhs yrhs' \<and> rhs_eq_cls yrhs' = rhs_eq_cls yrhs - {Y}" (is "\<exists> yrhs'. ?P yrhs'")proof (cases "(\<exists> reg. (Y, reg) \<in> yrhs)") case True thus ?thesis proof fix reg assume self_contained: "(Y, reg) \<in> 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 "\<And> 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 "\<And> 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 qednext 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 qedqedlemma equas_subst_f_del_no_other: assumes self_contained: "(Y, rhs) \<in> ES" shows "\<exists> rhs'. (Y, rhs') \<in> (equas_subst_f X xrhs ` ES)" (is "\<exists> rhs'. ?P rhs'")proof - have "\<exists> 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 blastqedlemma del_x_paired_del_only_x: "\<lbrakk>X \<noteq> Y; (X, rhs) \<in> ES\<rbrakk> \<Longrightarrow> (X, rhs) \<in> del_x_paired ES Y"by (auto simp:del_x_paired_def)lemma equas_subst_del_no_other: "\<lbrakk>(X, rhs) \<in> ES; X \<noteq> Y\<rbrakk> \<Longrightarrow> (\<exists>rhs. (X, rhs) \<in> equas_subst ES Y rhs')"unfolding equas_subst_defapply (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 \<Longrightarrow> 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) \<in> ES \<Longrightarrow> {Y. Y \<in> ES \<and> fst Y = X} \<inter> ES \<noteq> {}"by (auto)lemma del_x_paired_subset: "(X, rhs) \<in> ES \<Longrightarrow> ES - {Y. Y \<in> ES \<and> fst Y = X} \<subset> ES"apply (drule del_x_paired_dels)by autolemma del_x_paired_card_less: "\<lbrakk>finite ES; (X, rhs) \<in> ES\<rbrakk> \<Longrightarrow> 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: "\<lbrakk>finite ES; (Y, yrhs) \<in> ES\<rbrakk> \<Longrightarrow> 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: "\<forall>X xrhs. (X, xrhs) \<in> ES \<longrightarrow> ardenable (X, xrhs)" and X_in : "(X, xrhs) \<in> equas_subst ES Y yrhs'" shows "distinct_rhs xrhs"using X_in historyapply (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: "[] \<notin> L r \<Longrightarrow> 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 \<Longrightarrow> 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: "\<lbrakk>no_EMPTY_rhs yrhs; (Y, r) \<in> yrhs; Y \<noteq> {[]}\<rbrakk> \<Longrightarrow> 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 \<noteq> {[]}" and history: "\<forall>X xrhs. (X, xrhs) \<in> ES \<longrightarrow> ardenable (X, xrhs)" and X_in:"(X, xrhs) \<in> equas_subst ES Y yrhs'" shows "no_EMPTY_rhs xrhs"proof- from X_in have "\<exists> (Z, zrhs) \<in> 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) \<in> 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 "\<exists> r. (Y, r) \<in> zrhs") case True then obtain r where Y_in_zrhs: "(Y, r) \<in> zrhs" .. hence some: "(SOME r. (Y, r) \<in> 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) qedqedlemma equas_subst_f_holds_left_eq_right: assumes substor: "Y = L rhs'" and history: "\<forall>X xrhs. (X, xrhs) \<in> ES \<longrightarrow> distinct_rhs xrhs \<and> X = L xrhs" and subst: "(X, xrhs) = equas_subst_f Y rhs' (Z, zrhs)" and self_contained: "(Z, zrhs) \<in> ES" shows "X = L xrhs"proof (cases "\<exists> r. (Y, r) \<in> zrhs") case True from True obtain r where "(1)":"(Y, r) \<in> zrhs" .. show ?thesis proof - from history self_contained have dist: "distinct_rhs zrhs" by auto hence "(SOME r. (Y, r) \<in> 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) qednext case False thus ?thesis using history subst self_contained by (auto simp:equas_subst_f_def)qedlemma equas_subst_holds_left_eq_right: assumes history: "\<forall>X xrhs. (X, xrhs) \<in> ES \<longrightarrow> ardenable (X, xrhs)" and substor: "Y = L rhs'" and X_in : "(X, xrhs) \<in> equas_subst ES Y yrhs'" shows "\<forall>X xrhs. (X, xrhs) \<in> equas_subst ES Y rhs' \<longrightarrow> X = L xrhs"apply (clarsimp simp add:equas_subst_def del_x_paired_def)using substorapply (drule_tac equas_subst_f_holds_left_eq_right)using historyby (auto simp:ardenable_def)lemma equas_subst_holds_ardenable: assumes substor: "Y = L yrhs'" and history: "\<forall>X xrhs. (X, xrhs) \<in> ES \<longrightarrow> ardenable (X, xrhs)" and X_in:"(X, xrhs) \<in> equas_subst ES Y yrhs'" and dist': "distinct_rhs yrhs'" and not_lambda: "Y \<noteq> {[]}" 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 simpqedlemma equas_subst_holds_cls_defined: assumes X_in: "(X, xrhs) \<in> equas_subst ES Y yrhs'" and Inv_ES: "Inv X' ES" and subst: "(Y, yrhs) \<in> ES" and cls_holds_but_Y: "rhs_eq_cls yrhs' = rhs_eq_cls yrhs - {Y}" shows "rhs_eq_cls xrhs \<subseteq> insert {[]} (left_eq_cls (equas_subst ES Y yrhs'))"proof- have tac: "\<lbrakk> A \<subseteq> B; C \<subseteq> D; E \<subseteq> A \<union> B\<rbrakk> \<Longrightarrow> E \<subseteq> B \<union> D" by auto from X_in have "\<exists> (Z, zrhs) \<in> 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) \<in> ES" and X_Z: "(X, xrhs) = equas_subst_f Y yrhs' (Z, zrhs)" by blast hence "rhs_eq_cls zrhs \<subseteq> insert {[]} (left_eq_cls ES)" using Inv_ES by (auto simp:Inv_def) moreover have "rhs_eq_cls yrhs' \<subseteq> 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 \<subseteq> rhs_eq_cls zrhs \<union> 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 blastqedlemma iteration_step: assumes Inv_ES: "Inv X ES" and not_T: "\<not> TCon ES" shows "(\<exists> ES'. Inv X ES' \<and> (card ES', card ES) \<in> less_than)" proof - from Inv_ES not_T have another: "\<exists>Y yrhs. (Y, yrhs) \<in> ES \<and> X \<noteq> Y" unfolding Inv_def by (clarify, rule_tac exist_another_equa[where X = X], auto) then obtain Y yrhs where subst: "(Y, yrhs) \<in> ES" and not_X: " X \<noteq> Y" by blast show ?thesis (is "\<exists> ES'. ?P ES'") proof (cases "Y = {[]}") case True --"in this situation, we pick a \"\<lambda>\" 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) \<in> 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 "\<exists> yrhs'. Y = L yrhs' \<and> distinct_rhs yrhs' \<and> 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: "\<And> S x. finite S \<Longrightarrow> 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 "\<exists>rhs. (X, rhs) \<in> 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 "\<forall>X xrhs. (X, xrhs) \<in> equas_subst ES Y yrhs' \<longrightarrow> 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 "\<forall>X xrhs. (X, xrhs) \<in> equas_subst ES Y yrhs' \<longrightarrow> X \<noteq> {}" 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 "\<forall>X xrhs. (X, xrhs) \<in> equas_subst ES Y yrhs' \<longrightarrow> rhs_eq_cls xrhs \<subseteq> 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) \<in> 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 qedqedtext {* ***** END: proving every equation-system's iteration step satisfies Inv ************** *}lemma iteration_conc: assumes history: "Inv X ES" shows "\<exists> ES'. Inv X ES' \<and> TCon ES'" (is "\<exists> ES'. ?P ES'")proof (cases "TCon ES") case True hence "?P ES" using history by simp thus ?thesis by blastnext case False thus ?thesis using history iteration_step by (rule_tac f = card in wf_iter, simp_all)qedlemma eqset_imp_iff': "A = B \<Longrightarrow> \<forall> x. x \<in> A \<longleftrightarrow> x \<in> B"apply (auto simp:mem_def)donelemma set_cases2: "\<lbrakk>(A = {} \<Longrightarrow> R A); \<And> x. (A = {x}) \<Longrightarrow> R A; \<And> x y. \<lbrakk>x \<noteq> y; x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> R A\<rbrakk> \<Longrightarrow> R A"apply (case_tac "A = {}", simp)by (case_tac "\<exists> x. A = {x}", clarsimp, blast)lemma rhs_aux:"\<lbrakk>distinct_rhs rhs; {Y. \<exists>r. (Y, r) \<in> rhs} = {X}\<rbrakk> \<Longrightarrow> (\<exists>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 \<in> (UNIV Quo Lang)" shows "\<exists> (reg::rexp). L reg = X" (is "\<exists> r. ?E r")proof- have "\<exists>ES'. Inv X ES' \<and> 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 "\<exists> 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 \<noteq> {}" using Inv_ES' ES'_single_equa by (auto simp:Inv_def ardenable_def) have rhs_eq_cls: "rhs_eq_cls rhs \<subseteq> {X, {[]}}" using Inv_ES' ES'_single_equa by (auto simp:Inv_def ardenable_def left_eq_cls_def) have X_not_empty: "X \<noteq> {}" 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 "\<exists> rhs'. X = L rhs' \<and> rhs_eq_cls rhs' = rhs_eq_cls rhs - {X} \<and> 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' \<subseteq> {{[]}}" 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 "\<exists> 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 qedqedtext {* definition of a regular language *}abbreviation reg :: "string set => bool"where "reg L1 \<equiv> (\<exists>r::rexp. L r = L1)"theorem myhill_nerode: assumes finite_CS: "finite (UNIV Quo Lang)" shows "\<exists> (reg::rexp). Lang = L reg" (is "\<exists> r. ?P r")proof - have has_r_each: "\<forall>C\<in>{X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}. \<exists>(r::rexp). C = L r" using finite_CS by (auto dest:every_eqcl_has_reg) have "\<exists> (rS::rexp set). finite rS \<and> (\<forall> C \<in> {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}. \<exists> r \<in> rS. C = L r) \<and> (\<forall> r \<in> rS. \<exists> C \<in> {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}. C = L r)" (is "\<exists> rS. ?Q rS") proof- have "\<And> C. C \<in> {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang} \<Longrightarrow> 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 ((\<lambda> C. SOME r. C = L r) ` {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> 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': "\<forall> C \<in> {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}. \<exists> r \<in> (rS::rexp set). C = L r" and has_cl_each: "\<forall> r \<in> (rS::rexp set). \<exists> C \<in> {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}. C = L r" by blast have "?P (folds ALT NULL rS)" proof show "Lang \<subseteq> 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) \<subseteq> 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 blastqed text {* tests by Christian *}(* Alternative definition for Quo *)definition QUOT :: "string set \<Rightarrow> (string set) set" where "QUOT Lang \<equiv> (\<Union>x. {\<lbrakk>x\<rbrakk>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_CSunfolding test[symmetric]by (auto dest: myhill_nerode)lemma cons_one: "x @ y \<in> {z} \<Longrightarrow> x @ y = z"by simplemma quot_lambda: "QUOT {[]} = {{[]}, UNIV - {[]}}"proof show "QUOT {[]} \<subseteq> {{[]}, UNIV - {[]}}" proof fix x assume in_quot: "x \<in> QUOT {[]}" show "x \<in> {{[]}, UNIV - {[]}}" proof - from in_quot have "x = {[]} \<or> x = UNIV - {[]}" unfolding QUOT_def equiv_class_def proof fix xa assume in_univ: "xa \<in> UNIV" and in_eqiv: "x \<in> {{y. xa \<equiv>{[]}\<equiv> y}}" show "x = {[]} \<or> x = UNIV - {[]}" proof(cases "xa = []") case True hence "{y. xa \<equiv>{[]}\<equiv> 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 \<equiv>{[]}\<equiv> 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 qednext show "{{[]}, UNIV - {[]}} \<subseteq> QUOT {[]}" proof fix x assume in_res: "x \<in> {{[]}, (UNIV::string set) - {[]}}" show "x \<in> (QUOT {[]})" proof - have "x = {[]} \<Longrightarrow> x \<in> QUOT {[]}" apply (simp add:QUOT_def equiv_class_def equiv_str_def) by (rule_tac x = "[]" in exI, auto) moreover have "x = UNIV - {[]} \<Longrightarrow> x \<in> 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 qedqedlemma quot_single_aux: "\<lbrakk>x \<noteq> []; x \<noteq> [c]\<rbrakk> \<Longrightarrow> x @ z \<noteq> [c]"by (induct x, auto)lemma quot_single: "\<And> (c::char). QUOT {[c]} = {{[]}, {[c]}, UNIV - {[], [c]}}"proof - fix c::"char" have exist_another: "\<exists> a. a \<noteq> 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]} \<subseteq> {{[]},{[c]}, UNIV - {[], [c]}}" proof fix x assume in_quot: "x \<in> QUOT {[c]}" show "x \<in> {{[]}, {[c]}, UNIV - {[],[c]}}" proof - from in_quot have "x = {[]} \<or> x = {[c]} \<or> x = UNIV - {[],[c]}" unfolding QUOT_def equiv_class_def proof fix xa assume in_eqiv: "x \<in> {{y. xa \<equiv>{[c]}\<equiv> y}}" show "x = {[]} \<or> x = {[c]} \<or> x = UNIV - {[], [c]}" proof- have "xa = [] \<Longrightarrow> x = {[]}" using in_eqiv by (auto simp add:equiv_str_def) moreover have "xa = [c] \<Longrightarrow> x = {[c]}" proof - have "xa = [c] \<Longrightarrow> {y. xa \<equiv>{[c]}\<equiv> 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] \<Longrightarrow> x = {[c]}" using in_eqiv by simp qed moreover have "\<lbrakk>xa \<noteq> []; xa \<noteq> [c]\<rbrakk> \<Longrightarrow> x = UNIV - {[],[c]}" proof - have "\<lbrakk>xa \<noteq> []; xa \<noteq> [c]\<rbrakk> \<Longrightarrow> {y. xa \<equiv>{[c]}\<equiv> 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 "\<lbrakk>xa \<noteq> []; xa \<noteq> [c]\<rbrakk> \<Longrightarrow> 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]}} \<subseteq> QUOT {[c]}" proof fix x assume in_res: "x \<in> {{[]},{[c]}, (UNIV::string set) - {[],[c]}}" show "x \<in> (QUOT {[c]})" proof - have "x = {[]} \<Longrightarrow> x \<in> QUOT {[c]}" apply (simp add:QUOT_def equiv_class_def equiv_str_def) by (rule_tac x = "[]" in exI, auto) moreover have "x = {[c]} \<Longrightarrow> x \<in> 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]} \<Longrightarrow> x \<in> 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 qedqedlemma eq_class_imp_eq_str: "\<lbrakk>x\<rbrakk>lang = \<lbrakk>y\<rbrakk>lang \<Longrightarrow> x \<equiv>lang\<equiv> y"by (auto simp:equiv_class_def equiv_str_def)lemma finite_tag_image: "finite (range tag) \<Longrightarrow> 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: "\<And> m n. tag m = tag (n::string) \<Longrightarrow> m \<equiv>lang\<equiv> n" shows "inj_on ((op `) tag) (QUOT lang)"proof (clarsimp simp add:inj_on_def QUOT_def) fix x y assume eq_tag: "tag ` \<lbrakk>x\<rbrakk>lang = tag ` \<lbrakk>y\<rbrakk>lang" show "\<lbrakk>x\<rbrakk>lang = \<lbrakk>y\<rbrakk>lang" proof - have aux1:"\<And>a b. a \<in> (\<lbrakk>b\<rbrakk>lang) \<Longrightarrow> (a \<equiv>lang\<equiv> b)" by (simp add:equiv_class_def equiv_str_def) have aux2: "\<And> A B f. \<lbrakk>f ` A = f ` B; A \<noteq> {}\<rbrakk> \<Longrightarrow> \<exists> a b. a \<in> A \<and> b \<in> B \<and> f a = f b" by auto have aux3: "\<And> a l. \<lbrakk>a\<rbrakk>l \<noteq> {}" 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) qedqeddefinition tag_str_ALT :: "string set \<Rightarrow> string set \<Rightarrow> string \<Rightarrow> (string set \<times> string set)"where "tag_str_ALT L\<^isub>1 L\<^isub>2 x \<equiv> (\<lbrakk>x\<rbrakk>L\<^isub>1, \<lbrakk>x\<rbrakk>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. \<exists>x. y = (\<lbrakk>x\<rbrakk>L\<^isub>1, \<lbrakk>x\<rbrakk>L\<^isub>2)} \<subseteq> (QUOT L\<^isub>1) \<times> (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) \<times> (QUOT L\<^isub>2)"])qedlemma tag_str_alt_inj: "tag_str_ALT L\<^isub>1 L\<^isub>2 x = tag_str_ALT L\<^isub>1 L\<^isub>2 y \<Longrightarrow> x \<equiv>(L\<^isub>1 \<union> L\<^isub>2)\<equiv> y"apply (simp add:tag_str_ALT_def equiv_class_def equiv_str_def)by blastlemma quot_alt: assumes finite1: "finite (QUOT L\<^isub>1)" and finite2: "finite (QUOT L\<^isub>2)" shows "finite (QUOT (L\<^isub>1 \<union> 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 \<union> L\<^isub>2))" using finite_tag_image tag_str_alt_range_finite finite1 finite2 by autonext show "inj_on (op ` (tag_str_ALT L\<^isub>1 L\<^isub>2)) (QUOT (L\<^isub>1 \<union> 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 \<Rightarrow> 'a list \<Rightarrow> '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 \<Longrightarrow> x - xa = y "by (induct x, auto)lemma [simp]: "x - [] = x"by (induct x, auto)lemma [simp]: "xa \<le> x \<Longrightarrow> (x @ y) - xa = (x - xa) @ y"by (auto elim:prefixE)definition tag_str_SEQ:: "string set \<Rightarrow> string set \<Rightarrow> string \<Rightarrow> (string set \<times> string set set)"where "tag_str_SEQ L\<^isub>1 L\<^isub>2 x \<equiv> if (\<exists> xa \<le> x. xa \<in> L\<^isub>1) then (\<lbrakk>x\<rbrakk>L\<^isub>1, {\<lbrakk>(x - xa)\<rbrakk>L\<^isub>2 | xa. xa \<le> x \<and> xa \<in> L\<^isub>1}) else (\<lbrakk>x\<rbrakk>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 \<Longrightarrow> ((\<exists> xa \<le> x. xa \<in> L\<^isub>1) \<and> \<lbrakk>x\<rbrakk>L\<^isub>1 = \<lbrakk>y\<rbrakk>L\<^isub>1 \<and> {\<lbrakk>(x - xa)\<rbrakk>L\<^isub>2 | xa. xa \<le> x \<and> xa \<in> L\<^isub>1} = {\<lbrakk>(y - ya)\<rbrakk>L\<^isub>2 | ya. ya \<le> y \<and> ya \<in> L\<^isub>1} ) \<or> ((\<forall> xa \<le> x. xa \<notin> L\<^isub>1) \<and> \<lbrakk>x\<rbrakk>L\<^isub>1 = \<lbrakk>y\<rbrakk>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)) \<subseteq> (QUOT L\<^isub>1) \<times> (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) \<times> (Pow (QUOT L\<^isub>2))" in finite_subset, auto)qedlemma app_in_seq_decom[rule_format]: "\<forall> x. x @ z \<in> L\<^isub>1 ; L\<^isub>2 \<longrightarrow> (\<exists> xa \<le> x. xa \<in> L\<^isub>1 \<and> (x - xa) @ z \<in> L\<^isub>2) \<or> (\<exists> za \<le> z. (x @ za) \<in> L\<^isub>1 \<and> (z - za) \<in> 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)donelemma 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) \<equiv>(L\<^isub>1 ; L\<^isub>2)\<equiv> y"proof - have aux: "\<And> x y z. \<lbrakk>tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y; x @ z \<in> L\<^isub>1 ; L\<^isub>2\<rbrakk> \<Longrightarrow> y @ z \<in> 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: "\<exists>xa\<le>x. xa \<in> L\<^isub>1 \<and> (x - xa) @ z \<in> L\<^isub>2" from x_gets_l2 have "\<exists> xa \<le> x. xa \<in> L\<^isub>1" by blast hence xy_l2:"{\<lbrakk>(x - xa)\<rbrakk>L\<^isub>2 | xa. xa \<le> x \<and> xa \<in> L\<^isub>1} = {\<lbrakk>(y - ya)\<rbrakk>L\<^isub>2 | ya. ya \<le> y \<and> ya \<in> L\<^isub>1}" using tag_eq tag_seq_eq_E by blast from x_gets_l2 obtain xa where xa_le_x: "xa \<le> x" and xa_in_l1: "xa \<in> L\<^isub>1" and rest_in_l2: "(x - xa) @ z \<in> L\<^isub>2" by blast hence "\<exists> ya. \<lbrakk>(x - xa)\<rbrakk>L\<^isub>2 = \<lbrakk>(y - ya)\<rbrakk>L\<^isub>2 \<and> ya \<le> y \<and> ya \<in> L\<^isub>1" using xy_l2 by auto then obtain ya where ya_le_x: "ya \<le> y" and ya_in_l1: "ya \<in> L\<^isub>1" and rest_eq: "\<lbrakk>(x - xa)\<rbrakk>L\<^isub>2 = \<lbrakk>(y - ya)\<rbrakk>L\<^isub>2" by blast from rest_eq rest_in_l2 have "(y - ya) @ z \<in> L\<^isub>2" by (auto simp:equiv_class_def equiv_str_def) hence "ya @ ((y - ya) @ z) \<in> L\<^isub>1 ; L\<^isub>2" using ya_in_l1 by (auto simp:lang_seq_def) thus "y @ z \<in> 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: "\<exists>za\<le>z. x @ za \<in> L\<^isub>1 \<and> z - za \<in> L\<^isub>2" from tag_eq tag_seq_eq_E have x_y_eq: "\<lbrakk>x\<rbrakk>L\<^isub>1 = \<lbrakk>y\<rbrakk>L\<^isub>1" by blast from x_gets_l1 obtain za where za_le_z: "za \<le> z" and x_za_in_l1: "(x @ za) \<in> L\<^isub>1" and rest_in_l2: "z - za \<in> L\<^isub>2" by blast from x_y_eq x_za_in_l1 have y_za_in_l1: "y @ za \<in> L\<^isub>1" by (auto simp:equiv_class_def equiv_str_def) hence "(y @ za) @ (z - za) \<in> 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 \<in> 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)qedlemma 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 autonext 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]: "\<And> a. \<forall> b x y. a @ b = x @ y \<longrightarrow> (\<exists> aa ab. a = aa @ ab \<and> x = aa \<and> y = ab @ b) \<or> (\<exists> ba bb. b = ba @ bb \<and> x = a @ ba \<and> y = bb \<and> ba \<noteq> [])" apply (induct_tac a rule:List.induct, simp) apply (rule allI | rule impI)+ by (case_tac x, auto)definition tag_str_STAR:: "string set \<Rightarrow> string \<Rightarrow> string set set"where "tag_str_STAR L\<^isub>1 x \<equiv> if (x = []) then {} else {\<lbrakk>x\<^isub>2\<rbrakk>L\<^isub>1 | x\<^isub>1 x\<^isub>2. x = x\<^isub>1 @ x\<^isub>2 \<and> x\<^isub>1 \<in> L\<^isub>1\<star> \<and> x\<^isub>2 \<noteq> []}"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) \<subseteq> 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)qedlemma star_prop[rule_format]: "x \<in> lang\<star> \<Longrightarrow> \<forall> y. y \<in> lang\<star> \<longrightarrow> x @ y \<in> lang\<star>"by (erule Star.induct, auto)lemma star_prop2: "y \<in> lang \<Longrightarrow> y \<in> lang\<star>"by (drule step[of y lang "[]"], auto simp:start)lemma star_prop3[rule_format]: "x \<in> lang\<star> \<Longrightarrow> \<forall>y . y \<in> lang \<longrightarrow> x @ y \<in> lang\<star>"by (erule Star.induct, auto intro:star_prop2)lemma postfix_prop: "y >>= (x @ y) \<Longrightarrow> x = []"by (erule postfixE, induct x arbitrary:y, auto)lemma inj_aux: "\<lbrakk>(m @ z) \<in> L\<^isub>1\<star>; m \<equiv>L\<^isub>1\<equiv> yb; xa @ m = x; xa \<in> L\<^isub>1\<star>; m \<noteq> []; \<forall> xa xb. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star> \<longrightarrow> xb >>= m\<rbrakk> \<Longrightarrow> (yb @ z) \<in> L\<^isub>1\<star>"proof- have "\<And>s. s \<in> L\<^isub>1\<star> \<Longrightarrow> \<forall> m z yb. (s = m @ z \<and> m \<equiv>L\<^isub>1\<equiv> yb \<and> x = xa @ m \<and> xa \<in> L\<^isub>1\<star> \<and> m \<noteq> [] \<and> (\<forall> xa xb. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star> \<longrightarrow> xb >>= m) \<longrightarrow> (yb @ z) \<in> L\<^isub>1\<star>)" 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 "\<lbrakk>(m @ z) \<in> L\<^isub>1\<star>; m \<equiv>L\<^isub>1\<equiv> yb; xa @ m = x; xa \<in> L\<^isub>1\<star>; m \<noteq> []; \<forall> xa xb. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star> \<longrightarrow> xb >>= m\<rbrakk> \<Longrightarrow> (yb @ z) \<in> L\<^isub>1\<star>" by blastqedlemma min_postfix_exists[rule_format]: "finite A \<Longrightarrow> A \<noteq> {} \<and> (\<forall> a \<in> A. \<forall> b \<in> A. ((b >>= a) \<or> (a >>= b))) \<longrightarrow> (\<exists> min. (min \<in> A \<and> (\<forall> a \<in> A. a >>= min)))"apply (erule finite.induct)apply simpapply simpapply (case_tac "A = {}")apply simpapply clarsimpapply (case_tac "a >>= min")apply (rule_tac x = min in exI, simp)apply (rule_tac x = a in exI, simp)apply clarifyapply (rotate_tac 5)apply (erule_tac x = aa in ballE) defer apply simpapply (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) \<Longrightarrow> x \<equiv>L\<^isub>1\<star>\<equiv> y"proof - have aux: "\<And> x y z. \<lbrakk>tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 y; x @ z \<in> L\<^isub>1\<star>\<rbrakk> \<Longrightarrow> y @ z \<in> L\<^isub>1\<star>" 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 \<in> L\<^isub>1\<star>" show "y @ z \<in> L\<^isub>1\<star>" 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. \<exists> xa. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star>} \<noteq> {}" using x_z by (simp, rule_tac x = x in exI, rule_tac x = "[]" in exI, simp add:start) have finite_set: "finite {xb. \<exists> xa. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star>}" apply (rule_tac B = "{xb. \<exists> xa. x = xa @ xb}" in finite_subset) apply auto apply (induct x, simp) apply (subgoal_tac "{xb. \<exists>xa. a # x = xa @ xb} = insert (a # x) {xb. \<exists>xa. x = xa @ xb}") apply auto by (case_tac xaa, simp+) have comparable: "\<forall> a \<in> {xb. \<exists> xa. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star>}. \<forall> b \<in> {xb. \<exists> xa. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star>}. ((b >>= a) \<or> (a >>= b))" by (auto simp:postfix_def, drule app_eq_elim, blast) hence "\<exists> min. min \<in> {xb. \<exists> xa. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star>} \<and> (\<forall> xa xb. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star> \<longrightarrow> 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 \<and> xa \<in> L\<^isub>1\<star>" and min_not_empty: "min \<noteq> []" and min_z_in_star: "min @ z \<in> L\<^isub>1\<star>" and is_min: "\<forall> xa xb. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star> \<longrightarrow> xb >>= min" by blast from x_decom min_not_empty have "\<lbrakk>min\<rbrakk>L\<^isub>1 \<in> tag_str_STAR L\<^isub>1 x" by (auto simp:tag_str_STAR_def) hence "\<exists> yb. \<lbrakk>yb\<rbrakk>L\<^isub>1 \<in> tag_str_STAR L\<^isub>1 y \<and> \<lbrakk>min\<rbrakk>L\<^isub>1 = \<lbrakk>yb\<rbrakk>L\<^isub>1" using tag_eq by auto hence "\<exists> ya yb. y = ya @ yb \<and> ya \<in> L\<^isub>1\<star> \<and> min \<equiv>L\<^isub>1\<equiv> yb \<and> yb \<noteq> [] " 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 \<in> L\<^isub>1\<star>" and yb_not_empty: "yb \<noteq> []" and min_yb_eq: "min \<equiv>L\<^isub>1\<equiv> yb" by blast from min_z_in_star min_yb_eq min_not_empty is_min x_decom have "yb @ z \<in> L\<^isub>1\<star>" 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) \<Longrightarrow> x \<equiv>L\<^isub>1\<star>\<equiv> y" by (auto intro:aux simp:equiv_str_def)qedlemma quot_star: assumes finite1: "finite (QUOT L\<^isub>1)" shows "finite (QUOT (L\<^isub>1\<star>))"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\<star>))" using finite_tag_image tag_str_star_range_finite finite1 by autonext show "inj_on (op ` (tag_str_STAR L\<^isub>1)) (QUOT (L\<^isub>1\<star>))" apply (rule_tac str_inj_imps) by (erule_tac tag_str_star_inj)qedlemma other_direction: "Lang = L (r::rexp) \<Longrightarrow> 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