updated theories and itp-paper
authorurbanc
Thu, 02 Jun 2011 16:44:35 +0000
changeset 166 7743d2ad71d1
parent 165 b04cc5e4e84c
child 167 61d0a412a3ae
updated theories and itp-paper
Closure.thy
Derivs.thy
Myhill_1.thy
Myhill_2.thy
Paper/Paper.thy
Paper/ROOT.ML
Theories/Closure.thy
Theories/Derivs.thy
Theories/Folds.thy
Theories/Myhill_1.thy
Theories/Myhill_2.thy
Theories/Prefix_subtract.thy
Theories/Regular.thy
--- a/Closure.thy	Tue May 31 20:32:49 2011 +0000
+++ b/Closure.thy	Thu Jun 02 16:44:35 2011 +0000
@@ -1,5 +1,6 @@
+(* Author: Christian Urban, Xingyuan Zhang, Chunhan Wu *)
 theory Closure
-imports Myhill_2
+imports Derivs
 begin
 
 section {* Closure properties of regular languages *}
@@ -7,36 +8,40 @@
 abbreviation
   regular :: "lang \<Rightarrow> bool"
 where
-  "regular A \<equiv> \<exists>r::rexp. A = L r"
+  "regular A \<equiv> \<exists>r. A = L_rexp r"
 
+subsection {* Closure under set operations *}
 
 lemma closure_union[intro]:
   assumes "regular A" "regular B" 
   shows "regular (A \<union> B)"
 proof -
-  from assms obtain r1 r2::rexp where "L r1 = A" "L r2 = B" by auto
-  then have "A \<union> B = L (ALT r1 r2)" by simp
+  from assms obtain r1 r2::rexp where "L_rexp r1 = A" "L_rexp r2 = B" by auto
+  then have "A \<union> B = L_rexp (ALT r1 r2)" by simp
   then show "regular (A \<union> B)" by blast
 qed
 
 lemma closure_seq[intro]:
   assumes "regular A" "regular B" 
-  shows "regular (A ;; B)"
+  shows "regular (A \<cdot> B)"
 proof -
-  from assms obtain r1 r2::rexp where "L r1 = A" "L r2 = B" by auto
-  then have "A ;; B = L (SEQ r1 r2)" by simp
-  then show "regular (A ;; B)" by blast
+  from assms obtain r1 r2::rexp where "L_rexp r1 = A" "L_rexp r2 = B" by auto
+  then have "A \<cdot> B = L_rexp (SEQ r1 r2)" by simp
+  then show "regular (A \<cdot> B)" by blast
 qed
 
 lemma closure_star[intro]:
   assumes "regular A"
   shows "regular (A\<star>)"
 proof -
-  from assms obtain r::rexp where "L r = A" by auto
-  then have "A\<star> = L (STAR r)" by simp
+  from assms obtain r::rexp where "L_rexp r = A" by auto
+  then have "A\<star> = L_rexp (STAR r)" by simp
   then show "regular (A\<star>)" by blast
 qed
 
+text {* Closure under complementation is proved via the 
+  Myhill-Nerode theorem *}
+
 lemma closure_complement[intro]:
   assumes "regular A"
   shows "regular (- A)"
@@ -68,8 +73,7 @@
   ultimately show "regular (A \<inter> B)" by simp
 qed
 
-
-text {* closure under string reversal *}
+subsection {* Closure under string reversal *}
 
 fun
   Rev :: "rexp \<Rightarrow> rexp"
@@ -81,15 +85,12 @@
 | "Rev (SEQ r1 r2) = SEQ (Rev r2) (Rev r1)"
 | "Rev (STAR r) = STAR (Rev r)"
 
-lemma rev_Seq:
-  "(rev ` A) ;; (rev ` B) = rev ` (B ;; A)"
+lemma rev_seq[simp]:
+  shows "rev ` (B \<cdot> A) = (rev ` A) \<cdot> (rev ` B)"
 unfolding Seq_def image_def
-apply(auto)
-apply(rule_tac x="xb @ xa" in exI)
-apply(auto)
-done
+by (auto) (metis rev_append)+
 
-lemma rev_Star1:
+lemma rev_star1:
   assumes a: "s \<in> (rev ` A)\<star>"
   shows "s \<in> rev ` (A\<star>)"
 using a
@@ -104,7 +105,7 @@
   then show "s1 @ s2 \<in>  rev ` A\<star>" using eqs by simp
 qed (auto)
 
-lemma rev_Star2:
+lemma rev_star2:
   assumes a: "s \<in> A\<star>"
   shows "rev s \<in> (rev ` A)\<star>"
 using a
@@ -119,22 +120,39 @@
   ultimately show "rev (s1 @ s2) \<in>  (rev ` A)\<star>" by (auto intro: star_intro1)
 qed (auto)
 
-lemma rev_Star:
-  "(rev ` A)\<star> = rev ` (A\<star>)"
-using rev_Star1 rev_Star2 by auto
+lemma rev_star[simp]:
+  shows " rev ` (A\<star>) = (rev ` A)\<star>"
+using rev_star1 rev_star2 by auto
 
 lemma rev_lang:
-  "L (Rev r) = rev ` (L r)"
-by (induct r) (simp_all add: rev_Star rev_Seq image_Un)
+  shows "rev ` (L_rexp r) = L_rexp (Rev r)"
+by (induct r) (simp_all add: image_Un)
 
 lemma closure_reversal[intro]:
   assumes "regular A"
   shows "regular (rev ` A)"
 proof -
-  from assms obtain r::rexp where "L r = A" by auto
-  then have "L (Rev r) = rev ` A" by (simp add: rev_lang)
+  from assms obtain r::rexp where "A = L_rexp r" by auto
+  then have "L_rexp (Rev r) = rev ` A" by (simp add: rev_lang)
   then show "regular (rev` A)" by blast
 qed
   
+subsection {* Closure under left-quotients *}
+
+lemma closure_left_quotient:
+  assumes "regular A"
+  shows "regular (Ders_set B A)"
+proof -
+  from assms obtain r::rexp where eq: "L_rexp r = A" by auto
+  have fin: "finite (pders_set B r)" by (rule finite_pders_set)
+  
+  have "Ders_set B (L_rexp r) = (\<Union> L_rexp ` (pders_set B r))"
+    by (simp add: Ders_set_pders_set)
+  also have "\<dots> = L_rexp (\<Uplus>(pders_set B r))" using fin by simp
+  finally have "Ders_set B A = L_rexp (\<Uplus>(pders_set B r))" using eq
+    by simp
+  then show "regular (Ders_set B A)" by auto
+qed
+
 
 end
\ No newline at end of file
--- a/Derivs.thy	Tue May 31 20:32:49 2011 +0000
+++ b/Derivs.thy	Thu Jun 02 16:44:35 2011 +0000
@@ -1,8 +1,8 @@
 theory Derivs
-imports Closure
+imports Myhill_2
 begin
 
-section {* Experiments with Derivatives -- independent of Myhill-Nerode *}
+section {* Left-Quotients and Derivatives *}
 
 subsection {* Left-Quotients *}
 
@@ -52,30 +52,30 @@
 by auto
 
 lemma Der_seq [simp]:
-  shows "Der c (A ;; B) = (Der c A) ;; B \<union> (Delta A ;; Der c B)"
+  shows "Der c (A \<cdot> B) = (Der c A) \<cdot> B \<union> (Delta A \<cdot> Der c B)"
 unfolding Der_def Delta_def
 unfolding Seq_def
 by (auto simp add: Cons_eq_append_conv)
 
 lemma Der_star [simp]:
-  shows "Der c (A\<star>) = (Der c A) ;; A\<star>"
+  shows "Der c (A\<star>) = (Der c A) \<cdot> A\<star>"
 proof -
-  have incl: "Delta A ;; Der c (A\<star>) \<subseteq> (Der c A) ;; A\<star>"
+  have incl: "Delta A \<cdot> Der c (A\<star>) \<subseteq> (Der c A) \<cdot> A\<star>"
     unfolding Der_def Delta_def Seq_def
     apply(auto)
     apply(drule star_decom)
     apply(auto simp add: Cons_eq_append_conv)
     done
     
-  have "Der c (A\<star>) = Der c ({[]} \<union> A ;; A\<star>)"
+  have "Der c (A\<star>) = Der c ({[]} \<union> A \<cdot> A\<star>)"
     by (simp only: star_cases[symmetric])
-  also have "... = Der c (A ;; A\<star>)"
+  also have "... = Der c (A \<cdot> A\<star>)"
     by (simp only: Der_union Der_empty) (simp)
-  also have "... = (Der c A) ;; A\<star> \<union> (Delta A ;; Der c (A\<star>))"
+  also have "... = (Der c A) \<cdot> A\<star> \<union> (Delta A \<cdot> Der c (A\<star>))"
     by simp
-  also have "... =  (Der c A) ;; A\<star>"
+  also have "... =  (Der c A) \<cdot> A\<star>"
     using incl by auto
-  finally show "Der c (A\<star>) = (Der c A) ;; A\<star>" . 
+  finally show "Der c (A\<star>) = (Der c A) \<cdot> A\<star>" . 
 qed
 
 
@@ -126,18 +126,18 @@
   by (relation "measure (length o fst)") (auto)
 
 lemma Delta_nullable:
-  shows "Delta (L r) = (if nullable r then {[]} else {})"
+  shows "Delta (L_rexp r) = (if nullable r then {[]} else {})"
 unfolding Delta_def
 by (induct r) (auto simp add: Seq_def split: if_splits)
 
 lemma Der_der:
   fixes r::rexp
-  shows "Der c (L r) = L (der c r)"
+  shows "Der c (L_rexp r) = L_rexp (der c r)"
 by (induct r) (simp_all add: Delta_nullable)
 
 lemma Ders_ders:
   fixes r::rexp
-  shows "Ders s (L r) = L (ders s r)"
+  shows "Ders s (L_rexp r) = L_rexp (ders s r)"
 apply(induct s rule: rev_induct)
 apply(simp add: Ders_def)
 apply(simp only: ders.simps)
@@ -195,44 +195,44 @@
 done
 
 lemma pder_set_lang:
-  shows "(\<Union> (L ` pder_set c R)) = (\<Union>r \<in> R. (\<Union>L ` (pder c r)))"
+  shows "(\<Union> (L_rexp ` pder_set c R)) = (\<Union>r \<in> R. (\<Union>L_rexp ` (pder c r)))"
 unfolding image_def 
 by auto
 
 lemma
-  shows seq_UNION_left:  "B ;; (\<Union>n\<in>C. A n) = (\<Union>n\<in>C. B ;; A n)"
-  and   seq_UNION_right: "(\<Union>n\<in>C. A n) ;; B = (\<Union>n\<in>C. A n ;; B)"
+  shows seq_UNION_left:  "B \<cdot> (\<Union>n\<in>C. A n) = (\<Union>n\<in>C. B \<cdot> A n)"
+  and   seq_UNION_right: "(\<Union>n\<in>C. A n) \<cdot> B = (\<Union>n\<in>C. A n \<cdot> B)"
 unfolding Seq_def by auto
 
 lemma Der_pder:
   fixes r::rexp
-  shows "Der c (L r) = \<Union> L ` (pder c r)"
+  shows "Der c (L_rexp r) = \<Union> L_rexp ` (pder c r)"
 by (induct r) (auto simp add: Delta_nullable seq_UNION_right)
 
 lemma Ders_pders:
   fixes r::rexp
-  shows "Ders s (L r) = \<Union> L ` (pders s r)"
+  shows "Ders s (L_rexp r) = \<Union> L_rexp ` (pders s r)"
 proof (induct s rule: rev_induct)
   case (snoc c s)
-  have ih: "Ders s (L r) = \<Union> L ` (pders s r)" by fact
-  have "Ders (s @ [c]) (L r) = Ders [c] (Ders s (L r))"
+  have ih: "Ders s (L_rexp r) = \<Union> L_rexp ` (pders s r)" by fact
+  have "Ders (s @ [c]) (L_rexp r) = Ders [c] (Ders s (L_rexp r))"
     by (simp add: Ders_append)
-  also have "\<dots> = Der c (\<Union> L ` (pders s r))" using ih
+  also have "\<dots> = Der c (\<Union> L_rexp ` (pders s r))" using ih
     by (simp add: Ders_singleton)
-  also have "\<dots> = (\<Union>r\<in>pders s r. Der c (L r))" 
+  also have "\<dots> = (\<Union>r\<in>pders s r. Der c (L_rexp r))" 
     unfolding Der_def image_def by auto
-  also have "\<dots> = (\<Union>r\<in>pders s r. (\<Union> L `  (pder c r)))"
+  also have "\<dots> = (\<Union>r\<in>pders s r. (\<Union> L_rexp `  (pder c r)))"
     by (simp add: Der_pder)
-  also have "\<dots> = (\<Union>L ` (pder_set c (pders s r)))"
+  also have "\<dots> = (\<Union>L_rexp ` (pder_set c (pders s r)))"
     by (simp add: pder_set_lang)
-  also have "\<dots> = (\<Union>L ` (pders (s @ [c]) r))"
+  also have "\<dots> = (\<Union>L_rexp ` (pders (s @ [c]) r))"
     by simp
-  finally show "Ders (s @ [c]) (L r) = \<Union>L ` pders (s @ [c]) r" .
+  finally show "Ders (s @ [c]) (L_rexp r) = \<Union> L_rexp ` pders (s @ [c]) r" .
 qed (simp add: Ders_def)
 
 lemma Ders_set_pders_set:
   fixes r::rexp
-  shows "Ders_set A (L r) = (\<Union> L ` (pders_set A r))"
+  shows "Ders_set A (L_rexp r) = (\<Union> L_rexp ` (pders_set A r))"
 by (simp add: Ders_set_Ders Ders_pders)
 
 lemma pders_NULL [simp]:
@@ -255,12 +255,12 @@
   "Suf s \<equiv> {v. v \<noteq> [] \<and> (\<exists>u. u @ v = s)}"
 
 lemma Suf:
-  shows "Suf (s @ [c]) = (Suf s) ;; {[c]} \<union> {[c]}"
+  shows "Suf (s @ [c]) = (Suf s) \<cdot> {[c]} \<union> {[c]}"
 unfolding Suf_def Seq_def
 by (auto simp add: append_eq_append_conv2 append_eq_Cons_conv)
 
 lemma Suf_Union:
-  shows "(\<Union>v \<in> Suf s ;; {[c]}. P v) = (\<Union>v \<in> Suf s. P (v @ [c]))"
+  shows "(\<Union>v \<in> Suf s \<cdot> {[c]}. P v) = (\<Union>v \<in> Suf s. P (v @ [c]))"
 by (auto simp add: Seq_def)
 
 lemma inclusion1:
@@ -277,7 +277,7 @@
     by fact
   have "pders (s @ [c]) (SEQ r1 r2) = pder_set c (pders s (SEQ r1 r2))" by simp
   also have "\<dots> \<subseteq> pder_set c (SEQS (pders s r1) r2 \<union> (\<Union>v \<in> Suf s. pders v r2))"
-    using ih by auto 
+    using ih by (auto) (blast)
   also have "\<dots> = pder_set c (SEQS (pders s r1) r2) \<union> pder_set c (\<Union>v \<in> Suf s. pders v r2)"
     by (simp)
   also have "\<dots> = pder_set c (SEQS (pders s r1) r2) \<union> (\<Union>v \<in> Suf s. pder_set c (pders v r2))"
@@ -309,7 +309,7 @@
     also have "\<dots> = (\<Union>v\<in>Suf s. pder_set c (SEQS (pders v r) (STAR r)))"
       by simp
     also have "\<dots> \<subseteq> (\<Union>v\<in>Suf s. (SEQS (pder_set c (pders v r)) (STAR r) \<union> pder c (STAR r)))"
-      using inclusion1 by blast
+      using inclusion1 by (auto split: if_splits) 
     also have "\<dots> = (\<Union>v\<in>Suf s. (SEQS (pder_set c (pders v r)) (STAR r))) \<union> pder c (STAR r)"
       using asm by (auto simp add: Suf_def)
     also have "\<dots> = (\<Union>v\<in>Suf s. (SEQS (pders (v @ [c]) r) (STAR r))) \<union> (SEQS (pder c r) (STAR r))"
@@ -444,78 +444,49 @@
 
 lemma Myhill_Nerode3:
   fixes r::"rexp"
-  shows "finite (UNIV // \<approx>(L r))"
+  shows "finite (UNIV // \<approx>(L_rexp r))"
 proof -
   have "finite (UNIV // =(\<lambda>x. pders x r)=)"
   proof - 
-    have "range (\<lambda>x. pders x r) \<subseteq> {pders s r | s. s \<in> UNIV}" by auto
+    have "range (\<lambda>x. pders x r) = {pders s r | s. s \<in> UNIV}" by auto
     moreover 
     have "finite {pders s r | s. s \<in> UNIV}" by (rule finite_pders2)
     ultimately
     have "finite (range (\<lambda>x. pders x r))"
-      by (rule finite_subset)
+      by simp
     then show "finite (UNIV // =(\<lambda>x. pders x r)=)" 
       by (rule finite_eq_tag_rel)
   qed
   moreover 
-  have " =(\<lambda>x. pders x r)= \<subseteq> \<approx>(L r)"
+  have "=(\<lambda>x. pders x r)= \<subseteq> \<approx>(L_rexp r)"
     unfolding tag_eq_rel_def
-    by (auto simp add: str_eq_def[symmetric] MN_Rel_Ders Ders_pders)
+    unfolding str_eq_def2
+    unfolding MN_Rel_Ders
+    unfolding Ders_pders
+    by auto
   moreover 
   have "equiv UNIV =(\<lambda>x. pders x r)="
     unfolding equiv_def refl_on_def sym_def trans_def
     unfolding tag_eq_rel_def
     by auto
   moreover
-  have "equiv UNIV (\<approx>(L r))"
+  have "equiv UNIV (\<approx>(L_rexp r))"
     unfolding equiv_def refl_on_def sym_def trans_def
     unfolding str_eq_rel_def
     by auto
-  ultimately show "finite (UNIV // \<approx>(L r))" 
+  ultimately show "finite (UNIV // \<approx>(L_rexp r))" 
     by (rule refined_partition_finite)
 qed
 
 
-section {* Closure under Left-Quotients *}
-
-lemma closure_left_quotient:
-  assumes "regular A"
-  shows "regular (Ders_set B A)"
-proof -
-  from assms obtain r::rexp where eq: "L r = A" by auto
-  have fin: "finite (pders_set B r)" by (rule finite_pders_set)
-  
-  have "Ders_set B (L r) = (\<Union> L ` (pders_set B r))"
-    by (simp add: Ders_set_pders_set)
-  also have "\<dots> = L (\<Uplus>(pders_set B r))" using fin by simp
-  finally have "Ders_set B A = L (\<Uplus>(pders_set B r))" using eq
-    by simp
-  then show "regular (Ders_set B A)" by auto
-qed
-
-
-section {* Relating standard and partial derivations *}
+section {* Relating derivatives and partial derivatives *}
 
 lemma
-  shows "(\<Union> L ` (pder c r)) = L (der c r)"
+  shows "(\<Union> L_rexp ` (pder c r)) = L_rexp (der c r)"
 unfolding Der_der[symmetric] Der_pder by simp
 
 lemma
-  shows "(\<Union> L ` (pders s r)) = L (ders s r)"
+  shows "(\<Union> L_rexp ` (pders s r)) = L_rexp (ders s r)"
 unfolding Ders_ders[symmetric] Ders_pders by simp
 
-
-
-fun
-  width :: "rexp \<Rightarrow> nat"
-where
-  "width (NULL) = 0"
-| "width (EMPTY) = 0"
-| "width (CHAR c) = 1"
-| "width (ALT r1 r2) = width r1 + width r2"
-| "width (SEQ r1 r2) = width r1 + width r2"
-| "width (STAR r) = width r"
-
-
- 
 end
\ No newline at end of file
--- a/Myhill_1.thy	Tue May 31 20:32:49 2011 +0000
+++ b/Myhill_1.thy	Thu Jun 02 16:44:35 2011 +0000
@@ -1,5 +1,5 @@
 theory Myhill_1
