--- a/Myhill.thy Fri Dec 31 13:47:53 2010 +0000
+++ b/Myhill.thy Fri Jan 07 14:25:23 2011 +0000
@@ -3,8 +3,7 @@
begin
text {* sequential composition of languages *}
-definition
- Seq :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ;; _" [100,100] 100)
+definition Seq :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ;; _" [100,100] 100)
where
"L1 ;; L2 = {s1 @ s2 | s1 s2. s1 \<in> L1 \<and> s2 \<in> L2}"
@@ -19,25 +18,22 @@
"(A \<union> B) ;; C = (A ;; C) \<union> (B ;; C)"
by (auto simp:Seq_def)
+lemma seq_intro:
+ "\<lbrakk>x \<in> A; y \<in> B\<rbrakk> \<Longrightarrow> x @ y \<in> A ;; B "
+by (auto simp:Seq_def)
+
lemma seq_assoc:
"(A ;; B) ;; C = A ;; (B ;; C)"
-unfolding Seq_def
-apply(auto)
-apply(metis)
+apply(auto simp:Seq_def)
+apply blast
by (metis append_assoc)
-lemma union_seq:
- "\<Union> {f x y ;; z| x y. P x y } = (\<Union> {f x y|x y. P x y });; z"
-apply (auto simp add:Seq_def)
-apply metis
-done
-
theorem ardens_revised:
assumes nemp: "[] \<notin> A"
shows "(X = X ;; A \<union> B) \<longleftrightarrow> (X = B ;; A\<star>)"
proof
assume eq: "X = B ;; A\<star>"
- have "A\<star> = {[]} \<union> A\<star> ;; A" sorry
+ have "A\<star> = {[]} \<union> A\<star> ;; A" sorry
then have "B ;; A\<star> = B ;; ({[]} \<union> A\<star> ;; A)" unfolding Seq_def by simp
also have "\<dots> = B \<union> B ;; (A\<star> ;; A)" unfolding Seq_def by auto
also have "\<dots> = B \<union> (B ;; A\<star>) ;; A" unfolding Seq_def
@@ -46,7 +42,7 @@
next
assume "X = X ;; A \<union> B"
then have "B \<subseteq> X" "X ;; A \<subseteq> X" by auto
- show "X = B ;; A\<star>" sorry
+ thus "X = B ;; A\<star>" sorry
qed
datatype rexp =
@@ -778,286 +774,71 @@
"UNIV // (\<approx>{[]}) \<subseteq> {{[]}, UNIV - {[]}}"
proof
fix x
- assume h: "x \<in> UNIV // \<approx>{[]}"
- show "x \<in> {{[]}, UNIV - {[]}}"
-
-
- have "\<And> s. s \<approx>{[]} [] \<Longrightarrow> s = []"
- apply (auto simp add:str_eq_def)
- apply blast
-
- hence "False"
- apply (simp add:quotient_def)
-
-
-lemma other_direction:
- "Lang = L (r::rexp) \<Longrightarrow> finite (UNIV // (\<approx>Lang))"
-proof (induct arbitrary:Lang rule:rexp.induct)
- case NULL
- have "UNIV // (\<approx>{}) \<subseteq> {UNIV} "
- by (auto simp:quotient_def str_eq_rel_def str_eq_def)
- with prems show "?case" by (auto intro:finite_subset)
-next
- case EMPTY
- have "UNIV // (\<approx>{[]}) \<subseteq> {{[]}, UNIV - {[]}}"
- sorry
- with prems show ?case by (auto intro:finite_subset)
-next
- case (CHAR c)
- have "UNIV // (\<approx>{[c]}) \<subseteq> {{[]},{[c]}, UNIV - {[], [c]}}"
- sorry
- with prems show ?case by (auto intro:finite_subset)
-next
- case (SEQ r1 r2)
- show ?case sorry
-next
- case (ALT r1 r1)
- show ?case sorry
-next
- case (STAR r)
- show ?case sorry
-qed
-
-
-
-
-
-
-
-
-
-
-
-
-
-
-
-
-
-
-
-
-apply (induct arbitrary:Lang rule:rexp.induct)
-apply (simp add:QUOT_def equiv_class_def equiv_str_def)
-by (simp_all add:quot_lambda quot_single quot_seq quot_alt quot_star)
-
-(* Alternative definition for Quo *)
-definition
- QUOT :: "string set \<Rightarrow> (string set) set"
-where
- "QUOT Lang \<equiv> (\<Union>x. {\<lbrakk>x\<rbrakk>Lang})"
-
-lemma test:
- "UNIV Quo Lang = QUOT Lang"
-by (auto simp add: quot_def QUOT_def)
-
-lemma test1:
- assumes finite_CS: "finite (QUOT Lang)"
- shows "reg Lang"
-using finite_CS
-unfolding test[symmetric]
-by (auto dest: myhill_nerode)
-
-lemma cons_one: "x @ y \<in> {z} \<Longrightarrow> x @ y = z"
-by simp
-
-lemma quot_lambda: "QUOT {[]} = {{[]}, UNIV - {[]}}"
-proof
- show "QUOT {[]} \<subseteq> {{[]}, UNIV - {[]}}"
- proof
- fix x
- assume in_quot: "x \<in> QUOT {[]}"
- show "x \<in> {{[]}, UNIV - {[]}}"
- proof -
- from in_quot
- have "x = {[]} \<or> x = UNIV - {[]}"
- unfolding QUOT_def equiv_class_def
- proof
- fix xa
- assume in_univ: "xa \<in> UNIV"
- and in_eqiv: "x \<in> {{y. xa \<equiv>{[]}\<equiv> y}}"
- show "x = {[]} \<or> x = UNIV - {[]}"
- proof(cases "xa = []")
- case True
- hence "{y. xa \<equiv>{[]}\<equiv> y} = {[]}" using in_eqiv
- by (auto simp add:equiv_str_def)
- thus ?thesis using in_eqiv by (rule_tac disjI1, simp)
- next
- case False
- hence "{y. xa \<equiv>{[]}\<equiv> y} = UNIV - {[]}" using in_eqiv
- by (auto simp:equiv_str_def)
- thus ?thesis using in_eqiv by (rule_tac disjI2, simp)
- qed
- qed
- thus ?thesis by simp
- qed
- qed
-next
- show "{{[]}, UNIV - {[]}} \<subseteq> QUOT {[]}"
- proof
- fix x
- assume in_res: "x \<in> {{[]}, (UNIV::string set) - {[]}}"
- show "x \<in> (QUOT {[]})"
- proof -
- have "x = {[]} \<Longrightarrow> x \<in> QUOT {[]}"
- apply (simp add:QUOT_def equiv_class_def equiv_str_def)
- by (rule_tac x = "[]" in exI, auto)
- moreover have "x = UNIV - {[]} \<Longrightarrow> x \<in> QUOT {[]}"
- apply (simp add:QUOT_def equiv_class_def equiv_str_def)
- by (rule_tac x = "''a''" in exI, auto)
- ultimately show ?thesis using in_res by blast
- qed
+ assume "x \<in> UNIV // \<approx>{[]}"
+ then obtain y where h: "x = {z. (y, z) \<in> \<approx>{[]}}" unfolding quotient_def Image_def by blast
+ show "x \<in> {{[]}, UNIV - {[]}}"
+ proof (cases "y = []")
+ case True with h
+ have "x = {[]}" by (auto simp:str_eq_rel_def str_eq_def)
+ thus ?thesis by simp
+ next
+ case False with h
+ have "x = UNIV - {[]}" by (auto simp:str_eq_rel_def str_eq_def)
+ thus ?thesis by simp
qed
qed
-lemma quot_single_aux: "\<lbrakk>x \<noteq> []; x \<noteq> [c]\<rbrakk> \<Longrightarrow> x @ z \<noteq> [c]"
-by (induct x, auto)
-
-lemma quot_single: "\<And> (c::char). QUOT {[c]} = {{[]}, {[c]}, UNIV - {[], [c]}}"
-proof -
- fix c::"char"
- have exist_another: "\<exists> a. a \<noteq> c"
- apply (case_tac "c = CHR ''a''")
- apply (rule_tac x = "CHR ''b''" in exI, simp)
- by (rule_tac x = "CHR ''a''" in exI, simp)
- show "QUOT {[c]} = {{[]}, {[c]}, UNIV - {[], [c]}}"
- proof
- show "QUOT {[c]} \<subseteq> {{[]},{[c]}, UNIV - {[], [c]}}"
- proof
- fix x
- assume in_quot: "x \<in> QUOT {[c]}"
- show "x \<in> {{[]}, {[c]}, UNIV - {[],[c]}}"
- proof -
- from in_quot
- have "x = {[]} \<or> x = {[c]} \<or> x = UNIV - {[],[c]}"
- unfolding QUOT_def equiv_class_def
- proof
- fix xa
- assume in_eqiv: "x \<in> {{y. xa \<equiv>{[c]}\<equiv> y}}"
- show "x = {[]} \<or> x = {[c]} \<or> x = UNIV - {[], [c]}"
- proof-
- have "xa = [] \<Longrightarrow> x = {[]}" using in_eqiv
- by (auto simp add:equiv_str_def)
- moreover have "xa = [c] \<Longrightarrow> x = {[c]}"
- proof -
- have "xa = [c] \<Longrightarrow> {y. xa \<equiv>{[c]}\<equiv> y} = {[c]}" using in_eqiv
- apply (simp add:equiv_str_def)
- apply (rule set_ext, rule iffI, simp)
- apply (drule_tac x = "[]" in spec, auto)
- done
- thus "xa = [c] \<Longrightarrow> x = {[c]}" using in_eqiv by simp
- qed
- moreover have "\<lbrakk>xa \<noteq> []; xa \<noteq> [c]\<rbrakk> \<Longrightarrow> x = UNIV - {[],[c]}"
- proof -
- have "\<lbrakk>xa \<noteq> []; xa \<noteq> [c]\<rbrakk> \<Longrightarrow> {y. xa \<equiv>{[c]}\<equiv> y} = UNIV - {[],[c]}"
- apply (clarsimp simp add:equiv_str_def)
- apply (rule set_ext, rule iffI, simp)
- apply (rule conjI)
- apply (drule_tac x = "[c]" in spec, simp)
- apply (drule_tac x = "[]" in spec, simp)
- by (auto dest:quot_single_aux)
- thus "\<lbrakk>xa \<noteq> []; xa \<noteq> [c]\<rbrakk> \<Longrightarrow> x = UNIV - {[],[c]}" using in_eqiv by simp
- qed
- ultimately show ?thesis by blast
- qed
- qed
- thus ?thesis by simp
- qed
- qed
- next
- show "{{[]}, {[c]}, UNIV - {[],[c]}} \<subseteq> QUOT {[c]}"
- proof
- fix x
- assume in_res: "x \<in> {{[]},{[c]}, (UNIV::string set) - {[],[c]}}"
- show "x \<in> (QUOT {[c]})"
- proof -
- have "x = {[]} \<Longrightarrow> x \<in> QUOT {[c]}"
- apply (simp add:QUOT_def equiv_class_def equiv_str_def)
- by (rule_tac x = "[]" in exI, auto)
- moreover have "x = {[c]} \<Longrightarrow> x \<in> QUOT {[c]}"
- apply (simp add:QUOT_def equiv_class_def equiv_str_def)
- apply (rule_tac x = "[c]" in exI, simp)
- apply (rule set_ext, rule iffI, simp+)
- by (drule_tac x = "[]" in spec, simp)
- moreover have "x = UNIV - {[],[c]} \<Longrightarrow> x \<in> QUOT {[c]}"
- using exist_another
- apply (clarsimp simp add:QUOT_def equiv_class_def equiv_str_def)
- apply (rule_tac x = "[a]" in exI, simp)
- apply (rule set_ext, rule iffI, simp)
- apply (clarsimp simp:quot_single_aux, simp)
- apply (rule conjI)
- apply (drule_tac x = "[c]" in spec, simp)
- by (drule_tac x = "[]" in spec, simp)
- ultimately show ?thesis using in_res by blast
- qed
- qed
+lemma quot_char_subset:
+ "UNIV // (\<approx>{[c]}) \<subseteq> {{[]},{[c]}, UNIV - {[], [c]}}"
+proof
+ fix x
+ assume "x \<in> UNIV // \<approx>{[c]}"
+ then obtain y where h: "x = {z. (y, z) \<in> \<approx>{[c]}}" unfolding quotient_def Image_def by blast
+ show "x \<in> {{[]},{[c]}, UNIV - {[], [c]}}"
+ proof -
+ { assume "y = []" hence "x = {[]}" using h by (auto simp:str_eq_rel_def str_eq_def)
+ } moreover {
+ assume "y = [c]" hence "x = {[c]}" using h
+ by (auto dest!:spec[where x = "[]"] simp:str_eq_rel_def str_eq_def)
+ } moreover {
+ assume "y \<noteq> []" and "y \<noteq> [c]"
+ hence "\<forall> z. (y @ z) \<noteq> [c]" by (case_tac y, auto)
+ moreover have "\<And> p. (p \<noteq> [] \<and> p \<noteq> [c]) = (\<forall> q. p @ q \<noteq> [c])" by (case_tac p, auto)
+ ultimately have "x = UNIV - {[],[c]}" using h
+ by (auto simp add:str_eq_rel_def str_eq_def)
+ } ultimately show ?thesis by blast
qed
qed
-lemma eq_class_imp_eq_str:
- "\<lbrakk>x\<rbrakk>lang = \<lbrakk>y\<rbrakk>lang \<Longrightarrow> x \<equiv>lang\<equiv> y"
-by (auto simp:equiv_class_def equiv_str_def)
+text {* *************** Some common lemmas for following ALT, SEQ & STAR cases ******************* *}
-lemma finite_tag_image:
+lemma finite_tag_imageI:
"finite (range tag) \<Longrightarrow> finite (((op `) tag) ` S)"
apply (rule_tac B = "Pow (tag ` UNIV)" in finite_subset)
by (auto simp add:image_def Pow_def)
-lemma str_inj_imps:
- assumes str_inj: "\<And> m n. tag m = tag (n::string) \<Longrightarrow> m \<equiv>lang\<equiv> n"
- shows "inj_on ((op `) tag) (QUOT lang)"
-proof (clarsimp simp add:inj_on_def QUOT_def)
- fix x y
- assume eq_tag: "tag ` \<lbrakk>x\<rbrakk>lang = tag ` \<lbrakk>y\<rbrakk>lang"
- show "\<lbrakk>x\<rbrakk>lang = \<lbrakk>y\<rbrakk>lang"
- proof -
- have aux1:"\<And>a b. a \<in> (\<lbrakk>b\<rbrakk>lang) \<Longrightarrow> (a \<equiv>lang\<equiv> b)"
- by (simp add:equiv_class_def equiv_str_def)
- have aux2: "\<And> A B f. \<lbrakk>f ` A = f ` B; A \<noteq> {}\<rbrakk> \<Longrightarrow> \<exists> a b. a \<in> A \<and> b \<in> B \<and> f a = f b"
- by auto
- have aux3: "\<And> a l. \<lbrakk>a\<rbrakk>l \<noteq> {}"
- by (auto simp:equiv_class_def equiv_str_def)
- show ?thesis using eq_tag
- apply (drule_tac aux2, simp add:aux3, clarsimp)
- apply (drule_tac str_inj, (drule_tac aux1)+)
- by (simp add:equiv_str_def equiv_class_def)
- qed
+lemma eq_class_equalI:
+ "\<lbrakk>X \<in> UNIV // \<approx>lang; Y \<in> UNIV // \<approx>lang; x \<in> X; y \<in> Y; x \<approx>lang y\<rbrakk> \<Longrightarrow> X = Y"
+by (auto simp:quotient_def str_eq_rel_def str_eq_def)
+
+lemma tag_image_injI:
+ assumes str_inj: "\<And> m n. tag m = tag (n::string) \<Longrightarrow> m \<approx>lang n"
+ shows "inj_on ((op `) tag) (UNIV // \<approx>lang)"
+proof-
+ { fix X Y
+ assume X_in: "X \<in> UNIV // \<approx>lang"
+ and Y_in: "Y \<in> UNIV // \<approx>lang"
+ and tag_eq: "tag ` X = tag ` Y"
+ then obtain x y where "x \<in> X" and "y \<in> Y" and "tag x = tag y"
+ unfolding quotient_def Image_def str_eq_rel_def str_eq_def image_def
+ apply simp by blast
+ with X_in Y_in str_inj
+ have "X = Y" by (rule_tac eq_class_equalI, simp+)
+ }
+ thus ?thesis unfolding inj_on_def by auto
qed
-definition tag_str_ALT :: "string set \<Rightarrow> string set \<Rightarrow> string \<Rightarrow> (string set \<times> string set)"
-where
- "tag_str_ALT L\<^isub>1 L\<^isub>2 x \<equiv> (\<lbrakk>x\<rbrakk>L\<^isub>1, \<lbrakk>x\<rbrakk>L\<^isub>2)"
-
-lemma tag_str_alt_range_finite:
- assumes finite1: "finite (QUOT L\<^isub>1)"
- and finite2: "finite (QUOT L\<^isub>2)"
- shows "finite (range (tag_str_ALT L\<^isub>1 L\<^isub>2))"
-proof -
- have "{y. \<exists>x. y = (\<lbrakk>x\<rbrakk>L\<^isub>1, \<lbrakk>x\<rbrakk>L\<^isub>2)} \<subseteq> (QUOT L\<^isub>1) \<times> (QUOT L\<^isub>2)"
- by (auto simp:QUOT_def)
- thus ?thesis using finite1 finite2
- by (auto simp: image_def tag_str_ALT_def UNION_def
- intro: finite_subset[where B = "(QUOT L\<^isub>1) \<times> (QUOT L\<^isub>2)"])
-qed
-
-lemma tag_str_alt_inj:
- "tag_str_ALT L\<^isub>1 L\<^isub>2 x = tag_str_ALT L\<^isub>1 L\<^isub>2 y \<Longrightarrow> x \<equiv>(L\<^isub>1 \<union> L\<^isub>2)\<equiv> y"
-apply (simp add:tag_str_ALT_def equiv_class_def equiv_str_def)
-by blast
-
-lemma quot_alt:
- assumes finite1: "finite (QUOT L\<^isub>1)"
- and finite2: "finite (QUOT L\<^isub>2)"
- shows "finite (QUOT (L\<^isub>1 \<union> L\<^isub>2))"
-proof (rule_tac f = "(op `) (tag_str_ALT L\<^isub>1 L\<^isub>2)" in finite_imageD)
- show "finite (op ` (tag_str_ALT L\<^isub>1 L\<^isub>2) ` QUOT (L\<^isub>1 \<union> L\<^isub>2))"
- using finite_tag_image tag_str_alt_range_finite finite1 finite2
- by auto
-next
- show "inj_on (op ` (tag_str_ALT L\<^isub>1 L\<^isub>2)) (QUOT (L\<^isub>1 \<union> L\<^isub>2))"
- apply (rule_tac str_inj_imps)
- by (erule_tac tag_str_alt_inj)
-qed
+text {* **************** the SEQ case ************************ *}
(* list_diff:: list substract, once different return tailer *)
fun list_diff :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infix "-" 51)
@@ -1079,273 +860,297 @@
lemma [simp]: "x - [] = x"
by (induct x, auto)
-lemma [simp]: "xa \<le> x \<Longrightarrow> (x @ y) - xa = (x - xa) @ y"
+lemma [simp]: "(x - y = []) \<Longrightarrow> (x \<le> y)"
+proof-
+ have "\<exists>xa. x = xa @ (x - y) \<and> xa \<le> y"
+ apply (rule list_diff.induct[of _ x y], simp+)
+ by (clarsimp, rule_tac x = "y # xa" in exI, simp+)
+ thus "(x - y = []) \<Longrightarrow> (x \<le> y)" by simp
+qed
+
+lemma diff_prefix:
+ "\<lbrakk>c \<le> a - b; b \<le> a\<rbrakk> \<Longrightarrow> b @ c \<le> a"
by (auto elim:prefixE)
-definition tag_str_SEQ:: "string set \<Rightarrow> string set \<Rightarrow> string \<Rightarrow> (string set \<times> string set set)"
-where
- "tag_str_SEQ L\<^isub>1 L\<^isub>2 x \<equiv> if (\<exists> xa \<le> x. xa \<in> L\<^isub>1)
- then (\<lbrakk>x\<rbrakk>L\<^isub>1, {\<lbrakk>(x - xa)\<rbrakk>L\<^isub>2 | xa. xa \<le> x \<and> xa \<in> L\<^isub>1})
- else (\<lbrakk>x\<rbrakk>L\<^isub>1, {})"
+lemma diff_diff_appd:
+ "\<lbrakk>c < a - b; b < a\<rbrakk> \<Longrightarrow> (a - b) - c = a - (b @ c)"
+apply (clarsimp simp:strict_prefix_def)
+by (drule diff_prefix, auto elim:prefixE)
-lemma tag_seq_eq_E:
- "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y \<Longrightarrow>
- ((\<exists> xa \<le> x. xa \<in> L\<^isub>1) \<and> \<lbrakk>x\<rbrakk>L\<^isub>1 = \<lbrakk>y\<rbrakk>L\<^isub>1 \<and>
- {\<lbrakk>(x - xa)\<rbrakk>L\<^isub>2 | xa. xa \<le> x \<and> xa \<in> L\<^isub>1} = {\<lbrakk>(y - ya)\<rbrakk>L\<^isub>2 | ya. ya \<le> y \<and> ya \<in> L\<^isub>1} ) \<or>
- ((\<forall> xa \<le> x. xa \<notin> L\<^isub>1) \<and> \<lbrakk>x\<rbrakk>L\<^isub>1 = \<lbrakk>y\<rbrakk>L\<^isub>1)"
-by (simp add:tag_str_SEQ_def split:if_splits, blast)
+lemma app_eq_cases[rule_format]:
+ "\<forall> x . x @ y = m @ n \<longrightarrow> (x \<le> m \<or> m \<le> x)"
+apply (induct y, simp)
+apply (clarify, drule_tac x = "x @ [a]" in spec)
+by (clarsimp, auto simp:prefix_def)
+
+lemma app_eq_dest:
+ "x @ y = m @ n \<Longrightarrow> (x \<le> m \<and> (m - x) @ n = y) \<or> (m \<le> x \<and> (x - m) @ y = n)"
+by (frule_tac app_eq_cases, auto elim:prefixE)
+
+definition
+ "tag_str_SEQ L\<^isub>1 L\<^isub>2 x \<equiv> ((\<approx>L\<^isub>1) `` {x}, {(\<approx>L\<^isub>2) `` {x - xa}| xa. xa \<le> x \<and> xa \<in> L\<^isub>1})"
lemma tag_str_seq_range_finite:
- assumes finite1: "finite (QUOT L\<^isub>1)"
- and finite2: "finite (QUOT L\<^isub>2)"
- shows "finite (range (tag_str_SEQ L\<^isub>1 L\<^isub>2))"
-proof -
- have "(range (tag_str_SEQ L\<^isub>1 L\<^isub>2)) \<subseteq> (QUOT L\<^isub>1) \<times> (Pow (QUOT L\<^isub>2))"
- by (auto simp:image_def tag_str_SEQ_def QUOT_def)
- thus ?thesis using finite1 finite2
- by (rule_tac B = "(QUOT L\<^isub>1) \<times> (Pow (QUOT L\<^isub>2))" in finite_subset, auto)
-qed
-
-lemma app_in_seq_decom[rule_format]:
- "\<forall> x. x @ z \<in> L\<^isub>1 ; L\<^isub>2 \<longrightarrow> (\<exists> xa \<le> x. xa \<in> L\<^isub>1 \<and> (x - xa) @ z \<in> L\<^isub>2) \<or>
- (\<exists> za \<le> z. (x @ za) \<in> L\<^isub>1 \<and> (z - za) \<in> L\<^isub>2)"
-apply (induct z)
-apply (simp, rule allI, rule impI, rule disjI1)
-apply (clarsimp simp add:lang_seq_def)
-apply (rule_tac x = s1 in exI, simp)
-apply (rule allI | rule impI)+
-apply (drule_tac x = "x @ [a]" in spec, simp)
-apply (erule exE | erule conjE | erule disjE)+
-apply (rule disjI2, rule_tac x = "[a]" in exI, simp)
-apply (rule disjI1, rule_tac x = xa in exI, simp)
-apply (erule exE | erule conjE)+
-apply (rule disjI2, rule_tac x = "a # za" in exI, simp)
-done
+ "\<lbrakk>finite (UNIV // \<approx>L\<^isub>1); finite (UNIV // \<approx>L\<^isub>2)\<rbrakk> \<Longrightarrow> finite (range (tag_str_SEQ L\<^isub>1 L\<^isub>2))"
+apply (rule_tac B = "(UNIV // \<approx>L\<^isub>1) \<times> (Pow (UNIV // \<approx>L\<^isub>2))" in finite_subset)
+by (auto simp:tag_str_SEQ_def Image_def quotient_def split:if_splits)
-lemma tag_str_seq_inj:
- assumes tag_eq: "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y"
- shows "(x::string) \<equiv>(L\<^isub>1 ; L\<^isub>2)\<equiv> y"
-proof -
- have aux: "\<And> x y z. \<lbrakk>tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y; x @ z \<in> L\<^isub>1 ; L\<^isub>2\<rbrakk>
- \<Longrightarrow> y @ z \<in> L\<^isub>1 ; L\<^isub>2"
- proof (drule app_in_seq_decom, erule disjE)
- fix x y z
- assume tag_eq: "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y"
- and x_gets_l2: "\<exists>xa\<le>x. xa \<in> L\<^isub>1 \<and> (x - xa) @ z \<in> L\<^isub>2"
- from x_gets_l2 have "\<exists> xa \<le> x. xa \<in> L\<^isub>1" by blast
- hence xy_l2:"{\<lbrakk>(x - xa)\<rbrakk>L\<^isub>2 | xa. xa \<le> x \<and> xa \<in> L\<^isub>1} = {\<lbrakk>(y - ya)\<rbrakk>L\<^isub>2 | ya. ya \<le> y \<and> ya \<in> L\<^isub>1}"
- using tag_eq tag_seq_eq_E by blast
- from x_gets_l2 obtain xa where xa_le_x: "xa \<le> x"
- and xa_in_l1: "xa \<in> L\<^isub>1"
- and rest_in_l2: "(x - xa) @ z \<in> L\<^isub>2" by blast
- hence "\<exists> ya. \<lbrakk>(x - xa)\<rbrakk>L\<^isub>2 = \<lbrakk>(y - ya)\<rbrakk>L\<^isub>2 \<and> ya \<le> y \<and> ya \<in> L\<^isub>1" using xy_l2 by auto
- then obtain ya where ya_le_x: "ya \<le> y"
- and ya_in_l1: "ya \<in> L\<^isub>1"
- and rest_eq: "\<lbrakk>(x - xa)\<rbrakk>L\<^isub>2 = \<lbrakk>(y - ya)\<rbrakk>L\<^isub>2" by blast
- from rest_eq rest_in_l2 have "(y - ya) @ z \<in> L\<^isub>2"
- by (auto simp:equiv_class_def equiv_str_def)
- hence "ya @ ((y - ya) @ z) \<in> L\<^isub>1 ; L\<^isub>2" using ya_in_l1
- by (auto simp:lang_seq_def)
- thus "y @ z \<in> L\<^isub>1 ; L\<^isub>2" using ya_le_x
- by (erule_tac prefixE, simp)
- next
- fix x y z
- assume tag_eq: "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y"
- and x_gets_l1: "\<exists>za\<le>z. x @ za \<in> L\<^isub>1 \<and> z - za \<in> L\<^isub>2"
- from tag_eq tag_seq_eq_E have x_y_eq: "\<lbrakk>x\<rbrakk>L\<^isub>1 = \<lbrakk>y\<rbrakk>L\<^isub>1" by blast
- from x_gets_l1 obtain za where za_le_z: "za \<le> z"
- and x_za_in_l1: "(x @ za) \<in> L\<^isub>1"
- and rest_in_l2: "z - za \<in> L\<^isub>2" by blast
- from x_y_eq x_za_in_l1 have y_za_in_l1: "y @ za \<in> L\<^isub>1"
- by (auto simp:equiv_class_def equiv_str_def)
- hence "(y @ za) @ (z - za) \<in> L\<^isub>1 ; L\<^isub>2" using rest_in_l2
- apply (simp add:lang_seq_def)
- by (rule_tac x = "y @ za" in exI, rule_tac x = "z - za" in exI, simp)
- thus "y @ z \<in> L\<^isub>1 ; L\<^isub>2" using za_le_z
- by (erule_tac prefixE, simp)
- qed
- show ?thesis using tag_eq
- apply (simp add:equiv_str_def)
- by (auto intro:aux)
+lemma append_seq_elim:
+ assumes "x @ y \<in> L\<^isub>1 ;; L\<^isub>2"
+ shows "(\<exists> xa \<le> x. xa \<in> L\<^isub>1 \<and> (x - xa) @ y \<in> L\<^isub>2) \<or> (\<exists> ya \<le> y. (x @ ya) \<in> L\<^isub>1 \<and> (y - ya) \<in> L\<^isub>2)"
+proof-
+ from assms obtain s\<^isub>1 s\<^isub>2 where "x @ y = s\<^isub>1 @ s\<^isub>2" and in_seq: "s\<^isub>1 \<in> L\<^isub>1 \<and> s\<^isub>2 \<in> L\<^isub>2"
+ by (auto simp:Seq_def)
+ hence "(x \<le> s\<^isub>1 \<and> (s\<^isub>1 - x) @ s\<^isub>2 = y) \<or> (s\<^isub>1 \<le> x \<and> (x - s\<^isub>1) @ y = s\<^isub>2)"
+ using app_eq_dest by auto
+ moreover have "\<lbrakk>x \<le> s\<^isub>1; (s\<^isub>1 - x) @ s\<^isub>2 = y\<rbrakk> \<Longrightarrow> \<exists> ya \<le> y. (x @ ya) \<in> L\<^isub>1 \<and> (y - ya) \<in> L\<^isub>2" using in_seq
+ by (rule_tac x = "s\<^isub>1 - x" in exI, auto elim:prefixE)
+ moreover have "\<lbrakk>s\<^isub>1 \<le> x; (x - s\<^isub>1) @ y = s\<^isub>2\<rbrakk> \<Longrightarrow> \<exists> xa \<le> x. xa \<in> L\<^isub>1 \<and> (x - xa) @ y \<in> L\<^isub>2" using in_seq
+ by (rule_tac x = s\<^isub>1 in exI, auto)
+ ultimately show ?thesis by blast
qed
-lemma quot_seq:
- assumes finite1: "finite (QUOT L\<^isub>1)"
- and finite2: "finite (QUOT L\<^isub>2)"
- shows "finite (QUOT (L\<^isub>1;L\<^isub>2))"
-proof (rule_tac f = "(op `) (tag_str_SEQ L\<^isub>1 L\<^isub>2)" in finite_imageD)
- show "finite (op ` (tag_str_SEQ L\<^isub>1 L\<^isub>2) ` QUOT (L\<^isub>1 ; L\<^isub>2))"
- using finite_tag_image tag_str_seq_range_finite finite1 finite2
- by auto
+lemma tag_str_SEQ_injI:
+ "tag_str_SEQ L\<^isub>1 L\<^isub>2 m = tag_str_SEQ L\<^isub>1 L\<^isub>2 n \<Longrightarrow> m \<approx>(L\<^isub>1 ;; L\<^isub>2) n"
+proof-
+ { fix x y z
+ assume xz_in_seq: "x @ z \<in> L\<^isub>1 ;; L\<^isub>2"
+ and tag_xy: "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y"
+ have"y @ z \<in> L\<^isub>1 ;; L\<^isub>2"
+ proof-
+ have "(\<exists> xa \<le> x. xa \<in> L\<^isub>1 \<and> (x - xa) @ z \<in> L\<^isub>2) \<or> (\<exists> za \<le> z. (x @ za) \<in> L\<^isub>1 \<and> (z - za) \<in> L\<^isub>2)"
+ using xz_in_seq append_seq_elim by simp
+ moreover {
+ fix xa
+ assume h1: "xa \<le> x" and h2: "xa \<in> L\<^isub>1" and h3: "(x - xa) @ z \<in> L\<^isub>2"
+ obtain ya where "ya \<le> y" and "ya \<in> L\<^isub>1" and "(y - ya) @ z \<in> L\<^isub>2"
+ proof -
+ have "\<exists> ya. ya \<le> y \<and> ya \<in> L\<^isub>1 \<and> (x - xa) \<approx>L\<^isub>2 (y - ya)"
+ proof -
+ have "{\<approx>L\<^isub>2 `` {x - xa} |xa. xa \<le> x \<and> xa \<in> L\<^isub>1} =
+ {\<approx>L\<^isub>2 `` {y - xa} |xa. xa \<le> y \<and> xa \<in> L\<^isub>1}" (is "?Left = ?Right")
+ using h1 tag_xy by (auto simp:tag_str_SEQ_def)
+ moreover have "\<approx>L\<^isub>2 `` {x - xa} \<in> ?Left" using h1 h2 by auto
+ ultimately have "\<approx>L\<^isub>2 `` {x - xa} \<in> ?Right" by simp
+ thus ?thesis by (auto simp:Image_def str_eq_rel_def str_eq_def)
+ qed
+ with prems show ?thesis by (auto simp:str_eq_rel_def str_eq_def)
+ qed
+ hence "y @ z \<in> L\<^isub>1 ;; L\<^isub>2" by (erule_tac prefixE, auto simp:Seq_def)
+ } moreover {
+ fix za
+ assume h1: "za \<le> z" and h2: "(x @ za) \<in> L\<^isub>1" and h3: "z - za \<in> L\<^isub>2"
+ hence "y @ za \<in> L\<^isub>1"
+ proof-
+ have "\<approx>L\<^isub>1 `` {x} = \<approx>L\<^isub>1 `` {y}" using h1 tag_xy by (auto simp:tag_str_SEQ_def)
+ with h2 show ?thesis by (auto simp:Image_def str_eq_rel_def str_eq_def)
+ qed
+ with h1 h3 have "y @ z \<in> L\<^isub>1 ;; L\<^isub>2" by (drule_tac A = L\<^isub>1 in seq_intro, auto elim:prefixE)
+ }
+ ultimately show ?thesis by blast
+ qed
+ } thus "tag_str_SEQ L\<^isub>1 L\<^isub>2 m = tag_str_SEQ L\<^isub>1 L\<^isub>2 n \<Longrightarrow> m \<approx>(L\<^isub>1 ;; L\<^isub>2) n"
+ by (auto simp add: str_eq_def str_eq_rel_def)
+qed
+
+lemma quot_seq_finiteI:
+ assumes finite1: "finite (UNIV // \<approx>(L\<^isub>1::string set))"
+ and finite2: "finite (UNIV // \<approx>L\<^isub>2)"
+ shows "finite (UNIV // \<approx>(L\<^isub>1 ;; L\<^isub>2))"
+proof(rule_tac f = "(op `) (tag_str_SEQ L\<^isub>1 L\<^isub>2)" in finite_imageD)
+ show "finite (op ` (tag_str_SEQ L\<^isub>1 L\<^isub>2) ` UNIV // \<approx>L\<^isub>1 ;; L\<^isub>2)" using finite1 finite2
+ by (auto intro:finite_tag_imageI tag_str_seq_range_finite)
next
- show "inj_on (op ` (tag_str_SEQ L\<^isub>1 L\<^isub>2)) (QUOT (L\<^isub>1 ; L\<^isub>2))"
- apply (rule_tac str_inj_imps)
- by (erule_tac tag_str_seq_inj)
+ show "inj_on (op ` (tag_str_SEQ L\<^isub>1 L\<^isub>2)) (UNIV // \<approx>L\<^isub>1 ;; L\<^isub>2)"
+ apply (rule tag_image_injI)
+ apply (rule tag_str_SEQ_injI)
+ by (auto intro:tag_image_injI tag_str_SEQ_injI simp:)
qed
-(****************** the STAR case ************************)
+text {* **************** the ALT case ************************ *}
+
+definition
+ "tag_str_ALT L\<^isub>1 L\<^isub>2 (x::string) \<equiv> ((\<approx>L\<^isub>1) `` {x}, (\<approx>L\<^isub>2) `` {x})"
-lemma app_eq_elim[rule_format]:
- "\<And> a. \<forall> b x y. a @ b = x @ y \<longrightarrow> (\<exists> aa ab. a = aa @ ab \<and> x = aa \<and> y = ab @ b) \<or>
- (\<exists> ba bb. b = ba @ bb \<and> x = a @ ba \<and> y = bb \<and> ba \<noteq> [])"
- apply (induct_tac a rule:List.induct, simp)
- apply (rule allI | rule impI)+
- by (case_tac x, auto)
+lemma tag_str_alt_range_finite:
+ "\<lbrakk>finite (UNIV // \<approx>L\<^isub>1); finite (UNIV // \<approx>L\<^isub>2)\<rbrakk> \<Longrightarrow> finite (range (tag_str_ALT L\<^isub>1 L\<^isub>2))"
+apply (rule_tac B = "(UNIV // \<approx>L\<^isub>1) \<times> (UNIV // \<approx>L\<^isub>2)" in finite_subset)
+by (auto simp:tag_str_ALT_def Image_def quotient_def)
-definition tag_str_STAR:: "string set \<Rightarrow> string \<Rightarrow> string set set"
-where
- "tag_str_STAR L\<^isub>1 x \<equiv> if (x = [])
- then {}
- else {\<lbrakk>x\<^isub>2\<rbrakk>L\<^isub>1 | x\<^isub>1 x\<^isub>2. x = x\<^isub>1 @ x\<^isub>2 \<and> x\<^isub>1 \<in> L\<^isub>1\<star> \<and> x\<^isub>2 \<noteq> []}"
-
-lemma tag_str_star_range_finite:
- assumes finite1: "finite (QUOT L\<^isub>1)"
- shows "finite (range (tag_str_STAR L\<^isub>1))"
-proof -
- have "range (tag_str_STAR L\<^isub>1) \<subseteq> Pow (QUOT L\<^isub>1)"
- by (auto simp:image_def tag_str_STAR_def QUOT_def)
- thus ?thesis using finite1
- by (rule_tac B = "Pow (QUOT L\<^isub>1)" in finite_subset, auto)
+lemma quot_union_finiteI:
+ assumes finite1: "finite (UNIV // \<approx>(L\<^isub>1::string set))"
+ and finite2: "finite (UNIV // \<approx>L\<^isub>2)"
+ shows "finite (UNIV // \<approx>(L\<^isub>1 \<union> L\<^isub>2))"
+proof(rule_tac f = "(op `) (tag_str_ALT L\<^isub>1 L\<^isub>2)" in finite_imageD)
+ show "finite (op ` (tag_str_ALT L\<^isub>1 L\<^isub>2) ` UNIV // \<approx>L\<^isub>1 \<union> L\<^isub>2)" using finite1 finite2
+ by (auto intro:finite_tag_imageI tag_str_alt_range_finite)
+next
+ show "inj_on (op ` (tag_str_ALT L\<^isub>1 L\<^isub>2)) (UNIV // \<approx>L\<^isub>1 \<union> L\<^isub>2)"
+ proof-
+ have "\<And>m n. tag_str_ALT L\<^isub>1 L\<^isub>2 m = tag_str_ALT L\<^isub>1 L\<^isub>2 n \<Longrightarrow> m \<approx>(L\<^isub>1 \<union> L\<^isub>2) n"
+ unfolding tag_str_ALT_def str_eq_def Image_def str_eq_rel_def by auto
+ thus ?thesis by (auto intro:tag_image_injI)
+ qed
qed
-lemma star_prop[rule_format]: "x \<in> lang\<star> \<Longrightarrow> \<forall> y. y \<in> lang\<star> \<longrightarrow> x @ y \<in> lang\<star>"
+text {* **************** the Star case ****************** *}
+
+lemma finite_set_has_max: "\<lbrakk>finite A; A \<noteq> {}\<rbrakk> \<Longrightarrow> (\<exists> max \<in> A. \<forall> a \<in> 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 prems obtain max where h1: "max \<in> A" and h2: "\<forall>a\<in>A. f a \<le> f max" by blast
+ show ?thesis
+ proof (cases "f a \<le> f max")
+ assume "f a \<le> f max"
+ with h1 h2 show ?thesis by (rule_tac x = max in bexI, auto)
+ next
+ assume "\<not> (f a \<le> f max)"
+ thus ?thesis using h2 by (rule_tac x = a in bexI, auto)
+ qed
+ qed
+qed
+
+lemma star_intro1[rule_format]: "x \<in> lang\<star> \<Longrightarrow> \<forall> y. y \<in> lang\<star> \<longrightarrow> x @ y \<in> lang\<star>"
by (erule Star.induct, auto)
-lemma star_prop2: "y \<in> lang \<Longrightarrow> y \<in> lang\<star>"
+lemma star_intro2: "y \<in> lang \<Longrightarrow> y \<in> lang\<star>"
by (drule step[of y lang "[]"], auto simp:start)
-lemma star_prop3[rule_format]: "x \<in> lang\<star> \<Longrightarrow> \<forall>y . y \<in> lang \<longrightarrow> x @ y \<in> lang\<star>"
-by (erule Star.induct, auto intro:star_prop2)
+lemma star_intro3[rule_format]: "x \<in> lang\<star> \<Longrightarrow> \<forall>y . y \<in> lang \<longrightarrow> x @ y \<in> lang\<star>"
+by (erule Star.induct, auto intro:star_intro2)
+
+lemma star_decom: "\<lbrakk>x \<in> lang\<star>; x \<noteq> []\<rbrakk> \<Longrightarrow>(\<exists> a b. x = a @ b \<and> a \<noteq> [] \<and> a \<in> lang \<and> b \<in> lang\<star>)"
+by (induct x rule: Star.induct, simp, blast)
-lemma postfix_prop: "y >>= (x @ y) \<Longrightarrow> x = []"
-by (erule postfixE, induct x arbitrary:y, auto)
+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} \<union> {xs}")
+by (auto simp:strict_prefix_def)
+
+definition
+ "tag_str_STAR L\<^isub>1 x \<equiv> {(\<approx>L\<^isub>1) `` {x - xa} | xa. xa < x \<and> xa \<in> L\<^isub>1\<star>}"
+
+lemma tag_str_star_range_finite:
+ "finite (UNIV // \<approx>L\<^isub>1) \<Longrightarrow> finite (range (tag_str_STAR L\<^isub>1))"
+apply (rule_tac B = "Pow (UNIV // \<approx>L\<^isub>1)" in finite_subset)
+by (auto simp:tag_str_STAR_def Image_def quotient_def split:if_splits)
-lemma inj_aux:
- "\<lbrakk>(m @ z) \<in> L\<^isub>1\<star>; m \<equiv>L\<^isub>1\<equiv> yb; xa @ m = x; xa \<in> L\<^isub>1\<star>; m \<noteq> [];
- \<forall> xa xb. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star> \<longrightarrow> xb >>= m\<rbrakk>
- \<Longrightarrow> (yb @ z) \<in> L\<^isub>1\<star>"
-proof-
- have "\<And>s. s \<in> L\<^isub>1\<star> \<Longrightarrow> \<forall> m z yb. (s = m @ z \<and> m \<equiv>L\<^isub>1\<equiv> yb \<and> x = xa @ m \<and> xa \<in> L\<^isub>1\<star> \<and> m \<noteq> [] \<and>
- (\<forall> xa xb. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star> \<longrightarrow> xb >>= m) \<longrightarrow> (yb @ z) \<in> L\<^isub>1\<star>)"
- apply (erule Star.induct, simp)
- apply (rule allI | rule impI | erule conjE)+
- apply (drule app_eq_elim)
- apply (erule disjE | erule exE | erule conjE)+
- apply simp
- apply (simp (no_asm) only:append_assoc[THEN sym])
- apply (rule step)
- apply (simp add:equiv_str_def)
- apply simp
-
- apply (erule exE | erule conjE)+
- apply (rotate_tac 3)
- apply (frule_tac x = "xa @ s1" in spec)
- apply (rotate_tac 12)
- apply (drule_tac x = ba in spec)
- apply (erule impE)
- apply ( simp add:star_prop3)
- apply (simp)
- apply (drule postfix_prop)
- apply simp
- done
- thus "\<lbrakk>(m @ z) \<in> L\<^isub>1\<star>; m \<equiv>L\<^isub>1\<equiv> yb; xa @ m = x; xa \<in> L\<^isub>1\<star>; m \<noteq> [];
- \<forall> xa xb. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star> \<longrightarrow> xb >>= m\<rbrakk>
- \<Longrightarrow> (yb @ z) \<in> L\<^isub>1\<star>" by blast
+lemma tag_str_STAR_injI:
+ "tag_str_STAR L\<^isub>1 m = tag_str_STAR L\<^isub>1 n \<Longrightarrow> m \<approx>(L\<^isub>1\<star>) n"
+proof-
+ { fix x y z
+ assume xz_in_star: "x @ z \<in> L\<^isub>1\<star>"
+ and tag_xy: "tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 y"
+ have "y @ z \<in> L\<^isub>1\<star>"
+ proof(cases "x = []")
+ case True
+ with tag_xy have "y = []" by (auto simp:tag_str_STAR_def strict_prefix_def)
+ thus ?thesis using xz_in_star True by simp
+ next
+ case False
+ obtain x_max where h1: "x_max < x" and h2: "x_max \<in> L\<^isub>1\<star>" and h3: "(x - x_max) @ z \<in> L\<^isub>1\<star>"
+ and h4:"\<forall> xa < x. xa \<in> L\<^isub>1\<star> \<and> (x - xa) @ z \<in> L\<^isub>1\<star> \<longrightarrow> length xa \<le> length x_max"
+ proof-
+ let ?S = "{xa. xa < x \<and> xa \<in> L\<^isub>1\<star> \<and> (x - xa) @ z \<in> L\<^isub>1\<star>}"
+ have "finite ?S"
+ by (rule_tac B = "{xa. xa < x}" in finite_subset, auto simp:finite_strict_prefix_set)
+ moreover have "?S \<noteq> {}" using False xz_in_star
+ by (simp, rule_tac x = "[]" in exI, auto simp:strict_prefix_def)
+ ultimately have "\<exists> max \<in> ?S. \<forall> a \<in> ?S. length a \<le> length max" using finite_set_has_max by blast
+ with prems show ?thesis by blast
+ qed
+ obtain ya where h5: "ya < y" and h6: "ya \<in> L\<^isub>1\<star>" and h7: "(x - x_max) \<approx>L\<^isub>1 (y - ya)"
+ proof-
+ from tag_xy have "{\<approx>L\<^isub>1 `` {x - xa} |xa. xa < x \<and> xa \<in> L\<^isub>1\<star>} =
+ {\<approx>L\<^isub>1 `` {y - xa} |xa. xa < y \<and> xa \<in> L\<^isub>1\<star>}" (is "?left = ?right")
+ by (auto simp:tag_str_STAR_def)
+ moreover have "\<approx>L\<^isub>1 `` {x - x_max} \<in> ?left" using h1 h2 by auto
+ ultimately have "\<approx>L\<^isub>1 `` {x - x_max} \<in> ?right" by simp
+ with prems show ?thesis apply (simp add:Image_def str_eq_rel_def str_eq_def) by blast
+ qed
+ have "(y - ya) @ z \<in> L\<^isub>1\<star>"
+ proof-
+ from h3 h1 obtain a b where a_in: "a \<in> L\<^isub>1" and a_neq: "a \<noteq> []" and b_in: "b \<in> L\<^isub>1\<star>"
+ and ab_max: "(x - x_max) @ z = a @ b"
+ by (drule_tac star_decom, auto simp:strict_prefix_def elim:prefixE)
+ have "(x - x_max) \<le> a \<and> (a - (x - x_max)) @ b = z"
+ proof -
+ have "((x - x_max) \<le> a \<and> (a - (x - x_max)) @ b = z) \<or> (a < (x - x_max) \<and> ((x - x_max) - a) @ z = b)"
+ using app_eq_dest[OF ab_max] by (auto simp:strict_prefix_def)
+ moreover {
+ assume np: "a < (x - x_max)" and b_eqs: " ((x - x_max) - a) @ z = b"
+ have "False"
+ proof -
+ let ?x_max' = "x_max @ a"
+ have "?x_max' < x" using np h1 by (clarsimp simp:strict_prefix_def diff_prefix)
+ moreover have "?x_max' \<in> L\<^isub>1\<star>" using a_in h2 by (simp add:star_intro3)
+ moreover have "(x - ?x_max') @ z \<in> L\<^isub>1\<star>" using b_eqs b_in np h1 by (simp add:diff_diff_appd)
+ moreover have "\<not> (length ?x_max' \<le> length x_max)" using a_neq by simp
+ ultimately show ?thesis using h4 by blast
+ qed
+ } ultimately show ?thesis by blast
+ qed
+ then obtain za where z_decom: "z = za @ b" and x_za: "(x - x_max) @ za \<in> L\<^isub>1"
+ using a_in by (auto elim:prefixE)
+ from x_za h7 have "(y - ya) @ za \<in> L\<^isub>1" by (auto simp:str_eq_def)
+ with z_decom b_in show ?thesis by (auto dest!:step[of "(y - ya) @ za"])
+ qed
+ with h5 h6 show ?thesis by (drule_tac star_intro1, auto simp:strict_prefix_def elim:prefixE)
+ qed
+ } thus "tag_str_STAR L\<^isub>1 m = tag_str_STAR L\<^isub>1 n \<Longrightarrow> m \<approx>(L\<^isub>1\<star>) n"
+ by (auto simp add:str_eq_def str_eq_rel_def)
qed
-
-lemma min_postfix_exists[rule_format]:
- "finite A \<Longrightarrow> A \<noteq> {} \<and> (\<forall> a \<in> A. \<forall> b \<in> A. ((b >>= a) \<or> (a >>= b)))
- \<longrightarrow> (\<exists> min. (min \<in> A \<and> (\<forall> a \<in> A. a >>= min)))"
-apply (erule finite.induct)
-apply simp
-apply simp
-apply (case_tac "A = {}")
-apply simp
-apply clarsimp
-apply (case_tac "a >>= min")
-apply (rule_tac x = min in exI, simp)
-apply (rule_tac x = a in exI, simp)
-apply clarify
-apply (rotate_tac 5)
-apply (erule_tac x = aa in ballE) defer apply simp
-apply (erule_tac ys = min in postfix_trans)
-apply (erule_tac x = min in ballE)
-by simp+
-
-lemma tag_str_star_inj:
- "tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 (y::string) \<Longrightarrow> x \<equiv>L\<^isub>1\<star>\<equiv> y"
-proof -
- have aux: "\<And> x y z. \<lbrakk>tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 y; x @ z \<in> L\<^isub>1\<star>\<rbrakk> \<Longrightarrow> y @ z \<in> L\<^isub>1\<star>"
- proof-
- fix x y z
- assume tag_eq: "tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 y"
- and x_z: "x @ z \<in> L\<^isub>1\<star>"
- show "y @ z \<in> L\<^isub>1\<star>"
- proof (cases "x = []")
- case True
- with tag_eq have "y = []" by (simp add:tag_str_STAR_def split:if_splits, blast)
- thus ?thesis using x_z True by simp
- next
- case False
- hence not_empty: "{xb. \<exists> xa. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star>} \<noteq> {}" using x_z
- by (simp, rule_tac x = x in exI, rule_tac x = "[]" in exI, simp add:start)
- have finite_set: "finite {xb. \<exists> xa. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star>}"
- apply (rule_tac B = "{xb. \<exists> xa. x = xa @ xb}" in finite_subset)
- apply auto
- apply (induct x, simp)
- apply (subgoal_tac "{xb. \<exists>xa. a # x = xa @ xb} = insert (a # x) {xb. \<exists>xa. x = xa @ xb}")
- apply auto
- by (case_tac xaa, simp+)
- have comparable: "\<forall> a \<in> {xb. \<exists> xa. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star>}.
- \<forall> b \<in> {xb. \<exists> xa. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star>}.
- ((b >>= a) \<or> (a >>= b))"
- by (auto simp:postfix_def, drule app_eq_elim, blast)
- hence "\<exists> min. min \<in> {xb. \<exists> xa. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star>}
- \<and> (\<forall> xa xb. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star> \<longrightarrow> xb >>= min)"
- using finite_set not_empty comparable
- apply (drule_tac min_postfix_exists, simp)
- by (erule exE, rule_tac x = min in exI, auto)
- then obtain min xa where x_decom: "x = xa @ min \<and> xa \<in> L\<^isub>1\<star>"
- and min_not_empty: "min \<noteq> []"
- and min_z_in_star: "min @ z \<in> L\<^isub>1\<star>"
- and is_min: "\<forall> xa xb. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star> \<longrightarrow> xb >>= min" by blast
- from x_decom min_not_empty have "\<lbrakk>min\<rbrakk>L\<^isub>1 \<in> tag_str_STAR L\<^isub>1 x" by (auto simp:tag_str_STAR_def)
- hence "\<exists> yb. \<lbrakk>yb\<rbrakk>L\<^isub>1 \<in> tag_str_STAR L\<^isub>1 y \<and> \<lbrakk>min\<rbrakk>L\<^isub>1 = \<lbrakk>yb\<rbrakk>L\<^isub>1" using tag_eq by auto
- hence "\<exists> ya yb. y = ya @ yb \<and> ya \<in> L\<^isub>1\<star> \<and> min \<equiv>L\<^isub>1\<equiv> yb \<and> yb \<noteq> [] "
- by (simp add:tag_str_STAR_def equiv_class_def equiv_str_def split:if_splits, blast)
- then obtain ya yb where y_decom: "y = ya @ yb"
- and ya_in_star: "ya \<in> L\<^isub>1\<star>"
- and yb_not_empty: "yb \<noteq> []"
- and min_yb_eq: "min \<equiv>L\<^isub>1\<equiv> yb" by blast
- from min_z_in_star min_yb_eq min_not_empty is_min x_decom
- have "yb @ z \<in> L\<^isub>1\<star>"
- by (rule_tac x = x and xa = xa in inj_aux, simp+)
- thus ?thesis using ya_in_star y_decom
- by (auto dest:star_prop)
- qed
- qed
- show "tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 (y::string) \<Longrightarrow> x \<equiv>L\<^isub>1\<star>\<equiv> y"
- by (auto intro:aux simp:equiv_str_def)
+lemma quot_star_finiteI:
+ assumes finite: "finite (UNIV // \<approx>(L\<^isub>1::string set))"
+ shows "finite (UNIV // \<approx>(L\<^isub>1\<star>))"
+proof(rule_tac f = "(op `) (tag_str_STAR L\<^isub>1)" in finite_imageD)
+ show "finite (op ` (tag_str_STAR L\<^isub>1) ` UNIV // \<approx>L\<^isub>1\<star>)" using finite
+ by (auto intro:finite_tag_imageI tag_str_star_range_finite)
+next
+ show "inj_on (op ` (tag_str_STAR L\<^isub>1)) (UNIV // \<approx>L\<^isub>1\<star>)"
+ by (auto intro:tag_image_injI tag_str_STAR_injI)
qed
-lemma quot_star:
- assumes finite1: "finite (QUOT L\<^isub>1)"
- shows "finite (QUOT (L\<^isub>1\<star>))"
-proof (rule_tac f = "(op `) (tag_str_STAR L\<^isub>1)" in finite_imageD)
- show "finite (op ` (tag_str_STAR L\<^isub>1) ` QUOT (L\<^isub>1\<star>))"
- using finite_tag_image tag_str_star_range_finite finite1
- by auto
-next
- show "inj_on (op ` (tag_str_STAR L\<^isub>1)) (QUOT (L\<^isub>1\<star>))"
- apply (rule_tac str_inj_imps)
- by (erule_tac tag_str_star_inj)
-qed
+text {* **************** the Other Direction ************ *}
lemma other_direction:
- "Lang = L (r::rexp) \<Longrightarrow> finite (QUOT Lang)"
-apply (induct arbitrary:Lang rule:rexp.induct)
-apply (simp add:QUOT_def equiv_class_def equiv_str_def)
-by (simp_all add:quot_lambda quot_single quot_seq quot_alt quot_star)
+ "Lang = L (r::rexp) \<Longrightarrow> finite (UNIV // (\<approx>Lang))"
+proof (induct arbitrary:Lang rule:rexp.induct)
+ case NULL
+ have "UNIV // (\<approx>{}) \<subseteq> {UNIV} "
+ by (auto simp:quotient_def str_eq_rel_def str_eq_def)
+ with prems show "?case" by (auto intro:finite_subset)
+next
+ case EMPTY
+ have "UNIV // (\<approx>{[]}) \<subseteq> {{[]}, UNIV - {[]}}" by (rule quot_empty_subset)
+ with prems show ?case by (auto intro:finite_subset)
+next
+ case (CHAR c)
+ have "UNIV // (\<approx>{[c]}) \<subseteq> {{[]},{[c]}, UNIV - {[], [c]}}" by (rule quot_char_subset)
+ with prems show ?case by (auto intro:finite_subset)
+next
+ case (SEQ r\<^isub>1 r\<^isub>2)
+ have "\<lbrakk>finite (UNIV // \<approx>(L r\<^isub>1)); finite (UNIV // \<approx>(L r\<^isub>2))\<rbrakk> \<Longrightarrow> finite (UNIV // \<approx>(L r\<^isub>1 ;; L r\<^isub>2))"
+ by (erule quot_seq_finiteI, simp)
+ with prems show ?case by simp
+next
+ case (ALT r\<^isub>1 r\<^isub>2)
+ have "\<lbrakk>finite (UNIV // \<approx>(L r\<^isub>1)); finite (UNIV // \<approx>(L r\<^isub>2))\<rbrakk> \<Longrightarrow> finite (UNIV // \<approx>(L r\<^isub>1 \<union> L r\<^isub>2))"
+ by (erule quot_union_finiteI, simp)
+ with prems show ?case by simp
+next
+ case (STAR r)
+ have "finite (UNIV // \<approx>(L r)) \<Longrightarrow> finite (UNIV // \<approx>((L r)\<star>))"
+ by (erule quot_star_finiteI)
+ with prems show ?case by simp
+qed
-end
+end
\ No newline at end of file