diff -r 54aa3b6dd71c -r 2409827d8eb8 Myhill_1.thy --- a/Myhill_1.thy Sun Feb 13 10:36:53 2011 +0000 +++ b/Myhill_1.thy Mon Feb 14 07:42:16 2011 +0000 @@ -487,14 +487,6 @@ definition "Solve X ES \ while Test (Iter X) ES" - -(* test *) -partial_function (tailrec) - solve -where - "solve X ES = (if (card ES = 1) then ES else solve X (Iter X ES))" - - text {* Since the @{text "while"} combinator from HOL library is used to implement @{text "Solve X ES"}, the induction principle @{thm [source] while_rule} is used to proved the desired properties @@ -666,19 +658,22 @@ *} lemma defined_by_str: - "\s \ X; X \ UNIV // (\Lang)\ \ X = (\Lang) `` {s}" -by (auto simp:quotient_def Image_def str_eq_rel_def) + assumes "s \ X" "X \ UNIV // \A" + shows "X = \A `` {s}" +using assms +unfolding quotient_def Image_def str_eq_rel_def +by auto lemma every_eqclass_has_transition: assumes has_str: "s @ [c] \ X" - and in_CS: "X \ UNIV // (\Lang)" - obtains Y where "Y \ UNIV // (\Lang)" and "Y ;; {[c]} \ X" and "s \ Y" + and in_CS: "X \ UNIV // \A" + obtains Y where "Y \ UNIV // \A" and "Y ;; {[c]} \ X" and "s \ Y" proof - - def Y \ "(\Lang) `` {s}" - have "Y \ UNIV // (\Lang)" + def Y \ "\A `` {s}" + have "Y \ UNIV // \A" unfolding Y_def quotient_def by auto moreover - have "X = (\Lang) `` {s @ [c]}" + have "X = \A `` {s @ [c]}" using has_str in_CS defined_by_str by blast then have "Y ;; {[c]} \ X" unfolding Y_def Image_def Seq_def @@ -687,18 +682,18 @@ moreover have "s \ Y" unfolding Y_def unfolding Image_def str_eq_rel_def by simp - ultimately show thesis by (blast intro: that) + ultimately show thesis using that by blast qed lemma l_eq_r_in_eqs: - assumes X_in_eqs: "(X, xrhs) \ (Init (UNIV // (\Lang)))" - shows "X = L xrhs" + assumes X_in_eqs: "(X, rhs) \ Init (UNIV // \A)" + shows "X = L rhs" proof - show "X \ L xrhs" + show "X \ L rhs" proof fix x assume "(1)": "x \ X" - show "x \ L xrhs" + show "x \ L rhs" proof (cases "x = []") assume empty: "x = []" thus ?thesis using X_in_eqs "(1)" @@ -707,14 +702,14 @@ assume not_empty: "x \ []" then obtain clist c where decom: "x = clist @ [c]" by (case_tac x rule:rev_cases, auto) - have "X \ UNIV // (\Lang)" using X_in_eqs by (auto simp:Init_def) + have "X \ UNIV // \A" using X_in_eqs by (auto simp:Init_def) then obtain Y - where "Y \ UNIV // (\Lang)" + where "Y \ UNIV // \A" and "Y ;; {[c]} \ X" and "clist \ Y" using decom "(1)" every_eqclass_has_transition by blast hence - "x \ L {Trn Y (CHAR c)| Y c. Y \ UNIV // (\Lang) \ Y \c\ X}" + "x \ L {Trn Y (CHAR c)| Y c. Y \ UNIV // \A \ Y \c\ X}" unfolding transition_def using "(1)" decom by (simp, rule_tac x = "Trn Y (CHAR c)" in exI, simp add:Seq_def) @@ -723,10 +718,17 @@ qed qed next - show "L xrhs \ X" using X_in_eqs + show "L rhs \ X" using X_in_eqs by (auto simp:Init_def Init_rhs_def transition_def) qed +lemma test: + assumes X_in_eqs: "(X, rhs) \ Init (UNIV // \A)" + shows "X = \ (L ` rhs)" +using assms +by (drule_tac l_eq_r_in_eqs) (simp) + + lemma finite_Init_rhs: assumes finite: "finite CS" shows "finite (Init_rhs CS X)"