Myhill.thy
changeset 30 f5db9e08effc
parent 29 c64241fa4dff
child 31 b6815473ee2e
--- a/Myhill.thy	Fri Jan 07 14:25:23 2011 +0000
+++ b/Myhill.thy	Mon Jan 24 11:29:55 2011 +0000
@@ -1,12 +1,15 @@
-theory MyhillNerode
-  imports "Main" "List_Prefix"
+theory Myhill
+  imports Main List_Prefix
 begin
 
-text {* sequential composition of languages *}
+section {* Preliminary definitions *}
+
+text {* Sequential composition of two languages @{text "L1"} and @{text "L2"} *}
 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}"
 
+text {* Transitive closure of language @{text "L"}. *}
 inductive_set
   Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102)
   for L :: "string set"
@@ -14,6 +17,8 @@
   start[intro]: "[] \<in> L\<star>"
 | step[intro]:  "\<lbrakk>s1 \<in> L; s2 \<in> L\<star>\<rbrakk> \<Longrightarrow> s1@s2 \<in> L\<star>" 
 
+text {* Some properties of operator @{text ";;"}.*}
+
 lemma seq_union_distrib:
   "(A \<union> B) ;; C = (A ;; C) \<union> (B ;; C)"
 by (auto simp:Seq_def)
@@ -28,23 +33,92 @@
 apply blast
 by (metis append_assoc)
 
+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_intro2: "y \<in> lang \<Longrightarrow> y \<in> lang\<star>"
+by (drule step[of y lang "[]"], auto simp:start)
+
+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 star_decom': 
+  "\<lbrakk>x \<in> lang\<star>; x \<noteq> []\<rbrakk> \<Longrightarrow> \<exists>a b. x = a @ b \<and> a \<in> lang\<star> \<and> b \<in> lang"
+apply (induct x rule:Star.induct, simp)
+apply (case_tac "s2 = []")
+apply (rule_tac x = "[]" in exI, rule_tac x = s1 in exI, simp add:start)
+apply (simp, (erule exE| erule conjE)+)
+by (rule_tac x = "s1 @ a" in exI, rule_tac x = b in exI, simp add:step)
+
+text {* Ardens lemma expressed at the level of language, rather than the level of regular expression. *}
+
 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 
