diff -r b794db0b79db -r b1258b7d2789 Attic/old/Myhill_2.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Attic/old/Myhill_2.thy Mon Jul 25 13:33:38 2011 +0000 @@ -0,0 +1,473 @@ +theory Myhill_2 + imports Myhill_1 Prefix_subtract + "~~/src/HOL/Library/List_Prefix" +begin + +section {* Direction @{text "regular language \ finite partition"} *} + +definition + str_eq :: "string \ lang \ string \ bool" ("_ \_ _") +where + "x \A y \ (x, y) \ (\A)" + +lemma str_eq_def2: + shows "\A = {(x, y) | x y. x \A y}" +unfolding str_eq_def +by simp + +definition + tag_eq_rel :: "(string \ 'b) \ (string \ string) set" ("=_=") +where + "=tag= \ {(x, y). tag x = tag y}" + +lemma finite_eq_tag_rel: + assumes rng_fnt: "finite (range tag)" + shows "finite (UNIV // =tag=)" +proof - + let "?f" = "\X. tag ` X" and ?A = "(UNIV // =tag=)" + have "finite (?f ` ?A)" + proof - + have "range ?f \ (Pow (range tag))" unfolding Pow_def by auto + moreover + have "finite (Pow (range tag))" using rng_fnt by simp + ultimately + have "finite (range ?f)" unfolding image_def by (blast intro: finite_subset) + moreover + have "?f ` ?A \ range ?f" by auto + ultimately show "finite (?f ` ?A)" by (rule rev_finite_subset) + qed + moreover + have "inj_on ?f ?A" + proof - + { fix X Y + assume X_in: "X \ ?A" + and Y_in: "Y \ ?A" + and tag_eq: "?f X = ?f Y" + then obtain x y + where "x \ X" "y \ Y" "tag x = tag y" + unfolding quotient_def Image_def image_def tag_eq_rel_def + by (simp) (blast) + with X_in Y_in + have "X = Y" + unfolding quotient_def tag_eq_rel_def by auto + } + then show "inj_on ?f ?A" unfolding inj_on_def by auto + qed + ultimately show "finite (UNIV // =tag=)" by (rule finite_imageD) +qed + +lemma refined_partition_finite: + assumes fnt: "finite (UNIV // R1)" + and refined: "R1 \ R2" + and eq1: "equiv UNIV R1" and eq2: "equiv UNIV R2" + shows "finite (UNIV // R2)" +proof - + let ?f = "\X. {R1 `` {x} | x. x \ X}" + and ?A = "UNIV // R2" and ?B = "UNIV // R1" + have "?f ` ?A \ Pow ?B" + unfolding image_def Pow_def quotient_def by auto + moreover + have "finite (Pow ?B)" using fnt by simp + ultimately + have "finite (?f ` ?A)" by (rule finite_subset) + moreover + have "inj_on ?f ?A" + proof - + { fix X Y + assume X_in: "X \ ?A" and Y_in: "Y \ ?A" and eq_f: "?f X = ?f Y" + from quotientE [OF X_in] + obtain x where "X = R2 `` {x}" by blast + with equiv_class_self[OF eq2] have x_in: "x \ X" by simp + then have "R1 ``{x} \ ?f X" by auto + with eq_f have "R1 `` {x} \ ?f Y" by simp + then obtain y + where y_in: "y \ Y" and eq_r1_xy: "R1 `` {x} = R1 `` {y}" by auto + with eq_equiv_class[OF _ eq1] + have "(x, y) \ R1" by blast + with refined have "(x, y) \ R2" by auto + with quotient_eqI [OF eq2 X_in Y_in x_in y_in] + have "X = Y" . + } + then show "inj_on ?f ?A" unfolding inj_on_def by blast + qed + ultimately show "finite (UNIV // R2)" by (rule finite_imageD) +qed + +lemma tag_finite_imageD: + assumes rng_fnt: "finite (range tag)" + and same_tag_eqvt: "\m n. tag m = tag n \ m \A n" + shows "finite (UNIV // \A)" +proof (rule_tac refined_partition_finite [of "=tag="]) + show "finite (UNIV // =tag=)" by (rule finite_eq_tag_rel[OF rng_fnt]) +next + from same_tag_eqvt + show "=tag= \ \A" unfolding tag_eq_rel_def str_eq_def + by auto +next + show "equiv UNIV =tag=" + unfolding equiv_def tag_eq_rel_def refl_on_def sym_def trans_def + by auto +next + show "equiv UNIV (\A)" + unfolding equiv_def str_eq_rel_def sym_def refl_on_def trans_def + by blast +qed + + +subsection {* The proof *} + +subsubsection {* The base case for @{const "NULL"} *} + +lemma quot_null_eq: + shows "UNIV // \{} = {UNIV}" +unfolding quotient_def Image_def str_eq_rel_def by auto + +lemma quot_null_finiteI [intro]: + shows "finite (UNIV // \{})" +unfolding quot_null_eq by simp + + +subsubsection {* The base case for @{const "EMPTY"} *} + +lemma quot_empty_subset: + shows "UNIV // \{[]} \ {{[]}, UNIV - {[]}}" +proof + fix x + assume "x \ UNIV // \{[]}" + then obtain y where h: "x = {z. (y, z) \ \{[]}}" + unfolding quotient_def Image_def by blast + show "x \ {{[]}, UNIV - {[]}}" + proof (cases "y = []") + case True with h + have "x = {[]}" by (auto simp: str_eq_rel_def) + thus ?thesis by simp + next + case False with h + have "x = UNIV - {[]}" by (auto simp: str_eq_rel_def) + thus ?thesis by simp + qed +qed + +lemma quot_empty_finiteI [intro]: + shows "finite (UNIV // \{[]})" +by (rule finite_subset[OF quot_empty_subset]) (simp) + + +subsubsection {* The base case for @{const "CHAR"} *} + +lemma quot_char_subset: + "UNIV // (\{[c]}) \ {{[]},{[c]}, UNIV - {[], [c]}}" +proof + fix x + assume "x \ UNIV // \{[c]}" + then obtain y where h: "x = {z. (y, z) \ \{[c]}}" + unfolding quotient_def Image_def by blast + show "x \ {{[]},{[c]}, UNIV - {[], [c]}}" + proof - + { assume "y = []" hence "x = {[]}" using h + by (auto simp:str_eq_rel_def) } + moreover + { assume "y = [c]" hence "x = {[c]}" using h + by (auto dest!:spec[where x = "[]"] simp:str_eq_rel_def) } + moreover + { assume "y \ []" and "y \ [c]" + hence "\ z. (y @ z) \ [c]" by (case_tac y, auto) + moreover have "\ p. (p \ [] \ p \ [c]) = (\ q. p @ q \ [c])" + by (case_tac p, auto) + ultimately have "x = UNIV - {[],[c]}" using h + by (auto simp add:str_eq_rel_def) + } + ultimately show ?thesis by blast + qed +qed + +lemma quot_char_finiteI [intro]: + shows "finite (UNIV // \{[c]})" +by (rule finite_subset[OF quot_char_subset]) (simp) + + +subsubsection {* The inductive case for @{const ALT} *} + +definition + tag_str_ALT :: "lang \ lang \ string \ (lang \ lang)" +where + "tag_str_ALT A B \ (\x. (\A `` {x}, \B `` {x}))" + +lemma quot_union_finiteI [intro]: + assumes finite1: "finite (UNIV // \A)" + and finite2: "finite (UNIV // \B)" + shows "finite (UNIV // \(A \ B))" +proof (rule_tac tag = "tag_str_ALT A B" in tag_finite_imageD) + have "finite ((UNIV // \A) \ (UNIV // \B))" + using finite1 finite2 by auto + then show "finite (range (tag_str_ALT A B))" + unfolding tag_str_ALT_def quotient_def + by (rule rev_finite_subset) (auto) +next + show "\x y. tag_str_ALT A B x = tag_str_ALT A B y \ x \(A \ B) y" + unfolding tag_str_ALT_def + unfolding str_eq_def + unfolding str_eq_rel_def + by auto +qed + + +subsubsection {* The inductive case for @{text "SEQ"}*} + +definition + tag_str_SEQ :: "lang \ lang \ string \ (lang \ lang set)" +where + "tag_str_SEQ L1 L2 \ + (\x. (\L1 `` {x}, {(\L2 `` {x - xa}) | xa. xa \ x \ xa \ L1}))" + +lemma Seq_in_cases: + assumes "x @ z \ A \ B" + shows "(\ x' \ x. x' \ A \ (x - x') @ z \ B) \ + (\ z' \ z. (x @ z') \ A \ (z - z') \ B)" +using assms +unfolding Seq_def prefix_def +by (auto simp add: append_eq_append_conv2) + +lemma tag_str_SEQ_injI: + assumes eq_tag: "tag_str_SEQ A B x = tag_str_SEQ A B y" + shows "x \(A \ B) y" +proof - + { fix x y z + assume xz_in_seq: "x @ z \ A \ B" + and tag_xy: "tag_str_SEQ A B x = tag_str_SEQ A B y" + have"y @ z \ A \ B" + proof - + { (* first case with x' in A and (x - x') @ z in B *) + fix x' + assume h1: "x' \ x" and h2: "x' \ A" and h3: "(x - x') @ z \ B" + obtain y' + where "y' \ y" + and "y' \ A" + and "(y - y') @ z \ B" + proof - + have "{\B `` {x - x'} |x'. x' \ x \ x' \ A} = + {\B `` {y - y'} |y'. y' \ y \ y' \ A}" (is "?Left = ?Right") + using tag_xy unfolding tag_str_SEQ_def by simp + moreover + have "\B `` {x - x'} \ ?Left" using h1 h2 by auto + ultimately + have "\B `` {x - x'} \ ?Right" by simp + then obtain y' + where eq_xy': "\B `` {x - x'} = \B `` {y - y'}" + and pref_y': "y' \ y" and y'_in: "y' \ A" + by simp blast + + have "(x - x') \B (y - y')" using eq_xy' + unfolding Image_def str_eq_rel_def str_eq_def by auto + with h3 have "(y - y') @ z \ B" + unfolding str_eq_rel_def str_eq_def by simp + with pref_y' y'_in + show ?thesis using that by blast + qed + then have "y @ z \ A \ B" by (erule_tac prefixE) (auto simp: Seq_def) + } + moreover + { (* second case with x @ z' in A and z - z' in B *) + fix z' + assume h1: "z' \ z" and h2: "(x @ z') \ A" and h3: "z - z' \ B" + have "\A `` {x} = \A `` {y}" + using tag_xy unfolding tag_str_SEQ_def by simp + with h2 have "y @ z' \ A" + unfolding Image_def str_eq_rel_def str_eq_def by auto + with h1 h3 have "y @ z \ A \ B" + unfolding prefix_def Seq_def + by (auto) (metis append_assoc) + } + ultimately show "y @ z \ A \ B" + using Seq_in_cases [OF xz_in_seq] by blast + qed + } + from this [OF _ eq_tag] and this [OF _ eq_tag [THEN sym]] + show "x \(A \ B) y" unfolding str_eq_def str_eq_rel_def by blast +qed + +lemma quot_seq_finiteI [intro]: + fixes L1 L2::"lang" + assumes fin1: "finite (UNIV // \L1)" + and fin2: "finite (UNIV // \L2)" + shows "finite (UNIV // \(L1 \ L2))" +proof (rule_tac tag = "tag_str_SEQ L1 L2" in tag_finite_imageD) + show "\x y. tag_str_SEQ L1 L2 x = tag_str_SEQ L1 L2 y \ x \(L1 \ L2) y" + by (rule tag_str_SEQ_injI) +next + have *: "finite ((UNIV // \L1) \ (Pow (UNIV // \L2)))" + using fin1 fin2 by auto + show "finite (range (tag_str_SEQ L1 L2))" + unfolding tag_str_SEQ_def + apply(rule finite_subset[OF _ *]) + unfolding quotient_def + by auto +qed + + +subsubsection {* The inductive case for @{const "STAR"} *} + +definition + tag_str_STAR :: "lang \ string \ lang set" +where + "tag_str_STAR L1 \ (\x. {\L1 `` {x - xa} | xa. xa < x \ xa \ L1\})" + +text {* A technical lemma. *} +lemma finite_set_has_max: "\finite A; A \ {}\ \ + (\ max \ A. \ a \ A. f a <= (f max :: nat))" +proof (induct rule:finite.induct) + case emptyI thus ?case by simp +next + case (insertI A a) + show ?case + proof (cases "A = {}") + case True thus ?thesis by (rule_tac x = a in bexI, auto) + next + case False + with insertI.hyps and False + obtain max + where h1: "max \ A" + and h2: "\a\A. f a \ f max" by blast + show ?thesis + proof (cases "f a \ f max") + assume "f a \ f max" + with h1 h2 show ?thesis by (rule_tac x = max in bexI, auto) + next + assume "\ (f a \ f max)" + thus ?thesis using h2 by (rule_tac x = a in bexI, auto) + qed + qed +qed + + +text {* The following is a technical lemma, which helps to show the range finiteness of tag function. *} + +lemma finite_strict_prefix_set: "finite {xa. xa < (x::string)}" +apply (induct x rule:rev_induct, simp) +apply (subgoal_tac "{xa. xa < xs @ [x]} = {xa. xa < xs} \ {xs}") +by (auto simp:strict_prefix_def) + + +lemma tag_str_STAR_injI: + assumes eq_tag: "tag_str_STAR L\<^isub>1 v = tag_str_STAR L\<^isub>1 w" + shows "v \(L\<^isub>1\) w" +proof- + { fix x y z + assume xz_in_star: "x @ z \ L\<^isub>1\" + and tag_xy: "tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 y" + have "y @ z \ L\<^isub>1\" + proof(cases "x = []") + case True + with tag_xy have "y = []" + by (auto simp add: tag_str_STAR_def strict_prefix_def) + thus ?thesis using xz_in_star True by simp + next + case False + let ?S = "{xa. xa < x \ xa \ L\<^isub>1\ \ (x - xa) @ z \ L\<^isub>1\}" + have "finite ?S" + by (rule_tac B = "{xa. xa < x}" in finite_subset, + auto simp:finite_strict_prefix_set) + moreover have "?S \ {}" using False xz_in_star + by (simp, rule_tac x = "[]" in exI, auto simp:strict_prefix_def) + ultimately have "\ xa_max \ ?S. \ xa \ ?S. length xa \ length xa_max" + using finite_set_has_max by blast + then obtain xa_max + where h1: "xa_max < x" + and h2: "xa_max \ L\<^isub>1\" + and h3: "(x - xa_max) @ z \ L\<^isub>1\" + and h4:"\ xa < x. xa \ L\<^isub>1\ \ (x - xa) @ z \ L\<^isub>1\ + \ length xa \ length xa_max" + by blast + obtain ya + where h5: "ya < y" and h6: "ya \ L\<^isub>1\" + and eq_xya: "(x - xa_max) \L\<^isub>1 (y - ya)" + proof- + from tag_xy have "{\L\<^isub>1 `` {x - xa} |xa. xa < x \ xa \ L\<^isub>1\} = + {\L\<^isub>1 `` {y - xa} |xa. xa < y \ xa \ L\<^isub>1\}" (is "?left = ?right") + by (auto simp:tag_str_STAR_def) + moreover have "\L\<^isub>1 `` {x - xa_max} \ ?left" using h1 h2 by auto + ultimately have "\L\<^isub>1 `` {x - xa_max} \ ?right" by simp + thus ?thesis using that + apply (simp add:Image_def str_eq_rel_def str_eq_def) by blast + qed + have "(y - ya) @ z \ L\<^isub>1\" + proof- + obtain za zb where eq_zab: "z = za @ zb" + and l_za: "(y - ya)@za \ L\<^isub>1" and ls_zb: "zb \ L\<^isub>1\" + proof - + from h1 have "(x - xa_max) @ z \ []" + by (auto simp:strict_prefix_def elim:prefixE) + from star_decom [OF h3 this] + obtain a b where a_in: "a \ L\<^isub>1" + and a_neq: "a \ []" and b_in: "b \ L\<^isub>1\" + and ab_max: "(x - xa_max) @ z = a @ b" by blast + let ?za = "a - (x - xa_max)" and ?zb = "b" + have pfx: "(x - xa_max) \ a" (is "?P1") + and eq_z: "z = ?za @ ?zb" (is "?P2") + proof - + have "((x - xa_max) \ a \ (a - (x - xa_max)) @ b = z) \ + (a < (x - xa_max) \ ((x - xa_max) - a) @ z = b)" + using append_eq_dest[OF ab_max] by (auto simp:strict_prefix_def) + moreover { + assume np: "a < (x - xa_max)" + and b_eqs: "((x - xa_max) - a) @ z = b" + have "False" + proof - + let ?xa_max' = "xa_max @ a" + have "?xa_max' < x" + using np h1 by (clarsimp simp:strict_prefix_def diff_prefix) + moreover have "?xa_max' \ L\<^isub>1\" + using a_in h2 by (simp add:star_intro3) + moreover have "(x - ?xa_max') @ z \ L\<^isub>1\" + using b_eqs b_in np h1 by (simp add:diff_diff_append) + moreover have "\ (length ?xa_max' \ length xa_max)" + using a_neq by simp + ultimately show ?thesis using h4 by blast + qed } + ultimately show ?P1 and ?P2 by auto + qed + hence "(x - xa_max)@?za \ L\<^isub>1" using a_in by (auto elim:prefixE) + with eq_xya have "(y - ya) @ ?za \ L\<^isub>1" + by (auto simp:str_eq_def str_eq_rel_def) + with eq_z and b_in + show ?thesis using that by blast + qed + have "((y - ya) @ za) @ zb \ L\<^isub>1\" using l_za ls_zb by blast + with eq_zab show ?thesis by simp + qed + with h5 h6 show ?thesis + by (drule_tac star_intro1) (auto simp:strict_prefix_def elim: prefixE) + qed + } + from this [OF _ eq_tag] and this [OF _ eq_tag [THEN sym]] + show ?thesis unfolding str_eq_def str_eq_rel_def by blast +qed + +lemma quot_star_finiteI [intro]: + assumes finite1: "finite (UNIV // \A)" + shows "finite (UNIV // \(A\))" +proof (rule_tac tag = "tag_str_STAR A" in tag_finite_imageD) + show "\x y. tag_str_STAR A x = tag_str_STAR A y \ x \(A\) y" + by (rule tag_str_STAR_injI) +next + have *: "finite (Pow (UNIV // \A))" + using finite1 by auto + show "finite (range (tag_str_STAR A))" + unfolding tag_str_STAR_def + apply(rule finite_subset[OF _ *]) + unfolding quotient_def + by auto +qed + +subsubsection{* The conclusion *} + +lemma Myhill_Nerode2: + shows "finite (UNIV // \(L_rexp r))" +by (induct r) (auto) + + +theorem Myhill_Nerode: + shows "(\r. A = L_rexp r) \ finite (UNIV // \A)" +using Myhill_Nerode1 Myhill_Nerode2 by auto + +end