Myhill.thy
changeset 29 c64241fa4dff
parent 28 cef2893f353b
child 30 f5db9e08effc
equal deleted inserted replaced
28:cef2893f353b 29:c64241fa4dff
     1 theory MyhillNerode
     1 theory MyhillNerode
     2   imports "Main" "List_Prefix"
     2   imports "Main" "List_Prefix"
     3 begin
     3 begin
     4 
     4 
     5 text {* sequential composition of languages *}
     5 text {* sequential composition of languages *}
     6 definition
     6 definition Seq :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ;; _" [100,100] 100)
     7   Seq :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ;; _" [100,100] 100)
       
     8 where 
     7 where 
     9   "L1 ;; L2 = {s1 @ s2 | s1 s2. s1 \<in> L1 \<and> s2 \<in> L2}"
     8   "L1 ;; L2 = {s1 @ s2 | s1 s2. s1 \<in> L1 \<and> s2 \<in> L2}"
    10 
     9 
    11 inductive_set
    10 inductive_set
    12   Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102)
    11   Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102)
    17 
    16 
    18 lemma seq_union_distrib:
    17 lemma seq_union_distrib:
    19   "(A \<union> B) ;; C = (A ;; C) \<union> (B ;; C)"
    18   "(A \<union> B) ;; C = (A ;; C) \<union> (B ;; C)"
    20 by (auto simp:Seq_def)
    19 by (auto simp:Seq_def)
    21 
    20 
       
    21 lemma seq_intro:
       
    22   "\<lbrakk>x \<in> A; y \<in> B\<rbrakk> \<Longrightarrow> x @ y \<in> A ;; B "
       
    23 by (auto simp:Seq_def)
       
    24 
    22 lemma seq_assoc:
    25 lemma seq_assoc:
    23   "(A ;; B) ;; C = A ;; (B ;; C)"
    26   "(A ;; B) ;; C = A ;; (B ;; C)"
    24 unfolding Seq_def
    27 apply(auto simp:Seq_def)
    25 apply(auto)
    28 apply blast
    26 apply(metis)
       
    27 by (metis append_assoc)
    29 by (metis append_assoc)
    28 
       
    29 lemma union_seq:
       
    30   "\<Union> {f x y ;; z| x y. P x y } = (\<Union> {f x y|x y. P x y });; z"
       
    31 apply (auto simp add:Seq_def)
       
    32 apply metis
       
    33 done
       
    34 
    30 
    35 theorem ardens_revised:
    31 theorem ardens_revised:
    36   assumes nemp: "[] \<notin> A"
    32   assumes nemp: "[] \<notin> A"
    37   shows "(X = X ;; A \<union> B) \<longleftrightarrow> (X = B ;; A\<star>)"
    33   shows "(X = X ;; A \<union> B) \<longleftrightarrow> (X = B ;; A\<star>)"
    38 proof
    34 proof
    39   assume eq: "X = B ;; A\<star>"
    35   assume eq: "X = B ;; A\<star>"
    40   have "A\<star> =  {[]} \<union> A\<star> ;; A" sorry
    36   have "A\<star> =  {[]} \<union> A\<star> ;; A" sorry 
    41   then have "B ;; A\<star> = B ;; ({[]} \<union> A\<star> ;; A)" unfolding Seq_def by simp
    37   then have "B ;; A\<star> = B ;; ({[]} \<union> A\<star> ;; A)" unfolding Seq_def by simp
    42   also have "\<dots> = B \<union> B ;; (A\<star> ;; A)"  unfolding Seq_def by auto
    38   also have "\<dots> = B \<union> B ;; (A\<star> ;; A)"  unfolding Seq_def by auto
    43   also have "\<dots> = B \<union> (B ;; A\<star>) ;; A"  unfolding Seq_def
    39   also have "\<dots> = B \<union> (B ;; A\<star>) ;; A"  unfolding Seq_def
    44     by (auto) (metis append_assoc)+
    40     by (auto) (metis append_assoc)+
    45   finally show "X = X ;; A \<union> B" using eq by auto
    41   finally show "X = X ;; A \<union> B" using eq by auto
    46 next
    42 next
    47   assume "X = X ;; A \<union> B"
    43   assume "X = X ;; A \<union> B"
    48   then have "B \<subseteq> X" "X ;; A \<subseteq> X" by auto
    44   then have "B \<subseteq> X" "X ;; A \<subseteq> X" by auto
    49   show "X = B ;; A\<star>" sorry
    45   thus "X = B ;; A\<star>" sorry
    50 qed
    46 qed
    51 
    47 
    52 datatype rexp =
    48 datatype rexp =
    53   NULL
    49   NULL
    54 | EMPTY
    50 | EMPTY
   776 
   772 
   777 lemma quot_empty_subset:
   773 lemma quot_empty_subset:
   778   "UNIV // (\<approx>{[]}) \<subseteq> {{[]}, UNIV - {[]}}"
   774   "UNIV // (\<approx>{[]}) \<subseteq> {{[]}, UNIV - {[]}}"
   779 proof
   775 proof
   780   fix x
   776   fix x
   781   assume h: "x \<in> UNIV // \<approx>{[]}"
   777   assume "x \<in> UNIV // \<approx>{[]}"
   782   show "x \<in> {{[]}, UNIV - {[]}}"
   778   then obtain y where h: "x = {z. (y, z) \<in> \<approx>{[]}}" unfolding quotient_def Image_def by blast
   783     
   779   show "x \<in> {{[]}, UNIV - {[]}}" 
   784  
   780   proof (cases "y = []")
   785   have "\<And> s. s \<approx>{[]} [] \<Longrightarrow> s = []"
   781     case True with h
   786     apply (auto simp add:str_eq_def)
   782     have "x = {[]}" by (auto simp:str_eq_rel_def str_eq_def)
   787     apply blast
   783     thus ?thesis by simp
   788   
   784   next
   789   hence "False"
   785     case False with h
   790     apply (simp add:quotient_def)
   786     have "x = UNIV - {[]}" by (auto simp:str_eq_rel_def str_eq_def)
   791 
   787     thus ?thesis by simp
       
   788   qed
       
   789 qed
       
   790 
       
   791 lemma quot_char_subset:
       
   792   "UNIV // (\<approx>{[c]}) \<subseteq> {{[]},{[c]}, UNIV - {[], [c]}}"
       
   793 proof 
       
   794   fix x 
       
   795   assume "x \<in> UNIV // \<approx>{[c]}"
       
   796   then obtain y where h: "x = {z. (y, z) \<in> \<approx>{[c]}}" unfolding quotient_def Image_def by blast
       
   797   show "x \<in> {{[]},{[c]}, UNIV - {[], [c]}}"
       
   798   proof -
       
   799     { assume "y = []" hence "x = {[]}" using h by (auto simp:str_eq_rel_def str_eq_def)
       
   800     } moreover {
       
   801       assume "y = [c]" hence "x = {[c]}" using h 
       
   802         by (auto dest!:spec[where x = "[]"] simp:str_eq_rel_def str_eq_def)
       
   803     } moreover {
       
   804       assume "y \<noteq> []" and "y \<noteq> [c]"
       
   805       hence "\<forall> z. (y @ z) \<noteq> [c]" by (case_tac y, auto)
       
   806       moreover have "\<And> p. (p \<noteq> [] \<and> p \<noteq> [c]) = (\<forall> q. p @ q \<noteq> [c])" by (case_tac p, auto)
       
   807       ultimately have "x = UNIV - {[],[c]}" using h
       
   808         by (auto simp add:str_eq_rel_def str_eq_def)
       
   809     } ultimately show ?thesis by blast
       
   810   qed
       
   811 qed
       
   812 
       
   813 text {* *************** Some common lemmas for following ALT, SEQ & STAR cases ******************* *}
       
   814 
       
   815 lemma finite_tag_imageI: 
       
   816   "finite (range tag) \<Longrightarrow> finite (((op `) tag) ` S)"
       
   817 apply (rule_tac B = "Pow (tag ` UNIV)" in finite_subset)
       
   818 by (auto simp add:image_def Pow_def)
       
   819 
       
   820 lemma eq_class_equalI:
       
   821   "\<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"
       
   822 by (auto simp:quotient_def str_eq_rel_def str_eq_def)
       
   823 
       
   824 lemma tag_image_injI:
       
   825   assumes str_inj: "\<And> m n. tag m = tag (n::string) \<Longrightarrow> m \<approx>lang n"
       
   826   shows "inj_on ((op `) tag) (UNIV // \<approx>lang)"
       
   827 proof-
       
   828   { fix X Y
       
   829     assume X_in: "X \<in> UNIV // \<approx>lang"
       
   830       and  Y_in: "Y \<in> UNIV // \<approx>lang"
       
   831       and  tag_eq: "tag ` X = tag ` Y"
       
   832     then obtain x y where "x \<in> X" and "y \<in> Y" and "tag x = tag y"
       
   833       unfolding quotient_def Image_def str_eq_rel_def str_eq_def image_def
       
   834       apply simp by blast
       
   835     with X_in Y_in str_inj
       
   836     have "X = Y" by (rule_tac eq_class_equalI, simp+)
       
   837   }
       
   838   thus ?thesis unfolding inj_on_def by auto
       
   839 qed
       
   840 
       
   841 text {* **************** the SEQ case ************************ *}
       
   842 
       
   843 (* list_diff:: list substract, once different return tailer *)
       
   844 fun list_diff :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infix "-" 51)
       
   845 where
       
   846   "list_diff []  xs = []" |
       
   847   "list_diff (x#xs) [] = x#xs" |
       
   848   "list_diff (x#xs) (y#ys) = (if x = y then list_diff xs ys else (x#xs))"
       
   849 
       
   850 lemma [simp]: "(x @ y) - x = y"
       
   851 apply (induct x)
       
   852 by (case_tac y, simp+)
       
   853 
       
   854 lemma [simp]: "x - x = []"
       
   855 by (induct x, auto)
       
   856 
       
   857 lemma [simp]: "x = xa @ y \<Longrightarrow> x - xa = y "
       
   858 by (induct x, auto)
       
   859 
       
   860 lemma [simp]: "x - [] = x"
       
   861 by (induct x, auto)
       
   862 
       
   863 lemma [simp]: "(x - y = []) \<Longrightarrow> (x \<le> y)"
       
   864 proof-   
       
   865   have "\<exists>xa. x = xa @ (x - y) \<and> xa \<le> y"
       
   866     apply (rule list_diff.induct[of _ x y], simp+)
       
   867     by (clarsimp, rule_tac x = "y # xa" in exI, simp+)
       
   868   thus "(x - y = []) \<Longrightarrow> (x \<le> y)" by simp
       
   869 qed
       
   870 
       
   871 lemma diff_prefix:
       
   872   "\<lbrakk>c \<le> a - b; b \<le> a\<rbrakk> \<Longrightarrow> b @ c \<le> a"
       
   873 by (auto elim:prefixE)
       
   874 
       
   875 lemma diff_diff_appd: 
       
   876   "\<lbrakk>c < a - b; b < a\<rbrakk> \<Longrightarrow> (a - b) - c = a - (b @ c)"
       
   877 apply (clarsimp simp:strict_prefix_def)
       
   878 by (drule diff_prefix, auto elim:prefixE)
       
   879 
       
   880 lemma app_eq_cases[rule_format]:
       
   881   "\<forall> x . x @ y = m @ n \<longrightarrow> (x \<le> m \<or> m \<le> x)"
       
   882 apply (induct y, simp)
       
   883 apply (clarify, drule_tac x = "x @ [a]" in spec)
       
   884 by (clarsimp, auto simp:prefix_def)
       
   885 
       
   886 lemma app_eq_dest:
       
   887   "x @ y = m @ n \<Longrightarrow> (x \<le> m \<and> (m - x) @ n = y) \<or> (m \<le> x \<and> (x - m) @ y = n)"
       
   888 by (frule_tac app_eq_cases, auto elim:prefixE)
       
   889 
       
   890 definition 
       
   891   "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})"
       
   892 
       
   893 lemma tag_str_seq_range_finite:
       
   894   "\<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))"
       
   895 apply (rule_tac B = "(UNIV // \<approx>L\<^isub>1) \<times> (Pow (UNIV // \<approx>L\<^isub>2))" in finite_subset)
       
   896 by (auto simp:tag_str_SEQ_def Image_def quotient_def split:if_splits)
       
   897 
       
   898 lemma append_seq_elim:
       
   899   assumes "x @ y \<in> L\<^isub>1 ;; L\<^isub>2"
       
   900   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)"
       
   901 proof-
       
   902   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" 
       
   903     by (auto simp:Seq_def)
       
   904   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)"
       
   905     using app_eq_dest by auto
       
   906   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
       
   907     by (rule_tac x = "s\<^isub>1 - x" in exI, auto elim:prefixE)
       
   908   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
       
   909     by (rule_tac x = s\<^isub>1 in exI, auto)
       
   910   ultimately show ?thesis by blast
       
   911 qed
       
   912 
       
   913 lemma tag_str_SEQ_injI:
       
   914   "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"
       
   915 proof-
       
   916   { fix x y z
       
   917     assume xz_in_seq: "x @ z \<in> L\<^isub>1 ;; L\<^isub>2"
       
   918     and tag_xy: "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y"
       
   919     have"y @ z \<in> L\<^isub>1 ;; L\<^isub>2" 
       
   920     proof-
       
   921       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)"
       
   922         using xz_in_seq append_seq_elim by simp
       
   923       moreover {
       
   924         fix xa
       
   925         assume h1: "xa \<le> x" and h2: "xa \<in> L\<^isub>1" and h3: "(x - xa) @ z \<in> L\<^isub>2"
       
   926         obtain ya where "ya \<le> y" and "ya \<in> L\<^isub>1" and "(y - ya) @ z \<in> L\<^isub>2" 
       
   927         proof -
       
   928           have "\<exists> ya.  ya \<le> y \<and> ya \<in> L\<^isub>1 \<and> (x - xa) \<approx>L\<^isub>2 (y - ya)"
       
   929           proof -
       
   930             have "{\<approx>L\<^isub>2 `` {x - xa} |xa. xa \<le> x \<and> xa \<in> L\<^isub>1} = 
       
   931                   {\<approx>L\<^isub>2 `` {y - xa} |xa. xa \<le> y \<and> xa \<in> L\<^isub>1}" (is "?Left = ?Right") 
       
   932               using h1 tag_xy by (auto simp:tag_str_SEQ_def)
       
   933             moreover have "\<approx>L\<^isub>2 `` {x - xa} \<in> ?Left" using h1 h2 by auto
       
   934             ultimately have "\<approx>L\<^isub>2 `` {x - xa} \<in> ?Right" by simp
       
   935             thus ?thesis by (auto simp:Image_def str_eq_rel_def str_eq_def)
       
   936           qed
       
   937           with prems show ?thesis by (auto simp:str_eq_rel_def str_eq_def)
       
   938         qed
       
   939         hence "y @ z \<in> L\<^isub>1 ;; L\<^isub>2" by (erule_tac prefixE, auto simp:Seq_def)          
       
   940       } moreover {
       
   941         fix za
       
   942         assume h1: "za \<le> z" and h2: "(x @ za) \<in> L\<^isub>1" and h3: "z - za \<in> L\<^isub>2"
       
   943         hence "y @ za \<in> L\<^isub>1"
       
   944         proof-
       
   945           have "\<approx>L\<^isub>1 `` {x} = \<approx>L\<^isub>1 `` {y}" using h1 tag_xy by (auto simp:tag_str_SEQ_def)
       
   946           with h2 show ?thesis by (auto simp:Image_def str_eq_rel_def str_eq_def) 
       
   947         qed
       
   948         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)
       
   949       }
       
   950       ultimately show ?thesis by blast
       
   951     qed
       
   952   } 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" 
       
   953     by (auto simp add: str_eq_def str_eq_rel_def)
       
   954 qed 
       
   955 
       
   956 lemma quot_seq_finiteI:
       
   957   assumes finite1: "finite (UNIV // \<approx>(L\<^isub>1::string set))"
       
   958   and finite2: "finite (UNIV // \<approx>L\<^isub>2)"
       
   959   shows "finite (UNIV // \<approx>(L\<^isub>1 ;; L\<^isub>2))"
       
   960 proof(rule_tac f = "(op `) (tag_str_SEQ L\<^isub>1 L\<^isub>2)" in finite_imageD)
       
   961   show "finite (op ` (tag_str_SEQ L\<^isub>1 L\<^isub>2) ` UNIV // \<approx>L\<^isub>1 ;; L\<^isub>2)" using finite1 finite2
       
   962     by (auto intro:finite_tag_imageI tag_str_seq_range_finite)
       
   963 next
       
   964   show  "inj_on (op ` (tag_str_SEQ L\<^isub>1 L\<^isub>2)) (UNIV // \<approx>L\<^isub>1 ;; L\<^isub>2)"
       
   965     apply (rule tag_image_injI)
       
   966     apply (rule tag_str_SEQ_injI)
       
   967     by (auto intro:tag_image_injI tag_str_SEQ_injI simp:)
       
   968 qed
       
   969 
       
   970 text {* **************** the ALT case ************************ *}
       
   971 
       
   972 definition 
       
   973   "tag_str_ALT L\<^isub>1 L\<^isub>2 (x::string) \<equiv> ((\<approx>L\<^isub>1) `` {x}, (\<approx>L\<^isub>2) `` {x})"
       
   974 
       
   975 lemma tag_str_alt_range_finite:
       
   976   "\<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))"
       
   977 apply (rule_tac B = "(UNIV // \<approx>L\<^isub>1) \<times> (UNIV // \<approx>L\<^isub>2)" in finite_subset)
       
   978 by (auto simp:tag_str_ALT_def Image_def quotient_def)
       
   979 
       
   980 lemma quot_union_finiteI:
       
   981   assumes finite1: "finite (UNIV // \<approx>(L\<^isub>1::string set))"
       
   982   and finite2: "finite (UNIV // \<approx>L\<^isub>2)"
       
   983   shows "finite (UNIV // \<approx>(L\<^isub>1 \<union> L\<^isub>2))"
       
   984 proof(rule_tac f = "(op `) (tag_str_ALT L\<^isub>1 L\<^isub>2)" in finite_imageD)
       
   985   show "finite (op ` (tag_str_ALT L\<^isub>1 L\<^isub>2) ` UNIV // \<approx>L\<^isub>1 \<union> L\<^isub>2)" using finite1 finite2
       
   986     by (auto intro:finite_tag_imageI tag_str_alt_range_finite)
       
   987 next
       
   988   show "inj_on (op ` (tag_str_ALT L\<^isub>1 L\<^isub>2)) (UNIV // \<approx>L\<^isub>1 \<union> L\<^isub>2)"
       
   989   proof-
       
   990     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"
       
   991       unfolding tag_str_ALT_def str_eq_def Image_def str_eq_rel_def by auto
       
   992     thus ?thesis by (auto intro:tag_image_injI)
       
   993   qed
       
   994 qed
       
   995 
       
   996 text {* **************** the Star case ****************** *}
       
   997 
       
   998 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))"
       
   999 proof (induct rule:finite.induct)
       
  1000   case emptyI thus ?case by simp
       
  1001 next
       
  1002   case (insertI A a)
       
  1003   show ?case
       
  1004   proof (cases "A = {}")
       
  1005     case True thus ?thesis by (rule_tac x = a in bexI, auto)
       
  1006   next
       
  1007     case False
       
  1008     with prems obtain max where h1: "max \<in> A" and h2: "\<forall>a\<in>A. f a \<le> f max" by blast
       
  1009     show ?thesis
       
  1010     proof (cases "f a \<le> f max")
       
  1011       assume "f a \<le> f max"
       
  1012       with h1 h2 show ?thesis by (rule_tac x = max in bexI, auto)
       
  1013     next
       
  1014       assume "\<not> (f a \<le> f max)"
       
  1015       thus ?thesis using h2 by (rule_tac x = a in bexI, auto)
       
  1016     qed
       
  1017   qed
       
  1018 qed
       
  1019 
       
  1020 lemma star_intro1[rule_format]: "x \<in> lang\<star> \<Longrightarrow> \<forall> y. y \<in> lang\<star> \<longrightarrow> x @ y \<in> lang\<star>"
       
  1021 by (erule Star.induct, auto)
       
  1022 
       
  1023 lemma star_intro2: "y \<in> lang \<Longrightarrow> y \<in> lang\<star>"
       
  1024 by (drule step[of y lang "[]"], auto simp:start)
       
  1025 
       
  1026 lemma star_intro3[rule_format]: "x \<in> lang\<star> \<Longrightarrow> \<forall>y . y \<in> lang \<longrightarrow> x @ y \<in> lang\<star>"
       
  1027 by (erule Star.induct, auto intro:star_intro2)
       
  1028 
       
  1029 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>)"
       
  1030 by (induct x rule: Star.induct, simp, blast)
       
  1031 
       
  1032 lemma finite_strict_prefix_set: "finite {xa. xa < (x::string)}"
       
  1033 apply (induct x rule:rev_induct, simp)
       
  1034 apply (subgoal_tac "{xa. xa < xs @ [x]} = {xa. xa < xs} \<union> {xs}")
       
  1035 by (auto simp:strict_prefix_def)
       
  1036 
       
  1037 definition 
       
  1038   "tag_str_STAR L\<^isub>1 x \<equiv> {(\<approx>L\<^isub>1) `` {x - xa} | xa. xa < x \<and> xa \<in> L\<^isub>1\<star>}"
       
  1039 
       
  1040 lemma tag_str_star_range_finite:
       
  1041   "finite (UNIV // \<approx>L\<^isub>1) \<Longrightarrow> finite (range (tag_str_STAR L\<^isub>1))"
       
  1042 apply (rule_tac B = "Pow (UNIV // \<approx>L\<^isub>1)" in finite_subset)
       
  1043 by (auto simp:tag_str_STAR_def Image_def quotient_def split:if_splits)
       
  1044 
       
  1045 lemma tag_str_STAR_injI:
       
  1046   "tag_str_STAR L\<^isub>1 m = tag_str_STAR L\<^isub>1 n \<Longrightarrow> m \<approx>(L\<^isub>1\<star>) n"
       
  1047 proof-
       
  1048   { fix x y z
       
  1049     assume xz_in_star: "x @ z \<in> L\<^isub>1\<star>"
       
  1050     and tag_xy: "tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 y"
       
  1051     have "y @ z \<in> L\<^isub>1\<star>"
       
  1052     proof(cases "x = []")
       
  1053       case True
       
  1054       with tag_xy have "y = []" by (auto simp:tag_str_STAR_def strict_prefix_def)
       
  1055       thus ?thesis using xz_in_star True by simp
       
  1056     next
       
  1057       case False
       
  1058       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>" 
       
  1059         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"
       
  1060       proof-
       
  1061         let ?S = "{xa. xa < x \<and> xa \<in> L\<^isub>1\<star> \<and> (x - xa) @ z \<in> L\<^isub>1\<star>}"
       
  1062         have "finite ?S"
       
  1063           by (rule_tac B = "{xa. xa < x}" in finite_subset, auto simp:finite_strict_prefix_set)
       
  1064         moreover have "?S \<noteq> {}" using False xz_in_star
       
  1065           by (simp, rule_tac x = "[]" in exI, auto simp:strict_prefix_def)
       
  1066         ultimately have "\<exists> max \<in> ?S. \<forall> a \<in> ?S. length a \<le> length max" using finite_set_has_max by blast
       
  1067         with prems show ?thesis by blast
       
  1068       qed
       
  1069       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)"
       
  1070       proof-
       
  1071         from tag_xy have "{\<approx>L\<^isub>1 `` {x - xa} |xa. xa < x \<and> xa \<in> L\<^isub>1\<star>} = 
       
  1072           {\<approx>L\<^isub>1 `` {y - xa} |xa. xa < y \<and> xa \<in> L\<^isub>1\<star>}" (is "?left = ?right")
       
  1073           by (auto simp:tag_str_STAR_def)
       
  1074         moreover have "\<approx>L\<^isub>1 `` {x - x_max} \<in> ?left" using h1 h2 by auto
       
  1075         ultimately have "\<approx>L\<^isub>1 `` {x - x_max} \<in> ?right" by simp
       
  1076         with prems show ?thesis apply (simp add:Image_def str_eq_rel_def str_eq_def) by blast
       
  1077       qed      
       
  1078       have "(y - ya) @ z \<in> L\<^isub>1\<star>" 
       
  1079       proof-
       
  1080         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>" 
       
  1081           and ab_max: "(x - x_max) @ z = a @ b" 
       
  1082           by (drule_tac star_decom, auto simp:strict_prefix_def elim:prefixE)
       
  1083         have "(x - x_max) \<le> a \<and> (a - (x - x_max)) @ b = z" 
       
  1084         proof -
       
  1085           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)" 
       
  1086             using app_eq_dest[OF ab_max] by (auto simp:strict_prefix_def)
       
  1087           moreover { 
       
  1088             assume np: "a < (x - x_max)" and b_eqs: " ((x - x_max) - a) @ z = b"
       
  1089             have "False"
       
  1090             proof -
       
  1091               let ?x_max' = "x_max @ a"
       
  1092               have "?x_max' < x" using np h1 by (clarsimp simp:strict_prefix_def diff_prefix) 
       
  1093               moreover have "?x_max' \<in> L\<^isub>1\<star>" using a_in h2 by (simp add:star_intro3) 
       
  1094               moreover have "(x - ?x_max') @ z \<in> L\<^isub>1\<star>" using b_eqs b_in np h1 by (simp add:diff_diff_appd)
       
  1095               moreover have "\<not> (length ?x_max' \<le> length x_max)" using a_neq by simp
       
  1096               ultimately show ?thesis using h4 by blast
       
  1097             qed 
       
  1098           } ultimately show ?thesis by blast
       
  1099         qed
       
  1100         then obtain za where z_decom: "z = za @ b" and x_za: "(x - x_max) @ za \<in> L\<^isub>1" 
       
  1101           using a_in by (auto elim:prefixE)        
       
  1102         from x_za h7 have "(y - ya) @ za \<in> L\<^isub>1" by (auto simp:str_eq_def)
       
  1103         with z_decom b_in show ?thesis by (auto dest!:step[of "(y - ya) @ za"])
       
  1104       qed
       
  1105       with h5 h6 show ?thesis by (drule_tac star_intro1, auto simp:strict_prefix_def elim:prefixE)
       
  1106     qed      
       
  1107   } thus "tag_str_STAR L\<^isub>1 m = tag_str_STAR L\<^isub>1 n \<Longrightarrow> m \<approx>(L\<^isub>1\<star>) n"
       
  1108     by (auto simp add:str_eq_def str_eq_rel_def)
       
  1109 qed
       
  1110 
       
  1111 lemma quot_star_finiteI:
       
  1112   assumes finite: "finite (UNIV // \<approx>(L\<^isub>1::string set))"
       
  1113   shows "finite (UNIV // \<approx>(L\<^isub>1\<star>))"
       
  1114 proof(rule_tac f = "(op `) (tag_str_STAR L\<^isub>1)" in finite_imageD)
       
  1115   show "finite (op ` (tag_str_STAR L\<^isub>1) ` UNIV // \<approx>L\<^isub>1\<star>)" using finite
       
  1116     by (auto intro:finite_tag_imageI tag_str_star_range_finite)
       
  1117 next
       
  1118   show "inj_on (op ` (tag_str_STAR L\<^isub>1)) (UNIV // \<approx>L\<^isub>1\<star>)"
       
  1119     by (auto intro:tag_image_injI tag_str_STAR_injI)
       
  1120 qed
       
  1121 
       
  1122 text {* **************** the Other Direction ************ *}
   792 
  1123 
   793 lemma other_direction:
  1124 lemma other_direction:
   794   "Lang = L (r::rexp) \<Longrightarrow> finite (UNIV // (\<approx>Lang))"
  1125   "Lang = L (r::rexp) \<Longrightarrow> finite (UNIV // (\<approx>Lang))"
   795 proof (induct arbitrary:Lang rule:rexp.induct)
  1126 proof (induct arbitrary:Lang rule:rexp.induct)
   796   case NULL
  1127   case NULL
   797   have "UNIV // (\<approx>{}) \<subseteq> {UNIV} "
  1128   have "UNIV // (\<approx>{}) \<subseteq> {UNIV} "
   798     by (auto simp:quotient_def str_eq_rel_def str_eq_def)
  1129     by (auto simp:quotient_def str_eq_rel_def str_eq_def)
   799   with prems show "?case" by (auto intro:finite_subset)
  1130   with prems show "?case" by (auto intro:finite_subset)
   800 next
  1131 next
   801   case EMPTY
  1132   case EMPTY
   802   have "UNIV // (\<approx>{[]}) \<subseteq> {{[]}, UNIV - {[]}}"
  1133   have "UNIV // (\<approx>{[]}) \<subseteq> {{[]}, UNIV - {[]}}" by (rule quot_empty_subset)
   803     sorry
       
   804   with prems show ?case by (auto intro:finite_subset)
  1134   with prems show ?case by (auto intro:finite_subset)
   805 next
  1135 next
   806   case (CHAR c)
  1136   case (CHAR c)
   807   have "UNIV // (\<approx>{[c]}) \<subseteq> {{[]},{[c]}, UNIV - {[], [c]}}"
  1137   have "UNIV // (\<approx>{[c]}) \<subseteq> {{[]},{[c]}, UNIV - {[], [c]}}" by (rule quot_char_subset)
   808     sorry
       
   809   with prems show ?case by (auto intro:finite_subset)
  1138   with prems show ?case by (auto intro:finite_subset)
   810 next
  1139 next
   811   case (SEQ r1 r2)
  1140   case (SEQ r\<^isub>1 r\<^isub>2)
   812   show ?case sorry
  1141   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))"
   813 next
  1142     by (erule quot_seq_finiteI, simp)
   814   case (ALT r1 r1)
  1143   with prems show ?case by simp
   815   show ?case sorry
  1144 next
       
  1145   case (ALT r\<^isub>1 r\<^isub>2)
       
  1146   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))"
       
  1147     by (erule quot_union_finiteI, simp)
       
  1148   with prems show ?case by simp  
   816 next
  1149 next
   817   case (STAR r)
  1150   case (STAR r)
   818   show ?case sorry
  1151   have "finite (UNIV // \<approx>(L r)) \<Longrightarrow> finite (UNIV // \<approx>((L r)\<star>))"
   819 qed
  1152     by (erule quot_star_finiteI)
   820     
  1153   with prems show ?case by simp
   821 
  1154 qed 
   822       
  1155 
   823 
  1156 end
   824 
       
   825 
       
   826 
       
   827 
       
   828 
       
   829 
       
   830 
       
   831 
       
   832 
       
   833 
       
   834 
       
   835 
       
   836 
       
   837 
       
   838 
       
   839 
       
   840 apply (induct arbitrary:Lang rule:rexp.induct)
       
   841 apply (simp add:QUOT_def equiv_class_def equiv_str_def)
       
   842 by (simp_all add:quot_lambda quot_single quot_seq quot_alt quot_star)  
       
   843 
       
   844 (* Alternative definition for Quo *)
       
   845 definition 
       
   846   QUOT :: "string set \<Rightarrow> (string set) set"  
       
   847 where
       
   848   "QUOT Lang \<equiv> (\<Union>x. {\<lbrakk>x\<rbrakk>Lang})"
       
   849 
       
   850 lemma test: 
       
   851   "UNIV Quo Lang = QUOT Lang"
       
   852 by (auto simp add: quot_def QUOT_def)
       
   853 
       
   854 lemma test1:
       
   855   assumes finite_CS: "finite (QUOT Lang)"
       
   856   shows "reg Lang"
       
   857 using finite_CS
       
   858 unfolding test[symmetric]
       
   859 by (auto dest: myhill_nerode)
       
   860 
       
   861 lemma cons_one: "x @ y \<in> {z} \<Longrightarrow> x @ y = z"
       
   862 by simp
       
   863 
       
   864 lemma quot_lambda: "QUOT {[]} = {{[]}, UNIV - {[]}}"
       
   865 proof 
       
   866   show "QUOT {[]} \<subseteq> {{[]}, UNIV - {[]}}"
       
   867   proof 
       
   868     fix x 
       
   869     assume in_quot: "x \<in> QUOT {[]}"
       
   870     show "x \<in> {{[]}, UNIV - {[]}}"
       
   871     proof -
       
   872       from in_quot 
       
   873       have "x = {[]} \<or> x = UNIV - {[]}"
       
   874         unfolding QUOT_def equiv_class_def
       
   875       proof 
       
   876         fix xa
       
   877         assume in_univ: "xa \<in> UNIV"
       
   878            and in_eqiv: "x \<in> {{y. xa \<equiv>{[]}\<equiv> y}}"
       
   879         show "x = {[]} \<or> x = UNIV - {[]}"
       
   880         proof(cases "xa = []")
       
   881           case True
       
   882           hence "{y. xa \<equiv>{[]}\<equiv> y} = {[]}" using in_eqiv
       
   883             by (auto simp add:equiv_str_def)
       
   884           thus ?thesis using in_eqiv by (rule_tac disjI1, simp)
       
   885         next
       
   886           case False
       
   887           hence "{y. xa \<equiv>{[]}\<equiv> y} = UNIV - {[]}" using in_eqiv
       
   888             by (auto simp:equiv_str_def)
       
   889           thus ?thesis using in_eqiv by (rule_tac disjI2, simp)
       
   890         qed
       
   891       qed
       
   892       thus ?thesis by simp
       
   893     qed
       
   894   qed
       
   895 next
       
   896   show "{{[]}, UNIV - {[]}} \<subseteq> QUOT {[]}"
       
   897   proof
       
   898     fix x
       
   899     assume in_res: "x \<in> {{[]}, (UNIV::string set) - {[]}}"
       
   900     show "x \<in> (QUOT {[]})"
       
   901     proof -
       
   902       have "x = {[]} \<Longrightarrow> x \<in> QUOT {[]}"
       
   903         apply (simp add:QUOT_def equiv_class_def equiv_str_def)
       
   904         by (rule_tac x = "[]" in exI, auto)
       
   905       moreover have "x = UNIV - {[]} \<Longrightarrow> x \<in> QUOT {[]}"
       
   906         apply (simp add:QUOT_def equiv_class_def equiv_str_def)
       
   907         by (rule_tac x = "''a''" in exI, auto)
       
   908       ultimately show ?thesis using in_res by blast
       
   909     qed
       
   910   qed
       
   911 qed
       
   912 
       
   913 lemma quot_single_aux: "\<lbrakk>x \<noteq> []; x \<noteq> [c]\<rbrakk> \<Longrightarrow> x @ z \<noteq> [c]"
       
   914 by (induct x, auto)
       
   915 
       
   916 lemma quot_single: "\<And> (c::char). QUOT {[c]} = {{[]}, {[c]}, UNIV - {[], [c]}}"
       
   917 proof - 
       
   918   fix c::"char" 
       
   919   have exist_another: "\<exists> a. a \<noteq> c" 
       
   920     apply (case_tac "c = CHR ''a''")
       
   921     apply (rule_tac x = "CHR ''b''" in exI, simp)
       
   922     by (rule_tac x = "CHR ''a''" in exI, simp)  
       
   923   show "QUOT {[c]} = {{[]}, {[c]}, UNIV - {[], [c]}}"
       
   924   proof
       
   925     show "QUOT {[c]} \<subseteq> {{[]},{[c]}, UNIV - {[], [c]}}"
       
   926     proof 
       
   927       fix x 
       
   928       assume in_quot: "x \<in> QUOT {[c]}"
       
   929       show "x \<in> {{[]}, {[c]}, UNIV - {[],[c]}}"
       
   930       proof -
       
   931         from in_quot 
       
   932         have "x = {[]} \<or> x = {[c]} \<or> x = UNIV - {[],[c]}"
       
   933           unfolding QUOT_def equiv_class_def
       
   934         proof 
       
   935           fix xa
       
   936           assume in_eqiv: "x \<in> {{y. xa \<equiv>{[c]}\<equiv> y}}"
       
   937           show "x = {[]} \<or> x = {[c]} \<or> x = UNIV - {[], [c]}"
       
   938           proof-
       
   939             have "xa = [] \<Longrightarrow> x = {[]}" using in_eqiv 
       
   940               by (auto simp add:equiv_str_def)
       
   941             moreover have "xa = [c] \<Longrightarrow> x = {[c]}"
       
   942             proof -
       
   943               have "xa = [c] \<Longrightarrow> {y. xa \<equiv>{[c]}\<equiv> y} = {[c]}" using in_eqiv
       
   944                 apply (simp add:equiv_str_def)
       
   945                 apply (rule set_ext, rule iffI, simp)
       
   946                 apply (drule_tac x = "[]" in spec, auto)
       
   947                 done   
       
   948               thus "xa = [c] \<Longrightarrow> x = {[c]}" using in_eqiv by simp 
       
   949             qed
       
   950             moreover have "\<lbrakk>xa \<noteq> []; xa \<noteq> [c]\<rbrakk> \<Longrightarrow> x = UNIV - {[],[c]}"
       
   951             proof -
       
   952               have "\<lbrakk>xa \<noteq> []; xa \<noteq> [c]\<rbrakk> \<Longrightarrow> {y. xa \<equiv>{[c]}\<equiv> y} = UNIV - {[],[c]}" 
       
   953                 apply (clarsimp simp add:equiv_str_def)
       
   954                 apply (rule set_ext, rule iffI, simp)
       
   955                 apply (rule conjI)
       
   956                 apply (drule_tac x = "[c]" in spec, simp)
       
   957                 apply (drule_tac x = "[]" in spec, simp)
       
   958                 by (auto dest:quot_single_aux)
       
   959               thus "\<lbrakk>xa \<noteq> []; xa \<noteq> [c]\<rbrakk> \<Longrightarrow> x = UNIV - {[],[c]}" using in_eqiv by simp
       
   960             qed
       
   961             ultimately show ?thesis by blast
       
   962           qed
       
   963         qed
       
   964         thus ?thesis by simp
       
   965       qed
       
   966     qed
       
   967   next
       
   968     show "{{[]}, {[c]}, UNIV - {[],[c]}} \<subseteq> QUOT {[c]}"
       
   969     proof
       
   970       fix x
       
   971       assume in_res: "x \<in> {{[]},{[c]}, (UNIV::string set) - {[],[c]}}"
       
   972       show "x \<in> (QUOT {[c]})"
       
   973       proof -
       
   974         have "x = {[]} \<Longrightarrow> x \<in> QUOT {[c]}"
       
   975           apply (simp add:QUOT_def equiv_class_def equiv_str_def)
       
   976           by (rule_tac x = "[]" in exI, auto)
       
   977         moreover have "x = {[c]} \<Longrightarrow> x \<in> QUOT {[c]}"
       
   978           apply (simp add:QUOT_def equiv_class_def equiv_str_def)
       
   979           apply (rule_tac x = "[c]" in exI, simp)
       
   980           apply (rule set_ext, rule iffI, simp+)
       
   981           by (drule_tac x = "[]" in spec, simp)
       
   982         moreover have "x = UNIV - {[],[c]} \<Longrightarrow> x \<in> QUOT {[c]}"
       
   983           using exist_another
       
   984           apply (clarsimp simp add:QUOT_def equiv_class_def equiv_str_def)        
       
   985           apply (rule_tac x = "[a]" in exI, simp)
       
   986           apply (rule set_ext, rule iffI, simp)
       
   987           apply (clarsimp simp:quot_single_aux, simp)
       
   988           apply (rule conjI)
       
   989           apply (drule_tac x = "[c]" in spec, simp)
       
   990           by (drule_tac x = "[]" in spec, simp)     
       
   991         ultimately show ?thesis using in_res by blast
       
   992       qed
       
   993     qed
       
   994   qed
       
   995 qed
       
   996 
       
   997 lemma eq_class_imp_eq_str:
       
   998   "\<lbrakk>x\<rbrakk>lang = \<lbrakk>y\<rbrakk>lang \<Longrightarrow> x \<equiv>lang\<equiv> y"
       
   999 by (auto simp:equiv_class_def equiv_str_def)
       
  1000 
       
  1001 lemma finite_tag_image: 
       
  1002   "finite (range tag) \<Longrightarrow> finite (((op `) tag) ` S)"
       
  1003 apply (rule_tac B = "Pow (tag ` UNIV)" in finite_subset)
       
  1004 by (auto simp add:image_def Pow_def)
       
  1005 
       
  1006 lemma str_inj_imps:
       
  1007   assumes str_inj: "\<And> m n. tag m = tag (n::string) \<Longrightarrow> m \<equiv>lang\<equiv> n"
       
  1008   shows "inj_on ((op `) tag) (QUOT lang)"
       
  1009 proof (clarsimp simp add:inj_on_def QUOT_def)
       
  1010   fix x y
       
  1011   assume eq_tag: "tag ` \<lbrakk>x\<rbrakk>lang = tag ` \<lbrakk>y\<rbrakk>lang"
       
  1012   show "\<lbrakk>x\<rbrakk>lang = \<lbrakk>y\<rbrakk>lang"
       
  1013   proof -
       
  1014     have aux1:"\<And>a b. a \<in> (\<lbrakk>b\<rbrakk>lang) \<Longrightarrow> (a \<equiv>lang\<equiv> b)"
       
  1015       by (simp add:equiv_class_def equiv_str_def)
       
  1016     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"
       
  1017       by auto
       
  1018     have aux3: "\<And> a l. \<lbrakk>a\<rbrakk>l \<noteq> {}" 
       
  1019       by (auto simp:equiv_class_def equiv_str_def)
       
  1020     show ?thesis using eq_tag
       
  1021       apply (drule_tac aux2, simp add:aux3, clarsimp)
       
  1022       apply (drule_tac str_inj, (drule_tac aux1)+)
       
  1023       by (simp add:equiv_str_def equiv_class_def)
       
  1024   qed
       
  1025 qed
       
  1026 
       
  1027 definition tag_str_ALT :: "string set \<Rightarrow> string set \<Rightarrow> string \<Rightarrow> (string set \<times> string set)"
       
  1028 where
       
  1029   "tag_str_ALT L\<^isub>1 L\<^isub>2 x \<equiv> (\<lbrakk>x\<rbrakk>L\<^isub>1, \<lbrakk>x\<rbrakk>L\<^isub>2)"
       
  1030 
       
  1031 lemma tag_str_alt_range_finite:
       
  1032   assumes finite1: "finite (QUOT L\<^isub>1)"
       
  1033   and finite2: "finite (QUOT L\<^isub>2)"
       
  1034   shows "finite (range (tag_str_ALT L\<^isub>1 L\<^isub>2))"
       
  1035 proof -
       
  1036   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)"
       
  1037     by (auto simp:QUOT_def)
       
  1038   thus ?thesis using finite1 finite2
       
  1039     by (auto simp: image_def tag_str_ALT_def UNION_def 
       
  1040             intro: finite_subset[where B = "(QUOT L\<^isub>1) \<times> (QUOT L\<^isub>2)"])
       
  1041 qed
       
  1042 
       
  1043 lemma tag_str_alt_inj:
       
  1044   "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"
       
  1045 apply (simp add:tag_str_ALT_def equiv_class_def equiv_str_def)
       
  1046 by blast
       
  1047   
       
  1048 lemma quot_alt:
       
  1049   assumes finite1: "finite (QUOT L\<^isub>1)"
       
  1050   and finite2: "finite (QUOT L\<^isub>2)"
       
  1051   shows "finite (QUOT (L\<^isub>1 \<union> L\<^isub>2))"
       
  1052 proof (rule_tac f = "(op `) (tag_str_ALT L\<^isub>1 L\<^isub>2)" in finite_imageD)
       
  1053   show "finite (op ` (tag_str_ALT L\<^isub>1 L\<^isub>2) ` QUOT (L\<^isub>1 \<union> L\<^isub>2))"
       
  1054     using finite_tag_image tag_str_alt_range_finite finite1 finite2
       
  1055     by auto
       
  1056 next
       
  1057   show "inj_on (op ` (tag_str_ALT L\<^isub>1 L\<^isub>2)) (QUOT (L\<^isub>1 \<union> L\<^isub>2))"
       
  1058     apply (rule_tac str_inj_imps)
       
  1059     by (erule_tac tag_str_alt_inj)
       
  1060 qed
       
  1061 
       
  1062 (* list_diff:: list substract, once different return tailer *)
       
  1063 fun list_diff :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infix "-" 51)
       
  1064 where
       
  1065   "list_diff []  xs = []" |
       
  1066   "list_diff (x#xs) [] = x#xs" |
       
  1067   "list_diff (x#xs) (y#ys) = (if x = y then list_diff xs ys else (x#xs))"
       
  1068 
       
  1069 lemma [simp]: "(x @ y) - x = y"
       
  1070 apply (induct x)
       
  1071 by (case_tac y, simp+)
       
  1072 
       
  1073 lemma [simp]: "x - x = []"
       
  1074 by (induct x, auto)
       
  1075 
       
  1076 lemma [simp]: "x = xa @ y \<Longrightarrow> x - xa = y "
       
  1077 by (induct x, auto)
       
  1078 
       
  1079 lemma [simp]: "x - [] = x"
       
  1080 by (induct x, auto)
       
  1081 
       
  1082 lemma [simp]: "xa \<le> x \<Longrightarrow> (x @ y) - xa = (x - xa) @ y"
       
  1083 by (auto elim:prefixE)
       
  1084 
       
  1085 definition tag_str_SEQ:: "string set \<Rightarrow> string set \<Rightarrow> string \<Rightarrow> (string set \<times> string set set)"
       
  1086 where
       
  1087   "tag_str_SEQ L\<^isub>1 L\<^isub>2 x \<equiv> if (\<exists> xa \<le> x. xa \<in> L\<^isub>1)
       
  1088                          then (\<lbrakk>x\<rbrakk>L\<^isub>1, {\<lbrakk>(x - xa)\<rbrakk>L\<^isub>2 | xa.  xa \<le> x \<and> xa \<in> L\<^isub>1})
       
  1089                          else (\<lbrakk>x\<rbrakk>L\<^isub>1, {})"
       
  1090 
       
  1091 lemma tag_seq_eq_E:
       
  1092   "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y \<Longrightarrow>
       
  1093    ((\<exists> xa \<le> x. xa \<in> L\<^isub>1) \<and> \<lbrakk>x\<rbrakk>L\<^isub>1 = \<lbrakk>y\<rbrakk>L\<^isub>1 \<and> 
       
  1094     {\<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>
       
  1095    ((\<forall> xa \<le> x. xa \<notin> L\<^isub>1) \<and> \<lbrakk>x\<rbrakk>L\<^isub>1 = \<lbrakk>y\<rbrakk>L\<^isub>1)"
       
  1096 by (simp add:tag_str_SEQ_def split:if_splits, blast)
       
  1097 
       
  1098 lemma tag_str_seq_range_finite:
       
  1099   assumes finite1: "finite (QUOT L\<^isub>1)"
       
  1100   and finite2: "finite (QUOT L\<^isub>2)"
       
  1101   shows "finite (range (tag_str_SEQ L\<^isub>1 L\<^isub>2))"
       
  1102 proof -
       
  1103   have "(range (tag_str_SEQ L\<^isub>1 L\<^isub>2)) \<subseteq> (QUOT L\<^isub>1) \<times> (Pow (QUOT L\<^isub>2))"
       
  1104     by (auto simp:image_def tag_str_SEQ_def QUOT_def)
       
  1105   thus ?thesis using finite1 finite2 
       
  1106     by (rule_tac B = "(QUOT L\<^isub>1) \<times> (Pow (QUOT L\<^isub>2))" in finite_subset, auto)
       
  1107 qed
       
  1108   
       
  1109 lemma app_in_seq_decom[rule_format]:
       
  1110   "\<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> 
       
  1111                             (\<exists> za \<le> z. (x @ za) \<in> L\<^isub>1 \<and> (z - za) \<in> L\<^isub>2)"
       
  1112 apply (induct z)
       
  1113 apply (simp, rule allI, rule impI, rule disjI1)
       
  1114 apply (clarsimp simp add:lang_seq_def)
       
  1115 apply (rule_tac x = s1 in exI, simp)
       
  1116 apply (rule allI | rule impI)+
       
  1117 apply (drule_tac x = "x @ [a]" in spec, simp)
       
  1118 apply (erule exE | erule conjE | erule disjE)+
       
  1119 apply (rule disjI2, rule_tac x = "[a]" in exI, simp)
       
  1120 apply (rule disjI1, rule_tac x = xa in exI, simp) 
       
  1121 apply (erule exE | erule conjE)+
       
  1122 apply (rule disjI2, rule_tac x = "a # za" in exI, simp)
       
  1123 done
       
  1124 
       
  1125 lemma tag_str_seq_inj:
       
  1126   assumes tag_eq: "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y"
       
  1127   shows "(x::string) \<equiv>(L\<^isub>1 ; L\<^isub>2)\<equiv> y"
       
  1128 proof -
       
  1129   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> 
       
  1130                        \<Longrightarrow> y @ z \<in> L\<^isub>1 ; L\<^isub>2"
       
  1131   proof (drule app_in_seq_decom, erule disjE)
       
  1132     fix x y z
       
  1133     assume tag_eq: "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y"
       
  1134       and x_gets_l2: "\<exists>xa\<le>x. xa \<in> L\<^isub>1 \<and> (x - xa) @ z \<in> L\<^isub>2"
       
  1135     from x_gets_l2 have "\<exists> xa \<le> x. xa \<in> L\<^isub>1" by blast
       
  1136     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}"
       
  1137       using tag_eq tag_seq_eq_E by blast
       
  1138     from x_gets_l2 obtain xa where xa_le_x: "xa \<le> x"
       
  1139                                and xa_in_l1: "xa \<in> L\<^isub>1"
       
  1140                                and rest_in_l2: "(x - xa) @ z \<in> L\<^isub>2" by blast
       
  1141     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
       
  1142     then obtain ya where ya_le_x: "ya \<le> y"
       
  1143                      and ya_in_l1: "ya \<in> L\<^isub>1"
       
  1144                      and rest_eq: "\<lbrakk>(x - xa)\<rbrakk>L\<^isub>2 = \<lbrakk>(y - ya)\<rbrakk>L\<^isub>2" by blast
       
  1145     from rest_eq rest_in_l2 have "(y - ya) @ z \<in> L\<^isub>2" 
       
  1146       by (auto simp:equiv_class_def equiv_str_def)
       
  1147     hence "ya @ ((y - ya) @ z) \<in> L\<^isub>1 ; L\<^isub>2" using ya_in_l1
       
  1148       by (auto simp:lang_seq_def)
       
  1149     thus "y @ z \<in> L\<^isub>1 ; L\<^isub>2" using ya_le_x 
       
  1150       by (erule_tac prefixE, simp)
       
  1151   next
       
  1152     fix x y z
       
  1153     assume tag_eq: "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y"
       
  1154       and x_gets_l1: "\<exists>za\<le>z. x @ za \<in> L\<^isub>1 \<and> z - za \<in> L\<^isub>2"
       
  1155     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
       
  1156     from x_gets_l1 obtain za where za_le_z: "za \<le> z"
       
  1157                                and x_za_in_l1: "(x @ za) \<in> L\<^isub>1"
       
  1158                                and rest_in_l2: "z - za \<in> L\<^isub>2" by blast
       
  1159     from x_y_eq x_za_in_l1 have y_za_in_l1: "y @ za \<in> L\<^isub>1"
       
  1160       by (auto simp:equiv_class_def equiv_str_def)
       
  1161     hence "(y @ za) @ (z - za) \<in> L\<^isub>1 ; L\<^isub>2" using rest_in_l2
       
  1162       apply (simp add:lang_seq_def)
       
  1163       by (rule_tac x = "y @ za" in exI, rule_tac x = "z - za" in exI, simp)
       
  1164     thus "y @ z \<in> L\<^isub>1 ; L\<^isub>2" using za_le_z
       
  1165       by (erule_tac prefixE, simp)
       
  1166   qed
       
  1167   show ?thesis using tag_eq
       
  1168     apply (simp add:equiv_str_def)
       
  1169     by (auto intro:aux)
       
  1170 qed
       
  1171 
       
  1172 lemma quot_seq: 
       
  1173   assumes finite1: "finite (QUOT L\<^isub>1)"
       
  1174   and finite2: "finite (QUOT L\<^isub>2)"
       
  1175   shows "finite (QUOT (L\<^isub>1;L\<^isub>2))"
       
  1176 proof (rule_tac f = "(op `) (tag_str_SEQ L\<^isub>1 L\<^isub>2)" in finite_imageD)
       
  1177   show "finite (op ` (tag_str_SEQ L\<^isub>1 L\<^isub>2) ` QUOT (L\<^isub>1 ; L\<^isub>2))"
       
  1178     using finite_tag_image tag_str_seq_range_finite finite1 finite2
       
  1179     by auto
       
  1180 next
       
  1181   show "inj_on (op ` (tag_str_SEQ L\<^isub>1 L\<^isub>2)) (QUOT (L\<^isub>1 ; L\<^isub>2))"
       
  1182     apply (rule_tac str_inj_imps)
       
  1183     by (erule_tac tag_str_seq_inj)
       
  1184 qed
       
  1185 
       
  1186 (****************** the STAR case ************************)
       
  1187 
       
  1188 lemma app_eq_elim[rule_format]:
       
  1189   "\<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>
       
  1190                                    (\<exists> ba bb. b = ba @ bb \<and> x = a @ ba \<and> y = bb \<and> ba \<noteq> [])"
       
  1191   apply (induct_tac a rule:List.induct, simp)
       
  1192   apply (rule allI | rule impI)+
       
  1193   by (case_tac x, auto)
       
  1194 
       
  1195 definition tag_str_STAR:: "string set \<Rightarrow> string \<Rightarrow> string set set"
       
  1196 where
       
  1197   "tag_str_STAR L\<^isub>1 x \<equiv> if (x = []) 
       
  1198                        then {}
       
  1199                        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> []}"
       
  1200 
       
  1201 lemma tag_str_star_range_finite:
       
  1202   assumes finite1: "finite (QUOT L\<^isub>1)"
       
  1203   shows "finite (range (tag_str_STAR L\<^isub>1))"
       
  1204 proof -
       
  1205   have "range (tag_str_STAR L\<^isub>1) \<subseteq> Pow (QUOT L\<^isub>1)"
       
  1206     by (auto simp:image_def tag_str_STAR_def QUOT_def)
       
  1207   thus ?thesis using finite1
       
  1208     by (rule_tac B = "Pow (QUOT L\<^isub>1)" in finite_subset, auto)
       
  1209 qed
       
  1210 
       
  1211 lemma star_prop[rule_format]: "x \<in> lang\<star> \<Longrightarrow> \<forall> y. y \<in> lang\<star> \<longrightarrow> x @ y \<in> lang\<star>"
       
  1212 by (erule Star.induct, auto)
       
  1213 
       
  1214 lemma star_prop2: "y \<in> lang \<Longrightarrow> y \<in> lang\<star>"
       
  1215 by (drule step[of y lang "[]"], auto simp:start)
       
  1216 
       
  1217 lemma star_prop3[rule_format]: "x \<in> lang\<star> \<Longrightarrow> \<forall>y . y \<in> lang \<longrightarrow> x @ y \<in> lang\<star>"
       
  1218 by (erule Star.induct, auto intro:star_prop2)
       
  1219 
       
  1220 lemma postfix_prop: "y >>= (x @ y) \<Longrightarrow> x = []"
       
  1221 by (erule postfixE, induct x arbitrary:y, auto)
       
  1222 
       
  1223 lemma inj_aux:
       
  1224   "\<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> [];
       
  1225     \<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> 
       
  1226   \<Longrightarrow> (yb @ z) \<in> L\<^isub>1\<star>"
       
  1227 proof- 
       
  1228   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>  
       
  1229     (\<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>)"    
       
  1230     apply (erule Star.induct, simp)
       
  1231     apply (rule allI | rule impI | erule conjE)+
       
  1232     apply (drule app_eq_elim)
       
  1233     apply (erule disjE | erule exE | erule conjE)+
       
  1234     apply simp
       
  1235     apply (simp (no_asm) only:append_assoc[THEN sym])
       
  1236     apply (rule step) 
       
  1237     apply (simp add:equiv_str_def)
       
  1238     apply simp
       
  1239 
       
  1240     apply (erule exE | erule conjE)+    
       
  1241     apply (rotate_tac 3)
       
  1242     apply (frule_tac x = "xa @ s1" in spec)
       
  1243     apply (rotate_tac 12)
       
  1244     apply (drule_tac x = ba in spec)
       
  1245     apply (erule impE)
       
  1246     apply ( simp add:star_prop3)
       
  1247     apply (simp)
       
  1248     apply (drule postfix_prop)
       
  1249     apply simp
       
  1250     done
       
  1251   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> [];
       
  1252          \<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> 
       
  1253         \<Longrightarrow> (yb @ z) \<in> L\<^isub>1\<star>" by blast
       
  1254 qed
       
  1255 
       
  1256 
       
  1257 lemma min_postfix_exists[rule_format]:
       
  1258   "finite A \<Longrightarrow> A \<noteq> {} \<and> (\<forall> a \<in> A. \<forall> b \<in> A. ((b >>= a) \<or> (a >>= b))) 
       
  1259                 \<longrightarrow> (\<exists> min. (min \<in> A \<and> (\<forall> a \<in> A. a >>= min)))"
       
  1260 apply (erule finite.induct)
       
  1261 apply simp
       
  1262 apply simp
       
  1263 apply (case_tac "A = {}")
       
  1264 apply simp
       
  1265 apply clarsimp
       
  1266 apply (case_tac "a >>= min")
       
  1267 apply (rule_tac x = min in exI, simp)
       
  1268 apply (rule_tac x = a in exI, simp)
       
  1269 apply clarify
       
  1270 apply (rotate_tac 5)
       
  1271 apply (erule_tac x = aa in ballE) defer apply simp
       
  1272 apply (erule_tac ys = min in postfix_trans)
       
  1273 apply (erule_tac x = min in ballE) 
       
  1274 by simp+
       
  1275 
       
  1276 lemma tag_str_star_inj:
       
  1277   "tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 (y::string) \<Longrightarrow> x \<equiv>L\<^isub>1\<star>\<equiv> y"
       
  1278 proof -
       
  1279   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>"
       
  1280   proof-
       
  1281     fix x y z
       
  1282     assume tag_eq: "tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 y"
       
  1283       and x_z: "x @ z \<in> L\<^isub>1\<star>"
       
  1284     show "y @ z \<in> L\<^isub>1\<star>"
       
  1285     proof (cases "x = []")
       
  1286       case True
       
  1287       with tag_eq have "y = []" by (simp add:tag_str_STAR_def split:if_splits, blast)
       
  1288       thus ?thesis using x_z True by simp
       
  1289     next
       
  1290       case False
       
  1291       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
       
  1292         by (simp, rule_tac x = x in exI, rule_tac x = "[]" in exI, simp add:start)
       
  1293       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>}"
       
  1294         apply (rule_tac B = "{xb. \<exists> xa. x = xa @ xb}" in finite_subset)
       
  1295         apply auto
       
  1296         apply (induct x, simp)
       
  1297         apply (subgoal_tac "{xb. \<exists>xa. a # x = xa @ xb} = insert (a # x) {xb. \<exists>xa. x = xa @ xb}")
       
  1298         apply auto
       
  1299         by (case_tac xaa, simp+)
       
  1300       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>}. 
       
  1301                         \<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>}.
       
  1302                           ((b >>= a) \<or> (a >>= b))"
       
  1303         by (auto simp:postfix_def, drule app_eq_elim, blast)
       
  1304       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>} 
       
  1305                 \<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)"
       
  1306         using finite_set not_empty comparable
       
  1307         apply (drule_tac min_postfix_exists, simp)
       
  1308         by (erule exE, rule_tac x = min in exI, auto)
       
  1309       then obtain min xa where x_decom: "x = xa @ min \<and> xa \<in> L\<^isub>1\<star>"
       
  1310         and min_not_empty: "min \<noteq> []"
       
  1311         and min_z_in_star: "min @ z \<in> L\<^isub>1\<star>"
       
  1312         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
       
  1313       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)
       
  1314       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
       
  1315       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> [] " 
       
  1316         by (simp add:tag_str_STAR_def equiv_class_def equiv_str_def split:if_splits, blast)        
       
  1317       then obtain ya yb where y_decom: "y = ya @ yb"
       
  1318                           and ya_in_star: "ya \<in> L\<^isub>1\<star>"
       
  1319                           and yb_not_empty: "yb \<noteq> []"
       
  1320                           and min_yb_eq: "min \<equiv>L\<^isub>1\<equiv> yb"  by blast
       
  1321       from min_z_in_star min_yb_eq min_not_empty is_min x_decom
       
  1322       have "yb @ z \<in> L\<^isub>1\<star>"
       
  1323         by (rule_tac x = x and xa = xa in inj_aux, simp+)
       
  1324       thus ?thesis using ya_in_star y_decom
       
  1325         by (auto dest:star_prop)        
       
  1326     qed
       
  1327   qed
       
  1328   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"
       
  1329     by (auto intro:aux simp:equiv_str_def)
       
  1330 qed
       
  1331 
       
  1332 lemma quot_star:  
       
  1333   assumes finite1: "finite (QUOT L\<^isub>1)"
       
  1334   shows "finite (QUOT (L\<^isub>1\<star>))"
       
  1335 proof (rule_tac f = "(op `) (tag_str_STAR L\<^isub>1)" in finite_imageD)
       
  1336   show "finite (op ` (tag_str_STAR L\<^isub>1) ` QUOT (L\<^isub>1\<star>))"
       
  1337     using finite_tag_image tag_str_star_range_finite finite1
       
  1338     by auto
       
  1339 next
       
  1340   show "inj_on (op ` (tag_str_STAR L\<^isub>1)) (QUOT (L\<^isub>1\<star>))"
       
  1341     apply (rule_tac str_inj_imps)
       
  1342     by (erule_tac tag_str_star_inj)
       
  1343 qed
       
  1344 
       
  1345 lemma other_direction:
       
  1346   "Lang = L (r::rexp) \<Longrightarrow> finite (QUOT Lang)"
       
  1347 apply (induct arbitrary:Lang rule:rexp.induct)
       
  1348 apply (simp add:QUOT_def equiv_class_def equiv_str_def)
       
  1349 by (simp_all add:quot_lambda quot_single quot_seq quot_alt quot_star)  
       
  1350 
       
  1351 end