-  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
-    by (auto) (metis append_assoc)+
-  finally show "X = X ;; A \<union> B" using eq by auto
+  have "A\<star> =  {[]} \<union> A\<star> ;; A" 
+    by (auto simp:Seq_def star_intro3 star_decom')  
+  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" 
+    by (simp only:seq_assoc)
+  finally show "X = X ;; A \<union> B" 
+    using eq by blast 
 next
-  assume "X = X ;; A \<union> B"
-  then have "B \<subseteq> X" "X ;; A \<subseteq> X" by auto
-  thus "X = B ;; A\<star>" sorry
+  assume eq': "X = X ;; A \<union> B"
+  hence c1': "\<And> x. x \<in> B \<Longrightarrow> x \<in> X" 
+    and c2': "\<And> x y. \<lbrakk>x \<in> X; y \<in> A\<rbrakk> \<Longrightarrow> x @ y \<in> X" 
+    using Seq_def by auto
+  show "X = B ;; A\<star>" 
+  proof
+    show "B ;; A\<star> \<subseteq> X"
+    proof-
+      { fix x y
+        have "\<lbrakk>y \<in> A\<star>; x \<in> X\<rbrakk> \<Longrightarrow> x @ y \<in> X "
+          apply (induct arbitrary:x rule:Star.induct, simp)
+          by (auto simp only:append_assoc[THEN sym] dest:c2')
+      } thus ?thesis using c1' by (auto simp:Seq_def) 
+    qed
+  next
+    show "X \<subseteq> B ;; A\<star>"
+    proof-
+      { fix x 
+        have "x \<in> X \<Longrightarrow> x \<in> B ;; A\<star>"
+        proof (induct x taking:length rule:measure_induct)
+          fix z
+          assume hyps: 
+            "\<forall>y. length y < length z \<longrightarrow> y \<in> X \<longrightarrow> y \<in> B ;; A\<star>" 
+            and z_in: "z \<in> X"
+          show "z \<in> B ;; A\<star>"
+          proof (cases "z \<in> B")
+            case True thus ?thesis by (auto simp:Seq_def start)
+          next
+            case False hence "z \<in> X ;; A" using eq' z_in by auto
+            then obtain za zb where za_in: "za \<in> X" 
+              and zab: "z = za @ zb \<and> zb \<in> A" and zbne: "zb \<noteq> []" 
+              using nemp unfolding Seq_def by blast
+            from zbne zab have "length za < length z" by auto
+            with za_in hyps have "za \<in> B ;; A\<star>" by blast
+            hence "za @ zb \<in> B ;; A\<star>" using zab 
+              by (clarsimp simp:Seq_def, blast dest:star_intro3)
+            thus ?thesis using zab by simp       
+          qed
+        qed 
+      } thus ?thesis by blast
+    qed
+  qed
 qed
 
+
+text {* The syntax of regular expressions is defined by the datatype @{text "rexp"}. *}
 datatype rexp =
   NULL
 | EMPTY
@@ -53,11 +127,20 @@
 | ALT rexp rexp
 | STAR rexp
 
+
+text {* 
+  The following @{text "L"} is an overloaded operator, where @{text "L(x)"} evaluates to 
+  the language represented by the syntactic object @{text "x"}.
+*}
 consts L:: "'a \<Rightarrow> string set"
 
+
+text {* 
+  The @{text "L(rexp)"} for regular expression @{text "rexp"} is defined by the 
+  following overloading function @{text "L_rexp"}.
+*}
 overloading L_rexp \<equiv> "L::  rexp \<Rightarrow> string set"
 begin
-
 fun
   L_rexp :: "rexp \<Rightarrow> string set"
 where
@@ -69,157 +152,258 @@
   | "L_rexp (STAR r) = (L_rexp r)\<star>"
 end
 
+text {*
+  To obtain equational system out of finite set of equivalent classes, a fold operation
+  on finite set @{text "folds"} is defined. The use of @{text "SOME"} makes @{text "fold"}
+  more robust than the @{text "fold"} in Isabelle library. The expression @{text "folds f"}
+  makes sense when @{text "f"} is not @{text "associative"} and @{text "commutitive"},
+  while @{text "fold f"} does not.  
+*}
+
 definition 
   folds :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b"
 where
   "folds f z S \<equiv> SOME x. fold_graph f z S x"
 
+text {* 
+  The following lemma assures that the arbitrary choice made by the @{text "SOME"} in @{text "folds"}
+  does not affect the @{text "L"}-value of the resultant regular expression. 
+  *}
 lemma folds_alt_simp [simp]:
   "finite rs \<Longrightarrow> L (folds ALT NULL rs) = \<Union> (L ` rs)"
 apply (rule set_ext, simp add:folds_def)
 apply (rule someI2_ex, erule finite_imp_fold_graph)
 by (erule fold_graph.induct, auto)
 
+(* Just a technical lemma. *)
 lemma [simp]:
   shows "(x, y) \<in> {(x, y). P x y} \<longleftrightarrow> P x y"
 by simp
 
-definition
-  str_eq ("_ \<approx>_ _")
-where
-  "x \<approx>Lang y \<equiv> (\<forall>z. x @ z \<in> Lang \<longleftrightarrow> y @ z \<in> Lang)"
 
+text {*
+  @{text "\<approx>L"} is an equivalent class defined by language @{text "Lang"}.
+*}
 definition
   str_eq_rel ("\<approx>_")
 where
-  "\<approx>Lang \<equiv> {(x, y). x \<approx>Lang y}"
+  "\<approx>Lang \<equiv> {(x, y).  (\<forall>z. x @ z \<in> Lang \<longleftrightarrow> y @ z \<in> Lang)}"
+
+text {* 
+  Among equivlant clases of @{text "\<approx>Lang"}, the set @{text "finals(Lang)"} singles out 
+  those which contains strings from @{text "Lang"}.
+*}
 
-definition
-  final :: "string set \<Rightarrow> string set \<Rightarrow> bool"
-where
-  "final X Lang \<equiv> (X \<in> UNIV // \<approx>Lang) \<and> (\<forall>s \<in> X. s \<in> Lang)"
+definition 
+   "finals Lang \<equiv> {\<approx>Lang `` {x} | x . x \<in> Lang}"
 
+text {* 
+  The following lemma show the relationshipt between @{text "finals(Lang)"} and @{text "Lang"}.
+*}
 lemma lang_is_union_of_finals: 
-  "Lang = \<Union> {X. final X Lang}"
+  "Lang = \<Union> finals(Lang)"
 proof 
-  show "Lang \<subseteq> \<Union> {X. final X Lang}"
+  show "Lang \<subseteq> \<Union> (finals Lang)"
   proof
     fix x
     assume "x \<in> Lang"   
-    thus "x \<in> \<Union> {X. final X Lang}"
-      apply (simp, rule_tac x = "(\<approx>Lang) `` {x}" in exI)      
-      apply (auto simp:final_def quotient_def Image_def str_eq_rel_def str_eq_def)
-      by (drule_tac x = "[]" in spec, simp)
+    thus "x \<in> \<Union> (finals Lang)"
+      apply (simp add:finals_def, rule_tac x = "(\<approx>Lang) `` {x}" in exI)
+      by (auto simp:Image_def str_eq_rel_def)    
   qed
 next
-  show "\<Union>{X. final X Lang} \<subseteq> Lang"
-    by (auto simp:final_def)
+  show "\<Union> (finals Lang) \<subseteq> Lang"
+    apply (clarsimp simp:finals_def str_eq_rel_def)
+    by (drule_tac x = "[]" in spec, auto)
 qed
 
-section {* finite \<Rightarrow> regular *}
+section {* Direction @{text "finite partition \<Rightarrow> regular language"}*}
+
+text {* 
+  The relationship between equivalent classes can be described by an
+  equational system.
+  For example, in equational system \eqref{example_eqns},  $X_0, X_1$ are equivalent 
+  classes. The first equation says every string in $X_0$ is obtained either by
+  appending one $b$ to a string in $X_0$ or by appending one $a$ to a string in
+  $X_1$ or just be an empty string (represented by the regular expression $\lambda$). Similary,
+  the second equation tells how the strings inside $X_1$ are composed.
+  \begin{equation}\label{example_eqns}
+    \begin{aligned}
+      X_0 & = X_0 b + X_1 a + \lambda \\
+      X_1 & = X_0 a + X_1 b
+    \end{aligned}
+  \end{equation}
+  The summands on the right hand side is represented by the following data type
+  @{text "rhs_item"}, mnemonic for 'right hand side item'.
+  Generally, there are two kinds of right hand side items, one kind corresponds to
+  pure regular expressions, like the $\lambda$ in \eqref{example_eqns}, the other kind corresponds to
+  transitions from one one equivalent class to another, like the $X_0 b, X_1 a$ etc.
+  *}
 
 datatype rhs_item = 
    Lam "rexp"                           (* Lambda *)
- | Trn "string set" "rexp"              (* Transition *)
+ | Trn "(string set)" "rexp"              (* Transition *)
 
-fun the_Trn:: "rhs_item \<Rightarrow> (string set \<times> rexp)"
-where "the_Trn (Trn Y r) = (Y, r)"
+text {*
+  In this formalization, pure regular expressions like $\lambda$ is 
+  repsented by @{text "Lam(EMPTY)"}, while transitions like $X_0 a$ is represented by $Trn~X_0~(CHAR~a)$.
+  *}
+
+text {*
+  The functions @{text "the_r"} and @{text "the_Trn"} are used to extract
+  subcomponents from right hand side items.
+  *}
 
 fun the_r :: "rhs_item \<Rightarrow> rexp"
 where "the_r (Lam r) = r"
 
+fun the_Trn:: "rhs_item \<Rightarrow> (string set \<times> rexp)"
+where "the_Trn (Trn Y r) = (Y, r)"
+
+text {*
+  Every right hand side item @{text "itm"} defines a string set given 
+  @{text "L(itm)"}, defined as:
+*}
 overloading L_rhs_e \<equiv> "L:: rhs_item \<Rightarrow> string set"
 begin
-fun L_rhs_e:: "rhs_item \<Rightarrow> string set"
-where
-  "L_rhs_e (Lam r) = L r" |
-  "L_rhs_e (Trn X r) = X ;; L r"
+  fun L_rhs_e:: "rhs_item \<Rightarrow> string set"
+  where
+     "L_rhs_e (Lam r) = L r" |
+     "L_rhs_e (Trn X r) = X ;; L r"
 end
 
+text {*
+  The right hand side of every equation is represented by a set of
+  items. The string set defined by such a set @{text "itms"} is given
+  by @{text "L(itms)"}, defined as:
+*}
+
 overloading L_rhs \<equiv> "L:: rhs_item set \<Rightarrow> string set"
 begin
-fun L_rhs:: "rhs_item set \<Rightarrow> string set"
-where
-  "L_rhs rhs = \<Union> (L ` rhs)"
+   fun L_rhs:: "rhs_item set \<Rightarrow> string set"
+   where "L_rhs rhs = \<Union> (L ` rhs)"
 end
 
+text {* 
+  Given a set of equivalent classses @{text "CS"} and one equivalent class @{text "X"} among
+  @{text "CS"}, the term @{text "init_rhs CS X"} is used to extract the right hand side of
+  the equation describing the formation of @{text "X"}. The definition of @{text "init_rhs"}
+  is:
+  *}
+
 definition
-  "init_rhs CS X \<equiv>  if ([] \<in> X)
-                    then {Lam EMPTY} \<union> {Trn Y (CHAR c)| Y c. Y \<in> CS \<and> Y ;; {[c]} \<subseteq> X}
-                    else {Trn Y (CHAR c)| Y c. Y \<in> CS \<and> Y ;; {[c]} \<subseteq> X}"
+  "init_rhs CS X \<equiv>  
+      if ([] \<in> X) then 
+          {Lam(EMPTY)} \<union> {Trn Y (CHAR c) | Y c. Y \<in> CS \<and> Y ;; {[c]} \<subseteq> X}
+      else 
+          {Trn Y (CHAR c)| Y c. Y \<in> CS \<and> Y ;; {[c]} \<subseteq> X}"
 
-definition
-  "eqs CS \<equiv> {(X, init_rhs CS X)|X.  X \<in> CS}"
+text {*
+  In the definition of @{text "init_rhs"}, the term 
+  @{text "{Trn Y (CHAR c)| Y c. Y \<in> CS \<and> Y ;; {[c]} \<subseteq> X}"} appearing on both branches
+  describes the formation of strings in @{text "X"} out of transitions, while 
+  the term @{text "{Lam(EMPTY)}"} describes the empty string which is intrinsically contained in
+  @{text "X"} rather than by transition. This @{text "{Lam(EMPTY)}"} corresponds to 
+  the $\lambda$ in \eqref{example_eqns}.
 
+  With the help of @{text "init_rhs"}, the equitional system descrbing the formation of every
+  equivalent class inside @{text "CS"} is given by the following @{text "eqs(CS)"}.
+  *}
+
+definition "eqs CS \<equiv> {(X, init_rhs CS X) | X.  X \<in> CS}"
 (************ arden's lemma variation ********************)
 
+text {* 
+  The following @{text "items_of rhs X"} returns all @{text "X"}-items in @{text "rhs"}.
+  *}
 definition
   "items_of rhs X \<equiv> {Trn X r | r. (Trn X r) \<in> rhs}"
 
+text {* 
+  The following @{text "rexp_of rhs X"} combines all regular expressions in @{text "X"}-items
+  using @{text "ALT"} to form a single regular expression. 
+  It will be used later to implement @{text "arden_variate"} and @{text "rhs_subst"}.
+  *}
+
+definition 
+  "rexp_of rhs X \<equiv> folds ALT NULL ((snd o the_Trn) ` items_of rhs X)"
+
+text {* 
+  The following @{text "lam_of rhs"} returns all pure regular expression items in @{text "rhs"}.
+  *}
+
 definition
   "lam_of rhs \<equiv> {Lam r | r. Lam r \<in> rhs}"
 
-definition 
-  "rexp_of rhs X \<equiv> folds ALT NULL ((snd o the_Trn) ` items_of rhs X)"
+text {*
+  The following @{text "rexp_of_lam rhs"} combines pure regular expression items in @{text "rhs"}
+  using @{text "ALT"} to form a single regular expression. 
+  When all variables inside @{text "rhs"} are eliminated, @{text "rexp_of_lam rhs"}
+  is used to compute compute the regular expression corresponds to @{text "rhs"}.
+  *}
 
 definition
   "rexp_of_lam rhs \<equiv> folds ALT NULL (the_r ` lam_of rhs)"
 
+text {*
+  The following @{text "attach_rexp rexp' itm"} attach 
+  the regular expression @{text "rexp'"} to
+  the right of right hand side item @{text "itm"}.
+  *}
+
 fun attach_rexp :: "rexp \<Rightarrow> rhs_item \<Rightarrow> rhs_item"
 where
-  "attach_rexp r' (Lam r)   = Lam (SEQ r r')"
-| "attach_rexp r' (Trn X r) = Trn X (SEQ r r')"
+  "attach_rexp rexp' (Lam rexp)   = Lam (SEQ rexp rexp')"
+| "attach_rexp rexp' (Trn X rexp) = Trn X (SEQ rexp rexp')"
+
+text {* 
+  The following @{text "append_rhs_rexp rhs rexp"} attaches 
+  @{text "rexp"} to every item in @{text "rhs"}.
+  *}
 
 definition
-  "append_rhs_rexp rhs r \<equiv> (attach_rexp r) ` rhs"
+  "append_rhs_rexp rhs rexp \<equiv> (attach_rexp rexp) ` rhs"
 
+text {*
+  With the help of the two functions immediately above, Ardens'
+  transformation on right hand side @{text "rhs"} is implemented
+  by the following function @{text "arden_variate X rhs"}.
+  After this transformation, the recursive occurent of @{text "X"}
+  in @{text "rhs"} will be eliminated, while the 
+  string set defined by @{text "rhs"} is kept unchanged.
+  *}
 definition 
-  "arden_variate X rhs \<equiv> append_rhs_rexp (rhs - items_of rhs X) (STAR (rexp_of rhs X))"
+  "arden_variate X rhs \<equiv> 
+        append_rhs_rexp (rhs - items_of rhs X) (STAR (rexp_of rhs X))"
 
 
 (*********** substitution of ES *************)
 
-text {* rhs_subst rhs X xrhs: substitude all occurence of X in rhs with xrhs *}
+text {* 
+  Suppose the equation defining @{text "X"} is $X = xrhs$,
+  the purpose of @{text "rhs_subst"} is to substitute all occurences of @{text "X"} in
+  @{text "rhs"} by @{text "xrhs"}.
+  A litte thought may reveal that the final result
+  should be: first append $(a_1 | a_2 | \ldots | a_n)$ to every item of @{text "xrhs"} and then
+  union the result with all non-@{text "X"}-items of @{text "rhs"}.
+ *}
 definition 
-  "rhs_subst rhs X xrhs \<equiv> (rhs - (items_of rhs X)) \<union> (append_rhs_rexp xrhs (rexp_of rhs X))"
+  "rhs_subst rhs X xrhs \<equiv> 
+        (rhs - (items_of rhs X)) \<union> (append_rhs_rexp xrhs (rexp_of rhs X))"
+
+text {*
+  Suppose the equation defining @{text "X"} is $X = xrhs$, the follwing
+  @{text "eqs_subst ES X xrhs"} substitute @{text "xrhs"} into every equation
+  of the equational system @{text "ES"}.
+  *}
 
 definition
   "eqs_subst ES X xrhs \<equiv> {(Y, rhs_subst yrhs X xrhs) | Y yrhs. (Y, yrhs) \<in> ES}"
 
 text {*
-  Inv: Invairance of the equation-system, during the decrease of the equation-system, Inv holds.
-*}
-
-definition 
-  "distinct_equas ES \<equiv> \<forall> X rhs rhs'. (X, rhs) \<in> ES \<and> (X, rhs') \<in> ES \<longrightarrow> rhs = rhs'"
-
-definition 
-  "valid_eqns ES \<equiv> \<forall> X rhs. (X, rhs) \<in> ES \<longrightarrow> (X = L rhs)"
-
-definition 
-  "rhs_nonempty rhs \<equiv> (\<forall> Y r. Trn Y r \<in> rhs \<longrightarrow> [] \<notin> L r)"
-
-definition 
-  "ardenable ES \<equiv> \<forall> X rhs. (X, rhs) \<in> ES \<longrightarrow> rhs_nonempty rhs"
-
-definition 
-  "non_empty ES \<equiv> \<forall> X rhs. (X, rhs) \<in> ES \<longrightarrow> X \<noteq> {}"
-
-definition
-  "finite_rhs ES \<equiv> \<forall> X rhs. (X, rhs) \<in> ES \<longrightarrow> finite rhs"
-
-definition 
-  "classes_of rhs \<equiv> {X. \<exists> r. Trn X r \<in> rhs}"
-
-definition
-  "lefts_of ES \<equiv> {Y | Y yrhs. (Y, yrhs) \<in> ES}"
-
-definition 
-  "self_contained ES \<equiv> \<forall> (X, xrhs) \<in> ES. classes_of xrhs \<subseteq> lefts_of ES"
-
-definition 
-  "Inv ES \<equiv> valid_eqns ES \<and> finite ES \<and> distinct_equas ES \<and> ardenable ES \<and> 
-            non_empty ES \<and> finite_rhs ES \<and> self_contained ES"
+  The computation of regular expressions for equivalent classes is accomplished
+  using a iteration principle given by the following lemma.
+  *}
 
 lemma wf_iter [rule_format]: 
   fixes f
@@ -248,7 +432,88 @@
   qed
 qed
 
-text {* ************* basic properties of definitions above ************************ *}
+text {*
+  The @{text "P"} in lemma @{text "wf_iter"} is an invaiant kept throughout the iteration procedure.
+  The particular invariant used to solve our problem is defined by function @{text "Inv(ES)"},
+  an invariant over equal system @{text "ES"}.
+  Every definition starting next till @{text "Inv"} stipulates a property to be satisfied by @{text "ES"}.
+*}
+
+text {* 
+  Every variable is defined at most onece in @{text "ES"}.
+  *}
+definition 
+  "distinct_equas ES \<equiv> 
+            \<forall> X rhs rhs'. (X, rhs) \<in> ES \<and> (X, rhs') \<in> ES \<longrightarrow> rhs = rhs'"
+text {* 
+  Every equation in @{text "ES"} (represented by @{text "(X, rhs)"}) is valid, i.e. @{text "(X = L rhs)"}.
+  *}
+definition 
+  "valid_eqns ES \<equiv> \<forall> X rhs. (X, rhs) \<in> ES \<longrightarrow> (X = L rhs)"
+
+text {*
+  @{text "rhs_nonempty rhs"} requires regular expressions occuring in transitional 
+  items of @{text "rhs"} does not contain empty string. This is necessary for
+  the application of Arden's transformation to @{text "rhs"}.
+  *}
+definition 
+  "rhs_nonempty rhs \<equiv> (\<forall> Y r. Trn Y r \<in> rhs \<longrightarrow> [] \<notin> L r)"
+
+text {*
+  @{text "ardenable ES"} requires that Arden's transformation is applicable
+  to every equation of equational system @{text "ES"}.
+  *}
+definition 
+  "ardenable ES \<equiv> \<forall> X rhs. (X, rhs) \<in> ES \<longrightarrow> rhs_nonempty rhs"
+
+(* The following non_empty seems useless. *)
+definition 
+  "non_empty ES \<equiv> \<forall> X rhs. (X, rhs) \<in> ES \<longrightarrow> X \<noteq> {}"
+
+text {*
+  The following @{text "finite_rhs ES"} requires every equation in @{text "rhs"} be finite.
+  *}
+definition
+  "finite_rhs ES \<equiv> \<forall> X rhs. (X, rhs) \<in> ES \<longrightarrow> finite rhs"
+
+text {*
+  The following @{text "classes_of rhs"} returns all variables (or equivalent classes)
+  occuring in @{text "rhs"}.
+  *}
+definition 
+  "classes_of rhs \<equiv> {X. \<exists> r. Trn X r \<in> rhs}"
+
+text {*
+  The following @{text "lefts_of ES"} returns all variables 
+  defined by equational system @{text "ES"}.
+  *}
+definition
+  "lefts_of ES \<equiv> {Y | Y yrhs. (Y, yrhs) \<in> ES}"
+
+text {*
+  The following @{text "self_contained ES"} requires that every
+  variable occuring on the right hand side of equations is already defined by some
+  equation in @{text "ES"}.
+  *}
+definition 
+  "self_contained ES \<equiv> \<forall> (X, xrhs) \<in> ES. classes_of xrhs \<subseteq> lefts_of ES"
+
+
+text {*
+  The invariant @{text "Inv(ES)"} is obtained by conjunctioning all the previous
+  defined constaints on @{text "ES"}.
+  *}
+definition 
+  "Inv ES \<equiv> valid_eqns ES \<and> finite ES \<and> distinct_equas ES \<and> ardenable ES \<and> 
+                non_empty ES \<and> finite_rhs ES \<and> self_contained ES"
+
+subsection {* Proof for this direction *}
+
+
+
+text {*
+  The following are some basic properties of the above definitions.
+*}
 
 lemma L_rhs_union_distrib:
   " L (A::rhs_item set) \<union> L B = L (A \<union> B)"
@@ -319,11 +584,14 @@
 by (auto simp:lefts_of_def)
 
 
-text {* ******BEGIN: proving the initial equation-system satisfies Inv ****** *}
+text {*
+  The following several lemmas until @{text "init_ES_satisfy_Inv"} are
+  to prove that initial equational system satisfies invariant @{text "Inv"}.
+  *}
 
 lemma defined_by_str:
   "\<lbrakk>s \<in> X; X \<in> UNIV // (\<approx>Lang)\<rbrakk> \<Longrightarrow> X = (\<approx>Lang) `` {s}"
-by (auto simp:quotient_def Image_def str_eq_rel_def str_eq_def)
+by (auto simp:quotient_def Image_def str_eq_rel_def)
 
 lemma every_eqclass_has_transition:
   assumes has_str: "s @ [c] \<in> X"
@@ -339,10 +607,10 @@
   then have "Y ;; {[c]} \<subseteq> X" 
     unfolding Y_def Image_def Seq_def
     unfolding str_eq_rel_def
-    by (auto) (simp add: str_eq_def)
+    by clarsimp
   moreover
   have "s \<in> Y" unfolding Y_def 
-    unfolding Image_def str_eq_rel_def str_eq_def by simp
+    unfolding Image_def str_eq_rel_def by simp
   ultimately show thesis by (blast intro: that)
 qed
 
@@ -369,7 +637,8 @@
         and "Y ;; {[c]} \<subseteq> X"
         and "clist \<in> Y"
         using decom "(1)" every_eqclass_has_transition by blast
-      hence "x \<in> L {Trn Y (CHAR c)| Y c. Y \<in> UNIV // (\<approx>Lang) \<and> Y ;; {[c]} \<subseteq> X}"
+      hence 
+        "x \<in> L {Trn Y (CHAR c)| Y c. Y \<in> UNIV // (\<approx>Lang) \<and> Y ;; {[c]} \<subseteq> X}"
         using "(1)" decom
         by (simp, rule_tac x = "Trn Y (CHAR c)" in exI, simp add:Seq_def)
       thus ?thesis using X_in_eqs "(1)"
@@ -412,7 +681,7 @@
   moreover have "valid_eqns (eqs (UNIV // (\<approx>Lang)))"
     using l_eq_r_in_eqs by (simp add:valid_eqns_def)
   moreover have "non_empty (eqs (UNIV // (\<approx>Lang)))"
-    by (auto simp:non_empty_def eqs_def quotient_def Image_def str_eq_rel_def str_eq_def)
+    by (auto simp:non_empty_def eqs_def quotient_def Image_def str_eq_rel_def)
   moreover have "finite_rhs (eqs (UNIV // (\<approx>Lang)))"
     using finite_init_rhs[OF finite_CS] 
     by (auto simp:finite_rhs_def eqs_def)
@@ -421,8 +690,11 @@
   ultimately show ?thesis by (simp add:Inv_def)
 qed
 
-text {* ****** BEGIN: proving every equation-system's iteration step satisfies Inv ***** *}
-
+text {*
+  From this point until @{text "iteration_step"}, we are trying to prove 
+  that there exists iteration steps which keep @{text "Inv(ES)"} while
+  decreasing the size of @{text "ES"} with every iteration.
+  *}
 lemma arden_variate_keeps_eq:
   assumes l_eq_r: "X = L rhs"
   and not_empty: "[] \<notin> L (rexp_of rhs X)"
@@ -506,7 +778,8 @@
   assumes finite:"finite (ES:: (string set \<times> rhs_item set) set)"
   shows "finite (eqs_subst ES Y yrhs)"
 proof -
-  have "finite {(Ya, rhs_subst yrhsa Y yrhs) |Ya yrhsa. (Ya, yrhsa) \<in> ES}" (is "finite ?A")
+  have "finite {(Ya, rhs_subst yrhsa Y yrhs) |Ya yrhsa. (Ya, yrhsa) \<in> ES}" 
+                                                                  (is "finite ?A")
   proof-
     def eqns' \<equiv> "{((Ya::string set), yrhsa)| Ya yrhsa. (Ya, yrhsa) \<in> ES}"
     def h \<equiv> "\<lambda> ((Ya::string set), yrhsa). (Ya, rhs_subst yrhsa Y yrhs)"
@@ -537,14 +810,17 @@
 by (auto simp:lefts_of_def eqs_subst_def)
 
 lemma rhs_subst_updates_cls:
-  "X \<notin> classes_of xrhs \<Longrightarrow> classes_of (rhs_subst rhs X xrhs) = classes_of rhs \<union> classes_of xrhs - {X}"
-apply (simp only:rhs_subst_def append_rhs_keeps_cls classes_of_union_distrib[THEN sym])
+  "X \<notin> classes_of xrhs \<Longrightarrow> 
+      classes_of (rhs_subst rhs X xrhs) = classes_of rhs \<union> classes_of xrhs - {X}"
+apply (simp only:rhs_subst_def append_rhs_keeps_cls 
+                              classes_of_union_distrib[THEN sym])
 by (auto simp:classes_of_def items_of_def)
 
 lemma eqs_subst_keeps_self_contained:
   fixes Y
   assumes sc: "self_contained (ES \<union> {(Y, yrhs)})" (is "self_contained ?A")
-  shows "self_contained (eqs_subst ES Y (arden_variate Y yrhs))" (is "self_contained ?B")
+  shows "self_contained (eqs_subst ES Y (arden_variate Y yrhs))" 
+                                                   (is "self_contained ?B")
 proof-
   { fix X xrhs'
     assume "(X, xrhs') \<in> ?B"
@@ -556,15 +832,18 @@
       have "lefts_of ?B = lefts_of ES" by (auto simp add:lefts_of_def eqs_subst_def)
       moreover have "classes_of xrhs' \<subseteq> lefts_of ES"
       proof-
-        have "classes_of xrhs' \<subseteq> classes_of xrhs \<union> classes_of (arden_variate Y yrhs) - {Y}"
+        have "classes_of xrhs' \<subseteq> 
+                        classes_of xrhs \<union> classes_of (arden_variate Y yrhs) - {Y}"
         proof-
-          have "Y \<notin> classes_of (arden_variate Y yrhs)" using arden_variate_removes_cl by simp
+          have "Y \<notin> classes_of (arden_variate Y yrhs)" 
+            using arden_variate_removes_cl by simp
           thus ?thesis using xrhs_xrhs' by (auto simp:rhs_subst_updates_cls)
         qed
         moreover have "classes_of xrhs \<subseteq> lefts_of ES \<union> {Y}" using X_in sc
           apply (simp only:self_contained_def lefts_of_union_distrib[THEN sym])
           by (drule_tac x = "(X, xrhs)" in bspec, auto simp:lefts_of_def)
-        moreover have "classes_of (arden_variate Y yrhs) \<subseteq> lefts_of ES \<union> {Y}" using sc
+        moreover have "classes_of (arden_variate Y yrhs) \<subseteq> lefts_of ES \<union> {Y}" 
+          using sc 
           by (auto simp add:arden_variate_removes_cl self_contained_def lefts_of_def)
         ultimately show ?thesis by auto
       qed
@@ -577,44 +856,57 @@
   assumes Inv_ES: "Inv (ES \<union> {(Y, yrhs)})"
   shows "Inv (eqs_subst ES Y (arden_variate Y yrhs))"
 proof -  
-  have finite_yrhs: "finite yrhs" using Inv_ES by (auto simp:Inv_def finite_rhs_def)
-  have nonempty_yrhs: "rhs_nonempty yrhs" using Inv_ES by (auto simp:Inv_def ardenable_def)
-  have Y_eq_yrhs: "Y = L yrhs" using Inv_ES by (simp only:Inv_def valid_eqns_def, blast)
-
-  have "distinct_equas (eqs_subst ES Y (arden_variate Y yrhs))" using Inv_ES
+  have finite_yrhs: "finite yrhs" 
+    using Inv_ES by (auto simp:Inv_def finite_rhs_def)
+  have nonempty_yrhs: "rhs_nonempty yrhs" 
+    using Inv_ES by (auto simp:Inv_def ardenable_def)
+  have Y_eq_yrhs: "Y = L yrhs" 
+    using Inv_ES by (simp only:Inv_def valid_eqns_def, blast)
+  have "distinct_equas (eqs_subst ES Y (arden_variate Y yrhs))" 
+    using Inv_ES
     by (auto simp:distinct_equas_def eqs_subst_def Inv_def)
-  moreover have "finite (eqs_subst ES Y (arden_variate Y yrhs))" using Inv_ES 
-    by (simp add:Inv_def eqs_subst_keeps_finite)
+  moreover have "finite (eqs_subst ES Y (arden_variate Y yrhs))" 
+    using Inv_ES by (simp add:Inv_def eqs_subst_keeps_finite)
   moreover have "finite_rhs (eqs_subst ES Y (arden_variate Y yrhs))"
   proof-
-    have "finite_rhs ES" using Inv_ES by (simp add:Inv_def finite_rhs_def)
+    have "finite_rhs ES" using Inv_ES 
+      by (simp add:Inv_def finite_rhs_def)
     moreover have "finite (arden_variate Y yrhs)"
     proof -
-      have "finite yrhs" using Inv_ES by (auto simp:Inv_def finite_rhs_def)
+      have "finite yrhs" using Inv_ES 
+        by (auto simp:Inv_def finite_rhs_def)
       thus ?thesis using arden_variate_keeps_finite by simp
     qed
-    ultimately show ?thesis by (simp add:eqs_subst_keeps_finite_rhs)
+    ultimately show ?thesis 
+      by (simp add:eqs_subst_keeps_finite_rhs)
   qed
   moreover have "ardenable (eqs_subst ES Y (arden_variate Y yrhs))"
   proof - 
     { fix X rhs
       assume "(X, rhs) \<in> ES"
-      hence "rhs_nonempty rhs"  using prems Inv_ES  by (simp add:Inv_def ardenable_def)
-      with nonempty_yrhs have "rhs_nonempty (rhs_subst rhs Y (arden_variate Y yrhs))"
-        by (simp add:nonempty_yrhs rhs_subst_keeps_nonempty arden_variate_keeps_nonempty)
+      hence "rhs_nonempty rhs"  using prems Inv_ES  
+        by (simp add:Inv_def ardenable_def)
+      with nonempty_yrhs 
+      have "rhs_nonempty (rhs_subst rhs Y (arden_variate Y yrhs))"
+        by (simp add:nonempty_yrhs 
+               rhs_subst_keeps_nonempty arden_variate_keeps_nonempty)
     } thus ?thesis by (auto simp add:ardenable_def eqs_subst_def)
   qed
   moreover have "valid_eqns (eqs_subst ES Y (arden_variate Y yrhs))"
   proof-
-    have "Y = L (arden_variate Y yrhs)" using Y_eq_yrhs Inv_ES finite_yrhs nonempty_yrhs      
-        by (rule_tac arden_variate_keeps_eq, (simp add:rexp_of_empty)+)
+    have "Y = L (arden_variate Y yrhs)" 
+      using Y_eq_yrhs Inv_ES finite_yrhs nonempty_yrhs      
+      by (rule_tac arden_variate_keeps_eq, (simp add:rexp_of_empty)+)
     thus ?thesis using Inv_ES 
-      by (clarsimp simp add:valid_eqns_def eqs_subst_def rhs_subst_keeps_eq Inv_def finite_rhs_def
+      by (clarsimp simp add:valid_eqns_def 
+              eqs_subst_def rhs_subst_keeps_eq Inv_def finite_rhs_def
                    simp del:L_rhs.simps)
   qed
-  moreover have non_empty_subst: "non_empty (eqs_subst ES Y (arden_variate Y yrhs))"
+  moreover have 
+    non_empty_subst: "non_empty (eqs_subst ES Y (arden_variate Y yrhs))"
     using Inv_ES by (auto simp:Inv_def non_empty_def eqs_subst_def)
-  moreover have self_subst: "self_contained (eqs_subst ES Y (arden_variate Y yrhs))"
+  moreover 
+  have self_subst: "self_contained (eqs_subst ES Y (arden_variate Y yrhs))"
     using Inv_ES eqs_subst_keeps_self_contained by (simp add:Inv_def)
   ultimately show ?thesis using Inv_ES by (simp add:Inv_def)
 qed
@@ -642,7 +934,8 @@
 proof-
   have "card (S - {e}) > 0"
   proof -
-    have "card S > 1" using card e_in finite  by (case_tac "card S", auto) 
+    have "card S > 1" using card e_in finite  
+      by (case_tac "card S", auto) 
     thus ?thesis using finite e_in by auto
   qed
   hence "S - {e} \<noteq> {}" using finite by (rule_tac notI, simp)
@@ -653,10 +946,12 @@
   assumes Inv_ES: "Inv ES"
   and    X_in_ES: "(X, xrhs) \<in> ES"
   and    not_T: "card ES \<noteq> 1"
-  shows "\<exists> ES'. (Inv ES' \<and> (\<exists> xrhs'.(X, xrhs') \<in> ES')) \<and> (card ES', card ES) \<in> less_than" (is "\<exists> ES'. ?P ES'")
+  shows "\<exists> ES'. (Inv ES' \<and> (\<exists> xrhs'.(X, xrhs') \<in> ES')) \<and> 
+                (card ES', card ES) \<in> less_than" (is "\<exists> ES'. ?P ES'")
 proof -
   have finite_ES: "finite ES" using Inv_ES by (simp add:Inv_def)
-  then obtain Y yrhs where Y_in_ES: "(Y, yrhs) \<in> ES" and not_eq: "(X, xrhs) \<noteq> (Y, yrhs)" 
+  then obtain Y yrhs 
+    where Y_in_ES: "(Y, yrhs) \<in> ES" and not_eq: "(X, xrhs) \<noteq> (Y, yrhs)" 
     using not_T X_in_ES by (drule_tac card_noteq_1_has_more, auto)
   def ES' == "ES - {(Y, yrhs)}"
   let ?ES'' = "eqs_subst ES' Y (arden_variate Y yrhs)"
@@ -679,12 +974,17 @@
   thus ?thesis by blast
 qed
 
-text {* ***** END: proving every equation-system's iteration step satisfies Inv ************** *}
+text {*
+  From this point until @{text "hard_direction"}, the hard direction is proved
+  through a simple application of the iteration principle.
+*}
 
 lemma iteration_conc: 
   assumes history: "Inv ES"
   and    X_in_ES: "\<exists> xrhs. (X, xrhs) \<in> ES"
-  shows "\<exists> ES'. (Inv ES' \<and> (\<exists> xrhs'. (X, xrhs') \<in> ES')) \<and> card ES' = 1" (is "\<exists> ES'. ?P ES'")
+  shows 
+  "\<exists> ES'. (Inv ES' \<and> (\<exists> xrhs'. (X, xrhs') \<in> ES')) \<and> card ES' = 1" 
+                                                          (is "\<exists> ES'. ?P ES'")
 proof (cases "card ES = 1")
   case True
   thus ?thesis using history X_in_ES
@@ -706,26 +1006,31 @@
     have "L (rexp_of_lam ?A) = L (lam_of ?A)"
     proof(rule rexp_of_lam_eq_lam_set)
       show "finite (arden_variate X xrhs)" using Inv_ES ES_single 
-        by (rule_tac arden_variate_keeps_finite, auto simp add:Inv_def finite_rhs_def)
+        by (rule_tac arden_variate_keeps_finite, 
+                       auto simp add:Inv_def finite_rhs_def)
     qed
     also have "\<dots> = L ?A"
     proof-
       have "lam_of ?A = ?A"
       proof-
         have "classes_of ?A = {}" using Inv_ES ES_single
-          by (simp add:arden_variate_removes_cl self_contained_def Inv_def lefts_of_def) 
-        thus ?thesis by (auto simp only:lam_of_def classes_of_def, case_tac x, auto)
+          by (simp add:arden_variate_removes_cl 
+                       self_contained_def Inv_def lefts_of_def) 
+        thus ?thesis 
+          by (auto simp only:lam_of_def classes_of_def, case_tac x, auto)
       qed
       thus ?thesis by simp
     qed
     also have "\<dots> = X"
     proof(rule arden_variate_keeps_eq [THEN sym])
-      show "X = L xrhs" using Inv_ES ES_single by (auto simp only:Inv_def valid_eqns_def)  
+      show "X = L xrhs" using Inv_ES ES_single 
+        by (auto simp only:Inv_def valid_eqns_def)  
     next
       from Inv_ES ES_single show "[] \<notin> L (rexp_of xrhs X)"
         by(simp add:Inv_def ardenable_def rexp_of_empty finite_rhs_def)
     next
-      from Inv_ES ES_single show "finite xrhs" by (simp add:Inv_def finite_rhs_def)
+      from Inv_ES ES_single show "finite xrhs" 
+        by (simp add:Inv_def finite_rhs_def)
     qed
     finally show ?thesis by simp
   qed
@@ -750,79 +1055,73 @@
     by (rule last_cl_exists_rexp)
 qed
 
+lemma finals_in_partitions:
+  "finals Lang \<subseteq> (UNIV // (\<approx>Lang))"
+  by (auto simp:finals_def quotient_def)   
+
 theorem hard_direction: 
   assumes finite_CS: "finite (UNIV // (\<approx>Lang))"
   shows   "\<exists> (reg::rexp). Lang = L reg"
 proof -
   have "\<forall> X \<in> (UNIV // (\<approx>Lang)). \<exists> (reg::rexp). X = L reg" 
     using finite_CS every_eqcl_has_reg by blast
-  then obtain f where f_prop: "\<forall> X \<in> (UNIV // (\<approx>Lang)). X = L ((f X)::rexp)" 
+  then obtain f 
+    where f_prop: "\<forall> X \<in> (UNIV // (\<approx>Lang)). X = L ((f X)::rexp)" 
     by (auto dest:bchoice)
-  def rs \<equiv> "f ` {X. final X Lang}"  
-  have "Lang = \<Union> {X. final X Lang}" using lang_is_union_of_finals by simp
+  def rs \<equiv> "f ` (finals Lang)"  
+  have "Lang = \<Union> (finals Lang)" using lang_is_union_of_finals by auto
   also have "\<dots> = L (folds ALT NULL rs)" 
   proof -
-    have "finite {X. final X Lang}" using finite_CS by (auto simp:final_def)
-    thus ?thesis  using f_prop by (auto simp:rs_def final_def)
+    have "finite rs"
+    proof -
+      have "finite (finals Lang)" 
+        using finite_CS finals_in_partitions[of "Lang"]   
+        by (erule_tac finite_subset, simp)
+      thus ?thesis using rs_def by auto
+    qed
+    thus ?thesis 
+      using f_prop rs_def finals_in_partitions[of "Lang"] by auto
   qed
   finally show ?thesis by blast
 qed 
 
-section {* regular \<Rightarrow> finite*}
+section {* Direction: @{text "regular language \<Rightarrow>finite partition"} *}
+
+subsection {* The scheme for this direction *}
 
-lemma quot_empty_subset:
-  "UNIV // (\<approx>{[]}) \<subseteq> {{[]}, UNIV - {[]}}"
-proof
-  fix x
-  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
+text {* 
+  The following convenient notation @{text "x \<approx>Lang y"} means:
+  string @{text "x"} and @{text "y"} are equivalent with respect to 
+  language @{text "Lang"}.
+  *}
+
+definition
+  str_eq ("_ \<approx>_ _")
+where
+  "x \<approx>Lang y \<equiv> (x, y) \<in> (\<approx>Lang)"
 
-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
+text {*
+  The very basic scheme to show the finiteness of the partion generated by a language @{text "Lang"}
+  is by attaching tags to every string. The set of tags are carfully choosen to make it finite.
+  If it can be proved that strings with the same tag are equivlent with respect @{text "Lang"},
+  then the partition given rise by @{text "Lang"} must be finite. The reason for this is a lemma 
+  in standard library (@{text "finite_imageD"}), which says: if the image of an injective 
+  function on a set @{text "A"} is finite, then @{text "A"} is finite. It can be shown that
+  the function obtained by llifting @{text "tag"}
+  to the level of equalent classes (i.e. @{text "((op `) tag)"}) is injective 
+  (by lemma @{text "tag_image_injI"}) and the image of this function is finite 
+  (with the help of lemma @{text "finite_tag_imageI"}).
 
-text {* *************** Some common lemmas for following ALT, SEQ & STAR cases ******************* *}
-
-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)
+  BUT, I think this argument can be encapsulated by one lemma instead of the current presentation.
+  *}
 
 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"
+  "\<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"
+  assumes str_inj: "\<And> x y. tag x = tag (y::string) \<Longrightarrow> x \<approx>lang y"
   shows "inj_on ((op `) tag) (UNIV // \<approx>lang)"
 proof-
   { fix X Y
@@ -838,7 +1137,17 @@
   thus ?thesis unfolding inj_on_def by auto
 qed
 
-text {* **************** the SEQ case ************************ *}
+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)
+
+
+subsection {* A small theory for list difference *}
+
+text {*
+  The notion of list diffrence is need to make proofs more readable.
+  *}
 
 (* list_diff:: list substract, once different return tailer *)
 fun list_diff :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infix "-" 51)
@@ -884,29 +1193,99 @@
 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)"
+  "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)
 
+subsection {* Lemmas for basic cases *}
+
+text {*
+  The the final result of this direction is in @{text "easier_direction"}, which
+  is an induction on the structure of regular expressions. There is one case 
+  for each regular expression operator. For basic operators such as @{text "NULL, EMPTY, CHAR c"},
+  the finiteness of their language partition can be established directly with no need
+  of taggiing. This section contains several technical lemma for these base cases.
+
+  The inductive cases involve operators @{text "ALT, SEQ"} and @{text "STAR"}. 
+  Tagging functions need to be defined individually for each of them. There will be one
+  dedicated section for each of these cases, and each section goes virtually the same way:
+  gives definition of the tagging function and prove that strings 
+  with the same tag are equivalent.
+  *}
+
+lemma quot_empty_subset:
+  "UNIV // (\<approx>{[]}) \<subseteq> {{[]}, UNIV - {[]}}"
+proof
+  fix x
+  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)
+    thus ?thesis by simp
+  next
+    case False with h
+    have "x = UNIV - {[]}" by (auto simp:str_eq_rel_def)
+    thus ?thesis by simp
+  qed
+qed
+
+lemma quot_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)
+    } moreover {
+      assume "y = [c]" hence "x = {[c]}" using h 
+        by (auto dest!:spec[where x = "[]"] simp:str_eq_rel_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)
+    } ultimately show ?thesis by blast
+  qed
+qed
+
+subsection {* The case for @{text "SEQ"}*}
+
 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})"
