diff -r 000000000000 -r 932f27fc6b55 MyhillNerode.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/MyhillNerode.thy Sun Oct 03 06:32:12 2010 +0000 @@ -0,0 +1,1826 @@ +theory RegExp + imports "Main" +begin + +text {* sequential composition of languages *} + +definition + lang_seq :: "string set \ string set \ string set" ("_ ; _" [100,100] 100) +where + "L1 ; L2 = {s1@s2 | s1 s2. s1 \ L1 \ s2 \ L2}" + +lemma lang_seq_empty: + shows "{[]} ; L = L" + and "L ; {[]} = L" +unfolding lang_seq_def by auto + +lemma lang_seq_null: + shows "{} ; L = {}" + and "L ; {} = {}" +unfolding lang_seq_def by auto + +lemma lang_seq_append: + assumes a: "s1 \ L1" + and b: "s2 \ L2" + shows "s1@s2 \ L1 ; L2" +unfolding lang_seq_def +using a b by auto + +lemma lang_seq_union: + shows "(L1 \ L2); L3 = (L1; L3) \ (L2; L3)" + and "L1; (L2 \ L3) = (L1; L2) \ (L1; L3)" +unfolding lang_seq_def by auto + +lemma lang_seq_assoc: + shows "(L1 ; L2) ; L3 = L1 ; (L2 ; L3)" +by (simp add: lang_seq_def Collect_def mem_def expand_fun_eq) + (metis append_assoc) + +lemma lang_seq_minus: + shows "(L1; L2) - {[]} = + (if [] \ L1 then L2 - {[]} else {}) \ + (if [] \ L2 then L1 - {[]} else {}) \ ((L1 - {[]}); (L2 - {[]}))" +apply(auto simp add: lang_seq_def) +apply(metis mem_def self_append_conv) +apply(metis mem_def self_append_conv2) +apply(metis mem_def self_append_conv2) +apply(metis mem_def self_append_conv) +done + +section {* Kleene star for languages defined as least fixed point *} + +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\" + +lemma lang_star_empty: + shows "{}\ = {[]}" +by (auto elim: Star.cases) + +lemma lang_star_cases: + shows "L\ = {[]} \ L ; L\" +proof + { fix x + have "x \ L\ \ x \ {[]} \ L ; L\" + unfolding lang_seq_def + by (induct rule: Star.induct) (auto) + } + then show "L\ \ {[]} \ L ; L\" by auto +next + show "{[]} \ L ; L\ \ L\" + unfolding lang_seq_def by auto +qed + +lemma lang_star_cases': + shows "L\ = {[]} \ L\ ; L" +proof + { fix x + have "x \ L\ \ x \ {[]} \ L\ ; 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\ \ {[]} \ L\ ; L" by auto +next + show "{[]} \ L\ ; L \ L\" + unfolding lang_seq_def + apply auto + apply (erule Star.induct) + apply auto + apply (drule step[of _ _ "[]"]) + by (auto intro:start) +qed + +lemma lang_star_simple: + shows "L \ L\" +by (subst lang_star_cases) + (auto simp only: lang_seq_def) + +lemma lang_star_prop0_aux: + "s2 \ L\ \ \ s1. s1 \ L \ (\ s3 s4. s3 \ L\ \ s4 \ L \ 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 auto + +lemma lang_star_prop0: + "\s1 \ L; s2 \ L\\ \ \ s3 s4. s3 \ L\ \ s4 \ L \ s1 @ s2 = s3 @ s4" +by (auto dest:lang_star_prop0_aux) + +lemma lang_star_prop1: + assumes asm: "L1; L2 \ L2" + shows "L1\; L2 \ L2" +proof - + { fix s1 s2 + assume minor: "s2 \ L2" + assume major: "s1 \ L1\" + then have "s1@s2 \ L2" + proof(induct rule: Star.induct) + case start + show "[]@s2 \ L2" using minor by simp + next + case (step s1 s1') + have "s1 \ L1" by fact + moreover + have "s1'@s2 \ L2" by fact + ultimately have "s1@(s1'@s2) \ L1; L2" by (auto simp add: lang_seq_def) + with asm have "s1@(s1'@s2) \ L2" by auto + then show "(s1@s1')@s2 \ L2" by simp + qed + } + then show "L1\; L2 \ L2" by (auto simp add: lang_seq_def) +qed + +lemma lang_star_prop2_aux: + "s2 \ L2\ \ \ s1. s1 \ L1 \ L1 ; L2 \ L1 \ s1 @ s2 \ 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 \ L1 \ L1 ; L2\ \ L1" +by (auto dest!:lang_star_prop2_aux simp:lang_seq_def) + +lemma lang_star_seq_subseteq: + shows "L ; L\ \ L\" +using lang_star_cases by blast + +lemma lang_star_double: + shows "L\; L\ = L\" +proof + show "L\ ; L\ \ L\" + using lang_star_prop1 lang_star_seq_subseteq by blast +next + have "L\ \ L\ \ L\; (L; L\)" by auto + also have "\ = L\;{[]} \ L\; (L; L\)" by (simp add: lang_seq_empty) + also have "\ = L\; ({[]} \ L; L\)" by (simp only: lang_seq_union) + also have "\ = L\; L\" using lang_star_cases by simp + finally show "L\ \ L\ ; L\" by simp +qed + +lemma lang_star_seq_subseteq': + shows "L\; L \ L\" +proof - + have "L \ L\" by (rule lang_star_simple) + then have "L\; L \ L\; L\" by (auto simp add: lang_seq_def) + then show "L\; L \ L\" using lang_star_double by blast +qed + +lemma + shows "L\ \ L\\" +by (rule lang_star_simple) + +section {* tricky section *} + +lemma k1: + assumes b: "s \ L\" + and a: "s \ []" + shows "s \ (L - {[]}); L\" +using b a +apply(induct rule: Star.induct) +apply(simp) +apply(case_tac "s1=[]") +apply(simp) +apply(simp add: lang_seq_def) +apply(blast) +done + +section {* (relies on lemma k1) *} + +lemma zzz: + shows "{s. c#s \ L1\} = {s. c#s \ L1} ; (L1\)" +apply(auto simp add: lang_seq_def Cons_eq_append_conv) +apply(drule k1) +apply(auto)[1] +apply(auto simp add: lang_seq_def)[1] +apply(rule_tac x="tl s1" in exI) +apply(rule_tac x="s2" in exI) +apply(auto)[1] +apply(auto simp add: Cons_eq_append_conv)[2] +apply(drule lang_seq_append) +apply(assumption) +apply(rotate_tac 1) +apply(drule rev_subsetD) +apply(rule lang_star_seq_subseteq) +apply(simp) +done + + + +section {* regular expressions *} + +datatype rexp = + NULL +| EMPTY +| CHAR char +| SEQ rexp rexp +| ALT rexp rexp +| STAR rexp + + +consts L:: "'a \ string set" + +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)\" + +defs (overloaded) + l_rexp_abs: "L rexp \ L_rexp rexp" + +declare L_rexp.simps [simp del] L_rexp.simps [folded l_rexp_abs, simp add] + +definition + Ls :: "rexp set \ string set" +where + "Ls R = (\r\R. (L r))" + +lemma Ls_union: + "Ls (R1 \ R2) = (Ls R1) \ (Ls R2)" +unfolding Ls_def by auto + +text {* helper function for termination proofs *} +fun + Left :: "rexp \ rexp" +where + "Left (SEQ r1 r2) = r1" + +text {* dagger function *} + +function + dagger :: "rexp \ char \ rexp list" ("_ \ _") +where + c1: "(NULL \ c) = []" +| c2: "(EMPTY) \ c = []" +| c3: "(CHAR c') \ c = (if c = c' then [EMPTY] else [])" +| c4: "(ALT r1 r2) \ c = r1 \ c @ r2 \ c" +| c5: "(SEQ NULL r2) \ c = []" +| c6: "(SEQ EMPTY r2) \ c = r2 \ c" +| c7: "(SEQ (CHAR c') r2) \ c = (if c = c' then [r2] else [])" +| c8: "(SEQ (SEQ r11 r12) r2) \ c = (SEQ r11 (SEQ r12 r2)) \ c" +| c9: "(SEQ (ALT r11 r12) r2) \ c = (SEQ r11 r2) \ c @ (SEQ r12 r2) \ c" +| c10: "(SEQ (STAR r1) r2) \ c = r2 \ c @ [SEQ r' (SEQ (STAR r1) r2). r' \ r1 \ c]" +| c11: "(STAR r) \ c = [SEQ r' (STAR r) . r' \ r \ c]" +by (pat_completeness) (auto) + +termination dagger + by (relation "measures [\(r, c). size r, \(r, c). size (Left r)]") (simp_all) + +lemma dagger_correctness: + "Ls (set r \ c) = {s. c#s \ L r}" +proof (induct rule: dagger.induct) + case (1 c) + show "Ls (set NULL \ c) = {s. c#s \ L NULL}" by (simp add: Ls_def) +next + case (2 c) + show "Ls (set EMPTY \ c) = {s. c#s \ L EMPTY}" by (simp add: Ls_def) +next + case (3 c' c) + show "Ls (set CHAR c' \ c) = {s. c#s \ L (CHAR c')}" by (simp add: Ls_def) +next + case (4 r1 r2 c) + have ih1: "Ls (set r1 \ c) = {s. c#s \ L r1}" by fact + have ih2: "Ls (set r2 \ c) = {s. c#s \ L r2}" by fact + show "Ls (set ALT r1 r2 \ c) = {s. c#s \ L (ALT r1 r2)}" + by (simp add: Ls_union ih1 ih2 Collect_disj_eq) +next + case (5 r2 c) + show "Ls (set SEQ NULL r2 \ c) = {s. c#s \ L (SEQ NULL r2)}" by (simp add: Ls_def lang_seq_null) +next + case (6 r2 c) + have ih: "Ls (set r2 \ c) = {s. c#s \ L r2}" by fact + show "Ls (set SEQ EMPTY r2 \ c) = {s. c#s \ L (SEQ EMPTY r2)}" + by (simp add: ih lang_seq_empty) +next + case (7 c' r2 c) + show "Ls (set SEQ (CHAR c') r2 \ c) = {s. c#s \ L (SEQ (CHAR c') r2)}" + by (simp add: Ls_def lang_seq_def) +next + case (8 r11 r12 r2 c) + have ih: "Ls (set SEQ r11 (SEQ r12 r2) \ c) = {s. c#s \ L (SEQ r11 (SEQ r12 r2))}" by fact + show "Ls (set SEQ (SEQ r11 r12) r2 \ c) = {s. c#s \ L (SEQ (SEQ r11 r12) r2)}" + by (simp add: ih lang_seq_assoc) +next + case (9 r11 r12 r2 c) + have ih1: "Ls (set SEQ r11 r2 \ c) = {s. c#s \ L (SEQ r11 r2)}" by fact + have ih2: "Ls (set SEQ r12 r2 \ c) = {s. c#s \ L (SEQ r12 r2)}" by fact + show "Ls (set SEQ (ALT r11 r12) r2 \ c) = {s. c#s \ L (SEQ (ALT r11 r12) r2)}" + by (simp add: Ls_union ih1 ih2 lang_seq_union Collect_disj_eq) +next + case (10 r1 r2 c) + have ih2: "Ls (set r2 \ c) = {s. c#s \ L r2}" by fact + have ih1: "Ls (set r1 \ c) = {s. c#s \ L r1}" by fact + have "Ls (set SEQ (STAR r1) r2 \ c) = Ls (set r2 \ c) \ (Ls (set r1 \ c); ((L r1)\ ; L r2))" + by (auto simp add: lang_seq_def Ls_def) + also have "\ = {s. c#s \ L r2} \ ({s. c#s \ L r1} ; ((L r1)\ ; L r2))" using ih1 ih2 by simp + also have "\ = {s. c#s \ L r2} \ ({s. c#s \ L r1} ; (L r1)\) ; L r2" by (simp add: lang_seq_assoc) + also have "\ = {s. c#s \ L r2} \ {s. c#s \ (L r1)\} ; L r2" by (simp add: zzz) + also have "\ = {s. c#s \ L r2} \ {s. c#s \ (L r1)\ ; L r2}" + by (auto simp add: lang_seq_def Cons_eq_append_conv) + also have "\ = {s. c#s \ (L r1)\ ; L r2}" + by (force simp add: lang_seq_def) + finally show "Ls (set SEQ (STAR r1) r2 \ c) = {s. c#s \ L (SEQ (STAR r1) r2)}" by simp +next + case (11 r c) + have ih: "Ls (set r \ c) = {s. c#s \ L r}" by fact + have "Ls (set (STAR r) \ c) = Ls (set r \ c) ; (L r)\" + by (auto simp add: lang_seq_def Ls_def) + also have "\ = {s. c#s \ L r} ; (L r)\" using ih by simp + also have "\ = {s. c#s \ (L r)\}" using zzz by simp + finally show "Ls (set (STAR r) \ c) = {s. c#s \ L (STAR r)}" by simp +qed + + +text {* matcher function (based on the "list"-dagger function) *} +fun + first_True :: "bool list \ bool" +where + "first_True [] = False" +| "first_True (x#xs) = (if x then True else first_True xs)" + +lemma not_first_True[simp]: + shows "(\(first_True xs)) = (\x \ set xs. \x)" +by (induct xs) (auto) + +lemma first_True: + shows "(first_True xs) = (\x \ set xs. x)" +by (induct xs) (auto) + +text {* matcher function *} + +function + matcher :: "rexp \ string \ bool" ("_ ! _") +where + "NULL ! s = False" +| "EMPTY ! s = (s =[])" +| "CHAR c ! s = (s = [c])" +| "ALT r1 r2 ! s = (r1 ! s \ r2 ! s)" +| "STAR r ! [] = True" +| "STAR r ! c#s = first_True [SEQ (r') (STAR r) ! s. r' \ r \ c]" +| "SEQ r1 r2 ! [] = (r1 ! [] \ r2 ! [])" +| "SEQ NULL r2 ! (c#s) = False" +| "SEQ EMPTY r2 ! (c#s) = (r2 ! c#s)" +| "SEQ (CHAR c') r2 ! (c#s) = (if c'=c then r2 ! s else False)" +| "SEQ (SEQ r11 r12) r2 ! (c#s) = (SEQ r11 (SEQ r12 r2) ! c#s)" +| "SEQ (ALT r11 r12) r2 ! (c#s) = ((SEQ r11 r2) ! (c#s) \ (SEQ r12 r2) ! (c#s))" +| "SEQ (STAR r1) r2 ! (c#s) = (r2 ! (c#s) \ first_True [SEQ r' (SEQ (STAR r1) r2) ! s. r' \ r1 \ c])" +by (pat_completeness) (auto) + +termination matcher + by(relation "measures [\(r,s). length s, \(r,s). size r, \(r,s). size (Left r)]") (simp_all) + +text {* positive correctness of the matcher *} +lemma matcher1: + shows "r ! s \ s \ L r" +proof (induct rule: matcher.induct) + case (1 s) + have "NULL ! s" by fact + then show "s \ L NULL" by simp +next + case (2 s) + have "EMPTY ! s" by fact + then show "s \ L EMPTY" by simp +next + case (3 c s) + have "CHAR c ! s" by fact + then show "s \ L (CHAR c)" by simp +next + case (4 r1 r2 s) + have ih1: "r1 ! s \ s \ L r1" by fact + have ih2: "r2 ! s \ s \ L r2" by fact + have "ALT r1 r2 ! s" by fact + with ih1 ih2 show "s \ L (ALT r1 r2)" by auto +next + case (5 r) + have "STAR r ! []" by fact + then show "[] \ L (STAR r)" by auto +next + case (6 r c s) + have ih1: "\rx. \rx \ set r \ c; SEQ rx (STAR r) ! s\ \ s \ L (SEQ rx (STAR r))" by fact + have as: "STAR r ! c#s" by fact + then obtain r' where imp1: "r' \ set r \ c" and imp2: "SEQ r' (STAR r) ! s" + by (auto simp add: first_True) + from imp2 imp1 have "s \ L (SEQ r' (STAR r))" using ih1 by simp + then have "s \ L r' ; (L r)\" by simp + then have "s \ Ls (set r \ c) ; (L r)\" using imp1 by (auto simp add: Ls_def lang_seq_def) + then have "s \ {s. c#s \ L r} ; (L r)\" by (auto simp add: dagger_correctness) + then have "s \ {s. c#s \ (L r)\}" by (simp add: zzz) + then have "c#s \ {[c]}; {s. c#s \ (L r)\}" by (auto simp add: lang_seq_def) + then have "c#s \ (L r)\" by (auto simp add: lang_seq_def) + then show "c#s \ L (STAR r)" by simp +next + case (7 r1 r2) + have ih1: "r1 ! [] \ [] \ L r1" by fact + have ih2: "r2 ! [] \ [] \ L r2" by fact + have as: "SEQ r1 r2 ! []" by fact + then have "r1 ! [] \ r2 ! []" by simp + then show "[] \ L (SEQ r1 r2)" using ih1 ih2 by (simp add: lang_seq_def) +next + case (8 r2 c s) + have "SEQ NULL r2 ! c#s" by fact + then show "c#s \ L (SEQ NULL r2)" by simp +next + case (9 r2 c s) + have ih1: "r2 ! c#s \ c#s \ L r2" by fact + have "SEQ EMPTY r2 ! c#s" by fact + then show "c#s \ L (SEQ EMPTY r2)" using ih1 by (simp add: lang_seq_def) +next + case (10 c' r2 c s) + have ih1: "\c' = c; r2 ! s\ \ s \ L r2" by fact + have "SEQ (CHAR c') r2 ! c#s" by fact + then show "c#s \ L (SEQ (CHAR c') r2)" + using ih1 by (auto simp add: lang_seq_def split: if_splits) +next + case (11 r11 r12 r2 c s) + have ih1: "SEQ r11 (SEQ r12 r2) ! c#s \ c#s \ L (SEQ r11 (SEQ r12 r2))" by fact + have "SEQ (SEQ r11 r12) r2 ! c#s" by fact + then have "c#s \ L (SEQ r11 (SEQ r12 r2))" using ih1 by simp + then show "c#s \ L (SEQ (SEQ r11 r12) r2)" by (simp add: lang_seq_assoc) +next + case (12 r11 r12 r2 c s) + have ih1: "SEQ r11 r2 ! c#s \ c#s \ L (SEQ r11 r2)" by fact + have ih2: "SEQ r12 r2 ! c#s \ c#s \ L (SEQ r12 r2)" by fact + have "SEQ (ALT r11 r12) r2 ! c#s" by fact + then show "c#s \ L (SEQ (ALT r11 r12) r2)" + using ih1 ih2 by (auto simp add: lang_seq_union) +next + case (13 r1 r2 c s) + have ih1: "r2 ! c#s \ c#s \ L r2" by fact + have ih2: "\r'. \r' \ set r1 \ c; SEQ r' (SEQ (STAR r1) r2) ! s\ \ + s \ L (SEQ r' (SEQ (STAR r1) r2))" by fact + have "SEQ (STAR r1) r2 ! c#s" by fact + then have "(r2 ! c#s) \ (\r' \ set r1 \ c. SEQ r' (SEQ (STAR r1) r2) ! s)" by (auto simp add: first_True) + moreover + { assume "r2 ! c#s" + with ih1 have "c#s \ L r2" by simp + then have "c # s \ L r1\ ; L r2" + by (auto simp add: lang_seq_def) + then have "c#s \ L (SEQ (STAR r1) r2)" by simp + } + moreover + { assume "\r' \ set r1 \ c. SEQ r' (SEQ (STAR r1) r2) ! s" + then obtain r' where imp1: "r' \ set r1 \ c" and imp2: "SEQ r' (SEQ (STAR r1) r2) ! s" by blast + from imp2 imp1 have "s \ L (SEQ r' (SEQ (STAR r1) r2))" using ih2 by simp + then have "s \ L r' ; ((L r1)\ ; L r2)" by simp + then have "s \ Ls (set r1 \ c) ; ((L r1)\ ; L r2)" using imp1 by (auto simp add: Ls_def lang_seq_def) + then have "s \ {s. c#s \ L r1} ; ((L r1)\ ; L r2)" by (simp add: dagger_correctness) + then have "s \ ({s. c#s \ L r1} ; (L r1)\) ; L r2" by (simp add: lang_seq_assoc) + then have "s \ {s. c#s \ (L r1)\} ; L r2" by (simp add: zzz) + then have "c#s \ {[c]}; ({s. c#s \ (L r1)\}; L r2)" by (auto simp add: lang_seq_def) + then have "c#s \ ({[c]}; {s. c#s \ (L r1)\}) ; L r2" by (simp add: lang_seq_assoc) + then have "c#s \ (L r1)\; L r2" by (auto simp add: lang_seq_def) + then have "c#s \ L (SEQ (STAR r1) r2)" by simp + } + ultimately show "c#s \ L (SEQ (STAR r1) r2)" by blast +qed + +text {* negative correctness of the matcher *} +lemma matcher2: + shows "\ r ! s \ s \ L r" +proof (induct rule: matcher.induct) + case (1 s) + have "\ NULL ! s" by fact + then show "s \ L NULL" by simp +next + case (2 s) + have "\ EMPTY ! s" by fact + then show "s \ L EMPTY" by simp +next + case (3 c s) + have "\ CHAR c ! s" by fact + then show "s \ L (CHAR c)" by simp +next + case (4 r1 r2 s) + have ih2: "\ r1 ! s \ s \ L r1" by fact + have ih4: "\ r2 ! s \ s \ L r2" by fact + have "\ ALT r1 r2 ! s" by fact + then show "s \ L (ALT r1 r2)" by (simp add: ih2 ih4) +next + case (5 r) + have "\ STAR r ! []" by fact + then show "[] \ L (STAR r)" by simp +next + case (6 r c s) + have ih: "\rx. \rx \ set r \ c; \SEQ rx (STAR r) ! s\ \ s \ L (SEQ rx (STAR r))" by fact + have as: "\ STAR r ! c#s" by fact + then have "\r'\ set r \ c. \ (SEQ r' (STAR r) ! s)" by simp + then have "\r'\ set r \ c. s \ L (SEQ r' (STAR r))" using ih by auto + then have "\r'\ set r \ c. s \ L r' ; ((L r)\)" by simp + then have "s \ (Ls (set r \ c)) ; ((L r)\)" by (auto simp add: Ls_def lang_seq_def) + then have "s \ {s. c#s \ L r} ; ((L r)\)" by (simp add: dagger_correctness) + then have "s \ {s. c#s \ (L r)\}" by (simp add: zzz) + then have "c#s \ {[c]} ; {s. c#s \ (L r)\}" by (auto simp add: lang_seq_assoc lang_seq_def) + then have "c#s \ (L r)\" by (simp add: lang_seq_def) + then show "c#s \ L (STAR r)" by simp +next + case (7 r1 r2) + have ih2: "\ r1 ! [] \ [] \ L r1" by fact + have ih4: "\ r2 ! [] \ [] \ L r2" by fact + have "\ SEQ r1 r2 ! []" by fact + then have "\ r1 ! [] \ \ r2 ! []" by simp + then show "[] \ L (SEQ r1 r2)" using ih2 ih4 + by (auto simp add: lang_seq_def) +next + case (8 r2 c s) + have "\ SEQ NULL r2 ! c#s" by fact + then show "c#s \ L (SEQ NULL r2)" by (simp add: lang_seq_null) +next + case (9 r2 c s) + have ih1: "\ r2 ! c#s \ c#s \ L r2" by fact + have "\ SEQ EMPTY r2 ! c#s" by fact + then show "c#s \ L (SEQ EMPTY r2)" + using ih1 by (simp add: lang_seq_def) +next + case (10 c' r2 c s) + have ih2: "\c' = c; \r2 ! s\ \ s \ L r2" by fact + have "\ SEQ (CHAR c') r2 ! c#s" by fact + then show "c#s \ L (SEQ (CHAR c') r2)" + using ih2 by (auto simp add: lang_seq_def) +next + case (11 r11 r12 r2 c s) + have ih2: "\ SEQ r11 (SEQ r12 r2) ! c#s \ c#s \ L (SEQ r11 (SEQ r12 r2))" by fact + have "\ SEQ (SEQ r11 r12) r2 ! c#s" by fact + then show "c#s \ L (SEQ (SEQ r11 r12) r2)" + using ih2 by (auto simp add: lang_seq_def) +next + case (12 r11 r12 r2 c s) + have ih2: "\ SEQ r11 r2 ! c#s \ c#s \ L (SEQ r11 r2)" by fact + have ih4: "\ SEQ r12 r2 ! c#s \ c#s \ L (SEQ r12 r2)" by fact + have "\ SEQ (ALT r11 r12) r2 ! c#s" by fact + then show " c#s \ L (SEQ (ALT r11 r12) r2)" + using ih2 ih4 by (simp add: lang_seq_union) +next + case (13 r1 r2 c s) + have ih1: "\r2 ! c#s \ c#s \ L r2" by fact + have ih2: "\rx. \rx \ set r1 \ c; \ SEQ rx (SEQ (STAR r1) r2) ! s\ + \ s \ L (SEQ rx (SEQ (STAR r1) r2))" by fact + have as: "\ SEQ (STAR r1) r2 ! c#s" by fact + then have as1: "\r2 ! c#s" and as2: "\r1'\set r1 \ c. \ SEQ r1' (SEQ (STAR r1) r2) ! s" by simp_all + from as1 have bs1: "c#s \ L r2" using ih1 by simp + from as2 have "\r1'\set r1 \ c. \ SEQ r1' (SEQ (STAR r1) r2) ! s" by simp + then have "\r1'\set r1 \ c. s \ L (SEQ r1' (SEQ (STAR r1) r2))" using ih2 by simp + then have "\r1'\set r1 \ c. s \ L r1'; ((L r1)\; L r2)" by simp + then have "s \ (Ls (set r1 \ c)) ; ((L r1)\; L r2)" by (auto simp add: Ls_def lang_seq_def) + then have "s \ {s. c#s \ L r1} ; ((L r1)\; L r2)" by (simp add: dagger_correctness) + then have "s \ ({s. c#s \ L r1} ; (L r1)\); L r2" by (simp add: lang_seq_assoc) + then have "s \ {s. c#s \ (L r1)\}; L r2" by (simp add: zzz) + then have "c#s \ {[c]}; ({s. c#s \ (L r1)\}; L r2)" by (auto simp add: lang_seq_def) + then have "c#s \ (L r1)\; L r2" using bs1 by (auto simp add: lang_seq_def Cons_eq_append_conv) + then show "c#s \ L (SEQ (STAR r1) r2)" by simp +qed + +section {* Questions *} + +text {* + - Why was the key lemma k1 omitted; is there an easy, non-induction + way for obtaining this property? + - Why was False included in the definition of the STAR-clause in + the matcher? Has this something to do with executing the code? + +*} + +section {* Code *} + +export_code dagger in SML module_name Dagger file - +export_code matcher in SML module_name Dagger file - + +section {* Examples *} + +text {* since now everything is based on lists, the evaluation is quite fast *} + +value "NULL ! []" +value "(CHAR (CHR ''a'')) ! [CHR ''a'']" +value "((CHAR a) ! [a,a])" +value "(STAR (CHAR a)) ! []" +value "(STAR (CHAR a)) ! [a,a]" +value "(SEQ (CHAR (CHR ''a'')) (SEQ (STAR (CHAR (CHR ''b''))) (CHAR (CHR ''c'')))) ! ''abbbbc''" +value "(SEQ (CHAR (CHR ''a'')) (SEQ (STAR (CHAR (CHR ''b''))) (CHAR (CHR ''c'')))) ! ''abbcbbc''" + +section {* Slind et al's matcher based on derivatives *} + +fun + nullable :: "rexp \ bool" +where + "nullable (NULL) = False" +| "nullable (EMPTY) = True" +| "nullable (CHAR c) = False" +| "nullable (ALT r1 r2) = ((nullable r1) \ (nullable r2))" +| "nullable (SEQ r1 r2) = ((nullable r1) \ (nullable r2))" +| "nullable (STAR r) = True" + +lemma nullable: + shows "([] \ L r) = nullable r" +by (induct r) + (auto simp add: lang_seq_def) + +fun + der :: "char \ rexp \ rexp" +where + "der c (NULL) = NULL" +| "der c (EMPTY) = NULL" +| "der c (CHAR c') = (if c=c' then EMPTY else NULL)" +| "der c (ALT r1 r2) = ALT (der c r1) (der c r2)" +| "der c (SEQ r1 r2) = ALT (SEQ (der c r1) r2) (if nullable r1 then der c r2 else NULL)" +| "der c (STAR r) = SEQ (der c r) (STAR r)" + +lemma k2: + assumes b: "s \ L1\" + and a: "s \ []" + shows "s \ (L1; (L1\))" +using b a +apply(induct rule: Star.induct) +apply(simp) +apply(case_tac "s1=[]") +apply(simp) +apply(simp add: lang_seq_def) +apply(blast) +done + + +lemma der_correctness: + shows "(s \ L (der c r)) = ((c#s) \ L r)" +proof (induct r arbitrary: s) + case (NULL s) + show "(s \ L (der c NULL)) = (c#s \ L NULL)" by simp +next + case (EMPTY s) + show "(s \ L (der c EMPTY)) = (c#s \ L EMPTY)" by simp +next + case (CHAR c' s) + show "(s \ L (der c (CHAR c'))) = (c#s \ L (CHAR c'))" by simp +next + case (SEQ r1 r2 s) + have ih1: "\s. (s \ L (der c r1)) = (c#s \ L r1)" by fact + have ih2: "\s. (s \ L (der c r2)) = (c#s \ L r2)" by fact + show "(s \ L (der c (SEQ r1 r2))) = (c#s \ L (SEQ r1 r2))" + using ih1 ih2 + by (auto simp add: nullable[symmetric] lang_seq_def Cons_eq_append_conv) +next + case (ALT r1 r2 s) + have ih1: "\s. (s \ L (der c r1)) = (c#s \ L r1)" by fact + have ih2: "\s. (s \ L (der c r2)) = (c#s \ L r2)" by fact + show "(s \ L (der c (ALT r1 r2))) = (c#s \ L (ALT r1 r2))" + using ih1 ih2 by (auto simp add: lang_seq_def) +next + case (STAR r s) + have ih1: "\s. (s \ L (der c r)) = (c#s \ L r)" by fact + show "(s \ L (der c (STAR r))) = (c#s \ L (STAR r))" + using ih1 + apply(simp) + apply(auto simp add: lang_seq_def) + apply(drule lang_seq_append) + apply(assumption) + apply(simp) + apply(subst lang_star_cases) + apply(simp) + thm k1 + apply(drule k2) + apply(simp) + apply(simp add: lang_seq_def) + apply(erule exE)+ + apply(erule conjE)+ + apply(auto simp add: lang_seq_def Cons_eq_append_conv) + apply(drule k1) + apply(simp) + apply(simp add: lang_seq_def) + apply(erule exE)+ + apply(erule conjE)+ + apply(auto simp add: lang_seq_def Cons_eq_append_conv) + done +qed + +fun + derivative :: "string \ rexp \ rexp" +where + "derivative [] r = r" +| "derivative (c#s) r = derivative s (der c r)" + +fun + slind_matcher :: "rexp \ string \ bool" +where + "slind_matcher r s = nullable (derivative s r)" + +lemma slind_matcher: + shows "slind_matcher r s = (s \ L r)" +by (induct s arbitrary: r) + (auto simp add: nullable der_correctness) + +export_code slind_matcher in SML module_name Slind file - + + +(* ******************************************** now is the codes writen by chunhan ************************************* *) + +section {* Arden's Lemma revised *} + +lemma arden_aux1: + assumes a: "X \ X ; A \ B" + and b: "[] \ A" + shows "x \ X \ x \ B ; A\" +apply (induct x taking:length rule:measure_induct) +apply (subgoal_tac "x \ X ; A \ B") +defer +using a +apply (auto)[1] +apply simp +apply (erule disjE) +defer +apply (auto simp add:lang_seq_def) [1] +apply (subgoal_tac "\ x1 x2. x = x1 @ x2 \ x1 \ X \ x2 \ A") +defer +apply (auto simp add:lang_seq_def) [1] +apply (erule exE | erule conjE)+ +apply simp +apply (drule_tac x = x1 in spec) +apply (simp) +using b +apply - +apply (auto)[1] +apply (subgoal_tac "x1 @ x2 \ (B ; A\) ; A") +defer +apply (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: "[] \ A" + shows "(X = X ; A \ B) \ (X = B ; A\)" +apply(rule iffI) +defer +apply(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 \ B" + then have as1: "X ; A \ B \ X" and as2: "X \ X ; A \ B" by simp_all + from as1 have a: "X ; A \ X" and b: "B \ X" by simp_all + from b have "B; A\ \ X ; A\" by (auto simp add: lang_seq_def) + moreover + from a have "X ; A\ \ X" + +by (rule lang_star_prop2) + ultimately have f1: "B ; A\ \ X" by simp + from as2 nemp + have f2: "X \ B; A\" using arden_aux1 by auto + from f1 f2 show "X = B; A\" by auto +qed + +section {* equiv class' definition *} + +definition + equiv_str :: "string \ string set \ string \ bool" ("_ \_\ _" [100, 100, 100] 100) +where + "x \L'\ y \ (\z. x@z \ L' \ y@z \ L')" + +definition + equiv_class :: "string \ (string set) \ string set" ("\_\_" [100, 100] 100) +where + "\x\L' \ {y. x \L'\ y}" + +text {* Chunhan modifies Q to Quo *} +definition + quot :: "string set \ (string set) \ (string set) set" ("_ Quo _" [100, 100] 100) +where + "L' Quo R \ { \x\R | x. x \ L'}" + +lemma lang_eqs_union_of_eqcls: + "Lang = \ {X. X \ (UNIV Quo Lang) \ (\ x \ X. x \ Lang)}" +proof + show "Lang \ \{X \ UNIV Quo Lang. \x\X. x \ Lang}" + proof + fix x + assume "x \ Lang" + thus "x \ \{X \ UNIV Quo Lang. \x\X. x \ Lang}" + proof (simp add:quot_def) + assume "(1)": "x \ Lang" + show "\xa. (\x. xa = \x\Lang) \ (\x\xa. x \ Lang) \ x \ xa" (is "\xa.?P xa") + proof - + have "?P (\x\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 "\{X \ UNIV Quo Lang. \x\X. x \ Lang} \ Lang" + by auto +qed + +lemma empty_notin_CS: "{} \ 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: + "\X \ UNIV Quo Lang; Y \ UNIV Quo Lang; X \ Y\ \ X \ Y = {}" +by (auto simp:quot_def equiv_class_def equiv_str_def) + +text {* equiv_class transition *} +definition + CT :: "string set \ char \ string set \ bool" ("_-_\_" [99,99]99) +where + "X-c\Y \ ((X;{[c]}) \ Y)" + +types t_equa_rhs = "(string set \ rexp) set" + +types t_equa = "(string set \ t_equa_rhs)" + +types t_equas = "t_equa set" + +text {* "empty_rhs" generates "\" for init-state, just like "\" for final states in Brzozowski method. + But if the init-state is "{[]}" ("\" itself) then empty set is returned, see definition of "equation_rhs" *} +definition + empty_rhs :: "string set \ t_equa_rhs" +where + "empty_rhs X \ if ([] \ X) then {({[]}, EMPTY)} else {}" + +definition + folds :: "('a \ 'b \ 'b) \ 'b \ 'a set \ 'b" +where + "folds f z S \ SOME x. fold_graph f z S x" + +definition + equation_rhs :: "(string set) set \ string set \ t_equa_rhs" +where + "equation_rhs CS X \ if (X = {[]}) then {({[]}, EMPTY)} + else {(S, folds ALT NULL {CHAR c| c. S-c\X} )| S. S \ CS} \ empty_rhs X" + +definition + equations :: "(string set) set \ t_equas" +where + "equations CS \ {(X, equation_rhs CS X) | X. X \ CS}" + +definition + L_rhs :: "t_equa_rhs \ string set" +where + "L_rhs rhs \ {x. \ X r. (X, r) \ rhs \ x \ X;(L r)}" + +defs (overloaded) + l_rhs_abs: "L rhs \ L_rhs rhs" + +lemmas L_def = L_rhs_def [folded l_rhs_abs] L_rexp.simps (* ??? is this OK ?? *) + +definition + distinct_rhs :: "t_equa_rhs \ bool" +where + "distinct_rhs rhs \ \ X reg\<^isub>1 reg\<^isub>2. (X, reg\<^isub>1) \ rhs \ (X, reg\<^isub>2) \ rhs \ reg\<^isub>1 = reg\<^isub>2" + +definition + distinct_equas_rhs :: "t_equas \ bool" +where + "distinct_equas_rhs equas \ \ X rhs. (X, rhs) \ equas \ distinct_rhs rhs" + +definition + distinct_equas :: "t_equas \ bool" +where + "distinct_equas equas \ \ X rhs rhs'. (X, rhs) \ equas \ (X, rhs') \ equas \ rhs = rhs'" + +definition + seq_rhs_r :: "t_equa_rhs \ rexp \ t_equa_rhs" +where + "seq_rhs_r rhs r \ (\(X, reg). (X, SEQ reg r)) ` rhs" + +definition + del_x_paired :: "('a \ 'b) set \ 'a \ ('a \ 'b) set" +where + "del_x_paired S x \ S - {X. X \ S \ fst X = x}" + +definition + arden_variate :: "string set \ rexp \ t_equa_rhs \ t_equa_rhs" +where + "arden_variate X r rhs \ seq_rhs_r (del_x_paired rhs X) (STAR r)" + +definition + no_EMPTY_rhs :: "t_equa_rhs \ bool" +where + "no_EMPTY_rhs rhs \ \ X r. (X, r) \ rhs \ X \ {[]} \ [] \ L r" + +definition + no_EMPTY_equas :: "t_equas \ bool" +where + "no_EMPTY_equas equas \ \ X rhs. (X, rhs) \ equas \ no_EMPTY_rhs rhs" + +lemma fold_alt_null_eqs: + "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) +by auto (*??? how do this be in Isar ?? *) + +lemma seq_rhs_r_prop1: + "L (seq_rhs_r rhs r) = (L rhs);(L r)" +apply (rule set_ext, rule iffI) +apply (auto simp:L_def 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: + "\distinct_rhs rhs; (X, r) \ rhs\ \ X ; L r \ L (del_x_paired rhs X) = L rhs" +apply (simp add:L_def del_x_paired_def) +apply (rule set_ext, rule iffI, simp) +apply (erule disjE, rule_tac x = X in exI, rule_tac x = r in exI, simp) +apply (clarify, rule_tac x = Xa in exI, rule_tac x = ra in exI, simp) +apply (clarsimp, drule_tac x = Xa in spec, drule_tac x = ra in spec) +apply (auto simp:distinct_rhs_def) +done + +lemma arden_variate_prop: + assumes "(X, rx) \ rhs" + shows "(\ Y. Y \ X \ (\ r. (Y, r) \ rhs) = (\ r. (Y, r) \ (arden_variate X rx rhs)))" +proof (rule allI, rule impI) + fix Y + assume "(1)": "Y \ X" + show "(\r. (Y, r) \ rhs) = (\r. (Y, r) \ arden_variate X rx rhs)" + proof + assume "(1_1)": "\r. (Y, r) \ rhs" + show "\r. (Y, r) \ arden_variate X rx rhs" (is "\r. ?P r") + proof - + from "(1_1)" obtain r where "(Y, r) \ rhs" .. + hence "?P (SEQ r (STAR rx))" + proof (simp add:arden_variate_def image_def) + have "(Y, r) \ rhs \ (Y, r) \ del_x_paired rhs X" + by (auto simp:del_x_paired_def "(1)") + thus "(Y, r) \ rhs \ (Y, SEQ r (STAR rx)) \ 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)": "\r. (Y, r) \ arden_variate X rx rhs" + thus "\r. (Y, r) \ rhs" (is "\ r. ?P r") + by (auto simp:arden_variate_def del_x_paired_def seq_rhs_r_def image_def) + qed +qed + +text {* + arden_variate_valid: proves variation from "X = X;r + Y;ry + \" to "X = Y;(SEQ ry (STAR r)) + \" holds the law of "language of left equiv language of right" +*} +lemma arden_variate_valid: + assumes X_not_empty: "X \ {[]}" + and l_eq_r: "X = L rhs" + and dist: "distinct_rhs rhs" + and no_empty: "no_EMPTY_rhs rhs" + and self_contained: "(X, r) \ rhs" + shows "X = L (arden_variate X r rhs)" +proof - + have "[] \ L r" using no_empty X_not_empty self_contained + by (auto simp:no_EMPTY_rhs_def) + hence ardens: "X = X;(L r) \ (L (del_x_paired rhs X)) \ X = (L (del_x_paired rhs X)) ; (L r)\" + by (rule ardens_revised) + have del_x: "X = X ; L r \ L (del_x_paired rhs X) \ X = L rhs" using dist l_eq_r self_contained + by (auto dest:del_x_paired_prop1) + show ?thesis + proof + show "X \ L (arden_variate X r rhs)" + proof + fix x + assume "(1_1)": "x \ X" with l_eq_r ardens del_x + show "x \ L (arden_variate X r rhs)" + by (simp add:arden_variate_def seq_rhs_r_prop1) + qed + next + show "L (arden_variate X r rhs) \ X" + proof + fix x + assume "(2_1)": "x \ L (arden_variate X r rhs)" with ardens del_x l_eq_r + show "x \ X" + by (simp add:arden_variate_def seq_rhs_r_prop1) + qed + qed +qed + +text {* merge_rhs {(X1, r1), (x2, r2}, (x4, r4), \} {(x1, r1'), (x3, r3'), \} = {(x1, ALT r1 r1'}, (x2, r2), (x3, r3'), (x4, r4), \} *} +definition + merge_rhs :: "t_equa_rhs \ t_equa_rhs \ t_equa_rhs" +where + "merge_rhs rhs rhs' \ {(X, r). (\ r1 r2. ((X,r1) \ rhs \ (X, r2) \ rhs') \ r = ALT r1 r2) \ + (\ r1. (X, r1) \ rhs \ (\ (\ r2. (X, r2) \ rhs')) \ r = r1) \ + (\ r2. (X, r2) \ rhs' \ (\ (\ r1. (X, r1) \ rhs)) \ r = r2) }" + + +text {* rhs_subst rhs X=xrhs r: substitude all occurence of X in rhs((X,r) \ rhs) with xrhs *} +definition + rhs_subst :: "t_equa_rhs \ string set \ t_equa_rhs \ rexp \ t_equa_rhs" +where + "rhs_subst rhs X xrhs r \ merge_rhs (del_x_paired rhs X) (seq_rhs_r xrhs r)" + +definition + equas_subst_f :: "string set \ t_equa_rhs \ t_equa \ t_equa" +where + "equas_subst_f X xrhs equa \ let (Y, rhs) = equa in + if (\ r. (X, r) \ rhs) + then (Y, rhs_subst rhs X xrhs (SOME r. (X, r) \ rhs)) + else equa" + +definition + equas_subst :: "t_equas \ string set \ t_equa_rhs \ t_equas" +where + "equas_subst ES X xrhs \ del_x_paired (equas_subst_f X xrhs ` ES) X" + +lemma lang_seq_prop1: + "x \ X ; L r \ x \ X ; (L r \ L r')" +by (auto simp:lang_seq_def) + +lemma lang_seq_prop1': + "x \ X; L r \ x \ X ; (L r' \ L r)" +by (auto simp:lang_seq_def) + +lemma lang_seq_prop2: + "x \ X; (L r \ L r') \ x \ X;L r \ x \ X;L r'" +by (auto simp:lang_seq_def) + +lemma merge_rhs_prop1: + shows "L (merge_rhs rhs rhs') = L rhs \ L rhs' " +apply (auto simp add:merge_rhs_def L_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 "\ r2. (X, r2) \ 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 "\ r1. (X, r1) \ 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) +done + +lemma no_EMPTY_rhss_imp_merge_no_EMPTY: + "\no_EMPTY_rhs rhs; no_EMPTY_rhs rhs'\ \ 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 auto + +lemma distinct_rhs_prop: + "\distinct_rhs rhs; (X, r1) \ rhs; (X, r2) \ rhs\ \ 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_rhs +apply (drule distinct_rhs_prop, simp+) +using dist_rhs' +apply (drule distinct_rhs_prop, simp+) +using dist_rhs +apply (drule distinct_rhs_prop, simp+) +using dist_rhs' +apply (drule distinct_rhs_prop, simp+) +done + +lemma seq_rhs_r_holds_distinct: + "distinct_rhs rhs \ 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_r +by (auto simp:seq_rhs_r_prop1) + +lemma rhs_subst_prop1: + assumes l_eq_r: "X = L xrhs" + and dist: "distinct_rhs rhs" + shows "(X, r) \ rhs \ L rhs = L (rhs_subst rhs X xrhs r)" +apply (simp add:rhs_subst_def merge_rhs_prop1) +using l_eq_r +apply (drule_tac r = r in seq_rhs_r_prop0, simp) +using dist +apply (auto dest:del_x_paired_prop1) +done + +lemma del_x_paired_holds_distinct_rhs: + "distinct_rhs rhs \ distinct_rhs (del_x_paired rhs x)" +by (auto simp:distinct_rhs_def del_x_paired_def) + +lemma rhs_subst_holds_distinct_rhs: + "\distinct_rhs rhs; distinct_rhs xrhs\ \ 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 \ (string set) set" +where + "left_eq_cls ES \ {X. \ rhs. (X, rhs) \ ES} " + +definition right_eq_cls :: "t_equas \ (string set) set" +where + "right_eq_cls ES \ {Y. \ X rhs r. (X, rhs) \ ES \ (Y, r) \ rhs }" + +definition rhs_eq_cls :: "t_equa_rhs \ (string set) set" +where + "rhs_eq_cls rhs \ {Y. \ r. (Y, r) \ rhs}" + +definition ardenable :: "t_equa \ bool" +where + "ardenable equa \ let (X, rhs) = equa in + distinct_rhs rhs \ no_EMPTY_rhs rhs \ X = L rhs" + +text {* + Inv: Invairance of the equation-system, during the decrease of the equation-system, Inv holds. +*} +definition Inv :: "string set \ t_equas \ bool" +where + "Inv X ES \ finite ES \ (\ rhs. (X, rhs) \ ES) \ distinct_equas ES \ + (\ X xrhs. (X, xrhs) \ ES \ ardenable (X, xrhs) \ X \ {} \ rhs_eq_cls xrhs \ insert {[]} (left_eq_cls ES))" + +text {* + TCon: Termination Condition of the equation-system decreasion. +*} +definition TCon:: "'a set \ bool" +where + "TCon ES \ 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 P and Q to prove the myhill-nerode theorem + +*} +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 + + +(* ********************************* BEGIN: proving the initial equation-system satisfies Inv **************************************** *) + +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 L_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 L_def lang_seq_def) + thus ?thesis using X_in_equas False empty "(1)" + unfolding equations_def equation_rhs_def by (auto simp:L_def) + 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 L_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:L_def 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 simp:L_def) + 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 L_def lang_seq_def) + qed + qed +qed + +lemma finite_CT_chars: + "finite {CHAR c |c. Xa-c\X}" +by (auto) + +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_CT_chars)+ +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) + +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 + + +(* ********************************* END: proving the initial equation-system satisfies Inv **************************************** *) + + +(* ***************************** 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 del_x_paired_del_only_x': + "(X, rhs) \ del_x_paired ES Y \ X \ Y \ (X, rhs) \ ES" +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) + 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:Inv_def) + 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 + +(* ****************************** 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 L_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 add:L_def) + 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:L_def 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:L_def lang_seq_def) + thus ?thesis by blast + qed +qed + +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 + +end + + + + + + + + + +