diff -r b04cc5e4e84c -r 7743d2ad71d1 Myhill_1.thy --- a/Myhill_1.thy Tue May 31 20:32:49 2011 +0000 +++ b/Myhill_1.thy Thu Jun 02 16:44:35 2011 +0000 @@ -1,5 +1,5 @@ theory Myhill_1 -imports Main Folds Regular +imports Regular "~~/src/HOL/Library/While_Combinator" begin @@ -37,43 +37,37 @@ text {* The two kinds of terms in the rhs of equations. *} -datatype rhs_trm = +datatype trm = Lam "rexp" (* Lambda-marker *) | Trn "lang" "rexp" (* Transition *) +fun + L_trm::"trm \ lang" +where + "L_trm (Lam r) = L_rexp r" +| "L_trm (Trn X r) = X \ L_rexp r" -overloading L_rhs_trm \ "L:: rhs_trm \ lang" -begin - fun L_rhs_trm:: "rhs_trm \ lang" - where - "L_rhs_trm (Lam r) = L r" - | "L_rhs_trm (Trn X r) = X ;; L r" -end - -overloading L_rhs \ "L:: rhs_trm set \ lang" -begin - fun L_rhs:: "rhs_trm set \ lang" - where - "L_rhs rhs = \ (L ` rhs)" -end +fun + L_rhs::"trm set \ lang" +where + "L_rhs rhs = \ (L_trm ` rhs)" lemma L_rhs_set: - shows "L {Trn X r | r. P r} = \{L (Trn X r) | r. P r}" -by (auto simp del: L_rhs_trm.simps) + shows "L_rhs {Trn X r | r. P r} = \{L_trm (Trn X r) | r. P r}" +by (auto) lemma L_rhs_union_distrib: - fixes A B::"rhs_trm set" - shows "L A \ L B = L (A \ B)" + fixes A B::"trm set" + shows "L_rhs A \ L_rhs B = L_rhs (A \ B)" by simp - text {* Transitions between equivalence classes *} definition transition :: "lang \ char \ lang \ bool" ("_ \_\_" [100,100,100] 100) where - "Y \c\ X \ Y ;; {[c]} \ X" + "Y \c\ X \ Y \ {[c]} \ X" text {* Initial equational system *} @@ -91,7 +85,7 @@ section {* Arden Operation on equations *} fun - Append_rexp :: "rexp \ rhs_trm \ rhs_trm" + Append_rexp :: "rexp \ trm \ trm" where "Append_rexp r (Lam rexp) = Lam (SEQ rexp r)" | "Append_rexp r (Trn X rexp) = Trn X (SEQ rexp r)" @@ -112,7 +106,7 @@ (rhs - {Trn X r | r. Trn X r \ rhs}) \ (Append_rexp_rhs xrhs (\ {r. Trn X r \ rhs}))" definition - Subst_all :: "(lang \ rhs_trm set) set \ lang \ rhs_trm set \ (lang \ rhs_trm set) set" + Subst_all :: "(lang \ trm set) set \ lang \ trm set \ (lang \ trm set) set" where "Subst_all ES X xrhs \ {(Y, Subst yrhs X xrhs) | Y yrhs. (Y, yrhs) \ ES}" @@ -149,10 +143,10 @@ \ X rhs rhs'. (X, rhs) \ ES \ (X, rhs') \ ES \ rhs = rhs'" definition - "soundness ES \ \(X, rhs) \ ES. X = L rhs" + "soundness ES \ \(X, rhs) \ ES. X = L_rhs rhs" definition - "ardenable rhs \ (\ Y r. Trn Y r \ rhs \ [] \ L r)" + "ardenable rhs \ (\ Y r. Trn Y r \ rhs \ [] \ L_rexp r)" definition "ardenable_all ES \ \(X, rhs) \ ES. ardenable rhs" @@ -227,29 +221,28 @@ done qed -lemma rhs_trm_soundness: +lemma trm_soundness: assumes finite:"finite rhs" - shows "L ({Trn X r| r. Trn X r \ rhs}) = X ;; (L (\{r. Trn X r \ rhs}))" + shows "L_rhs ({Trn X r| r. Trn X r \ rhs}) = X \ (L_rexp (\{r. Trn X r \ rhs}))" proof - have "finite {r. Trn X r \ rhs}" by (rule finite_Trn[OF finite]) - then show "L ({Trn X r| r. Trn X r \ rhs}) = X ;; (L (\{r. Trn X r \ rhs}))" - by (simp only: L_rhs_set L_rhs_trm.simps) (auto simp add: Seq_def) + then show "L_rhs ({Trn X r| r. Trn X r \ rhs}) = X \ (L_rexp (\{r. Trn X r \ rhs}))" + by (simp only: L_rhs_set L_trm.simps) (auto simp add: Seq_def) qed lemma lang_of_append_rexp: - "L (Append_rexp r rhs_trm) = L rhs_trm ;; L r" + "L_trm (Append_rexp r trm) = L_trm trm \ L_rexp r" by (induct rule: Append_rexp.induct) (auto simp add: seq_assoc) lemma lang_of_append_rexp_rhs: - "L (Append_rexp_rhs rhs r) = L rhs ;; L r" + "L_rhs (Append_rexp_rhs rhs r) = L_rhs rhs \ L_rexp r" unfolding Append_rexp_rhs_def by (auto simp add: Seq_def lang_of_append_rexp) - -subsubsection {* Intialization *} +subsubsection {* Intial Equational System *} lemma defined_by_str: assumes "s \ X" "X \ UNIV // \A" @@ -261,7 +254,7 @@ lemma every_eqclass_has_transition: assumes has_str: "s @ [c] \ X" and in_CS: "X \ UNIV // \A" - obtains Y where "Y \ UNIV // \A" and "Y ;; {[c]} \ X" and "s \ Y" + obtains Y where "Y \ UNIV // \A" and "Y \ {[c]} \ X" and "s \ Y" proof - def Y \ "\A `` {s}" have "Y \ UNIV // \A" @@ -269,7 +262,7 @@ moreover have "X = \A `` {s @ [c]}" using has_str in_CS defined_by_str by blast - then have "Y ;; {[c]} \ X" + then have "Y \ {[c]} \ X" unfolding Y_def Image_def Seq_def unfolding str_eq_rel_def by clarsimp @@ -281,14 +274,14 @@ lemma l_eq_r_in_eqs: assumes X_in_eqs: "(X, rhs) \ Init (UNIV // \A)" - shows "X = L rhs" + shows "X = L_rhs rhs" proof - show "X \ L rhs" + show "X \ L_rhs rhs" proof fix x assume in_X: "x \ X" { assume empty: "x = []" - then have "x \ L rhs" using X_in_eqs in_X + then have "x \ L_rhs rhs" using X_in_eqs in_X unfolding Init_def Init_rhs_def by auto } @@ -297,40 +290,40 @@ then obtain s c where decom: "x = s @ [c]" using rev_cases by blast have "X \ UNIV // \A" using X_in_eqs unfolding Init_def by auto - then obtain Y where "Y \ UNIV // \A" "Y ;; {[c]} \ X" "s \ Y" + then obtain Y where "Y \ UNIV // \A" "Y \ {[c]} \ X" "s \ Y" using decom in_X every_eqclass_has_transition by blast - then have "x \ L {Trn Y (CHAR c)| Y c. Y \ UNIV // \A \ Y \c\ X}" + then have "x \ L_rhs {Trn Y (CHAR c)| Y c. Y \ UNIV // \A \ Y \c\ X}" unfolding transition_def using decom by (force simp add: Seq_def) - then have "x \ L rhs" using X_in_eqs in_X + then have "x \ L_rhs rhs" using X_in_eqs in_X unfolding Init_def Init_rhs_def by simp } - ultimately show "x \ L rhs" by blast + ultimately show "x \ L_rhs rhs" by blast qed next - show "L rhs \ X" using X_in_eqs + show "L_rhs rhs \ X" using X_in_eqs unfolding Init_def Init_rhs_def transition_def by auto qed lemma test: assumes X_in_eqs: "(X, rhs) \ Init (UNIV // \A)" - shows "X = \ (L ` rhs)" + shows "X = \ (L_trm ` rhs)" using assms l_eq_r_in_eqs by (simp) lemma finite_Init_rhs: assumes finite: "finite CS" shows "finite (Init_rhs CS X)" proof- - def S \ "{(Y, c)| Y c. Y \ CS \ Y ;; {[c]} \ X}" + def S \ "{(Y, c)| Y c. Y \ CS \ Y \ {[c]} \ X}" def h \ "\ (Y, c). Trn Y (CHAR c)" have "finite (CS \ (UNIV::char set))" using finite by auto then have "finite S" using S_def by (rule_tac B = "CS \ UNIV" in finite_subset) (auto) - moreover have "{Trn Y (CHAR c) |Y c. Y \ CS \ Y ;; {[c]} \ X} = h ` S" + moreover have "{Trn Y (CHAR c) |Y c. Y \ CS \ Y \ {[c]} \ X} = h ` S" unfolding S_def h_def image_def by auto ultimately - have "finite {Trn Y (CHAR c) |Y c. Y \ CS \ Y ;; {[c]} \ X}" by auto + have "finite {Trn Y (CHAR c) |Y c. Y \ CS \ Y \ {[c]} \ X}" by auto then show "finite (Init_rhs CS X)" unfolding Init_rhs_def transition_def by simp qed @@ -359,32 +352,32 @@ subsubsection {* Interation step *} lemma Arden_keeps_eq: - assumes l_eq_r: "X = L rhs" + assumes l_eq_r: "X = L_rhs rhs" and not_empty: "ardenable rhs" and finite: "finite rhs" - shows "X = L (Arden X rhs)" + shows "X = L_rhs (Arden X rhs)" proof - - def A \ "L (\{r. Trn X r \ rhs})" + def A \ "L_rexp (\{r. Trn X r \ rhs})" def b \ "{Trn X r | r. Trn X r \ rhs}" - def B \ "L (rhs - b)" + def B \ "L_rhs (rhs - b)" have not_empty2: "[] \ A" using finite_Trn[OF finite] not_empty unfolding A_def ardenable_def by simp - have "X = L rhs" using l_eq_r by simp - also have "\ = L (b \ (rhs - b))" unfolding b_def by auto - also have "\ = L b \ B" unfolding B_def by (simp only: L_rhs_union_distrib) - also have "\ = X ;; A \ B" + have "X = L_rhs rhs" using l_eq_r by simp + also have "\ = L_rhs (b \ (rhs - b))" unfolding b_def by auto + also have "\ = L_rhs b \ B" unfolding B_def by (simp only: L_rhs_union_distrib) + also have "\ = X \ A \ B" unfolding b_def - unfolding rhs_trm_soundness[OF finite] + unfolding trm_soundness[OF finite] unfolding A_def by blast - finally have "X = X ;; A \ B" . - then have "X = B ;; A\" + finally have "X = X \ A \ B" . + then have "X = B \ A\" by (simp add: arden[OF not_empty2]) - also have "\ = L (Arden X rhs)" + also have "\ = L_rhs (Arden X rhs)" unfolding Arden_def A_def B_def b_def by (simp only: lang_of_append_rexp_rhs L_rexp.simps) - finally show "X = L (Arden X rhs)" by simp + finally show "X = L_rhs (Arden X rhs)" by simp qed lemma Append_keeps_finite: @@ -418,16 +411,16 @@ by (simp only: Subst_def Append_keeps_nonempty nonempty_set_union nonempty_set_sub) lemma Subst_keeps_eq: - assumes substor: "X = L xrhs" + assumes substor: "X = L_rhs xrhs" and finite: "finite rhs" - shows "L (Subst rhs X xrhs) = L rhs" (is "?Left = ?Right") + shows "L_rhs (Subst rhs X xrhs) = L_rhs rhs" (is "?Left = ?Right") proof- - def A \ "L (rhs - {Trn X r | r. Trn X r \ rhs})" - have "?Left = A \ L (Append_rexp_rhs xrhs (\{r. Trn X r \ rhs}))" + def A \ "L_rhs (rhs - {Trn X r | r. Trn X r \ rhs})" + have "?Left = A \ L_rhs (Append_rexp_rhs xrhs (\{r. Trn X r \ rhs}))" unfolding Subst_def unfolding L_rhs_union_distrib[symmetric] by (simp add: A_def) - moreover have "?Right = A \ L ({Trn X r | r. Trn X r \ rhs})" + moreover have "?Right = A \ L_rhs {Trn X r | r. Trn X r \ rhs}" proof- have "rhs = (rhs - {Trn X r | r. Trn X r \ rhs}) \ ({Trn X r | r. Trn X r \ rhs})" by auto thus ?thesis @@ -435,8 +428,8 @@ unfolding L_rhs_union_distrib by simp qed - moreover have "L (Append_rexp_rhs xrhs (\{r. Trn X r \ rhs})) = L ({Trn X r | r. Trn X r \ rhs})" - using finite substor by (simp only: lang_of_append_rexp_rhs rhs_trm_soundness) + moreover have "L_rhs (Append_rexp_rhs xrhs (\{r. Trn X r \ rhs})) = L_rhs {Trn X r | r. Trn X r \ rhs}" + using finite substor by (simp only: lang_of_append_rexp_rhs trm_soundness) ultimately show ?thesis by simp qed @@ -519,7 +512,7 @@ assumes invariant_ES: "invariant (ES \ {(Y, yrhs)})" shows "invariant (Subst_all ES Y (Arden Y yrhs))" proof (rule invariantI) - have Y_eq_yrhs: "Y = L yrhs" + have Y_eq_yrhs: "Y = L_rhs yrhs" using invariant_ES by (simp only:invariant_def soundness_def, blast) have finite_yrhs: "finite yrhs" using invariant_ES by (auto simp:invariant_def finite_rhs_def) @@ -527,7 +520,7 @@ using invariant_ES by (auto simp:invariant_def ardenable_all_def) show "soundness (Subst_all ES Y (Arden Y yrhs))" proof - - have "Y = L (Arden Y yrhs)" + have "Y = L_rhs (Arden Y yrhs)" using Y_eq_yrhs invariant_ES finite_yrhs using finite_Trn[OF finite_yrhs] apply(rule_tac Arden_keeps_eq) @@ -723,7 +716,7 @@ lemma every_eqcl_has_reg: assumes finite_CS: "finite (UNIV // \A)" and X_in_CS: "X \ (UNIV // \A)" - shows "\r::rexp. X = L r" + shows "\r. X = L_rexp r" proof - from finite_CS X_in_CS obtain xrhs where Inv_ES: "invariant {(X, xrhs)}" @@ -742,14 +735,14 @@ using Arden_keeps_finite by auto then have fin: "finite {r. Lam r \ A}" by (rule finite_Lam) - have "X = L xrhs" using Inv_ES unfolding invariant_def soundness_def + have "X = L_rhs xrhs" using Inv_ES unfolding invariant_def soundness_def by simp - then have "X = L A" using Inv_ES + then have "X = L_rhs A" using Inv_ES unfolding A_def invariant_def ardenable_all_def finite_rhs_def by (rule_tac Arden_keeps_eq) (simp_all add: finite_Trn) - then have "X = L {Lam r | r. Lam r \ A}" using eq by simp - then have "X = L (\{r. Lam r \ A})" using fin by auto - then show "\r::rexp. X = L r" by blast + then have "X = L_rhs {Lam r | r. Lam r \ A}" using eq by simp + then have "X = L_rexp (\{r. Lam r \ A})" using fin by auto + then show "\r. X = L_rexp r" by blast qed lemma bchoice_finite_set: @@ -764,19 +757,19 @@ theorem Myhill_Nerode1: assumes finite_CS: "finite (UNIV // \A)" - shows "\r::rexp. A = L r" + shows "\r. A = L_rexp r" proof - have fin: "finite (finals A)" using finals_in_partitions finite_CS by (rule finite_subset) - have "\X \ (UNIV // \A). \r::rexp. X = L r" + have "\X \ (UNIV // \A). \r. X = L_rexp r" using finite_CS every_eqcl_has_reg by blast - then have a: "\X \ finals A. \r::rexp. X = L r" + then have a: "\X \ finals A. \r. X = L_rexp r" using finals_in_partitions by auto - then obtain rs::"rexp set" where "\ (finals A) = \(L ` rs)" "finite rs" + then obtain rs::"rexp set" where "\ (finals A) = \(L_rexp ` rs)" "finite rs" using fin by (auto dest: bchoice_finite_set) - then have "A = L (\rs)" + then have "A = L_rexp (\rs)" unfolding lang_is_union_of_finals[symmetric] by simp - then show "\r::rexp. A = L r" by blast + then show "\r. A = L_rexp r" by blast qed