+  "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:
-  "\<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))"
+  "\<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 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)"
+  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" 
+  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)
+  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
 
@@ -918,7 +1297,8 @@
     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)"
+      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
@@ -928,7 +1308,8 @@
           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") 
+                  {\<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
@@ -942,10 +1323,13 @@
         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) 
+          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)
+        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
@@ -958,7 +1342,8 @@
   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
+  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)) (UNIV // \<approx>L\<^isub>1 ;; L\<^isub>2)"
@@ -967,13 +1352,14 @@
     by (auto intro:tag_image_injI tag_str_SEQ_injI simp:)
 qed
 
-text {* **************** the ALT case ************************ *}
+subsection {* The case for @{text "ALT"} *}
 
 definition 
   "tag_str_ALT L\<^isub>1 L\<^isub>2 (x::string) \<equiv> ((\<approx>L\<^isub>1) `` {x}, (\<approx>L\<^isub>2) `` {x})"
 
 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))"
+  "\<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)
 
@@ -982,20 +1368,33 @@
   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
+  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"
+    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
 
-text {* **************** the Star case ****************** *}
+
+subsection {*
+  The case for @{text "STAR"}
+  *}
 
-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))"
+text {* 
+  This turned out to be the most tricky case. 
+  *} (* I will make some illustrations for it. *)
+
+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 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
@@ -1005,7 +1404,9 @@
     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
