diff -r 09e6f3719cbc -r 5d724fe0e096 Myhill_1.thy --- a/Myhill_1.thy Fri Aug 19 20:39:07 2011 +0000 +++ b/Myhill_1.thy Mon Aug 22 12:49:27 2011 +0000 @@ -1,9 +1,13 @@ theory Myhill_1 -imports More_Regular_Set +imports "Folds" "~~/src/HOL/Library/While_Combinator" begin -section {* Direction @{text "finite partition \ regular language"} *} +section {* First direction of MN: @{text "finite partition \ regular language"} *} + +notation + conc (infixr "\" 100) and + star ("_\" [101] 102) lemma Pair_Collect [simp]: shows "(x, y) \ {(x, y). P x y} \ P x y" @@ -11,7 +15,6 @@ text {* Myhill-Nerode relation *} - definition str_eq :: "'a lang \ ('a list \ 'a list) set" ("\_" [100] 100) where @@ -39,7 +42,7 @@ unfolding finals_def quotient_def by auto -section {* Equational systems *} +subsection {* Equational systems *} text {* The two kinds of terms in the rhs of equations. *} @@ -87,7 +90,7 @@ "Init CS \ {(X, Init_rhs CS X) | X. X \ CS}" -section {* Arden Operation on equations *} +subsection {* Arden Operation on equations *} fun Append_rexp :: "'a rexp \ 'a trm \ 'a trm" @@ -104,7 +107,7 @@ Append_rexp_rhs (rhs - {Trn X r | r. Trn X r \ rhs}) (Star (\ {r. Trn X r \ rhs}))" -section {* Substitution Operation on equations *} +subsection {* Substitution Operation on equations *} definition "Subst rhs X xrhs \ @@ -120,7 +123,7 @@ Subst_all (ES - {(X, xrhs)}) X (Arden X xrhs)" -section {* While-combinator *} +subsection {* While-combinator and invariants *} definition "Iter X ES \ (let (Y, yrhs) = SOME (Y, yrhs). (Y, yrhs) \ ES \ X \ Y @@ -141,8 +144,6 @@ "Solve X ES \ while Cond (Iter X) ES" -section {* Invariants *} - definition "distinctness ES \ \ X rhs rhs'. (X, rhs) \ ES \ (X, rhs') \ ES \ rhs = rhs'" @@ -197,8 +198,6 @@ using assms by (simp add: invariant_def) -subsection {* The proof of this direction *} - lemma finite_Trn: assumes fin: "finite rhs" shows "finite {r. Trn Y r \ rhs}" @@ -247,7 +246,7 @@ by (auto simp add: conc_def lang_of_append_rexp) -subsubsection {* Intial Equational System *} +subsection {* Intial Equational Systems *} lemma defined_by_str: assumes "s \ X" "X \ UNIV // \A" @@ -353,7 +352,7 @@ by auto qed -subsubsection {* Interation step *} +subsection {* Interations *} lemma Arden_preserves_soundness: assumes l_eq_r: "X = lang_rhs rhs" @@ -377,7 +376,7 @@ by blast finally have "X = X \ A \ B" . then have "X = B \ A\" - by (simp add: arden[OF not_empty2]) + by (simp add: reversed_Arden[OF not_empty2]) also have "\ = lang_rhs (Arden X rhs)" unfolding Arden_def A_def B_def b_def by (simp only: lang_of_append_rexp_rhs lang.simps) @@ -676,7 +675,7 @@ qed -subsubsection {* Conclusion of the proof *} +subsection {* The conclusion of the first direction *} lemma Solve: fixes A::"('a::finite) lang"