-imports Main Folds Regular
+imports Regular
         "~~/src/HOL/Library/While_Combinator" 
 begin
 
@@ -37,43 +37,37 @@
 
 text {* The two kinds of terms in the rhs of equations. *}
 
-datatype rhs_trm = 
+datatype trm = 
    Lam "rexp"            (* Lambda-marker *)
  | Trn "lang" "rexp"     (* Transition *)
 
+fun 
+  L_trm::"trm \<Rightarrow> lang"
+where
+  "L_trm (Lam r) = L_rexp r" 
+| "L_trm (Trn X r) = X \<cdot> L_rexp r"
 
-overloading L_rhs_trm \<equiv> "L:: rhs_trm \<Rightarrow> lang"
-begin
-  fun L_rhs_trm:: "rhs_trm \<Rightarrow> lang"
-  where
-    "L_rhs_trm (Lam r) = L r" 
-  | "L_rhs_trm (Trn X r) = X ;; L r"
-end
-
-overloading L_rhs \<equiv> "L:: rhs_trm set \<Rightarrow> lang"
-begin
-   fun L_rhs:: "rhs_trm set \<Rightarrow> lang"
-   where 
-     "L_rhs rhs = \<Union> (L ` rhs)"
-end
+fun 
+  L_rhs::"trm set \<Rightarrow> lang"
+where 
+  "L_rhs rhs = \<Union> (L_trm ` rhs)"
 
 lemma L_rhs_set:
-  shows "L {Trn X r | r. P r} = \<Union>{L (Trn X r) | r. P r}"
-by (auto simp del: L_rhs_trm.simps)
+  shows "L_rhs {Trn X r | r. P r} = \<Union>{L_trm (Trn X r) | r. P r}"
+by (auto)
 
 lemma L_rhs_union_distrib:
-  fixes A B::"rhs_trm set"
-  shows "L A \<union> L B = L (A \<union> B)"
+  fixes A B::"trm set"
+  shows "L_rhs A \<union> L_rhs B = L_rhs (A \<union> B)"
 by simp
 
 
-
 text {* Transitions between equivalence classes *}
 
 definition 
   transition :: "lang \<Rightarrow> char \<Rightarrow> lang \<Rightarrow> bool" ("_ \<Turnstile>_\<Rightarrow>_" [100,100,100] 100)
 where
-  "Y \<Turnstile>c\<Rightarrow> X \<equiv> Y ;; {[c]} \<subseteq> X"
+  "Y \<Turnstile>c\<Rightarrow> X \<equiv> Y \<cdot> {[c]} \<subseteq> X"
 
 text {* Initial equational system *}
 
@@ -91,7 +85,7 @@
 section {* Arden Operation on equations *}
 
 fun 
-  Append_rexp :: "rexp \<Rightarrow> rhs_trm \<Rightarrow> rhs_trm"
+  Append_rexp :: "rexp \<Rightarrow> trm \<Rightarrow> trm"
 where
   "Append_rexp r (Lam rexp)   = Lam (SEQ rexp r)"
 | "Append_rexp r (Trn X rexp) = Trn X (SEQ rexp r)"
@@ -112,7 +106,7 @@
         (rhs - {Trn X r | r. Trn X r \<in> rhs}) \<union> (Append_rexp_rhs xrhs (\<Uplus> {r. Trn X r \<in> rhs}))"
 
 definition
-  Subst_all :: "(lang \<times> rhs_trm set) set \<Rightarrow> lang \<Rightarrow> rhs_trm set \<Rightarrow> (lang \<times> rhs_trm set) set"
+  Subst_all :: "(lang \<times> trm set) set \<Rightarrow> lang \<Rightarrow> trm set \<Rightarrow> (lang \<times> trm set) set"
 where
   "Subst_all ES X xrhs \<equiv> {(Y, Subst yrhs X xrhs) | Y yrhs. (Y, yrhs) \<in> ES}"
 
@@ -149,10 +143,10 @@
      \<forall> X rhs rhs'. (X, rhs) \<in> ES \<and> (X, rhs') \<in> ES \<longrightarrow> rhs = rhs'"
 
 definition 
-  "soundness ES \<equiv> \<forall>(X, rhs) \<in> ES. X = L rhs"
+  "soundness ES \<equiv> \<forall>(X, rhs) \<in> ES. X = L_rhs rhs"
 
 definition 
-  "ardenable rhs \<equiv> (\<forall> Y r. Trn Y r \<in> rhs \<longrightarrow> [] \<notin> L r)"
+  "ardenable rhs \<equiv> (\<forall> Y r. Trn Y r \<in> rhs \<longrightarrow> [] \<notin> L_rexp r)"
 
 definition 
   "ardenable_all ES \<equiv> \<forall>(X, rhs) \<in> ES. ardenable rhs"
@@ -227,29 +221,28 @@
     done
 qed
 
-lemma rhs_trm_soundness:
+lemma trm_soundness:
   assumes finite:"finite rhs"
-  shows "L ({Trn X r| r. Trn X r \<in> rhs}) = X ;; (L (\<Uplus>{r. Trn X r \<in> rhs}))"
+  shows "L_rhs ({Trn X r| r. Trn X r \<in> rhs}) = X \<cdot> (L_rexp (\<Uplus>{r. Trn X r \<in> rhs}))"
 proof -
   have "finite {r. Trn X r \<in> rhs}" 
     by (rule finite_Trn[OF finite]) 
-  then show "L ({Trn X r| r. Trn X r \<in> rhs}) = X ;; (L (\<Uplus>{r. Trn X r \<in> rhs}))"
-    by (simp only: L_rhs_set L_rhs_trm.simps) (auto simp add: Seq_def)
+  then show "L_rhs ({Trn X r| r. Trn X r \<in> rhs}) = X \<cdot> (L_rexp (\<Uplus>{r. Trn X r \<in> rhs}))"
+    by (simp only: L_rhs_set L_trm.simps) (auto simp add: Seq_def)
 qed
 
 lemma lang_of_append_rexp:
-  "L (Append_rexp r rhs_trm) = L rhs_trm ;; L r"
+  "L_trm (Append_rexp r trm) = L_trm trm \<cdot> L_rexp r"
 by (induct rule: Append_rexp.induct)
    (auto simp add: seq_assoc)
 
 lemma lang_of_append_rexp_rhs:
-  "L (Append_rexp_rhs rhs r) = L rhs ;; L r"
+  "L_rhs (Append_rexp_rhs rhs r) = L_rhs rhs \<cdot> L_rexp r"
 unfolding Append_rexp_rhs_def
 by (auto simp add: Seq_def lang_of_append_rexp)
 
 
-
-subsubsection {* Intialization *}
+subsubsection {* Intial Equational System *}
 
 lemma defined_by_str:
   assumes "s \<in> X" "X \<in> UNIV // \<approx>A" 
@@ -261,7 +254,7 @@
 lemma every_eqclass_has_transition:
   assumes has_str: "s @ [c] \<in> X"
   and     in_CS:   "X \<in> UNIV // \<approx>A"
-  obtains Y where "Y \<in> UNIV // \<approx>A" and "Y ;; {[c]} \<subseteq> X" and "s \<in> Y"
+  obtains Y where "Y \<in> UNIV // \<approx>A" and "Y \<cdot> {[c]} \<subseteq> X" and "s \<in> Y"
 proof -
   def Y \<equiv> "\<approx>A `` {s}"
   have "Y \<in> UNIV // \<approx>A" 
@@ -269,7 +262,7 @@
   moreover
   have "X = \<approx>A `` {s @ [c]}" 
     using has_str in_CS defined_by_str by blast
-  then have "Y ;; {[c]} \<subseteq> X" 
+  then have "Y \<cdot> {[c]} \<subseteq> X" 
     unfolding Y_def Image_def Seq_def
     unfolding str_eq_rel_def
     by clarsimp
@@ -281,14 +274,14 @@
 
 lemma l_eq_r_in_eqs:
   assumes X_in_eqs: "(X, rhs) \<in> Init (UNIV // \<approx>A)"
-  shows "X = L rhs"
+  shows "X = L_rhs rhs"
 proof 
-  show "X \<subseteq> L rhs"
+  show "X \<subseteq> L_rhs rhs"
   proof
     fix x
     assume in_X: "x \<in> X"
     { assume empty: "x = []"
-      then have "x \<in> L rhs" using X_in_eqs in_X
+      then have "x \<in> L_rhs rhs" using X_in_eqs in_X
 	unfolding Init_def Init_rhs_def
         by auto
     }
@@ -297,40 +290,40 @@
       then obtain s c where decom: "x = s @ [c]"
 	using rev_cases by blast
       have "X \<in> UNIV // \<approx>A" using X_in_eqs unfolding Init_def by auto
-      then obtain Y where "Y \<in> UNIV // \<approx>A" "Y ;; {[c]} \<subseteq> X" "s \<in> Y"
+      then obtain Y where "Y \<in> UNIV // \<approx>A" "Y \<cdot> {[c]} \<subseteq> X" "s \<in> Y"
         using decom in_X every_eqclass_has_transition by blast
-      then have "x \<in> L {Trn Y (CHAR c)| Y c. Y \<in> UNIV // \<approx>A \<and> Y \<Turnstile>c\<Rightarrow> X}"
+      then have "x \<in> L_rhs {Trn Y (CHAR c)| Y c. Y \<in> UNIV // \<approx>A \<and> Y \<Turnstile>c\<Rightarrow> X}"
         unfolding transition_def
 	using decom by (force simp add: Seq_def)
-      then have "x \<in> L rhs" using X_in_eqs in_X
+      then have "x \<in> L_rhs rhs" using X_in_eqs in_X
 	unfolding Init_def Init_rhs_def by simp
     }
-    ultimately show "x \<in> L rhs" by blast
+    ultimately show "x \<in> L_rhs rhs" by blast
   qed
 next
-  show "L rhs \<subseteq> X" using X_in_eqs
+  show "L_rhs rhs \<subseteq> X" using X_in_eqs
     unfolding Init_def Init_rhs_def transition_def
     by auto 
 qed
 
 lemma test:
   assumes X_in_eqs: "(X, rhs) \<in> Init (UNIV // \<approx>A)"
-  shows "X = \<Union> (L `  rhs)"
+  shows "X = \<Union> (L_trm `  rhs)"
 using assms l_eq_r_in_eqs by (simp)
 
 lemma finite_Init_rhs: 
   assumes finite: "finite CS"
   shows "finite (Init_rhs CS X)"
 proof-
-  def S \<equiv> "{(Y, c)| Y c. Y \<in> CS \<and> Y ;; {[c]} \<subseteq> X}" 
+  def S \<equiv> "{(Y, c)| Y c. Y \<in> CS \<and> Y \<cdot> {[c]} \<subseteq> X}" 
   def h \<equiv> "\<lambda> (Y, c). Trn Y (CHAR c)"
   have "finite (CS \<times> (UNIV::char set))" using finite by auto
   then have "finite S" using S_def 
     by (rule_tac B = "CS \<times> UNIV" in finite_subset) (auto)
-  moreover have "{Trn Y (CHAR c) |Y c. Y \<in> CS \<and> Y ;; {[c]} \<subseteq> X} = h ` S"
+  moreover have "{Trn Y (CHAR c) |Y c. Y \<in> CS \<and> Y \<cdot> {[c]} \<subseteq> X} = h ` S"
     unfolding S_def h_def image_def by auto
   ultimately
-  have "finite {Trn Y (CHAR c) |Y c. Y \<in> CS \<and> Y ;; {[c]} \<subseteq> X}" by auto
+  have "finite {Trn Y (CHAR c) |Y c. Y \<in> CS \<and> Y \<cdot> {[c]} \<subseteq> X}" by auto
   then show "finite (Init_rhs CS X)" unfolding Init_rhs_def transition_def by simp
 qed
 
@@ -359,32 +352,32 @@
 subsubsection {* Interation step *}
 
 lemma Arden_keeps_eq:
-  assumes l_eq_r: "X = L rhs"
+  assumes l_eq_r: "X = L_rhs rhs"
   and not_empty: "ardenable rhs"
   and finite: "finite rhs"
-  shows "X = L (Arden X rhs)"
+  shows "X = L_rhs (Arden X rhs)"
 proof -
-  def A \<equiv> "L (\<Uplus>{r. Trn X r \<in> rhs})"
+  def A \<equiv> "L_rexp (\<Uplus>{r. Trn X r \<in> rhs})"
   def b \<equiv> "{Trn X r | r. Trn X r \<in> rhs}"
-  def B \<equiv> "L (rhs - b)"
+  def B \<equiv> "L_rhs (rhs - b)"
   have not_empty2: "[] \<notin> A" 
     using finite_Trn[OF finite] not_empty
     unfolding A_def ardenable_def by simp
-  have "X = L rhs" using l_eq_r by simp
-  also have "\<dots> = L (b \<union> (rhs - b))" unfolding b_def by auto
-  also have "\<dots> = L b \<union> B" unfolding B_def by (simp only: L_rhs_union_distrib)
-  also have "\<dots> = X ;; A \<union> B"
+  have "X = L_rhs rhs" using l_eq_r by simp
+  also have "\<dots> = L_rhs (b \<union> (rhs - b))" unfolding b_def by auto
+  also have "\<dots> = L_rhs b \<union> B" unfolding B_def by (simp only: L_rhs_union_distrib)
+  also have "\<dots> = X \<cdot> A \<union> B"
     unfolding b_def
-    unfolding rhs_trm_soundness[OF finite]
+    unfolding trm_soundness[OF finite]
     unfolding A_def
     by blast
-  finally have "X = X ;; A \<union> B" . 
-  then have "X = B ;; A\<star>"
+  finally have "X = X \<cdot> A \<union> B" . 
+  then have "X = B \<cdot> A\<star>"
     by (simp add: arden[OF not_empty2])
-  also have "\<dots> = L (Arden X rhs)"
+  also have "\<dots> = L_rhs (Arden X rhs)"
     unfolding Arden_def A_def B_def b_def
     by (simp only: lang_of_append_rexp_rhs L_rexp.simps)
-  finally show "X = L (Arden X rhs)" by simp
+  finally show "X = L_rhs (Arden X rhs)" by simp
 qed 
 
 lemma Append_keeps_finite:
@@ -418,16 +411,16 @@
 by (simp only: Subst_def Append_keeps_nonempty nonempty_set_union nonempty_set_sub)
 
 lemma Subst_keeps_eq:
-  assumes substor: "X = L xrhs"
+  assumes substor: "X = L_rhs xrhs"
   and finite: "finite rhs"
-  shows "L (Subst rhs X xrhs) = L rhs" (is "?Left = ?Right")
+  shows "L_rhs (Subst rhs X xrhs) = L_rhs rhs" (is "?Left = ?Right")
 proof-
-  def A \<equiv> "L (rhs - {Trn X r | r. Trn X r \<in> rhs})"
-  have "?Left = A \<union> L (Append_rexp_rhs xrhs (\<Uplus>{r. Trn X r \<in> rhs}))"
+  def A \<equiv> "L_rhs (rhs - {Trn X r | r. Trn X r \<in> rhs})"
+  have "?Left = A \<union> L_rhs (Append_rexp_rhs xrhs (\<Uplus>{r. Trn X r \<in> rhs}))"
     unfolding Subst_def
     unfolding L_rhs_union_distrib[symmetric]
     by (simp add: A_def)
-  moreover have "?Right = A \<union> L ({Trn X r | r. Trn X r \<in> rhs})"
+  moreover have "?Right = A \<union> L_rhs {Trn X r | r. Trn X r \<in> rhs}"
   proof-
     have "rhs = (rhs - {Trn X r | r. Trn X r \<in> rhs}) \<union> ({Trn X r | r. Trn X r \<in> rhs})" by auto
     thus ?thesis 
@@ -435,8 +428,8 @@
       unfolding L_rhs_union_distrib
       by simp
   qed
-  moreover have "L (Append_rexp_rhs xrhs (\<Uplus>{r. Trn X r \<in> rhs})) = L ({Trn X r | r. Trn X r \<in> rhs})" 
-    using finite substor by (simp only: lang_of_append_rexp_rhs rhs_trm_soundness)
+  moreover have "L_rhs (Append_rexp_rhs xrhs (\<Uplus>{r. Trn X r \<in> rhs})) = L_rhs {Trn X r | r. Trn X r \<in> rhs}" 
+    using finite substor by (simp only: lang_of_append_rexp_rhs trm_soundness)
   ultimately show ?thesis by simp
 qed
 
@@ -519,7 +512,7 @@
   assumes invariant_ES: "invariant (ES \<union> {(Y, yrhs)})"
   shows "invariant (Subst_all ES Y (Arden Y yrhs))"
 proof (rule invariantI)
-  have Y_eq_yrhs: "Y = L yrhs" 
+  have Y_eq_yrhs: "Y = L_rhs yrhs" 
     using invariant_ES by (simp only:invariant_def soundness_def, blast)
    have finite_yrhs: "finite yrhs" 
     using invariant_ES by (auto simp:invariant_def finite_rhs_def)
@@ -527,7 +520,7 @@
     using invariant_ES by (auto simp:invariant_def ardenable_all_def)
   show "soundness (Subst_all ES Y (Arden Y yrhs))"
   proof -
-    have "Y = L (Arden Y yrhs)" 
+    have "Y = L_rhs (Arden Y yrhs)" 
       using Y_eq_yrhs invariant_ES finite_yrhs
       using finite_Trn[OF finite_yrhs]
       apply(rule_tac Arden_keeps_eq)
@@ -723,7 +716,7 @@
 lemma every_eqcl_has_reg:
   assumes finite_CS: "finite (UNIV // \<approx>A)"
   and X_in_CS: "X \<in> (UNIV // \<approx>A)"
-  shows "\<exists>r::rexp. X = L r" 
+  shows "\<exists>r. X = L_rexp r" 
 proof -
   from finite_CS X_in_CS 
   obtain xrhs where Inv_ES: "invariant {(X, xrhs)}"
@@ -742,14 +735,14 @@
     using Arden_keeps_finite by auto
   then have fin: "finite {r. Lam r \<in> A}" by (rule finite_Lam)
 
-  have "X = L xrhs" using Inv_ES unfolding invariant_def soundness_def
+  have "X = L_rhs xrhs" using Inv_ES unfolding invariant_def soundness_def
     by simp
-  then have "X = L A" using Inv_ES 
+  then have "X = L_rhs A" using Inv_ES 
     unfolding A_def invariant_def ardenable_all_def finite_rhs_def 
     by (rule_tac Arden_keeps_eq) (simp_all add: finite_Trn)
-  then have "X = L {Lam r | r. Lam r \<in> A}" using eq by simp
-  then have "X = L (\<Uplus>{r. Lam r \<in> A})" using fin by auto
-  then show "\<exists>r::rexp. X = L r" by blast
+  then have "X = L_rhs {Lam r | r. Lam r \<in> A}" using eq by simp
+  then have "X = L_rexp (\<Uplus>{r. Lam r \<in> A})" using fin by auto
+  then show "\<exists>r. X = L_rexp r" by blast
 qed
 
 lemma bchoice_finite_set:
@@ -764,19 +757,19 @@
 
 theorem Myhill_Nerode1:
   assumes finite_CS: "finite (UNIV // \<approx>A)"
-  shows   "\<exists>r::rexp. A = L r"
+  shows   "\<exists>r. A = L_rexp r"
 proof -
   have fin: "finite (finals A)" 
     using finals_in_partitions finite_CS by (rule finite_subset)
-  have "\<forall>X \<in> (UNIV // \<approx>A). \<exists>r::rexp. X = L r" 
+  have "\<forall>X \<in> (UNIV // \<approx>A). \<exists>r. X = L_rexp r" 
     using finite_CS every_eqcl_has_reg by blast
-  then have a: "\<forall>X \<in> finals A. \<exists>r::rexp. X = L r"
+  then have a: "\<forall>X \<in> finals A. \<exists>r. X = L_rexp r"
     using finals_in_partitions by auto
-  then obtain rs::"rexp set" where "\<Union> (finals A) = \<Union>(L ` rs)" "finite rs"
+  then obtain rs::"rexp set" where "\<Union> (finals A) = \<Union>(L_rexp ` rs)" "finite rs"
     using fin by (auto dest: bchoice_finite_set)