+    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"
@@ -1017,30 +1418,17 @@
   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_intro2: "y \<in> lang \<Longrightarrow> y \<in> lang\<star>"
-by (drule step[of y lang "[]"], auto simp:start)
-
-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 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)
+by (auto simp:tag_str_STAR_def Image_def 
+                       quotient_def split:if_splits)
 
 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"
@@ -1051,58 +1439,76 @@
     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)
+      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"
+      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)
+          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
+        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)"
+      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
+        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>" 
+        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)" 
+          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
+              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" 
+        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)
+        from x_za h7 have "(y - ya) @ za \<in> L\<^isub>1" 
+          by (auto simp:str_eq_def str_eq_rel_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)
+      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)
@@ -1119,9 +1525,11 @@
     by (auto intro:tag_image_injI tag_str_STAR_injI)
 qed
 
-text {* **************** the Other Direction ************ *}
+subsection {*
+  The main lemma
+  *}
 
-lemma other_direction:
+lemma easier_directio\<nu>:
   "Lang = L (r::rexp) \<Longrightarrow> finite (UNIV // (\<approx>Lang))"
 proof (induct arbitrary:Lang rule:rexp.induct)
   case NULL
@@ -1130,27 +1538,33 @@
   with prems show "?case" by (auto intro:finite_subset)
 next
   case EMPTY
-  have "UNIV // (\<approx>{[]}) \<subseteq> {{[]}, UNIV - {[]}}" by (rule quot_empty_subset)
+  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)
+  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))"
+  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))"
+  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>))"
+  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
\ No newline at end of file
+end
+