diff -r b04cc5e4e84c -r 7743d2ad71d1 Myhill_2.thy --- a/Myhill_2.thy Tue May 31 20:32:49 2011 +0000 +++ b/Myhill_2.thy Thu Jun 02 16:44:35 2011 +0000 @@ -3,17 +3,22 @@ "~~/src/HOL/Library/List_Prefix" begin -section {* Direction @{text "regular language \finite partition"} *} +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) | x y. tag x = tag y}" + "=tag= \ {(x, y). tag x = tag y}" lemma finite_eq_tag_rel: assumes rng_fnt: "finite (range tag)" @@ -216,7 +221,7 @@ (\x. (\L1 `` {x}, {(\L2 `` {x - xa}) | xa. xa \ x \ xa \ L1}))" lemma Seq_in_cases: - assumes "x @ z \ A ;; B" + assumes "x @ z \ A \ B" shows "(\ x' \ x. x' \ A \ (x - x') @ z \ B) \ (\ z' \ z. (x @ z') \ A \ (z - z') \ B)" using assms @@ -225,12 +230,12 @@ 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" + shows "x \(A \ B) y" proof - { fix x y z - assume xz_in_seq: "x @ z \ A ;; B" + 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" + have"y @ z \ A \ B" proof - { (* first case with x' in A and (x - x') @ z in B *) fix x' @@ -259,7 +264,7 @@ 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) + 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 *) @@ -269,25 +274,25 @@ 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" + with h1 h3 have "y @ z \ A \ B" unfolding prefix_def Seq_def by (auto) (metis append_assoc) } - ultimately show "y @ z \ A ;; B" + 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 + 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))" + 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" + 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)))" @@ -431,7 +436,7 @@ 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) + 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]] @@ -439,15 +444,15 @@ qed lemma quot_star_finiteI [intro]: - assumes finite1: "finite (UNIV // \L1)" - shows "finite (UNIV // \(L1\))" -proof (rule_tac tag = "tag_str_STAR L1" in tag_finite_imageD) - show "\x y. tag_str_STAR L1 x = tag_str_STAR L1 y \ x \(L1\) y" + 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 // \L1))" + have *: "finite (Pow (UNIV // \A))" using finite1 by auto - show "finite (range (tag_str_STAR L1))" + show "finite (range (tag_str_STAR A))" unfolding tag_str_STAR_def apply(rule finite_subset[OF _ *]) unfolding quotient_def @@ -457,13 +462,12 @@ subsubsection{* The conclusion *} lemma Myhill_Nerode2: - fixes r::"rexp" - shows "finite (UNIV // \(L r))" + shows "finite (UNIV // \(L_rexp r))" by (induct r) (auto) theorem Myhill_Nerode: - shows "(\r::rexp. A = L r) \ finite (UNIV // \A)" + shows "(\r. A = L_rexp r) \ finite (UNIV // \A)" using Myhill_Nerode1 Myhill_Nerode2 by auto end