-  then have "A = L (\<Uplus>rs)" 
+  then have "A = L_rexp (\<Uplus>rs)" 
     unfolding lang_is_union_of_finals[symmetric] by simp
-  then show "\<exists>r::rexp. A = L r" by blast
+  then show "\<exists>r. A = L_rexp r" by blast
 qed 
 
 
--- a/Myhill_2.thy	Tue May 31 20:32:49 2011 +0000
+++ b/Myhill_2.thy	Thu Jun 02 16:44:35 2011 +0000
@@ -3,17 +3,22 @@
           "~~/src/HOL/Library/List_Prefix"
 begin
 
-section {* Direction @{text "regular language \<Rightarrow>finite partition"} *}
+section {* Direction @{text "regular language \<Rightarrow> finite partition"} *}
 
 definition
   str_eq :: "string \<Rightarrow> lang \<Rightarrow> string \<Rightarrow> bool" ("_ \<approx>_ _")
 where
   "x \<approx>A y \<equiv> (x, y) \<in> (\<approx>A)"
 
+lemma str_eq_def2:
+  shows "\<approx>A = {(x, y) | x y. x \<approx>A y}"
+unfolding str_eq_def
+by simp
+
 definition 
    tag_eq_rel :: "(string \<Rightarrow> 'b) \<Rightarrow> (string \<times> string) set" ("=_=")
 where
-   "=tag= \<equiv> {(x, y) | x y. tag x = tag y}"
+   "=tag= \<equiv> {(x, y). tag x = tag y}"
 
 lemma finite_eq_tag_rel:
   assumes rng_fnt: "finite (range tag)"
@@ -216,7 +221,7 @@
      (\<lambda>x. (\<approx>L1 `` {x}, {(\<approx>L2 `` {x - xa}) | xa.  xa \<le> x \<and> xa \<in> L1}))"
 
 lemma Seq_in_cases:
-  assumes "x @ z \<in> A ;; B"
+  assumes "x @ z \<in> A \<cdot> B"
   shows "(\<exists> x' \<le> x. x' \<in> A \<and> (x - x') @ z \<in> B) \<or> 
          (\<exists> z' \<le> z. (x @ z') \<in> A \<and> (z - z') \<in> B)"
 using assms
@@ -225,12 +230,12 @@
 
 lemma tag_str_SEQ_injI:
   assumes eq_tag: "tag_str_SEQ A B x = tag_str_SEQ A B y" 
-  shows "x \<approx>(A ;; B) y"
+  shows "x \<approx>(A \<cdot> B) y"
 proof -
   { fix x y z
-    assume xz_in_seq: "x @ z \<in> A ;; B"
+    assume xz_in_seq: "x @ z \<in> A \<cdot> B"
     and tag_xy: "tag_str_SEQ A B x = tag_str_SEQ A B y"
-    have"y @ z \<in> A ;; B" 
+    have"y @ z \<in> A \<cdot> B" 
     proof -
       { (* first case with x' in A and (x - x') @ z in B *)
         fix x'
@@ -259,7 +264,7 @@
           with pref_y' y'_in 
           show ?thesis using that by blast
         qed
-	then have "y @ z \<in> A ;; B" by (erule_tac prefixE) (auto simp: Seq_def)
+	then have "y @ z \<in> A \<cdot> B" by (erule_tac prefixE) (auto simp: Seq_def)
       } 
       moreover 
       { (* second case with x @ z' in A and z - z' in B *)
@@ -269,25 +274,25 @@
            using tag_xy unfolding tag_str_SEQ_def by simp
          with h2 have "y @ z' \<in> A"
             unfolding Image_def str_eq_rel_def str_eq_def by auto
-        with h1 h3 have "y @ z \<in> A ;; B"
+        with h1 h3 have "y @ z \<in> A \<cdot> B"
 	  unfolding prefix_def Seq_def
 	  by (auto) (metis append_assoc)
       }
-      ultimately show "y @ z \<in> A ;; B" 
+      ultimately show "y @ z \<in> A \<cdot> B" 
 	using Seq_in_cases [OF xz_in_seq] by blast
     qed
   }
   from this [OF _ eq_tag] and this [OF _ eq_tag [THEN sym]]
-    show "x \<approx>(A ;; B) y" unfolding str_eq_def str_eq_rel_def by blast
+    show "x \<approx>(A \<cdot> B) y" unfolding str_eq_def str_eq_rel_def by blast
 qed 
 
 lemma quot_seq_finiteI [intro]:
   fixes L1 L2::"lang"
   assumes fin1: "finite (UNIV // \<approx>L1)" 
   and     fin2: "finite (UNIV // \<approx>L2)" 
-  shows "finite (UNIV // \<approx>(L1 ;; L2))"
+  shows "finite (UNIV // \<approx>(L1 \<cdot> L2))"
 proof (rule_tac tag = "tag_str_SEQ L1 L2" in tag_finite_imageD)
-  show "\<And>x y. tag_str_SEQ L1 L2 x = tag_str_SEQ L1 L2 y \<Longrightarrow> x \<approx>(L1 ;; L2) y"
+  show "\<And>x y. tag_str_SEQ L1 L2 x = tag_str_SEQ L1 L2 y \<Longrightarrow> x \<approx>(L1 \<cdot> L2) y"
     by (rule tag_str_SEQ_injI)
 next
   have *: "finite ((UNIV // \<approx>L1) \<times> (Pow (UNIV // \<approx>L2)))" 
@@ -431,7 +436,7 @@
         with eq_zab show ?thesis by simp
       qed
       with h5 h6 show ?thesis 
-        by (drule_tac star_intro1) (auto simp:strict_prefix_def elim:prefixE)
+        by (drule_tac star_intro1) (auto simp:strict_prefix_def elim: prefixE)
     qed
   } 
   from this [OF _ eq_tag] and this [OF _ eq_tag [THEN sym]]
@@ -439,15 +444,15 @@
 qed
 
 lemma quot_star_finiteI [intro]:
-  assumes finite1: "finite (UNIV // \<approx>L1)"
-  shows "finite (UNIV // \<approx>(L1\<star>))"
-proof (rule_tac tag = "tag_str_STAR L1" in tag_finite_imageD)
-  show "\<And>x y. tag_str_STAR L1 x = tag_str_STAR L1 y \<Longrightarrow> x \<approx>(L1\<star>) y"
+  assumes finite1: "finite (UNIV // \<approx>A)"
+  shows "finite (UNIV // \<approx>(A\<star>))"
+proof (rule_tac tag = "tag_str_STAR A" in tag_finite_imageD)
+  show "\<And>x y. tag_str_STAR A x = tag_str_STAR A y \<Longrightarrow> x \<approx>(A\<star>) y"
     by (rule tag_str_STAR_injI)
 next
-  have *: "finite (Pow (UNIV // \<approx>L1))" 
+  have *: "finite (Pow (UNIV // \<approx>A))" 
     using finite1 by auto
-  show "finite (range (tag_str_STAR L1))"
+  show "finite (range (tag_str_STAR A))"
     unfolding tag_str_STAR_def
     apply(rule finite_subset[OF _ *])
     unfolding quotient_def
@@ -457,13 +462,12 @@
 subsubsection{* The conclusion *}
 
 lemma Myhill_Nerode2:
-  fixes r::"rexp"
-  shows "finite (UNIV // \<approx>(L r))"
+  shows "finite (UNIV // \<approx>(L_rexp r))"
 by (induct r) (auto)
 
 
 theorem Myhill_Nerode:
-  shows "(\<exists>r::rexp. A = L r) \<longleftrightarrow> finite (UNIV // \<approx>A)"
+  shows "(\<exists>r. A = L_rexp r) \<longleftrightarrow> finite (UNIV // \<approx>A)"
 using Myhill_Nerode1 Myhill_Nerode2 by auto
 
 end
--- a/Paper/Paper.thy	Tue May 31 20:32:49 2011 +0000
+++ b/Paper/Paper.thy	Thu Jun 02 16:44:35 2011 +0000
@@ -1,6 +1,6 @@
 (*<*)
 theory Paper
-imports "../Myhill_2" "~~/src/HOL/Library/LaTeXsugar" 
+imports "../Myhill_2"
 begin
 
 declare [[show_question_marks = false]]
@@ -25,7 +25,7 @@
   quotient ("_ \<^raw:\ensuremath{\!\sslash\!}> _" [90, 90] 90) and
   REL ("\<approx>") and
   UPLUS ("_ \<^raw:\ensuremath{\uplus}> _" [90, 90] 90) and
-  L ("\<^raw:\ensuremath{\cal{L}}>'(_')" [0] 101) and
+  L_rexp ("\<^raw:\ensuremath{\cal{L}}>'(_')" [0] 101) and
   Lam ("\<lambda>'(_')" [100] 100) and 
   Trn ("'(_, _')" [100, 100] 100) and 
   EClass ("\<lbrakk>_\<rbrakk>\<^bsub>_\<^esub>" [100, 100] 100) and
@@ -41,10 +41,44 @@
   tag_str_SEQ ("tag\<^isub>S\<^isub>E\<^isub>Q _ _ _" [100, 100, 100] 100) and
   tag_str_STAR ("tag\<^isub>S\<^isub>T\<^isub>A\<^isub>R _" [100] 100) and
   tag_str_STAR ("tag\<^isub>S\<^isub>T\<^isub>A\<^isub>R _ _" [100, 100] 100)
+
 lemma meta_eq_app:
   shows "f \<equiv> \<lambda>x. g x \<Longrightarrow> f x \<equiv> g x"
   by auto
 
+(* THEOREMS *)
+notation (Rule output)
+  "==>"  ("\<^raw:\mbox{}\inferrule{\mbox{>_\<^raw:}}>\<^raw:{\mbox{>_\<^raw:}}>")
+
+syntax (Rule output)
+  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
+  ("\<^raw:\mbox{}\inferrule{>_\<^raw:}>\<^raw:{\mbox{>_\<^raw:}}>")
+
+  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" 
+  ("\<^raw:\mbox{>_\<^raw:}\\>/ _")
+
+  "_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>")
+
+notation (Axiom output)
+  "Trueprop"  ("\<^raw:\mbox{}\inferrule{\mbox{}}{\mbox{>_\<^raw:}}>")
+
+notation (IfThen output)
+  "==>"  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
+syntax (IfThen output)
+  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
+  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
+  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}> /\<^raw:{\normalsize \,>and\<^raw:\,}>/ _")
+  "_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>")
+
+notation (IfThenNoBox output)
+  "==>"  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
+syntax (IfThenNoBox output)
+  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
+  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
+  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("_ /\<^raw:{\normalsize \,>and\<^raw:\,}>/ _")
+  "_asm" :: "prop \<Rightarrow> asms" ("_")
+
+
 (*>*)
 
 
@@ -239,7 +273,7 @@
   being represented by the empty list, written @{term "[]"}.  \emph{Languages}
   are sets of strings. The language containing all strings is written in
   Isabelle/HOL as @{term "UNIV::string set"}. The concatenation of two languages 
-  is written @{term "A ;; B"} and a language raised to the power @{text n} is written 
+  is written @{term "A \<cdot> B"} and a language raised to the power @{text n} is written 
   @{term "A \<up> n"}. They are defined as usual
 
   \begin{center}
@@ -278,10 +312,10 @@
 
   Central to our proof will be the solution of equational systems
   involving equivalence classes of languages. For this we will use Arden's Lemma \cite{Brzozowski64},
-  which solves equations of the form @{term "X = A ;; X \<union> B"} provided
+  which solves equations of the form @{term "X = A \<cdot> X \<union> B"} provided
   @{term "[] \<notin> A"}. However we will need the following `reverse' 
-  version of Arden's Lemma (`reverse' in the sense of changing the order of @{term "A ;; X"} to
-  \mbox{@{term "X ;; A"}}).
+  version of Arden's Lemma (`reverse' in the sense of changing the order of @{term "A \<cdot> X"} to
+  \mbox{@{term "X \<cdot> A"}}).
 
   \begin{lemma}[Reverse Arden's Lemma]\label{arden}\mbox{}\\
   If @{thm (prem 1) arden} then
@@ -292,10 +326,10 @@
   \begin{proof}
   For the right-to-left direction we assume @{thm (rhs) arden} and show
   that @{thm (lhs) arden} holds. From Prop.~\ref{langprops}@{text "(i)"} 
-  we have @{term "A\<star> = {[]} \<union> A ;; A\<star>"},
-  which is equal to @{term "A\<star> = {[]} \<union> A\<star> ;; A"}. Adding @{text B} to both 
-  sides gives @{term "B ;; A\<star> = B ;; ({[]} \<union> A\<star> ;; A)"}, whose right-hand side
-  is equal to @{term "(B ;; A\<star>) ;; A \<union> B"}. This completes this direction. 
+  we have @{term "A\<star> = {[]} \<union> A \<cdot> A\<star>"},
+  which is equal to @{term "A\<star> = {[]} \<union> A\<star> \<cdot> A"}. Adding @{text B} to both 
+  sides gives @{term "B \<cdot> A\<star> = B \<cdot> ({[]} \<union> A\<star> \<cdot> A)"}, whose right-hand side
+  is equal to @{term "(B \<cdot> A\<star>) \<cdot> A \<union> B"}. This completes this direction. 
 
   For the other direction we assume @{thm (lhs) arden}. By a simple induction
   on @{text n}, we can establish the property
@@ -305,18 +339,18 @@
   \end{center}
   
   \noindent
-  Using this property we can show that @{term "B ;; (A \<up> n) \<subseteq> X"} holds for
-  all @{text n}. From this we can infer @{term "B ;; A\<star> \<subseteq> X"} using the definition
+  Using this property we can show that @{term "B \<cdot> (A \<up> n) \<subseteq> X"} holds for
+  all @{text n}. From this we can infer @{term "B \<cdot> A\<star> \<subseteq> X"} using the definition
   of @{text "\<star>"}.
   For the inclusion in the other direction we assume a string @{text s}
   with length @{text k} is an element in @{text X}. Since @{thm (prem 1) arden}
   we know by Prop.~\ref{langprops}@{text "(ii)"} that 
-  @{term "s \<notin> X ;; (A \<up> Suc k)"} since its length is only @{text k}
-  (the strings in @{term "X ;; (A \<up> Suc k)"} are all longer). 
+  @{term "s \<notin> X \<cdot> (A \<up> Suc k)"} since its length is only @{text k}
+  (the strings in @{term "X \<cdot> (A \<up> Suc k)"} are all longer). 
   From @{text "(*)"} it follows then that
-  @{term s} must be an element in @{term "(\<Union>m\<in>{0..k}. B ;; (A \<up> m))"}. This in turn
-  implies that @{term s} is in @{term "(\<Union>n. B ;; (A \<up> n))"}. Using Prop.~\ref{langprops}@{text "(iii)"} 
-  this is equal to @{term "B ;; A\<star>"}, as we needed to show.\qed
+  @{term s} must be an element in @{term "(\<Union>m\<in>{0..k}. B \<cdot> (A \<up> m))"}. This in turn
+  implies that @{term s} is in @{term "(\<Union>n. B \<cdot> (A \<up> n))"}. Using Prop.~\ref{langprops}@{text "(iii)"} 
+  this is equal to @{term "B \<cdot> A\<star>"}, as we needed to show.\qed
   \end{proof}
 
   \noindent
@@ -494,8 +528,8 @@
   
   \begin{center}
   @{text "\<calL>(Y, r) \<equiv>"} %
-  @{thm (rhs) L_rhs_trm.simps(2)[where X="Y" and r="r", THEN eq_reflection]}\hspace{10mm}
-  @{thm L_rhs_trm.simps(1)[where r="r", THEN eq_reflection]}
+  @{thm (rhs) L_trm.simps(2)[where X="Y" and r="r", THEN eq_reflection]}\hspace{10mm}
+  @{thm L_trm.simps(1)[where r="r", THEN eq_reflection]}
   \end{center}
 
   \noindent
@@ -1026,12 +1060,12 @@
   to be able to infer that 
   %
   \begin{center}
-  @{text "\<dots>"}@{term "x @ z \<in> A ;; B \<longrightarrow> y @ z \<in> A ;; B"}
+  @{text "\<dots>"}@{term "x @ z \<in> A \<cdot> B \<longrightarrow> y @ z \<in> A \<cdot> B"}
   \end{center}
   %
   \noindent
   using the information given by the appropriate tagging-function. The complication 
-  is to find out what the possible splits of @{text "x @ z"} are to be in @{term "A ;; B"}
+  is to find out what the possible splits of @{text "x @ z"} are to be in @{term "A \<cdot> B"}
   (this was easy in case of @{term "A \<union> B"}). To deal with this complication we define the
   notions of \emph{string prefixes} 
   %
@@ -1052,8 +1086,8 @@
   \noindent
   where @{text c} and @{text d} are characters, and @{text x} and @{text y} are strings. 
   
-  Now assuming  @{term "x @ z \<in> A ;; B"} there are only two possible ways of how to `split' 
-  this string to be in @{term "A ;; B"}:
+  Now assuming  @{term "x @ z \<in> A \<cdot> B"} there are only two possible ways of how to `split' 
+  this string to be in @{term "A \<cdot> B"}:
   %
   \begin{center}
   \begin{tabular}{@ {}c@ {\hspace{10mm}}c@ {}}
@@ -1073,7 +1107,7 @@
 
     \draw[decoration={brace,transform={yscale=3}},decorate]
            ($(xa.north west)+(0em,3ex)$) -- ($(z.north east)+(0em,3ex)$)
-               node[midway, above=0.8em]{@{term "x @ z \<in> A ;; B"}};
+               node[midway, above=0.8em]{@{term "x @ z \<in> A \<cdot> B"}};
 
     \draw[decoration={brace,transform={yscale=3}},decorate]
            ($(z.south east)+(0em,0ex)$) -- ($(xxa.south west)+(0em,0ex)$)
@@ -1100,7 +1134,7 @@
 
     \draw[decoration={brace,transform={yscale=3}},decorate]
            ($(x.north west)+(0em,3ex)$) -- ($(zza.north east)+(0em,3ex)$)
-               node[midway, above=0.8em]{@{term "x @ z \<in> A ;; B"}};
+               node[midway, above=0.8em]{@{term "x @ z \<in> A \<cdot> B"}};
 
     \draw[decoration={brace,transform={yscale=3}},decorate]
            ($(za.south east)+(0em,0ex)$) -- ($(x.south west)+(0em,0ex)$)
@@ -1116,7 +1150,7 @@
   \noindent
   Either there is a prefix of @{text x} in @{text A} and the rest is in @{text B} (first picture),
   or @{text x} and a prefix of @{text "z"} is in @{text A} and the rest in @{text B} (second picture).
-  In both cases we have to show that @{term "y @ z \<in> A ;; B"}. For this we use the 
+  In both cases we have to show that @{term "y @ z \<in> A \<cdot> B"}. For this we use the 
   following tagging-function
   %
   \begin{center}
@@ -1134,7 +1168,7 @@
   We have to show injectivity of this tagging-function as
   %
   \begin{center}
-  @{term "\<forall>z. tag_str_SEQ A B x = tag_str_SEQ A B y \<and> x @ z \<in> A ;; B \<longrightarrow> y @ z \<in> A ;; B"}
+  @{term "\<forall>z. tag_str_SEQ A B x = tag_str_SEQ A B y \<and> x @ z \<in> A \<cdot> B \<longrightarrow> y @ z \<in> A \<cdot> B"}
   \end{center}
   %
   \noindent
@@ -1159,13 +1193,13 @@
   @{term "(x - x') \<approx>B (y - y')"} holds. Unfolding the Myhill-Nerode
   relation and together with the fact that @{text "(x - x') @ z \<in> B"}, we
   have @{text "(y - y') @ z \<in> B"}. We already know @{text "y' \<in> A"}, therefore
-  @{term "y @ z \<in> A ;; B"}, as needed in this case.
+  @{term "y @ z \<in> A \<cdot> B"}, as needed in this case.
 
   Second, there exists a @{text "z'"} such that @{term "x @ z' \<in> A"} and @{text "z - z' \<in> B"}.
   By the assumption about @{term "tag_str_SEQ A B"} we have
   @{term "\<approx>A `` {x} = \<approx>A `` {y}"} and thus @{term "x \<approx>A y"}. Which means by the Myhill-Nerode
   relation that @{term "y @ z' \<in> A"} holds. Using @{text "z - z' \<in> B"}, we can conclude also in this case
-  with @{term "y @ z \<in> A ;; B"}. We again can complete the @{const SEQ}-case
+  with @{term "y @ z \<in> A \<cdot> B"}. We again can complete the @{const SEQ}-case
   by setting @{text A} to @{term "L r\<^isub>1"} and @{text B} to @{term "L r\<^isub>2"}.\qed 
   \end{proof}
   
--- a/Paper/ROOT.ML	Tue May 31 20:32:49 2011 +0000
+++ b/Paper/ROOT.ML	Thu Jun 02 16:44:35 2011 +0000
@@ -1,4 +1,3 @@
 no_document use_thy "../Myhill";
-no_document use_thy "~~/src/HOL/Library/LaTeXsugar";
 
 use_thy "Paper"
\ No newline at end of file
--- a/Theories/Closure.thy	Tue May 31 20:32:49 2011 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,140 +0,0 @@
-theory Closure
-imports Myhill_2
-begin
-
-section {* Closure properties of regular languages *}
-
-abbreviation
-  regular :: "lang \<Rightarrow> bool"
-where
-  "regular A \<equiv> \<exists>r::rexp. A = L r"
-
-
-lemma closure_union[intro]:
-  assumes "regular A" "regular B" 
-  shows "regular (A \<union> B)"
-proof -
-  from assms obtain r1 r2::rexp where "L r1 = A" "L r2 = B" by auto
-  then have "A \<union> B = L (ALT r1 r2)" by simp
-  then show "regular (A \<union> B)" by blast
-qed
-
-lemma closure_seq[intro]:
-  assumes "regular A" "regular B" 
-  shows "regular (A ;; B)"
-proof -
-  from assms obtain r1 r2::rexp where "L r1 = A" "L r2 = B" by auto
-  then have "A ;; B = L (SEQ r1 r2)" by simp
-  then show "regular (A ;; B)" by blast
-qed
-
-lemma closure_star[intro]:
-  assumes "regular A"
-  shows "regular (A\<star>)"
-proof -
-  from assms obtain r::rexp where "L r = A" by auto
-  then have "A\<star> = L (STAR r)" by simp
-  then show "regular (A\<star>)" by blast
-qed
-
-lemma closure_complement[intro]:
-  assumes "regular A"
-  shows "regular (- A)"
-proof -
-  from assms have "finite (UNIV // \<approx>A)" by (simp add: Myhill_Nerode)
-  then have "finite (UNIV // \<approx>(-A))" by (simp add: str_eq_rel_def)
-  then show "regular (- A)" by (simp add: Myhill_Nerode)
-qed
-
-lemma closure_difference[intro]:
-  assumes "regular A" "regular B" 
-  shows "regular (A - B)"
-proof -
-  have "A - B = - (- A \<union> B)" by blast
-  moreover
-  have "regular (- (- A \<union> B))" 
-    using assms by blast
-  ultimately show "regular (A - B)" by simp
-qed
-
-lemma closure_intersection[intro]:
-  assumes "regular A" "regular B" 
-  shows "regular (A \<inter> B)"
-proof -
-  have "A \<inter> B = - (- A \<union> - B)" by blast
-  moreover
-  have "regular (- (- A \<union> - B))" 
-    using assms by blast
-  ultimately show "regular (A \<inter> B)" by simp
-qed
-
-
-text {* closure under string reversal *}
-
-fun
-  Rev :: "rexp \<Rightarrow> rexp"
-where
-  "Rev NULL = NULL"
-| "Rev EMPTY = EMPTY"
-| "Rev (CHAR c) = CHAR c"
-| "Rev (ALT r1 r2) = ALT (Rev r1) (Rev r2)"
-| "Rev (SEQ r1 r2) = SEQ (Rev r2) (Rev r1)"
-| "Rev (STAR r) = STAR (Rev r)"
-
-lemma rev_Seq:
-  "(rev ` A) ;; (rev ` B) = rev ` (B ;; A)"
-unfolding Seq_def image_def
-apply(auto)
-apply(rule_tac x="xb @ xa" in exI)
-apply(auto)
-done
-
-lemma rev_Star1:
-  assumes a: "s \<in> (rev ` A)\<star>"
-  shows "s \<in> rev ` (A\<star>)"
-using a
-proof(induct rule: star_induct)
-  case (step s1 s2)
-  have inj: "inj (rev::string \<Rightarrow> string)" unfolding inj_on_def by auto
-  have "s1 \<in> rev ` A" "s2 \<in> rev ` (A\<star>)" by fact+
-  then obtain x1 x2 where "x1 \<in> A" "x2 \<in> A\<star>" and eqs: "s1 = rev x1" "s2 = rev x2" by auto
-  then have "x1 \<in> A\<star>" "x2 \<in> A\<star>" by (auto intro: star_intro2)
-  then have "x2 @ x1 \<in> A\<star>" by (auto intro: star_intro1)
-  then have "rev (x2 @ x1) \<in> rev ` A\<star>" using inj by (simp only: inj_image_mem_iff)
-  then show "s1 @ s2 \<in>  rev ` A\<star>" using eqs by simp
-qed (auto)
-
-lemma rev_Star2:
-  assumes a: "s \<in> A\<star>"
-  shows "rev s \<in> (rev ` A)\<star>"
-using a
-proof(induct rule: star_induct)
-  case (step s1 s2)
-  have inj: "inj (rev::string \<Rightarrow> string)" unfolding inj_on_def by auto
-  have "s1 \<in> A"by fact
-  then have "rev s1 \<in> rev ` A" using inj by (simp only: inj_image_mem_iff)
-  then have "rev s1 \<in> (rev ` A)\<star>" by (auto intro: star_intro2)
-  moreover
-  have "rev s2 \<in> (rev ` A)\<star>" by fact
-  ultimately show "rev (s1 @ s2) \<in>  (rev ` A)\<star>" by (auto intro: star_intro1)
-qed (auto)
-
-lemma rev_Star:
-  "(rev ` A)\<star> = rev ` (A\<star>)"
-using rev_Star1 rev_Star2 by auto
-
-lemma rev_lang:
-  "L (Rev r) = rev ` (L r)"
-by (induct r) (simp_all add: rev_Star rev_Seq image_Un)
-
-lemma closure_reversal[intro]:
-  assumes "regular A"
-  shows "regular (rev ` A)"
-proof -
-  from assms obtain r::rexp where "L r = A" by auto
-  then have "L (Rev r) = rev ` A" by (simp add: rev_lang)
-  then show "regular (rev` A)" by blast
-qed
-  
-
-end
\ No newline at end of file
--- a/Theories/Derivs.thy	Tue May 31 20:32:49 2011 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,521 +0,0 @@
-theory Derivs
-imports Closure
-begin
-
-section {* Experiments with Derivatives -- independent of Myhill-Nerode *}
-
-subsection {* Left-Quotients *}
-
-definition
-  Delta :: "lang \<Rightarrow> lang"
-where
-  "Delta A = (if [] \<in> A then {[]} else {})"
-
-definition
-  Der :: "char \<Rightarrow> lang \<Rightarrow> lang"
-where
-  "Der c A \<equiv> {s. [c] @ s \<in> A}"
-
-definition
-  Ders :: "string \<Rightarrow> lang \<Rightarrow> lang"
-where
-  "Ders s A \<equiv> {s'. s @ s' \<in> A}"
-
-definition
-  Ders_set :: "lang \<Rightarrow> lang \<Rightarrow> lang"
-where
-  "Ders_set A B \<equiv> {s' | s s'. s @ s' \<in> B \<and> s \<in> A}"
-
-lemma Ders_set_Ders:
-  shows "Ders_set A B = (\<Union>s \<in> A. Ders s B)"
-unfolding Ders_set_def Ders_def
-by auto
-
-lemma Der_null [simp]:
-  shows "Der c {} = {}"
-unfolding Der_def
-by auto
-
-lemma Der_empty [simp]:
-  shows "Der c {[]} = {}"
-unfolding Der_def
-by auto
-
-lemma Der_char [simp]:
-  shows "Der c {[d]} = (if c = d then {[]} else {})"
-unfolding Der_def
-by auto
-
-lemma Der_union [simp]:
-  shows "Der c (A \<union> B) = Der c A \<union> Der c B"
-unfolding Der_def
-by auto
-
-lemma Der_seq [simp]:
-  shows "Der c (A ;; B) = (Der c A) ;; B \<union> (Delta A ;; Der c B)"
-unfolding Der_def Delta_def
-unfolding Seq_def
-by (auto simp add: Cons_eq_append_conv)
-
-lemma Der_star [simp]:
-  shows "Der c (A\<star>) = (Der c A) ;; A\<star>"
-apply(subst star_cases)
-apply(simp only: Delta_def Der_union Der_seq Der_empty)
-apply(simp add: Der_def Seq_def)
-apply(auto)
-apply(drule star_decom)
-apply(auto simp add: Cons_eq_append_conv)
-done
-
-lemma Ders_singleton:
-  shows "Ders [c] A = Der c A"
-unfolding Der_def Ders_def
-by simp
-
-lemma Ders_append:
-  shows "Ders (s1 @ s2) A = Ders s2 (Ders s1 A)"
-unfolding Ders_def by simp 
-
-lemma MN_Rel_Ders:
-  shows "x \<approx>A y \<longleftrightarrow> Ders x A = Ders y A"
-unfolding Ders_def str_eq_def str_eq_rel_def
-by auto
-
-subsection {* Brozowsky's derivatives of regular expressions *}
-
-fun
-  nullable :: "rexp \<Rightarrow> bool"
-where
-  "nullable (NULL) = False"
-| "nullable (EMPTY) = True"
-| "nullable (CHAR c) = False"
-| "nullable (ALT r1 r2) = (nullable r1 \<or> nullable r2)"
-| "nullable (SEQ r1 r2) = (nullable r1 \<and> nullable r2)"
-| "nullable (STAR r) = True"
-
-fun
-  der :: "char \<Rightarrow> rexp \<Rightarrow> rexp"
-where
-  "der c (NULL) = NULL"
-| "der c (EMPTY) = NULL"
-| "der c (CHAR c') = (if c = c' then EMPTY else NULL)"
-| "der c (ALT r1 r2) = ALT (der c r1) (der c r2)"
-| "der c (SEQ r1 r2) = ALT (SEQ (der c r1) r2) (if nullable r1 then der c r2 else NULL)"
-| "der c (STAR r) = SEQ (der c r) (STAR r)"
-
-function 
-  ders :: "string \<Rightarrow> rexp \<Rightarrow> rexp"
-where
-  "ders [] r = r"
-| "ders (s @ [c]) r = der c (ders s r)"
-by (auto) (metis rev_cases)
-
-termination
-  by (relation "measure (length o fst)") (auto)
-
-lemma Delta_nullable:
-  shows "Delta (L r) = (if nullable r then {[]} else {})"
-unfolding Delta_def
-by (induct r) (auto simp add: Seq_def split: if_splits)
-
-lemma Der_der:
-  fixes r::rexp
-  shows "Der c (L r) = L (der c r)"
-by (induct r) (simp_all add: Delta_nullable)
-
-lemma Ders_ders:
-  fixes r::rexp
-  shows "Ders s (L r) = L (ders s r)"
-apply(induct s rule: rev_induct)
-apply(simp add: Ders_def)
-apply(simp only: ders.simps)
-apply(simp only: Ders_append)
-apply(simp only: Ders_singleton)
-apply(simp only: Der_der)
-done
-
-
-subsection {* Antimirov's Partial Derivatives *}
-
-abbreviation
-  "SEQS R r \<equiv> {SEQ r' r | r'. r' \<in> R}"
-
-fun
-  pder :: "char \<Rightarrow> rexp \<Rightarrow> rexp set"
-where
-  "pder c NULL = {NULL}"
-| "pder c EMPTY = {NULL}"
-| "pder c (CHAR c') = (if c = c' then {EMPTY} else {NULL})"
-| "pder c (ALT r1 r2) = (pder c r1) \<union> (pder c r2)"
-| "pder c (SEQ r1 r2) = SEQS (pder c r1) r2 \<union> (if nullable r1 then pder c r2 else {})"
-| "pder c (STAR r) = SEQS (pder c r) (STAR r)"
-
-abbreviation
-  "pder_set c R \<equiv> \<Union>r \<in> R. pder c r"
-
-function 
-  pders :: "string \<Rightarrow> rexp \<Rightarrow> rexp set"
-where
-  "pders [] r = {r}"
-| "pders (s @ [c]) r = pder_set c (pders s r)"
-by (auto) (metis rev_cases)
-
-termination
-  by (relation "measure (length o fst)") (auto)
-
-abbreviation
-  "pders_set A r \<equiv> \<Union>s \<in> A. pders s r"
-
-
-lemma pders_append:
-  "pders (s1 @ s2) r = \<Union> (pders s2) ` (pders s1 r)"
-apply(induct s2 arbitrary: s1 r rule: rev_induct)
-apply(simp)
-apply(subst append_assoc[symmetric])
-apply(simp only: pders.simps)
-apply(auto)
-done
-
-lemma pders_singleton:
-  "pders [c] r = pder c r"
-apply(subst append_Nil[symmetric])
-apply(simp only: pders.simps)
-apply(simp)
-done
-
-lemma pder_set_lang:
-  shows "(\<Union> (L ` pder_set c R)) = (\<Union>r \<in> R. (\<Union>L ` (pder c r)))"
-unfolding image_def 
-by auto
-
-lemma
-  shows seq_UNION_left:  "B ;; (\<Union>n\<in>C. A n) = (\<Union>n\<in>C. B ;; A n)"
-  and   seq_UNION_right: "(\<Union>n\<in>C. A n) ;; B = (\<Union>n\<in>C. A n ;; B)"
-unfolding Seq_def by auto
-
-lemma Der_pder:
-  fixes r::rexp
-  shows "Der c (L r) = \<Union> L ` (pder c r)"
-by (induct r) (auto simp add: Delta_nullable seq_UNION_right)
-
-lemma Ders_pders:
-  fixes r::rexp
-  shows "Ders s (L r) = \<Union> L ` (pders s r)"
-proof (induct s rule: rev_induct)
-  case (snoc c s)
-  have ih: "Ders s (L r) = \<Union> L ` (pders s r)" by fact
-  have "Ders (s @ [c]) (L r) = Ders [c] (Ders s (L r))"
-    by (simp add: Ders_append)
-  also have "\<dots> = Der c (\<Union> L ` (pders s r))" using ih
-    by (simp add: Ders_singleton)
-  also have "\<dots> = (\<Union>r\<in>pders s r. Der c (L r))" 
-    unfolding Der_def image_def by auto
-  also have "\<dots> = (\<Union>r\<in>pders s r. (\<Union> L `  (pder c r)))"
-    by (simp add: Der_pder)
-  also have "\<dots> = (\<Union>L ` (pder_set c (pders s r)))"
-    by (simp add: pder_set_lang)
-  also have "\<dots> = (\<Union>L ` (pders (s @ [c]) r))"
-    by simp
-  finally show "Ders (s @ [c]) (L r) = \<Union>L ` pders (s @ [c]) r" .
-qed (simp add: Ders_def)
-
-lemma Ders_set_pders_set:
-  fixes r::rexp
-  shows "Ders_set A (L r) = (\<Union> L ` (pders_set A r))"
-by (simp add: Ders_set_Ders Ders_pders)
-
-lemma pders_NULL [simp]:
-  shows "pders s NULL = {NULL}"
-by (induct s rule: rev_induct) (simp_all)
-
-lemma pders_EMPTY [simp]:
-  shows "pders s EMPTY = (if s = [] then {EMPTY} else {NULL})"
-by (induct s rule: rev_induct) (auto)
-
-lemma pders_CHAR [simp]:
-  shows "pders s (CHAR c) = (if s = [] then {CHAR c} else (if s = [c] then {EMPTY} else {NULL}))"
-by (induct s rule: rev_induct) (auto)
-
-lemma pders_ALT [simp]:
-  shows "pders s (ALT r1 r2) = (if s = [] then {ALT r1 r2} else (pders s r1) \<union> (pders s r2))"
-by (induct s rule: rev_induct) (auto)
-
-definition
-  "Suf s \<equiv> {v. v \<noteq> [] \<and> (\<exists>u. u @ v = s)}"
-
-lemma Suf:
-  shows "Suf (s @ [c]) = (Suf s) ;; {[c]} \<union> {[c]}"
-unfolding Suf_def Seq_def
-by (auto simp add: append_eq_append_conv2 append_eq_Cons_conv)
-
-lemma Suf_Union:
-  shows "(\<Union>v \<in> Suf s ;; {[c]}. P v) = (\<Union>v \<in> Suf s. P (v @ [c]))"
-by (auto simp add: Seq_def)
-
-lemma inclusion1:
-  shows "pder_set c (SEQS R r2) \<subseteq> SEQS (pder_set c R) r2 \<union> (pder c r2)"
-apply(auto simp add: if_splits)
-apply(blast)
-done
-
-lemma pders_SEQ:
-  shows "pders s (SEQ r1 r2) \<subseteq> SEQS (pders s r1) r2 \<union> (\<Union>v \<in> Suf s. pders v r2)"
-proof (induct s rule: rev_induct)
-  case (snoc c s)
-  have ih: "pders s (SEQ r1 r2) \<subseteq> SEQS (pders s r1) r2 \<union> (\<Union>v \<in> Suf s. pders v r2)" 
-    by fact
-  have "pders (s @ [c]) (SEQ r1 r2) = pder_set c (pders s (SEQ r1 r2))" by simp
-  also have "\<dots> \<subseteq> pder_set c (SEQS (pders s r1) r2 \<union> (\<Union>v \<in> Suf s. pders v r2))"
-    using ih by auto 
-  also have "\<dots> = pder_set c (SEQS (pders s r1) r2) \<union> pder_set c (\<Union>v \<in> Suf s. pders v r2)"
-    by (simp)
-  also have "\<dots> = pder_set c (SEQS (pders s r1) r2) \<union> (\<Union>v \<in> Suf s. pder_set c (pders v r2))"
-    by (simp)
-  also have "\<dots> \<subseteq> pder_set c (SEQS (pders s r1) r2) \<union> (pder c r2) \<union> (\<Union>v \<in> Suf s. pders (v @ [c]) r2)"
-    by (auto)
-  also have "\<dots> \<subseteq> SEQS (pder_set c (pders s r1)) r2 \<union> (pder c r2) \<union> (\<Union>v \<in> Suf s. pders (v @ [c]) r2)"
-    using inclusion1 by blast
-  also have "\<dots> = SEQS (pders (s @ [c]) r1) r2 \<union> (\<Union>v \<in> Suf (s @ [c]). pders v r2)"
-    apply(subst (2) pders.simps)
-    apply(simp only: Suf)
-    apply(simp add: Suf_Union pders_singleton)
-    apply(auto)
-    done
-  finally show ?case .
-qed (simp)
-
-lemma pders_STAR:
-  assumes a: "s \<noteq> []"
-  shows "pders s (STAR r) \<subseteq> (\<Union>v \<in> Suf s. SEQS (pders v r) (STAR r))"
-using a
-proof (induct s rule: rev_induct)
-  case (snoc c s)
-  have ih: "s \<noteq> [] \<Longrightarrow> pders s (STAR r) \<subseteq> (\<Union>v\<in>Suf s. SEQS (pders v r) (STAR r))" by fact
-  { assume asm: "s \<noteq> []"
-    have "pders (s @ [c]) (STAR r) = pder_set c (pders s (STAR r))" by simp
-    also have "\<dots> \<subseteq> (pder_set c (\<Union>v\<in>Suf s. SEQS (pders v r) (STAR r)))"
-      using ih[OF asm] by blast
-    also have "\<dots> = (\<Union>v\<in>Suf s. pder_set c (SEQS (pders v r) (STAR r)))"
-      by simp
-    also have "\<dots> \<subseteq> (\<Union>v\<in>Suf s. (SEQS (pder_set c (pders v r)) (STAR r) \<union> pder c (STAR r)))"
-      using inclusion1 by blast
-    also have "\<dots> = (\<Union>v\<in>Suf s. (SEQS (pder_set c (pders v r)) (STAR r))) \<union> pder c (STAR r)"
-      using asm by (auto simp add: Suf_def)
-    also have "\<dots> = (\<Union>v\<in>Suf s. (SEQS (pders (v @ [c]) r) (STAR r))) \<union> (SEQS (pder c r) (STAR r))"
-      by simp
-    also have "\<dots> = (\<Union>v\<in>Suf (s @ [c]). (SEQS (pders v r) (STAR r)))"
-      apply(simp only: Suf)
-      apply(simp add: Suf_Union pders_singleton)
-      apply(auto)
-      done
-    finally have ?case .
-  }
-  moreover
-  { assume asm: "s = []"
-    then have ?case
-      apply(simp add: pders_singleton Suf_def)
-      apply(auto)
-      apply(rule_tac x="[c]" in exI)
-      apply(simp add: pders_singleton)
-      done
-  }
-  ultimately show ?case by blast
-qed (simp)
-
-abbreviation 
-  "UNIV1 \<equiv> UNIV - {[]}"
-
-lemma pders_set_NULL:
-  shows "pders_set UNIV1 NULL = {NULL}"
-by auto
-
-lemma pders_set_EMPTY:
-  shows "pders_set UNIV1 EMPTY = {NULL}"
-by (auto split: if_splits)
-
-lemma pders_set_CHAR:
-  shows "pders_set UNIV1 (CHAR c) \<subseteq> {EMPTY, NULL}"
-by (auto split: if_splits)
-
-lemma pders_set_ALT:
-  shows "pders_set UNIV1 (ALT r1 r2) = pders_set UNIV1 r1 \<union> pders_set UNIV1 r2"
-by auto
-
-lemma pders_set_SEQ_aux:
-  assumes a: "s \<in> UNIV1"
-  shows "pders_set (Suf s) r2 \<subseteq> pders_set UNIV1 r2"
-using a by (auto simp add: Suf_def)
-
-lemma pders_set_SEQ:
-  shows "pders_set UNIV1 (SEQ r1 r2) \<subseteq> SEQS (pders_set UNIV1 r1) r2 \<union> pders_set UNIV1 r2"
-apply(rule UN_least)
-apply(rule subset_trans)
-apply(rule pders_SEQ)
-apply(simp)
-apply(rule conjI) 
-apply(auto)[1]
-apply(rule subset_trans)
-apply(rule pders_set_SEQ_aux)
-apply(auto)
-done
-
-lemma pders_set_STAR:
-  shows "pders_set UNIV1 (STAR r) \<subseteq> SEQS (pders_set UNIV1 r) (STAR r)"
-apply(rule UN_least)
-apply(rule subset_trans)
-apply(rule pders_STAR)
-apply(simp)
-apply(simp add: Suf_def)
-apply(auto)
-done
-
-lemma finite_SEQS:
-  assumes a: "finite A"
-  shows "finite (SEQS A r)"
-using a by (auto)
-
-lemma finite_pders_set_UNIV1:
-  shows "finite (pders_set UNIV1 r)"
-apply(induct r)
-apply(simp)
-apply(simp only: pders_set_EMPTY)
-apply(simp)
-apply(rule finite_subset)
-apply(rule pders_set_CHAR)
-apply(simp)
-apply(rule finite_subset)
-apply(rule pders_set_SEQ)
-apply(simp only: finite_SEQS finite_Un)
-apply(simp)
-apply(simp only: pders_set_ALT)
-apply(simp)
-apply(rule finite_subset)
-apply(rule pders_set_STAR)
-apply(simp only: finite_SEQS)
-done
-    
-lemma pders_set_UNIV_UNIV1:
-  shows "pders_set UNIV r = pders [] r \<union> pders_set UNIV1 r"
-apply(auto)
-apply(rule_tac x="[]" in exI)
-apply(simp)
-done
-
-lemma finite_pders_set_UNIV:
-  shows "finite (pders_set UNIV r)"
-unfolding pders_set_UNIV_UNIV1
-by (simp add: finite_pders_set_UNIV1)
-
-lemma finite_pders_set:
-  shows "finite (pders_set A r)"
-apply(rule rev_finite_subset)
-apply(rule_tac r="r" in finite_pders_set_UNIV)
-apply(auto)
-done
-
-lemma finite_pders:
-  shows "finite (pders s r)"
-using finite_pders_set[where A="{s}" and r="r"]
-by simp
-
-
-lemma test: 
-  shows "pders_set A r = (\<Union> {pders s r | s. s \<in> A})"
-by auto
-
-lemma finite_pow_pders:
-  shows "finite (Pow (\<Union> {pders s r | s. s \<in> A}))"
-using finite_pders_set
-by (simp add: test)
-
-lemma test2:
-  shows "{pders s r | s. s \<in> A} \<subseteq> Pow (\<Union> {pders s r | s. s \<in> A})"
-  by auto
-
-lemma test3:
-  shows "finite ({pders s r | s. s \<in> A})"
-apply(rule finite_subset)
-apply(rule test2)
-apply(rule finite_pow_pders)
-done
-
-lemma Myhill_Nerode_aux:
-  fixes r::"rexp"
-  shows "finite (UNIV // =(\<lambda>x. pders x r)=)"
-apply(rule finite_eq_tag_rel)
-apply(rule rev_finite_subset)
-apply(rule test3)
-apply(auto)
-done
-
-lemma Myhill_Nerode3:
-  fixes r::"rexp"
-  shows "finite (UNIV // \<approx>(L r))"
-apply(rule refined_partition_finite)
-apply(rule Myhill_Nerode_aux[where r="r"])
-apply(simp add: tag_eq_rel_def)
-apply(auto)[1]
-unfolding str_eq_def[symmetric]
-unfolding MN_Rel_Ders
-apply(simp add: Ders_pders)
-apply(rule equivI)
-apply(auto simp add: tag_eq_rel_def refl_on_def sym_def trans_def)
-apply(rule equivI)
-apply(auto simp add: str_eq_rel_def refl_on_def sym_def trans_def)
-done
-
-
-section {* Closure under Left-Quotients *}
-
-lemma closure_left_quotient:
-  assumes "regular A"
-  shows "regular (Ders_set B A)"
-proof -
-  from assms obtain r::rexp where eq: "L r = A" by auto
-  have fin: "finite (pders_set B r)" by (rule finite_pders_set)
-  
-  have "Ders_set B (L r) = (\<Union> L ` (pders_set B r))"
-    by (simp add: Ders_set_pders_set)
-  also have "\<dots> = L (\<Uplus>(pders_set B r))" using fin by simp
-  finally have "Ders_set B A = L (\<Uplus>(pders_set B r))" using eq
-    by simp
-  then show "regular (Ders_set B A)" by auto
-qed
-
-
-section {* Relating standard and partial derivations *}
-
-lemma
-  shows "(\<Union> L ` (pder c r)) = L (der c r)"
-unfolding Der_der[symmetric] Der_pder by simp
-
-lemma
-  shows "(\<Union> L ` (pders s r)) = L (ders s r)"
-unfolding Ders_ders[symmetric] Ders_pders by simp
-
-
-section {* attempt to prove MN *}
-(*
-lemma Myhill_Nerode3:
-  fixes r::"rexp"
-  shows "finite (UNIV // =(\<lambda>x. pders x r)=)"
-apply(rule finite_eq_tag_rel)
-apply(rule finite_pders_set)
-apply(simp add: Range_def)
-unfolding Quotien_def
-by (induct r) (auto)
-*)
-
-fun
-  width :: "rexp \<Rightarrow> nat"
-where
-  "width (NULL) = 0"
-| "width (EMPTY) = 0"
-| "width (CHAR c) = 1"
-| "width (ALT r1 r2) = width r1 + width r2"
-| "width (SEQ r1 r2) = width r1 + width r2"
-| "width (STAR r) = width r"
-
-
- 
-end
\ No newline at end of file
--- a/Theories/Folds.thy	Tue May 31 20:32:49 2011 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,22 +0,0 @@
-theory Folds
-imports Main
-begin
-
-section {* Folds for Sets *}
-
-text {*
-  To obtain equational system out of finite set of equivalence classes, a fold operation
-  on finite sets @{text "folds"} is defined. The use of @{text "SOME"} makes @{text "folds"}
-  more robust than the @{text "fold"} in the 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"
-
-
-end
\ No newline at end of file
--- a/Theories/Myhill_1.thy	Tue May 31 20:32:49 2011 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,783 +0,0 @@
-theory Myhill_1
-imports Main Folds Regular
-        "~~/src/HOL/Library/While_Combinator" 
-begin
-
-section {* Direction @{text "finite partition \<Rightarrow> regular language"} *}
-
-lemma Pair_Collect[simp]:
-  shows "(x, y) \<in> {(x, y). P x y} \<longleftrightarrow> P x y"
-by simp
-
-text {* Myhill-Nerode relation *}
-
-definition
-  str_eq_rel :: "lang \<Rightarrow> (string \<times> string) set" ("\<approx>_" [100] 100)
-where
-  "\<approx>A \<equiv> {(x, y).  (\<forall>z. x @ z \<in> A \<longleftrightarrow> y @ z \<in> A)}"
-
-definition 
-  finals :: "lang \<Rightarrow> lang set"
-where
-  "finals A \<equiv> {\<approx>A `` {s} | s . s \<in> A}"
-
-lemma lang_is_union_of_finals: 
-  shows "A = \<Union> finals A"
-unfolding finals_def
-unfolding Image_def
-unfolding str_eq_rel_def
-by (auto) (metis append_Nil2)
-
-lemma finals_in_partitions:
-  shows "finals A \<subseteq> (UNIV // \<approx>A)"
-unfolding finals_def quotient_def
-by auto
-
-section {* Equational systems *}
-
-text {* The two kinds of terms in the rhs of equations. *}
-
-datatype rhs_trm = 
-   Lam "rexp"            (* Lambda-marker *)
- | Trn "lang" "rexp"     (* Transition *)
-
-
-overloading L_rhs_trm \<equiv> "L:: rhs_trm \<Rightarrow> lang"
-begin
-  fun L_rhs_trm:: "rhs_trm \<Rightarrow> lang"
-  where
-    "L_rhs_trm (Lam r) = L r" 
-  | "L_rhs_trm (Trn X r) = X ;; L r"
-end
-
-overloading L_rhs \<equiv> "L:: rhs_trm set \<Rightarrow> lang"
-begin
-   fun L_rhs:: "rhs_trm set \<Rightarrow> lang"
-   where 
-     "L_rhs rhs = \<Union> (L ` rhs)"
-end
-
-lemma L_rhs_set:
-  shows "L {Trn X r | r. P r} = \<Union>{L (Trn X r) | r. P r}"
-by (auto simp del: L_rhs_trm.simps)
-
-lemma L_rhs_union_distrib:
-  fixes A B::"rhs_trm set"
-  shows "L A \<union> L B = L (A \<union> B)"
-by simp
-
-
-
-text {* Transitions between equivalence classes *}
-
-definition 
-  transition :: "lang \<Rightarrow> char \<Rightarrow> lang \<Rightarrow> bool" ("_ \<Turnstile>_\<Rightarrow>_" [100,100,100] 100)
-where
-  "Y \<Turnstile>c\<Rightarrow> X \<equiv> Y ;; {[c]} \<subseteq> X"
-
-text {* Initial equational system *}
-
-definition
-  "Init_rhs CS X \<equiv>  
-      if ([] \<in> X) then 
-          {Lam EMPTY} \<union> {Trn Y (CHAR c) | Y c. Y \<in> CS \<and> Y \<Turnstile>c\<Rightarrow> X}
-      else 
-          {Trn Y (CHAR c)| Y c. Y \<in> CS \<and> Y \<Turnstile>c\<Rightarrow> X}"
-
-definition 
-  "Init CS \<equiv> {(X, Init_rhs CS X) | X.  X \<in> CS}"
-
-
-section {* Arden Operation on equations *}
-
-fun 
-  Append_rexp :: "rexp \<Rightarrow> rhs_trm \<Rightarrow> rhs_trm"
-where
-  "Append_rexp r (Lam rexp)   = Lam (SEQ rexp r)"
-| "Append_rexp r (Trn X rexp) = Trn X (SEQ rexp r)"
-
-
-definition
-  "Append_rexp_rhs rhs rexp \<equiv> (Append_rexp rexp) ` rhs"
-
-definition 
-  "Arden X rhs \<equiv> 
-     Append_rexp_rhs (rhs - {Trn X r | r. Trn X r \<in> rhs}) (STAR (\<Uplus> {r. Trn X r \<in> rhs}))"
-
-
-section {* Substitution Operation on equations *}
-
-definition 
-  "Subst rhs X xrhs \<equiv> 
-        (rhs - {Trn X r | r. Trn X r \<in> rhs}) \<union> (Append_rexp_rhs xrhs (\<Uplus> {r. Trn X r \<in> rhs}))"
-
-definition
-  Subst_all :: "(lang \<times> rhs_trm set) set \<Rightarrow> lang \<Rightarrow> rhs_trm set \<Rightarrow> (lang \<times> rhs_trm set) set"
-where
-  "Subst_all ES X xrhs \<equiv> {(Y, Subst yrhs X xrhs) | Y yrhs. (Y, yrhs) \<in> ES}"
-
-definition
-  "Remove ES X xrhs \<equiv> 
-      Subst_all  (ES - {(X, xrhs)}) X (Arden X xrhs)"
-
-
-section {* While-combinator *}
-
-definition 
-  "Iter X ES \<equiv> (let (Y, yrhs) = SOME (Y, yrhs). (Y, yrhs) \<in> ES \<and> X \<noteq> Y
-                in Remove ES Y yrhs)"
-
-lemma IterI2:
-  assumes "(Y, yrhs) \<in> ES"
-  and     "X \<noteq> Y"
-  and     "\<And>Y yrhs. \<lbrakk>(Y, yrhs) \<in> ES; X \<noteq> Y\<rbrakk> \<Longrightarrow> Q (Remove ES Y yrhs)"
-  shows "Q (Iter X ES)"
-unfolding Iter_def using assms
-by (rule_tac a="(Y, yrhs)" in someI2) (auto)
-
-abbreviation
-  "Cond ES \<equiv> card ES \<noteq> 1"
-
-definition 
-  "Solve X ES \<equiv> while Cond (Iter X) ES"
-
-
-section {* Invariants *}
-
-definition 
-  "distinctness ES \<equiv> 
-     \<forall> X rhs rhs'. (X, rhs) \<in> ES \<and> (X, rhs') \<in> ES \<longrightarrow> rhs = rhs'"
-
-definition 
-  "soundness ES \<equiv> \<forall>(X, rhs) \<in> ES. X = L rhs"
-
-definition 
-  "ardenable rhs \<equiv> (\<forall> Y r. Trn Y r \<in> rhs \<longrightarrow> [] \<notin> L r)"
-
-definition 
-  "ardenable_all ES \<equiv> \<forall>(X, rhs) \<in> ES. ardenable rhs"
-
-definition
-  "finite_rhs ES \<equiv> \<forall>(X, rhs) \<in> ES. finite rhs"
-
-lemma finite_rhs_def2:
-  "finite_rhs ES = (\<forall> X rhs. (X, rhs) \<in> ES \<longrightarrow> finite rhs)"
-unfolding finite_rhs_def by auto
-
-definition 
-  "rhss rhs \<equiv> {X | X r. Trn X r \<in> rhs}"
-
-definition
-  "lhss ES \<equiv> {Y | Y yrhs. (Y, yrhs) \<in> ES}"
-
-definition 
-  "validity ES \<equiv> \<forall>(X, rhs) \<in> ES. rhss rhs \<subseteq> lhss ES"
-
-lemma rhss_union_distrib:
-  shows "rhss (A \<union> B) = rhss A \<union> rhss B"
-by (auto simp add: rhss_def)
-
-lemma lhss_union_distrib:
-  shows "lhss (A \<union> B) = lhss A \<union> lhss B"
-by (auto simp add: lhss_def)
-
-
-definition 
-  "invariant ES \<equiv> finite ES
-                \<and> finite_rhs ES
-                \<and> soundness ES 
-                \<and> distinctness ES 
-                \<and> ardenable_all ES 
-                \<and> validity ES"
-
-
-lemma invariantI:
-  assumes "soundness ES" "finite ES" "distinctness ES" "ardenable_all ES" 
-          "finite_rhs ES" "validity ES"
-  shows "invariant ES"
-using assms by (simp add: invariant_def)
-
-
-subsection {* The proof of this direction *}
-
-lemma finite_Trn:
-  assumes fin: "finite rhs"
-  shows "finite {r. Trn Y r \<in> rhs}"
-proof -
-  have "finite {Trn Y r | Y r. Trn Y r \<in> rhs}"
-    by (rule rev_finite_subset[OF fin]) (auto)
-  then have "finite ((\<lambda>(Y, r). Trn Y r) ` {(Y, r) | Y r. Trn Y r \<in> rhs})"
-    by (simp add: image_Collect)
-  then have "finite {(Y, r) | Y r. Trn Y r \<in> rhs}"
-    by (erule_tac finite_imageD) (simp add: inj_on_def)
-  then show "finite {r. Trn Y r \<in> rhs}"
-    by (erule_tac f="snd" in finite_surj) (auto simp add: image_def)
-qed
-
-lemma finite_Lam:
-  assumes fin: "finite rhs"
-  shows "finite {r. Lam r \<in> rhs}"
-proof -
-  have "finite {Lam r | r. Lam r \<in> rhs}"
-    by (rule rev_finite_subset[OF fin]) (auto)
-  then show "finite {r. Lam r \<in> rhs}"
-    apply(simp add: image_Collect[symmetric])
-    apply(erule finite_imageD)
-    apply(auto simp add: inj_on_def)
-    done
-qed
-
-lemma rhs_trm_soundness:
-  assumes finite:"finite rhs"
-  shows "L ({Trn X r| r. Trn X r \<in> rhs}) = X ;; (L (\<Uplus>{r. Trn X r \<in> rhs}))"
-proof -
-  have "finite {r. Trn X r \<in> rhs}" 
-    by (rule finite_Trn[OF finite]) 
-  then show "L ({Trn X r| r. Trn X r \<in> rhs}) = X ;; (L (\<Uplus>{r. Trn X r \<in> rhs}))"
-    by (simp only: L_rhs_set L_rhs_trm.simps) (auto simp add: Seq_def)
-qed
-
-lemma lang_of_append_rexp:
-  "L (Append_rexp r rhs_trm) = L rhs_trm ;; L r"
-by (induct rule: Append_rexp.induct)
-   (auto simp add: seq_assoc)
-
-lemma lang_of_append_rexp_rhs:
-  "L (Append_rexp_rhs rhs r) = L rhs ;; L r"
-unfolding Append_rexp_rhs_def
-by (auto simp add: Seq_def lang_of_append_rexp)
-
-
-
-subsubsection {* Intialization *}
-
-lemma defined_by_str:
-  assumes "s \<in> X" "X \<in> UNIV // \<approx>A" 
-  shows "X = \<approx>A `` {s}"
-using assms
-unfolding quotient_def Image_def str_eq_rel_def
-by auto
-
-lemma every_eqclass_has_transition:
-  assumes has_str: "s @ [c] \<in> X"
-  and     in_CS:   "X \<in> UNIV // \<approx>A"
-  obtains Y where "Y \<in> UNIV // \<approx>A" and "Y ;; {[c]} \<subseteq> X" and "s \<in> Y"
-proof -
-  def Y \<equiv> "\<approx>A `` {s}"
-  have "Y \<in> UNIV // \<approx>A" 
-    unfolding Y_def quotient_def by auto
-  moreover
-  have "X = \<approx>A `` {s @ [c]}" 
-    using has_str in_CS defined_by_str by blast
-  then have "Y ;; {[c]} \<subseteq> X" 
-    unfolding Y_def Image_def Seq_def
-    unfolding str_eq_rel_def
-    by clarsimp
-  moreover
-  have "s \<in> Y" unfolding Y_def 
-    unfolding Image_def str_eq_rel_def by simp
-  ultimately show thesis using that by blast
-qed
-
-lemma l_eq_r_in_eqs:
-  assumes X_in_eqs: "(X, rhs) \<in> Init (UNIV // \<approx>A)"
-  shows "X = L rhs"
-proof 
-  show "X \<subseteq> L rhs"
-  proof
-    fix x
-    assume in_X: "x \<in> X"
-    { assume empty: "x = []"
-      then have "x \<in> L rhs" using X_in_eqs in_X
-	unfolding Init_def Init_rhs_def
-        by auto
-    }
-    moreover
-    { assume not_empty: "x \<noteq> []"
-      then obtain s c where decom: "x = s @ [c]"
-	using rev_cases by blast
-      have "X \<in> UNIV // \<approx>A" using X_in_eqs unfolding Init_def by auto
-      then obtain Y where "Y \<in> UNIV // \<approx>A" "Y ;; {[c]} \<subseteq> X" "s \<in> Y"
-        using decom in_X every_eqclass_has_transition by blast
-      then have "x \<in> L {Trn Y (CHAR c)| Y c. Y \<in> UNIV // \<approx>A \<and> Y \<Turnstile>c\<Rightarrow> X}"
-        unfolding transition_def
-	using decom by (force simp add: Seq_def)
-      then have "x \<in> L rhs" using X_in_eqs in_X
-	unfolding Init_def Init_rhs_def by simp
-    }
-    ultimately show "x \<in> L rhs" by blast
-  qed
-next
-  show "L rhs \<subseteq> X" using X_in_eqs
-    unfolding Init_def Init_rhs_def transition_def
-    by auto 
-qed
-
-lemma test:
-  assumes X_in_eqs: "(X, rhs) \<in> Init (UNIV // \<approx>A)"
-  shows "X = \<Union> (L `  rhs)"
-using assms l_eq_r_in_eqs by (simp)
-
-lemma finite_Init_rhs: 
-  assumes finite: "finite CS"
-  shows "finite (Init_rhs CS X)"
-proof-
-  def S \<equiv> "{(Y, c)| Y c. Y \<in> CS \<and> Y ;; {[c]} \<subseteq> X}" 
-  def h \<equiv> "\<lambda> (Y, c). Trn Y (CHAR c)"
-  have "finite (CS \<times> (UNIV::char set))" using finite by auto
-  then have "finite S" using S_def 
-    by (rule_tac B = "CS \<times> UNIV" in finite_subset) (auto)
-  moreover have "{Trn Y (CHAR c) |Y c. Y \<in> CS \<and> Y ;; {[c]} \<subseteq> X} = h ` S"
-    unfolding S_def h_def image_def by auto
-  ultimately
-  have "finite {Trn Y (CHAR c) |Y c. Y \<in> CS \<and> Y ;; {[c]} \<subseteq> X}" by auto
-  then show "finite (Init_rhs CS X)" unfolding Init_rhs_def transition_def by simp
-qed
-
-lemma Init_ES_satisfies_invariant:
-  assumes finite_CS: "finite (UNIV // \<approx>A)"
-  shows "invariant (Init (UNIV // \<approx>A))"
-proof (rule invariantI)
-  show "soundness (Init (UNIV // \<approx>A))"
-    unfolding soundness_def 
-    using l_eq_r_in_eqs by auto
-  show "finite (Init (UNIV // \<approx>A))" using finite_CS
-    unfolding Init_def by simp
-  show "distinctness (Init (UNIV // \<approx>A))"     
-    unfolding distinctness_def Init_def by simp
-  show "ardenable_all (Init (UNIV // \<approx>A))"
-    unfolding ardenable_all_def Init_def Init_rhs_def ardenable_def
-   by auto 
-  show "finite_rhs (Init (UNIV // \<approx>A))"
-    using finite_Init_rhs[OF finite_CS]
-    unfolding finite_rhs_def Init_def by auto
-  show "validity (Init (UNIV // \<approx>A))"
-    unfolding validity_def Init_def Init_rhs_def rhss_def lhss_def
-    by auto
-qed
-
-subsubsection {* Interation step *}
-
-lemma Arden_keeps_eq:
-  assumes l_eq_r: "X = L rhs"
-  and not_empty: "ardenable rhs"
-  and finite: "finite rhs"
-  shows "X = L (Arden X rhs)"
-proof -
-  def A \<equiv> "L (\<Uplus>{r. Trn X r \<in> rhs})"
-  def b \<equiv> "{Trn X r | r. Trn X r \<in> rhs}"
-  def B \<equiv> "L (rhs - b)"
-  have not_empty2: "[] \<notin> A" 
-    using finite_Trn[OF finite] not_empty
-    unfolding A_def ardenable_def by simp
-  have "X = L rhs" using l_eq_r by simp
-  also have "\<dots> = L (b \<union> (rhs - b))" unfolding b_def by auto
-  also have "\<dots> = L b \<union> B" unfolding B_def by (simp only: L_rhs_union_distrib)
-  also have "\<dots> = X ;; A \<union> B"
-    unfolding b_def
-    unfolding rhs_trm_soundness[OF finite]
-    unfolding A_def
-    by blast
-  finally have "X = X ;; A \<union> B" . 
-  then have "X = B ;; A\<star>"
-    by (simp add: arden[OF not_empty2])
-  also have "\<dots> = L (Arden X rhs)"
-    unfolding Arden_def A_def B_def b_def
-    by (simp only: lang_of_append_rexp_rhs L_rexp.simps)
-  finally show "X = L (Arden X rhs)" by simp
-qed 
-
-lemma Append_keeps_finite:
-  "finite rhs \<Longrightarrow> finite (Append_rexp_rhs rhs r)"
-by (auto simp:Append_rexp_rhs_def)
-
-lemma Arden_keeps_finite:
-  "finite rhs \<Longrightarrow> finite (Arden X rhs)"
-by (auto simp:Arden_def Append_keeps_finite)
-
-lemma Append_keeps_nonempty:
-  "ardenable rhs \<Longrightarrow> ardenable (Append_rexp_rhs rhs r)"
-apply (auto simp:ardenable_def Append_rexp_rhs_def)
-by (case_tac x, auto simp:Seq_def)
-
-lemma nonempty_set_sub:
-  "ardenable rhs \<Longrightarrow> ardenable (rhs - A)"
-by (auto simp:ardenable_def)
-
-lemma nonempty_set_union:
-  "\<lbrakk>ardenable rhs; ardenable rhs'\<rbrakk> \<Longrightarrow> ardenable (rhs \<union> rhs')"
-by (auto simp:ardenable_def)
-
-lemma Arden_keeps_nonempty:
-  "ardenable rhs \<Longrightarrow> ardenable (Arden X rhs)"
-by (simp only:Arden_def Append_keeps_nonempty nonempty_set_sub)
-
-
-lemma Subst_keeps_nonempty:
-  "\<lbrakk>ardenable rhs; ardenable xrhs\<rbrakk> \<Longrightarrow> ardenable (Subst rhs X xrhs)"
-by (simp only: Subst_def Append_keeps_nonempty nonempty_set_union nonempty_set_sub)
-
-lemma Subst_keeps_eq:
-  assumes substor: "X = L xrhs"
-  and finite: "finite rhs"
-  shows "L (Subst rhs X xrhs) = L rhs" (is "?Left = ?Right")
-proof-
-  def A \<equiv> "L (rhs - {Trn X r | r. Trn X r \<in> rhs})"
-  have "?Left = A \<union> L (Append_rexp_rhs xrhs (\<Uplus>{r. Trn X r \<in> rhs}))"
-    unfolding Subst_def
-    unfolding L_rhs_union_distrib[symmetric]
-    by (simp add: A_def)
-  moreover have "?Right = A \<union> L ({Trn X r | r. Trn X r \<in> rhs})"
-  proof-
-    have "rhs = (rhs - {Trn X r | r. Trn X r \<in> rhs}) \<union> ({Trn X r | r. Trn X r \<in> rhs})" by auto
-    thus ?thesis 
-      unfolding A_def
-      unfolding L_rhs_union_distrib
-      by simp
-  qed
-  moreover have "L (Append_rexp_rhs xrhs (\<Uplus>{r. Trn X r \<in> rhs})) = L ({Trn X r | r. Trn X r \<in> rhs})" 
-    using finite substor by (simp only: lang_of_append_rexp_rhs rhs_trm_soundness)
-  ultimately show ?thesis by simp
-qed
-
-lemma Subst_keeps_finite_rhs:
-  "\<lbrakk>finite rhs; finite yrhs\<rbrakk> \<Longrightarrow> finite (Subst rhs Y yrhs)"
-by (auto simp: Subst_def Append_keeps_finite)
-
-lemma Subst_all_keeps_finite:
-  assumes finite: "finite ES"
-  shows "finite (Subst_all ES Y yrhs)"
-proof -
-  def eqns \<equiv> "{(X::lang, rhs) |X rhs. (X, rhs) \<in> ES}"
-  def h \<equiv> "\<lambda>(X::lang, rhs). (X, Subst rhs Y yrhs)"
-  have "finite (h ` eqns)" using finite h_def eqns_def by auto
-  moreover 
-  have "Subst_all ES Y yrhs = h ` eqns" unfolding h_def eqns_def Subst_all_def by auto
-  ultimately
-  show "finite (Subst_all ES Y yrhs)" by simp
-qed
-
-lemma Subst_all_keeps_finite_rhs:
-  "\<lbrakk>finite_rhs ES; finite yrhs\<rbrakk> \<Longrightarrow> finite_rhs (Subst_all ES Y yrhs)"
-by (auto intro:Subst_keeps_finite_rhs simp add:Subst_all_def finite_rhs_def)
-
-lemma append_rhs_keeps_cls:
-  "rhss (Append_rexp_rhs rhs r) = rhss rhs"
-apply (auto simp:rhss_def Append_rexp_rhs_def)
-apply (case_tac xa, auto simp:image_def)
-by (rule_tac x = "SEQ ra r" in exI, rule_tac x = "Trn x ra" in bexI, simp+)
-
-lemma Arden_removes_cl:
-  "rhss (Arden Y yrhs) = rhss yrhs - {Y}"
-apply (simp add:Arden_def append_rhs_keeps_cls)
-by (auto simp:rhss_def)
-
-lemma lhss_keeps_cls:
-  "lhss (Subst_all ES Y yrhs) = lhss ES"
-by (auto simp:lhss_def Subst_all_def)
-
-lemma Subst_updates_cls:
-  "X \<notin> rhss xrhs \<Longrightarrow> 
-      rhss (Subst rhs X xrhs) = rhss rhs \<union> rhss xrhs - {X}"
-apply (simp only:Subst_def append_rhs_keeps_cls rhss_union_distrib)
-by (auto simp:rhss_def)
-
-lemma Subst_all_keeps_validity:
-  assumes sc: "validity (ES \<union> {(Y, yrhs)})"        (is "validity ?A")
-  shows "validity (Subst_all ES Y (Arden Y yrhs))"  (is "validity ?B")
-proof -
-  { fix X xrhs'
-    assume "(X, xrhs') \<in> ?B"
-    then obtain xrhs 
-      where xrhs_xrhs': "xrhs' = Subst xrhs Y (Arden Y yrhs)"
-      and X_in: "(X, xrhs) \<in> ES" by (simp add:Subst_all_def, blast)    
-    have "rhss xrhs' \<subseteq> lhss ?B"
-    proof-
-      have "lhss ?B = lhss ES" by (auto simp add:lhss_def Subst_all_def)
-      moreover have "rhss xrhs' \<subseteq> lhss ES"
-      proof-
-        have "rhss xrhs' \<subseteq>  rhss xrhs \<union> rhss (Arden Y yrhs) - {Y}"
-        proof-
-          have "Y \<notin> rhss (Arden Y yrhs)" 
-            using Arden_removes_cl by simp
-          thus ?thesis using xrhs_xrhs' by (auto simp:Subst_updates_cls)
-        qed
-        moreover have "rhss xrhs \<subseteq> lhss ES \<union> {Y}" using X_in sc
-          apply (simp only:validity_def lhss_union_distrib)
-          by (drule_tac x = "(X, xrhs)" in bspec, auto simp:lhss_def)
-        moreover have "rhss (Arden Y yrhs) \<subseteq> lhss ES \<union> {Y}" 
-          using sc 
-          by (auto simp add:Arden_removes_cl validity_def lhss_def)
-        ultimately show ?thesis by auto
-      qed
-      ultimately show ?thesis by simp
-    qed
-  } thus ?thesis by (auto simp only:Subst_all_def validity_def)
-qed
-
-lemma Subst_all_satisfies_invariant:
-  assumes invariant_ES: "invariant (ES \<union> {(Y, yrhs)})"
-  shows "invariant (Subst_all ES Y (Arden Y yrhs))"
-proof (rule invariantI)
-  have Y_eq_yrhs: "Y = L yrhs" 
-    using invariant_ES by (simp only:invariant_def soundness_def, blast)
-   have finite_yrhs: "finite yrhs" 
-    using invariant_ES by (auto simp:invariant_def finite_rhs_def)
-  have nonempty_yrhs: "ardenable yrhs" 
-    using invariant_ES by (auto simp:invariant_def ardenable_all_def)
-  show "soundness (Subst_all ES Y (Arden Y yrhs))"
-  proof -
-    have "Y = L (Arden Y yrhs)" 
-      using Y_eq_yrhs invariant_ES finite_yrhs
-      using finite_Trn[OF finite_yrhs]
-      apply(rule_tac Arden_keeps_eq)
-      apply(simp_all)
-      unfolding invariant_def ardenable_all_def ardenable_def
-      apply(auto)
-      done
-    thus ?thesis using invariant_ES
-      unfolding invariant_def finite_rhs_def2 soundness_def Subst_all_def
-      by (auto simp add: Subst_keeps_eq simp del: L_rhs.simps)
-  qed
-  show "finite (Subst_all ES Y (Arden Y yrhs))" 
-    using invariant_ES by (simp add:invariant_def Subst_all_keeps_finite)
-  show "distinctness (Subst_all ES Y (Arden Y yrhs))" 
-    using invariant_ES 
-    unfolding distinctness_def Subst_all_def invariant_def by auto
-  show "ardenable_all (Subst_all ES Y (Arden Y yrhs))"
-  proof - 
-    { fix X rhs
-      assume "(X, rhs) \<in> ES"
-      hence "ardenable rhs"  using invariant_ES  
-        by (auto simp add:invariant_def ardenable_all_def)
-      with nonempty_yrhs 
-      have "ardenable (Subst rhs Y (Arden Y yrhs))"
-        by (simp add:nonempty_yrhs 
-               Subst_keeps_nonempty Arden_keeps_nonempty)
-    } thus ?thesis by (auto simp add:ardenable_all_def Subst_all_def)
-  qed
-  show "finite_rhs (Subst_all ES Y (Arden Y yrhs))"
-  proof-
-    have "finite_rhs ES" using invariant_ES 
-      by (simp add:invariant_def finite_rhs_def)
-    moreover have "finite (Arden Y yrhs)"
-    proof -
-      have "finite yrhs" using invariant_ES 
-        by (auto simp:invariant_def finite_rhs_def)
-      thus ?thesis using Arden_keeps_finite by simp
-    qed
-    ultimately show ?thesis 
-      by (simp add:Subst_all_keeps_finite_rhs)
-  qed
-  show "validity (Subst_all ES Y (Arden Y yrhs))"
-    using invariant_ES Subst_all_keeps_validity by (simp add:invariant_def)
-qed
-
-lemma Remove_in_card_measure:
-  assumes finite: "finite ES"
-  and     in_ES: "(X, rhs) \<in> ES"
-  shows "(Remove ES X rhs, ES) \<in> measure card"
-proof -
-  def f \<equiv> "\<lambda> x. ((fst x)::lang, Subst (snd x) X (Arden X rhs))"
-  def ES' \<equiv> "ES - {(X, rhs)}"
-  have "Subst_all ES' X (Arden X rhs) = f ` ES'" 
-    apply (auto simp: Subst_all_def f_def image_def)
-    by (rule_tac x = "(Y, yrhs)" in bexI, simp+)
-  then have "card (Subst_all ES' X (Arden X rhs)) \<le> card ES'"
-    unfolding ES'_def using finite by (auto intro: card_image_le)
-  also have "\<dots> < card ES" unfolding ES'_def 
-    using in_ES finite by (rule_tac card_Diff1_less)
-  finally show "(Remove ES X rhs, ES) \<in> measure card" 
-    unfolding Remove_def ES'_def by simp
-qed
-    
-
-lemma Subst_all_cls_remains: 
-  "(X, xrhs) \<in> ES \<Longrightarrow> \<exists> xrhs'. (X, xrhs') \<in> (Subst_all ES Y yrhs)"
-by (auto simp: Subst_all_def)
-
-lemma card_noteq_1_has_more:
-  assumes card:"Cond ES"
-  and e_in: "(X, xrhs) \<in> ES"
-  and finite: "finite ES"
-  shows "\<exists>(Y, yrhs) \<in> ES. (X, xrhs) \<noteq> (Y, yrhs)"
-proof-
-  have "card ES > 1" using card e_in finite 
-    by (cases "card ES") (auto) 
-  then have "card (ES - {(X, xrhs)}) > 0"
-    using finite e_in by auto
-  then have "(ES - {(X, xrhs)}) \<noteq> {}" using finite by (rule_tac notI, simp)
-  then show "\<exists>(Y, yrhs) \<in> ES. (X, xrhs) \<noteq> (Y, yrhs)"
-    by auto
-qed
-
-lemma iteration_step_measure:
-  assumes Inv_ES: "invariant ES"
-  and    X_in_ES: "(X, xrhs) \<in> ES"
-  and    Cnd:     "Cond ES "
-  shows "(Iter X ES, ES) \<in> measure card"
-proof -
-  have fin: "finite ES" using Inv_ES unfolding invariant_def by simp
-  then obtain Y yrhs 
-    where Y_in_ES: "(Y, yrhs) \<in> ES" and not_eq: "(X, xrhs) \<noteq> (Y, yrhs)" 
-    using Cnd X_in_ES by (drule_tac card_noteq_1_has_more) (auto)
-  then have "(Y, yrhs) \<in> ES " "X \<noteq> Y"  
-    using X_in_ES Inv_ES unfolding invariant_def distinctness_def
-    by auto
-  then show "(Iter X ES, ES) \<in> measure card" 
-  apply(rule IterI2)
-  apply(rule Remove_in_card_measure)
-  apply(simp_all add: fin)
-  done
-qed
-
-lemma iteration_step_invariant:
-  assumes Inv_ES: "invariant ES"
-  and    X_in_ES: "(X, xrhs) \<in> ES"
-  and    Cnd: "Cond ES"
-  shows "invariant (Iter X ES)"
-proof -
-  have finite_ES: "finite ES" using Inv_ES by (simp add: invariant_def)
-  then obtain Y yrhs 
-    where Y_in_ES: "(Y, yrhs) \<in> ES" and not_eq: "(X, xrhs) \<noteq> (Y, yrhs)" 
-    using Cnd X_in_ES by (drule_tac card_noteq_1_has_more) (auto)
-  then have "(Y, yrhs) \<in> ES" "X \<noteq> Y" 
-    using X_in_ES Inv_ES unfolding invariant_def distinctness_def
-    by auto
-  then show "invariant (Iter X ES)" 
-  proof(rule IterI2)
-    fix Y yrhs
-    assume h: "(Y, yrhs) \<in> ES" "X \<noteq> Y"
-    then have "ES - {(Y, yrhs)} \<union> {(Y, yrhs)} = ES" by auto
-    then show "invariant (Remove ES Y yrhs)" unfolding Remove_def
-      using Inv_ES
-      by (rule_tac Subst_all_satisfies_invariant) (simp) 
-  qed
-qed
-
-lemma iteration_step_ex:
-  assumes Inv_ES: "invariant ES"
-  and    X_in_ES: "(X, xrhs) \<in> ES"
-  and    Cnd: "Cond ES"
-  shows "\<exists>xrhs'. (X, xrhs') \<in> (Iter X ES)"
-proof -
-  have finite_ES: "finite ES" using Inv_ES by (simp add: invariant_def)
-  then obtain Y yrhs 
-    where "(Y, yrhs) \<in> ES" "(X, xrhs) \<noteq> (Y, yrhs)" 
-    using Cnd X_in_ES by (drule_tac card_noteq_1_has_more) (auto)
-  then have "(Y, yrhs) \<in> ES " "X \<noteq> Y"  
-    using X_in_ES Inv_ES unfolding invariant_def distinctness_def
-    by auto
-  then show "\<exists>xrhs'. (X, xrhs') \<in> (Iter X ES)" 
-  apply(rule IterI2)
-  unfolding Remove_def
-  apply(rule Subst_all_cls_remains)
-  using X_in_ES
-  apply(auto)
-  done
-qed
-
-
-subsubsection {* Conclusion of the proof *}
-
-lemma Solve:
-  assumes fin: "finite (UNIV // \<approx>A)"
-  and     X_in: "X \<in> (UNIV // \<approx>A)"
-  shows "\<exists>rhs. Solve X (Init (UNIV // \<approx>A)) = {(X, rhs)} \<and> invariant {(X, rhs)}"
-proof -
-  def Inv \<equiv> "\<lambda>ES. invariant ES \<and> (\<exists>rhs. (X, rhs) \<in> ES)"
-  have "Inv (Init (UNIV // \<approx>A))" unfolding Inv_def
-      using fin X_in by (simp add: Init_ES_satisfies_invariant, simp add: Init_def)
-  moreover
-  { fix ES
-    assume inv: "Inv ES" and crd: "Cond ES"
-    then have "Inv (Iter X ES)"
-      unfolding Inv_def
-      by (auto simp add: iteration_step_invariant iteration_step_ex) }
-  moreover
-  { fix ES
-    assume inv: "Inv ES" and not_crd: "\<not>Cond ES"
-    from inv obtain rhs where "(X, rhs) \<in> ES" unfolding Inv_def by auto
-    moreover
-    from not_crd have "card ES = 1" by simp
-    ultimately 
-    have "ES = {(X, rhs)}" by (auto simp add: card_Suc_eq) 
-    then have "\<exists>rhs'. ES = {(X, rhs')} \<and> invariant {(X, rhs')}" using inv
-      unfolding Inv_def by auto }
-  moreover
-    have "wf (measure card)" by simp
-  moreover
-  { fix ES
-    assume inv: "Inv ES" and crd: "Cond ES"
-    then have "(Iter X ES, ES) \<in> measure card"
-      unfolding Inv_def
-      apply(clarify)
-      apply(rule_tac iteration_step_measure)
-      apply(auto)
-      done }
-  ultimately 
-  show "\<exists>rhs. Solve X (Init (UNIV // \<approx>A)) = {(X, rhs)} \<and> invariant {(X, rhs)}" 
-    unfolding Solve_def by (rule while_rule)
-qed
-
-lemma every_eqcl_has_reg:
-  assumes finite_CS: "finite (UNIV // \<approx>A)"
-  and X_in_CS: "X \<in> (UNIV // \<approx>A)"
-  shows "\<exists>r::rexp. X = L r" 
-proof -
-  from finite_CS X_in_CS 
-  obtain xrhs where Inv_ES: "invariant {(X, xrhs)}"
-    using Solve by metis
-
-  def A \<equiv> "Arden X xrhs"
-  have "rhss xrhs \<subseteq> {X}" using Inv_ES 
-    unfolding validity_def invariant_def rhss_def lhss_def
-    by auto
-  then have "rhss A = {}" unfolding A_def 
-    by (simp add: Arden_removes_cl)
-  then have eq: "{Lam r | r. Lam r \<in> A} = A" unfolding rhss_def
-    by (auto, case_tac x, auto)
-  
-  have "finite A" using Inv_ES unfolding A_def invariant_def finite_rhs_def
-    using Arden_keeps_finite by auto
-  then have fin: "finite {r. Lam r \<in> A}" by (rule finite_Lam)
-
-  have "X = L xrhs" using Inv_ES unfolding invariant_def soundness_def
-    by simp
-  then have "X = L A" using Inv_ES 
-    unfolding A_def invariant_def ardenable_all_def finite_rhs_def 
-    by (rule_tac Arden_keeps_eq) (simp_all add: finite_Trn)
-  then have "X = L {Lam r | r. Lam r \<in> A}" using eq by simp
-  then have "X = L (\<Uplus>{r. Lam r \<in> A})" using fin by auto
-  then show "\<exists>r::rexp. X = L r" by blast
-qed
-
-lemma bchoice_finite_set:
-  assumes a: "\<forall>x \<in> S. \<exists>y. x = f y" 
-  and     b: "finite S"
-  shows "\<exists>ys. (\<Union> S) = \<Union>(f ` ys) \<and> finite ys"
-using bchoice[OF a] b
-apply(erule_tac exE)
-apply(rule_tac x="fa ` S" in exI)
-apply(auto)
-done
-
-theorem Myhill_Nerode1:
-  assumes finite_CS: "finite (UNIV // \<approx>A)"
-  shows   "\<exists>r::rexp. A = L r"
-proof -
-  have fin: "finite (finals A)" 
-    using finals_in_partitions finite_CS by (rule finite_subset)
-  have "\<forall>X \<in> (UNIV // \<approx>A). \<exists>r::rexp. X = L r" 
-    using finite_CS every_eqcl_has_reg by blast
-  then have a: "\<forall>X \<in> finals A. \<exists>r::rexp. X = L r"
-    using finals_in_partitions by auto
-  then obtain rs::"rexp set" where "\<Union> (finals A) = \<Union>(L ` rs)" "finite rs"
-    using fin by (auto dest: bchoice_finite_set)
-  then have "A = L (\<Uplus>rs)" 
-    unfolding lang_is_union_of_finals[symmetric] by simp
-  then show "\<exists>r::rexp. A = L r" by blast
-qed 
-
-
-end
\ No newline at end of file
--- a/Theories/Myhill_2.thy	Tue May 31 20:32:49 2011 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,470 +0,0 @@
-theory Myhill_2
-  imports Myhill_1 Prefix_subtract
-          "~~/src/HOL/Library/List_Prefix"
-begin
-
-section {* Direction @{text "regular language \<Rightarrow>finite partition"} *}
-
-definition
-  str_eq :: "string \<Rightarrow> lang \<Rightarrow> string \<Rightarrow> bool" ("_ \<approx>_ _")
-where
-  "x \<approx>A y \<equiv> (x, y) \<in> (\<approx>A)"
-
-definition 
-   tag_eq_rel :: "(string \<Rightarrow> 'b) \<Rightarrow> (string \<times> string) set" ("=_=")
-where
-   "=tag= \<equiv> {(x, y) | x y. tag x = tag y}"
-
-lemma finite_eq_tag_rel:
-  assumes rng_fnt: "finite (range tag)"
-  shows "finite (UNIV // =tag=)"
-proof -
-  let "?f" =  "\<lambda>X. tag ` X" and ?A = "(UNIV // =tag=)"
-  have "finite (?f ` ?A)" 
-  proof -
-    have "range ?f \<subseteq> (Pow (range tag))" unfolding Pow_def by auto
-    moreover 
-    have "finite (Pow (range tag))" using rng_fnt by simp
-    ultimately 
-    have "finite (range ?f)" unfolding image_def by (blast intro: finite_subset)
-    moreover
-    have "?f ` ?A \<subseteq> range ?f" by auto
-    ultimately show "finite (?f ` ?A)" by (rule rev_finite_subset) 
-  qed
-  moreover
-  have "inj_on ?f ?A"
-  proof -
-    { fix X Y
-      assume X_in: "X \<in> ?A"
-        and  Y_in: "Y \<in> ?A"
-        and  tag_eq: "?f X = ?f Y"
-      then obtain x y 
-        where "x \<in> X" "y \<in> Y" "tag x = tag y"
-        unfolding quotient_def Image_def image_def tag_eq_rel_def
-        by (simp) (blast)
-      with X_in Y_in 
-      have "X = Y"
-	unfolding quotient_def tag_eq_rel_def by auto
-    } 
-    then show "inj_on ?f ?A" unfolding inj_on_def by auto
-  qed
-  ultimately show "finite (UNIV // =tag=)" by (rule finite_imageD)
-qed
-
-lemma refined_partition_finite:
-  assumes fnt: "finite (UNIV // R1)"
-  and refined: "R1 \<subseteq> R2"
-  and eq1: "equiv UNIV R1" and eq2: "equiv UNIV R2"
-  shows "finite (UNIV // R2)"
-proof -
-  let ?f = "\<lambda>X. {R1 `` {x} | x. x \<in> X}" 
-    and ?A = "UNIV // R2" and ?B = "UNIV // R1"
-  have "?f ` ?A \<subseteq> Pow ?B"
-    unfolding image_def Pow_def quotient_def by auto
-  moreover
-  have "finite (Pow ?B)" using fnt by simp
-  ultimately  
-  have "finite (?f ` ?A)" by (rule finite_subset)
-  moreover
-  have "inj_on ?f ?A"
-  proof -
-    { fix X Y
-      assume X_in: "X \<in> ?A" and Y_in: "Y \<in> ?A" and eq_f: "?f X = ?f Y"
-      from quotientE [OF X_in]
-      obtain x where "X = R2 `` {x}" by blast
-      with equiv_class_self[OF eq2] have x_in: "x \<in> X" by simp
-      then have "R1 ``{x} \<in> ?f X" by auto
-      with eq_f have "R1 `` {x} \<in> ?f Y" by simp
-      then obtain y 
-        where y_in: "y \<in> Y" and eq_r1_xy: "R1 `` {x} = R1 `` {y}" by auto
-      with eq_equiv_class[OF _ eq1] 
-      have "(x, y) \<in> R1" by blast
-      with refined have "(x, y) \<in> R2" by auto
-      with quotient_eqI [OF eq2 X_in Y_in x_in y_in]
-      have "X = Y" .
-    } 
-    then show "inj_on ?f ?A" unfolding inj_on_def by blast 
-  qed
-  ultimately show "finite (UNIV // R2)" by (rule finite_imageD)
-qed
-
-lemma tag_finite_imageD:
-  assumes rng_fnt: "finite (range tag)" 
-  and same_tag_eqvt: "\<And>m n. tag m = tag n \<Longrightarrow> m \<approx>A n"
-  shows "finite (UNIV // \<approx>A)"
-proof (rule_tac refined_partition_finite [of "=tag="])
-  show "finite (UNIV // =tag=)" by (rule finite_eq_tag_rel[OF rng_fnt])
-next
-  from same_tag_eqvt
-  show "=tag= \<subseteq> \<approx>A" unfolding tag_eq_rel_def str_eq_def
-    by auto
-next
-  show "equiv UNIV =tag="
-    unfolding equiv_def tag_eq_rel_def refl_on_def sym_def trans_def
-    by auto
-next
-  show "equiv UNIV (\<approx>A)" 
-    unfolding equiv_def str_eq_rel_def sym_def refl_on_def trans_def
-    by blast
-qed
-
-
-subsection {* The proof *}
-
-subsubsection {* The base case for @{const "NULL"} *}
-
-lemma quot_null_eq:
-  shows "UNIV // \<approx>{} = {UNIV}"
-unfolding quotient_def Image_def str_eq_rel_def by auto
-
-lemma quot_null_finiteI [intro]:
-  shows "finite (UNIV // \<approx>{})"
-unfolding quot_null_eq by simp
-
-
-subsubsection {* The base case for @{const "EMPTY"} *}
-
-lemma quot_empty_subset:
-  shows "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_empty_finiteI [intro]:
-  shows "finite (UNIV // \<approx>{[]})"
-by (rule finite_subset[OF quot_empty_subset]) (simp)
-
-
-subsubsection {* The base case for @{const "CHAR"} *}
-
-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
-
-lemma quot_char_finiteI [intro]:
-  shows "finite (UNIV // \<approx>{[c]})"
-by (rule finite_subset[OF quot_char_subset]) (simp)
-
-
-subsubsection {* The inductive case for @{const ALT} *}
-
-definition 
-  tag_str_ALT :: "lang \<Rightarrow> lang \<Rightarrow> string \<Rightarrow> (lang \<times> lang)"
-where
-  "tag_str_ALT A B \<equiv> (\<lambda>x. (\<approx>A `` {x}, \<approx>B `` {x}))"
-
-lemma quot_union_finiteI [intro]:
-  fixes L1 L2::"lang"
-  assumes finite1: "finite (UNIV // \<approx>A)"
-  and     finite2: "finite (UNIV // \<approx>B)"
-  shows "finite (UNIV // \<approx>(A \<union> B))"
-proof (rule_tac tag = "tag_str_ALT A B" in tag_finite_imageD)
-  have "finite ((UNIV // \<approx>A) \<times> (UNIV // \<approx>B))" 
-    using finite1 finite2 by auto
-  then show "finite (range (tag_str_ALT A B))"
-    unfolding tag_str_ALT_def quotient_def
-    by (rule rev_finite_subset) (auto)
-next
-  show "\<And>x y. tag_str_ALT A B x = tag_str_ALT A B y \<Longrightarrow> x \<approx>(A \<union> B) y"
-    unfolding tag_str_ALT_def
-    unfolding str_eq_def
-    unfolding str_eq_rel_def
-    by auto
-qed
-
-
-subsubsection {* The inductive case for @{text "SEQ"}*}
-
-definition 
-  tag_str_SEQ :: "lang \<Rightarrow> lang \<Rightarrow> string \<Rightarrow> (lang \<times> lang set)"
-where
-  "tag_str_SEQ L1 L2 \<equiv>
-     (\<lambda>x. (\<approx>L1 `` {x}, {(\<approx>L2 `` {x - xa}) | xa.  xa \<le> x \<and> xa \<in> L1}))"
-
-lemma Seq_in_cases:
-  assumes "x @ z \<in> A ;; B"
-  shows "(\<exists> x' \<le> x. x' \<in> A \<and> (x - x') @ z \<in> B) \<or> 
-         (\<exists> z' \<le> z. (x @ z') \<in> A \<and> (z - z') \<in> B)"
-using assms
-unfolding Seq_def prefix_def
-by (auto simp add: append_eq_append_conv2)
-
-lemma tag_str_SEQ_injI:
-  assumes eq_tag: "tag_str_SEQ A B x = tag_str_SEQ A B y" 
-  shows "x \<approx>(A ;; B) y"
-proof -
-  { fix x y z
-    assume xz_in_seq: "x @ z \<in> A ;; B"
-    and tag_xy: "tag_str_SEQ A B x = tag_str_SEQ A B y"
-    have"y @ z \<in> A ;; B" 
-    proof -
-      { (* first case with x' in A and (x - x') @ z in B *)
-        fix x'
-        assume h1: "x' \<le> x" and h2: "x' \<in> A" and h3: "(x - x') @ z \<in> B"
-        obtain y' 
-          where "y' \<le> y" 
-          and "y' \<in> A" 
-          and "(y - y') @ z \<in> B"
-        proof -
-          have "{\<approx>B `` {x - x'} |x'. x' \<le> x \<and> x' \<in> A} = 
-                {\<approx>B `` {y - y'} |y'. y' \<le> y \<and> y' \<in> A}" (is "?Left = ?Right")
-            using tag_xy unfolding tag_str_SEQ_def by simp
-          moreover 
-	  have "\<approx>B `` {x - x'} \<in> ?Left" using h1 h2 by auto
-          ultimately 
-	  have "\<approx>B `` {x - x'} \<in> ?Right" by simp
-          then obtain y' 
-            where eq_xy': "\<approx>B `` {x - x'} = \<approx>B `` {y - y'}" 
-            and pref_y': "y' \<le> y" and y'_in: "y' \<in> A"
-            by simp blast
-	  
-	  have "(x - x')  \<approx>B (y - y')" using eq_xy'
-            unfolding Image_def str_eq_rel_def str_eq_def by auto
-          with h3 have "(y - y') @ z \<in> B" 
-	    unfolding str_eq_rel_def str_eq_def by simp
-          with pref_y' y'_in 
-          show ?thesis using that by blast
-        qed
-	then have "y @ z \<in> A ;; B" by (erule_tac prefixE) (auto simp: Seq_def)
-      } 
-      moreover 
-      { (* second case with x @ z' in A and z - z' in B *)
-        fix z'
-        assume h1: "z' \<le> z" and h2: "(x @ z') \<in> A" and h3: "z - z' \<in> B"
-	 have "\<approx>A `` {x} = \<approx>A `` {y}" 
-           using tag_xy unfolding tag_str_SEQ_def by simp
-         with h2 have "y @ z' \<in> A"
-            unfolding Image_def str_eq_rel_def str_eq_def by auto
-        with h1 h3 have "y @ z \<in> A ;; B"
-	  unfolding prefix_def Seq_def
-	  by (auto) (metis append_assoc)
-      }
-      ultimately show "y @ z \<in> A ;; B" 
-	using Seq_in_cases [OF xz_in_seq] by blast
-    qed
-  }
-  from this [OF _ eq_tag] and this [OF _ eq_tag [THEN sym]]
-    show "x \<approx>(A ;; B) y" unfolding str_eq_def str_eq_rel_def by blast
-qed 
-
-lemma quot_seq_finiteI [intro]:
-  fixes L1 L2::"lang"
-  assumes fin1: "finite (UNIV // \<approx>L1)" 
-  and     fin2: "finite (UNIV // \<approx>L2)" 
-  shows "finite (UNIV // \<approx>(L1 ;; L2))"
-proof (rule_tac tag = "tag_str_SEQ L1 L2" in tag_finite_imageD)
-  show "\<And>x y. tag_str_SEQ L1 L2 x = tag_str_SEQ L1 L2 y \<Longrightarrow> x \<approx>(L1 ;; L2) y"
-    by (rule tag_str_SEQ_injI)
-next
-  have *: "finite ((UNIV // \<approx>L1) \<times> (Pow (UNIV // \<approx>L2)))" 
-    using fin1 fin2 by auto
-  show "finite (range (tag_str_SEQ L1 L2))" 
-    unfolding tag_str_SEQ_def
-    apply(rule finite_subset[OF _ *])
-    unfolding quotient_def
-    by auto
-qed
-
-
-subsubsection {* The inductive case for @{const "STAR"} *}
-
-definition 
-  tag_str_STAR :: "lang \<Rightarrow> string \<Rightarrow> lang set"
-where
-  "tag_str_STAR L1 \<equiv> (\<lambda>x. {\<approx>L1 `` {x - xa} | xa. xa < x \<and> xa \<in> L1\<star>})"
-
-text {* A technical lemma. *}
-lemma finite_set_has_max: "\<lbrakk>finite A; A \<noteq> {}\<rbrakk> \<Longrightarrow> 
-           (\<exists> max \<in> A. \<forall> a \<in> A. f a <= (f max :: nat))"
-proof (induct rule:finite.induct)
-  case emptyI thus ?case by simp
-next
-  case (insertI A a)
-  show ?case
-  proof (cases "A = {}")
-    case True thus ?thesis by (rule_tac x = a in bexI, auto)
-  next
-    case False
-    with insertI.hyps and False  
-    obtain max 
-      where h1: "max \<in> A" 
-      and h2: "\<forall>a\<in>A. f a \<le> f max" by blast
-    show ?thesis
-    proof (cases "f a \<le> f max")
-      assume "f a \<le> f max"
-      with h1 h2 show ?thesis by (rule_tac x = max in bexI, auto)
-    next
-      assume "\<not> (f a \<le> f max)"
-      thus ?thesis using h2 by (rule_tac x = a in bexI, auto)
-    qed
-  qed
-qed
-
-
-text {* The following is a technical lemma, which helps to show the range finiteness of tag function. *}
-
-lemma finite_strict_prefix_set: "finite {xa. xa < (x::string)}"
-apply (induct x rule:rev_induct, simp)
-apply (subgoal_tac "{xa. xa < xs @ [x]} = {xa. xa < xs} \<union> {xs}")
-by (auto simp:strict_prefix_def)
-
-
-lemma tag_str_STAR_injI:
-  assumes eq_tag: "tag_str_STAR L\<^isub>1 v = tag_str_STAR L\<^isub>1 w"
-  shows "v \<approx>(L\<^isub>1\<star>) w"
-proof-
-  { fix x y z
-    assume xz_in_star: "x @ z \<in> L\<^isub>1\<star>" 
-      and tag_xy: "tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 y"
-    have "y @ z \<in> L\<^isub>1\<star>"
-    proof(cases "x = []")
-      case True
-      with tag_xy have "y = []" 
-        by (auto simp add: tag_str_STAR_def strict_prefix_def)
-      thus ?thesis using xz_in_star True by simp
-    next
-      case False
-      let ?S = "{xa. xa < x \<and> xa \<in> L\<^isub>1\<star> \<and> (x - xa) @ z \<in> L\<^isub>1\<star>}"
-      have "finite ?S"
-        by (rule_tac B = "{xa. xa < x}" in finite_subset, 
-          auto simp:finite_strict_prefix_set)
-      moreover have "?S \<noteq> {}" using False xz_in_star
-        by (simp, rule_tac x = "[]" in exI, auto simp:strict_prefix_def)
-      ultimately have "\<exists> xa_max \<in> ?S. \<forall> xa \<in> ?S. length xa \<le> length xa_max" 
-        using finite_set_has_max by blast
-      then obtain xa_max 
-        where h1: "xa_max < x" 
-        and h2: "xa_max \<in> L\<^isub>1\<star>" 
-        and h3: "(x - xa_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 xa_max"
-        by blast
-      obtain ya 
-        where h5: "ya < y" and h6: "ya \<in> L\<^isub>1\<star>" 
-        and eq_xya: "(x - xa_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 - xa_max} \<in> ?left" using h1 h2 by auto
-        ultimately have "\<approx>L\<^isub>1 `` {x - xa_max} \<in> ?right" by simp
-        thus ?thesis using that 
-          apply (simp add:Image_def str_eq_rel_def str_eq_def) by blast
-      qed 
-      have "(y - ya) @ z \<in> L\<^isub>1\<star>" 
-      proof-
-        obtain za zb where eq_zab: "z = za @ zb" 
-          and l_za: "(y - ya)@za \<in> L\<^isub>1" and ls_zb: "zb \<in> L\<^isub>1\<star>"
-        proof -
-          from h1 have "(x - xa_max) @ z \<noteq> []" 
-            by (auto simp:strict_prefix_def elim:prefixE)
-          from star_decom [OF h3 this]
-          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 - xa_max) @ z = a @ b" by blast
-          let ?za = "a - (x - xa_max)" and ?zb = "b"
-          have pfx: "(x - xa_max) \<le> a" (is "?P1") 
-            and eq_z: "z = ?za @ ?zb" (is "?P2")
-          proof -
-            have "((x - xa_max) \<le> a \<and> (a - (x - xa_max)) @ b = z) \<or> 
-              (a < (x - xa_max) \<and> ((x - xa_max) - a) @ z = b)" 
-              using append_eq_dest[OF ab_max] by (auto simp:strict_prefix_def)
-            moreover {
-              assume np: "a < (x - xa_max)" 
-                and b_eqs: "((x - xa_max) - a) @ z = b"
-              have "False"
-              proof -
-                let ?xa_max' = "xa_max @ a"
-                have "?xa_max' < x" 
-                  using np h1 by (clarsimp simp:strict_prefix_def diff_prefix) 
-                moreover have "?xa_max' \<in> L\<^isub>1\<star>" 
-                  using a_in h2 by (simp add:star_intro3) 
-                moreover have "(x - ?xa_max') @ z \<in> L\<^isub>1\<star>" 
-                  using b_eqs b_in np h1 by (simp add:diff_diff_append)
-                moreover have "\<not> (length ?xa_max' \<le> length xa_max)" 
-                  using a_neq by simp
-                ultimately show ?thesis using h4 by blast
-              qed }
-            ultimately show ?P1 and ?P2 by auto
-          qed
-          hence "(x - xa_max)@?za \<in> L\<^isub>1" using a_in by (auto elim:prefixE)
-          with eq_xya have "(y - ya) @ ?za \<in> L\<^isub>1" 
-            by (auto simp:str_eq_def str_eq_rel_def)
-           with eq_z and b_in 
-          show ?thesis using that by blast
-        qed
-        have "((y - ya) @ za) @ zb \<in> L\<^isub>1\<star>" using  l_za ls_zb by blast
-        with eq_zab show ?thesis by simp
-      qed
-      with h5 h6 show ?thesis 
-        by (drule_tac star_intro1, auto simp:strict_prefix_def elim:prefixE)
-    qed
-  } 
-  from this [OF _ eq_tag] and this [OF _ eq_tag [THEN sym]]
-    show  ?thesis unfolding str_eq_def str_eq_rel_def by blast
-qed
-
-lemma quot_star_finiteI [intro]:
-  assumes finite1: "finite (UNIV // \<approx>L1)"
-  shows "finite (UNIV // \<approx>(L1\<star>))"
-proof (rule_tac tag = "tag_str_STAR L1" in tag_finite_imageD)
-  show "\<And>x y. tag_str_STAR L1 x = tag_str_STAR L1 y \<Longrightarrow> x \<approx>(L1\<star>) y"
-    by (rule tag_str_STAR_injI)
-next
-  have *: "finite (Pow (UNIV // \<approx>L1))" 
-    using finite1 by auto
-  show "finite (range (tag_str_STAR L1))"
-    unfolding tag_str_STAR_def
-    apply(rule finite_subset[OF _ *])
-    unfolding quotient_def
-    by auto
-qed
-
-subsubsection{* The conclusion *}
-
-lemma Myhill_Nerode2:
-  fixes r::"rexp"
-  shows "finite (UNIV // \<approx>(L r))"
-by (induct r) (auto)
-
-
-theorem Myhill_Nerode:
-  shows "(\<exists>r::rexp. A = L r) \<longleftrightarrow> finite (UNIV // \<approx>A)"
-using Myhill_Nerode1 Myhill_Nerode2 by auto
-
-end
--- a/Theories/Prefix_subtract.thy	Tue May 31 20:32:49 2011 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,60 +0,0 @@
-theory Prefix_subtract
-  imports Main "~~/src/HOL/Library/List_Prefix"
-begin
-
-
-section {* A small theory of prefix subtraction *}
-
-text {*
-  The notion of @{text "prefix_subtract"} makes 
-  the second direction of the Myhill-Nerode theorem 
-  more readable.
-*}
-
-instantiation list :: (type) minus
-begin
-
-fun minus_list :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" 
-where
-  "minus_list []     xs     = []" 
-| "minus_list (x#xs) []     = x#xs" 
-| "minus_list (x#xs) (y#ys) = (if x = y then minus_list xs ys else (x#xs))"
-
-instance by default
-
-end
-
-lemma [simp]: "x - [] = x"
-by (induct x) (auto)
-
-lemma [simp]: "(x @ y) - x = y"
-by (induct x) (auto)
-
-lemma [simp]: "x - x = []"
-by (induct x) (auto)
-
-lemma [simp]: "x = z @ y \<Longrightarrow> x - z = y "
-by (induct x) (auto)
-
-lemma diff_prefix:
-  "\<lbrakk>c \<le> a - b; b \<le> a\<rbrakk> \<Longrightarrow> b @ c \<le> a"
-by (auto elim: prefixE)
-
-lemma diff_diff_append: 
-  "\<lbrakk>c < a - b; b < a\<rbrakk> \<Longrightarrow> (a - b) - c = a - (b @ c)"
-apply (clarsimp simp:strict_prefix_def)
-by (drule diff_prefix, auto elim:prefixE)
-
-lemma append_eq_cases:
-  assumes a: "x @ y = m @ n"
-  shows "x \<le> m \<or> m \<le> x"
-unfolding prefix_def using a
-by (auto simp add: append_eq_append_conv2)
-
-lemma append_eq_dest:
-  assumes a: "x @ y = m @ n" 
-  shows "(x \<le> m \<and> (m - x) @ n = y) \<or> (m \<le> x \<and> (x - m) @ y = n)"
-using append_eq_cases[OF a] a
-by (auto elim: prefixE)
-
-end
--- a/Theories/Regular.thy	Tue May 31 20:32:49 2011 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,311 +0,0 @@
-theory Regular
-imports Main Folds
-begin
-
-section {* Preliminary definitions *}
-
-type_synonym lang = "string set"
-
-
-text {*  Sequential composition of two languages *}
-
-definition 
-  Seq :: "lang \<Rightarrow> lang \<Rightarrow> lang" (infixr ";;" 100)
-where 
-  "A ;; B = {s\<^isub>1 @ s\<^isub>2 | s\<^isub>1 s\<^isub>2. s\<^isub>1 \<in> A \<and> s\<^isub>2 \<in> B}"
-
-
-text {* Some properties of operator @{text ";;"}. *}
-
-lemma seq_add_left:
-  assumes a: "A = B"
-  shows "C ;; A = C ;; B"
-using a by simp
-
-lemma seq_union_distrib_right:
-  shows "(A \<union> B) ;; C = (A ;; C) \<union> (B ;; C)"
-unfolding Seq_def by auto
-
-lemma seq_union_distrib_left:
-  shows "C ;; (A \<union> B) = (C ;; A) \<union> (C ;; B)"
-unfolding Seq_def by  auto
-
-lemma seq_intro:
-  assumes a: "x \<in> A" "y \<in> B"
-  shows "x @ y \<in> A ;; B "
-using a by (auto simp: Seq_def)
-
-lemma seq_assoc:
-  shows "(A ;; B) ;; C = A ;; (B ;; C)"
-unfolding Seq_def
-apply(auto)
-apply(blast)
-by (metis append_assoc)
-
-lemma seq_empty [simp]:
-  shows "A ;; {[]} = A"
-  and   "{[]} ;; A = A"
-by (simp_all add: Seq_def)
-
-lemma seq_null [simp]:
-  shows "A ;; {} = {}"
-  and   "{} ;; A = {}"
-by (simp_all add: Seq_def)
-
-
-text {* Power and Star of a language *}
-
-fun 
-  pow :: "lang \<Rightarrow> nat \<Rightarrow> lang" (infixl "\<up>" 100)
-where
-  "A \<up> 0 = {[]}"
-| "A \<up> (Suc n) =  A ;; (A \<up> n)" 
-
-definition
-  Star :: "lang \<Rightarrow> lang" ("_\<star>" [101] 102)
-where
-  "A\<star> \<equiv> (\<Union>n. A \<up> n)"
-
-
-lemma star_start[intro]:
-  shows "[] \<in> A\<star>"
-proof -
-  have "[] \<in> A \<up> 0" by auto
-  then show "[] \<in> A\<star>" unfolding Star_def by blast
-qed
-
-lemma star_step [intro]:
-  assumes a: "s1 \<in> A" 
-  and     b: "s2 \<in> A\<star>"
-  shows "s1 @ s2 \<in> A\<star>"
-proof -
-  from b obtain n where "s2 \<in> A \<up> n" unfolding Star_def by auto
-  then have "s1 @ s2 \<in> A \<up> (Suc n)" using a by (auto simp add: Seq_def)
-  then show "s1 @ s2 \<in> A\<star>" unfolding Star_def by blast
-qed
-
-lemma star_induct[consumes 1, case_names start step]:
-  assumes a: "x \<in> A\<star>" 
-  and     b: "P []"
-  and     c: "\<And>s1 s2. \<lbrakk>s1 \<in> A; s2 \<in> A\<star>; P s2\<rbrakk> \<Longrightarrow> P (s1 @ s2)"
-  shows "P x"
-proof -
-  from a obtain n where "x \<in> A \<up> n" unfolding Star_def by auto
-  then show "P x"
-    by (induct n arbitrary: x)
-       (auto intro!: b c simp add: Seq_def Star_def)
-qed
-    
-lemma star_intro1:
-  assumes a: "x \<in> A\<star>"
-  and     b: "y \<in> A\<star>"
-  shows "x @ y \<in> A\<star>"
-using a b
-by (induct rule: star_induct) (auto)
-
-lemma star_intro2: 
-  assumes a: "y \<in> A"
-  shows "y \<in> A\<star>"
-proof -
-  from a have "y @ [] \<in> A\<star>" by blast
-  then show "y \<in> A\<star>" by simp
-qed
-
-lemma star_intro3:
-  assumes a: "x \<in> A\<star>"
-  and     b: "y \<in> A"
-  shows "x @ y \<in> A\<star>"
-using a b by (blast intro: star_intro1 star_intro2)
-
-lemma star_cases:
-  shows "A\<star> =  {[]} \<union> A ;; A\<star>"
-proof
-  { fix x
-    have "x \<in> A\<star> \<Longrightarrow> x \<in> {[]} \<union> A ;; A\<star>"
-      unfolding Seq_def
-    by (induct rule: star_induct) (auto)
-  }
-  then show "A\<star> \<subseteq> {[]} \<union> A ;; A\<star>" by auto
-next
-  show "{[]} \<union> A ;; A\<star> \<subseteq> A\<star>"
-    unfolding Seq_def by auto
-qed
-
-lemma star_decom: 
-  assumes a: "x \<in> A\<star>" "x \<noteq> []"
-  shows "\<exists>a b. x = a @ b \<and> a \<noteq> [] \<and> a \<in> A \<and> b \<in> A\<star>"
-using a
-by (induct rule: star_induct) (blast)+
-
-lemma
-  shows seq_Union_left:  "B ;; (\<Union>n. A \<up> n) = (\<Union>n. B ;; (A \<up> n))"
-  and   seq_Union_right: "(\<Union>n. A \<up> n) ;; B = (\<Union>n. (A \<up> n) ;; B)"
-unfolding Seq_def by auto
-
-lemma seq_pow_comm:
-  shows "A ;; (A \<up> n) = (A \<up> n) ;; A"
-by (induct n) (simp_all add: seq_assoc[symmetric])
-
-lemma seq_star_comm:
-  shows "A ;; A\<star> = A\<star> ;; A"
-unfolding Star_def seq_Union_left
-unfolding seq_pow_comm seq_Union_right 
-by simp
-
-
-text {* Two lemmas about the length of strings in @{text "A \<up> n"} *}
-
-lemma pow_length:
-  assumes a: "[] \<notin> A"
-  and     b: "s \<in> A \<up> Suc n"
-  shows "n < length s"
-using b
-proof (induct n arbitrary: s)
-  case 0
-  have "s \<in> A \<up> Suc 0" by fact
-  with a have "s \<noteq> []" by auto
-  then show "0 < length s" by auto
-next
-  case (Suc n)
-  have ih: "\<And>s. s \<in> A \<up> Suc n \<Longrightarrow> n < length s" by fact
-  have "s \<in> A \<up> Suc (Suc n)" by fact
-  then obtain s1 s2 where eq: "s = s1 @ s2" and *: "s1 \<in> A" and **: "s2 \<in> A \<up> Suc n"
-    by (auto simp add: Seq_def)
-  from ih ** have "n < length s2" by simp
-  moreover have "0 < length s1" using * a by auto
-  ultimately show "Suc n < length s" unfolding eq 
-    by (simp only: length_append)
-qed
-
-lemma seq_pow_length:
-  assumes a: "[] \<notin> A"
-  and     b: "s \<in> B ;; (A \<up> Suc n)"
-  shows "n < length s"
-proof -
-  from b obtain s1 s2 where eq: "s = s1 @ s2" and *: "s2 \<in> A \<up> Suc n"
-    unfolding Seq_def by auto
-  from * have " n < length s2" by (rule pow_length[OF a])
-  then show "n < length s" using eq by simp
-qed
-
-
-section {* A modified version of Arden's lemma *}
-
-text {*  A helper lemma for Arden *}
-
-lemma arden_helper:
-  assumes eq: "X = X ;; A \<union> B"
-  shows "X = X ;; (A \<up> Suc n) \<union> (\<Union>m\<in>{0..n}. B ;; (A \<up> m))"
-proof (induct n)
-  case 0 
-  show "X = X ;; (A \<up> Suc 0) \<union> (\<Union>(m::nat)\<in>{0..0}. B ;; (A \<up> m))"
-    using eq by simp
-next
-  case (Suc n)
-  have ih: "X = X ;; (A \<up> Suc n) \<union> (\<Union>m\<in>{0..n}. B ;; (A \<up> m))" by fact
-  also have "\<dots> = (X ;; A \<union> B) ;; (A \<up> Suc n) \<union> (\<Union>m\<in>{0..n}. B ;; (A \<up> m))" using eq by simp
-  also have "\<dots> = X ;; (A \<up> Suc (Suc n)) \<union> (B ;; (A \<up> Suc n)) \<union> (\<Union>m\<in>{0..n}. B ;; (A \<up> m))"
-    by (simp add: seq_union_distrib_right seq_assoc)
-  also have "\<dots> = X ;; (A \<up> Suc (Suc n)) \<union> (\<Union>m\<in>{0..Suc n}. B ;; (A \<up> m))"
-    by (auto simp add: le_Suc_eq)
-  finally show "X = X ;; (A \<up> Suc (Suc n)) \<union> (\<Union>m\<in>{0..Suc n}. B ;; (A \<up> m))" .
-qed
-
-theorem arden:
-  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" 
-    unfolding seq_star_comm[symmetric]
-    by (rule star_cases)
-  then have "B ;; A\<star> = B ;; ({[]} \<union> A\<star> ;; A)"
-    by (rule seq_add_left)
-  also have "\<dots> = B \<union> B ;; (A\<star> ;; A)"
-    unfolding seq_union_distrib_left by simp
-  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 eq: "X = X ;; A \<union> B"
-  { fix n::nat
-    have "B ;; (A \<up> n) \<subseteq> X" using arden_helper[OF eq, of "n"] by auto }
-  then have "B ;; A\<star> \<subseteq> X" 
-    unfolding Seq_def Star_def UNION_def by auto
-  moreover
-  { fix s::string
-    obtain k where "k = length s" by auto
-    then have not_in: "s \<notin> X ;; (A \<up> Suc k)" 
-      using seq_pow_length[OF nemp] by blast
-    assume "s \<in> X"
-    then have "s \<in> X ;; (A \<up> Suc k) \<union> (\<Union>m\<in>{0..k}. B ;; (A \<up> m))"
-      using arden_helper[OF eq, of "k"] by auto
-    then have "s \<in> (\<Union>m\<in>{0..k}. B ;; (A \<up> m))" using not_in by auto
-    moreover
-    have "(\<Union>m\<in>{0..k}. B ;; (A \<up> m)) \<subseteq> (\<Union>n. B ;; (A \<up> n))" by auto
-    ultimately 
-    have "s \<in> B ;; A\<star>" 
-      unfolding seq_Union_left Star_def by auto }
-  then have "X \<subseteq> B ;; A\<star>" by auto
-  ultimately 
-  show "X = B ;; A\<star>" by simp
-qed
-
-
-section {* Regular Expressions *}
-
-datatype rexp =
-  NULL
-| EMPTY
-| CHAR char
-| SEQ rexp rexp
-| ALT rexp rexp
-| STAR rexp
-
-
-text {* 
-  The function @{text L} is overloaded, with the idea that @{text "L x"} 
-  evaluates to the language represented by the object @{text x}.
-*}
-
-consts L:: "'a \<Rightarrow> lang"
-
-overloading L_rexp \<equiv> "L::  rexp \<Rightarrow> lang"
-begin
-fun
-  L_rexp :: "rexp \<Rightarrow> lang"
-where
-    "L_rexp (NULL) = {}"
-  | "L_rexp (EMPTY) = {[]}"
-  | "L_rexp (CHAR c) = {[c]}"
-  | "L_rexp (SEQ r1 r2) = (L_rexp r1) ;; (L_rexp r2)"
-  | "L_rexp (ALT r1 r2) = (L_rexp r1) \<union> (L_rexp r2)"
-  | "L_rexp (STAR r) = (L_rexp r)\<star>"
-end
-
-
-text {* ALT-combination for a set of regular expressions *}
-
-abbreviation
-  Setalt  ("\<Uplus>_" [1000] 999) 
-where
-  "\<Uplus>A \<equiv> folds ALT NULL A"
-
-text {* 
-  For finite sets, @{term Setalt} is preserved under @{term L}.
-*}
-
-lemma folds_alt_simp [simp]:
-  fixes rs::"rexp set"
-  assumes a: "finite rs"
-  shows "L (\<Uplus>rs) = \<Union> (L ` rs)"
-unfolding folds_def
-apply(rule set_eqI)
-apply(rule someI2_ex)
-apply(rule_tac finite_imp_fold_graph[OF a])
-apply(erule fold_graph.induct)
-apply(auto)
-done
-
-end
\ No newline at end of file