made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
authorurbanc
Mon, 25 Jul 2011 13:33:38 +0000
changeset 170 b1258b7d2789
parent 169 b794db0b79db
child 171 feb7b31d6bf1
made the theories compatible with the existing developments in the AFP; old theories are in the directory Attic
Attic/Myhill.thy
Attic/MyhillNerode.thy
Attic/Prelude.thy
Attic/old/Closure.thy
Attic/old/Derivs.thy
Attic/old/Folds.thy
Attic/old/My.thy
Attic/old/Myhill_1.thy
Attic/old/Myhill_2.thy
Attic/old/Prefix_subtract.thy
Attic/old/Regular.thy
Closure.thy
Closures.thy
Derivatives.thy
Derivs.thy
Journal/Paper.thy
More_Regular_Set.thy
My.thy
Myhill.thy
MyhillNerode.thy
Myhill_1.thy
Myhill_2.thy
Paper/Paper.thy
Paper/ROOT.ML
Prelude.thy
Regular.thy
Regular_Exp.thy
Regular_Set.thy
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Myhill.thy	Mon Jul 25 13:33:38 2011 +0000
@@ -0,0 +1,331 @@
+theory Myhill
+  imports Myhill_2
+begin
+
+section {* Preliminaries \label{sec_prelim}*}
+
+subsection {* Finite automata and \mht  \label{sec_fa_mh} *}
+
+text {*
+
+A {\em determinisitc finite automata (DFA)} $M$ is a 5-tuple 
+$(Q, \Sigma, \delta, s, F)$, where:
+\begin{enumerate}
+  \item $Q$ is a finite set of {\em states}, also denoted $Q_M$.
+  \item $\Sigma$ is a finite set of {\em alphabets}, also denoted $\Sigma_M$.
+  \item $\delta$ is a {\em transition function} of type @{text "Q \<times> \<Sigma> \<Rightarrow> Q"} (a total function),
+           also denoted $\delta_M$.
+  \item @{text "s \<in> Q"} is a state called {\em initial state}, also denoted $s_M$.
+  \item @{text "F \<subseteq> Q"} is a set of states named {\em accepting states}, also denoted $F_M$.
+\end{enumerate}
+Therefore, we have $M = (Q_M, \Sigma_M, \delta_M, s_M, F_M)$. Every DFA $M$ can be interpreted as 
+a function assigning states to strings, denoted $\dfa{M}$, the  definition of which is as the following:
+\begin{equation}
+\begin{aligned}
+  \dfa{M}([]) &\equiv s_M \\
+   \dfa{M}(xa) &\equiv  \delta_M(\dfa{M}(x), a)
+\end{aligned}
+\end{equation}
+A string @{text "x"} is said to be {\em accepted} (or {\em recognized}) by a DFA $M$ if 
+$\dfa{M}(x) \in F_M$. The language recoginzed by DFA $M$, denoted
+$L(M)$, is defined as:
+\begin{equation}
+  L(M) \equiv \{x~|~\dfa{M}(x) \in F_M\}
+\end{equation}
+The standard way of specifying a laugage $\Lang$ as {\em regular} is by stipulating that:
+$\Lang = L(M)$ for some DFA $M$.
+
+For any DFA $M$, the DFA obtained by changing initial state to another $p \in Q_M$ is denoted $M_p$, 
+which is defined as:
+\begin{equation}
+    M_p \ \equiv\ (Q_M, \Sigma_M, \delta_M, p, F_M) 
+\end{equation}
+Two states $p, q \in Q_M$ are said to be {\em equivalent}, denoted $p \approx_M q$, iff.
+\begin{equation}\label{m_eq_def}
+  L(M_p) = L(M_q)
+\end{equation}
+It is obvious that $\approx_M$ is an equivalent relation over $Q_M$. and 
+the partition induced by $\approx_M$ has $|Q_M|$ equivalent classes.
+By overloading $\approx_M$,  an equivalent relation over strings can be defined:
+\begin{equation}
+  x \approx_M y ~ ~ \equiv ~ ~ \dfa{M}(x) \approx_M \dfa{M}(y)
+\end{equation}
+It can be proved that the the partition induced by $\approx_M$ also has $|Q_M|$ equivalent classes.
+It is also easy to show that: if $x \approx_M y$, then $x \approx_{L(M)} y$, and this means
+$\approx_M$ is a more refined equivalent relation than $\approx_{L(M)}$. Since partition induced by
+$\approx_M$ is finite, the one induced by $\approx_{L(M)}$ must also be finite, and this is 
+one of the two directions of \mht:
+\begin{Lem}[\mht , Direction two]
+  If a language $\Lang$ is regular (i.e. $\Lang = L(M)$ for some DFA $M$), then 
+  the partition induced by $\approx_\Lang$ is finite.
+\end{Lem}
+
+The other direction is:
+\begin{Lem}[\mht , Direction one]\label{auto_mh_d1}
+  If  the partition induced by $\approx_\Lang$ is finite, then
+  $\Lang$ is regular (i.e. $\Lang = L(M)$ for some DFA $M$).
+\end{Lem}
+The $M$ we are seeking when prove lemma \ref{auto_mh_d2} can be constructed out of $\approx_\Lang$,
+denoted $M_\Lang$ and defined as the following:
+\begin{subequations}
+\begin{eqnarray}
+  Q_{M_\Lang} ~ & \equiv & ~ \{ \cls{x}{\approx_\Lang}~|~ x \in \Sigma^* \}\\
+  \Sigma_{M_\Lang} ~ & \equiv & ~ \Sigma_M \\
+  \delta_{M_\Lang} ~ & \equiv & ~ \left (\lambda (\cls{x}{\approx_\Lang}, a).  \cls{xa}{\approx_\Lang} \right) \\
+  s_{M_\Lang} ~ & \equiv & ~ \cls{[]}{\approx_\Lang} \\
+  F_{M_\Lang} ~ & \equiv & ~ \{ \cls{x}{\approx_\Lang}~|~ x \in \Lang \}
+\end{eqnarray}
+\end{subequations}
+It can be proved that $Q_{M_\Lang}$ is indeed finite and $\Lang = L(M_\Lang)$, so lemma \ref{auto_mh_d1} holds.
+It can also be proved that $M_\Lang$ is the minimal DFA (therefore unique) which recoginzes $\Lang$.
+
+
+
+*}
+
+subsection {* The objective and the underlying intuition *}
+
+text {*
+  It is now obvious from section \ref{sec_fa_mh} that \mht\ can be established easily when
+  {\em reglar languages} are defined as ones recognized by finite automata. 
+  Under the context where the use of finite automata is forbiden, the situation is quite different.
+  The theorem now has to be expressed as:
+  \begin{Thm}[\mht , Regular expression version]
+      A language $\Lang$ is regular (i.e. $\Lang = L(\re)$ for some {\em regular expression} $\re$)
+      iff. the partition induced by $\approx_\Lang$ is finite.
+  \end{Thm}
+  The proof of this version consists of two directions (if the use of automata are not allowed):
+  \begin{description}
+    \item [Direction one:] 
+      generating a regular expression $\re$ out of the finite partition induced by $\approx_\Lang$,
+      such that $\Lang = L(\re)$.
+    \item [Direction two:]
+          showing the finiteness of the partition induced by $\approx_\Lang$, under the assmption
+          that $\Lang$ is recognized by some regular expression $\re$ (i.e. $\Lang = L(\re)$).
+  \end{description}
+  The development of these two directions consititutes the body of this paper.
+
+*}
+
+section {* Direction @{text "regular language \<Rightarrow>finite partition"} *}
+
+text {*
+  Although not used explicitly, the notion of finite autotmata and its relationship with 
+  language partition, as outlined in section \ref{sec_fa_mh}, still servers as important intuitive 
+  guides in the development of this paper.
+  For example, {\em Direction one} follows the {\em Brzozowski algebraic method}
+  used to convert finite autotmata to regular expressions, under the intuition that every 
+  partition member $\cls{x}{\approx_\Lang}$ is a state in the DFA $M_\Lang$  constructed to prove 
+  lemma \ref{auto_mh_d1} of section \ref{sec_fa_mh}.
+
+  The basic idea of Brzozowski method is to extract an equational system out of the 
+  transition relationship of the automaton in question. In the equational system, every
+  automaton state is represented by an unknown, the solution of which is expected to be 
+  a regular expresion characterizing the state in a certain sense. There are two choices of 
+  how a automaton state can be characterized.  The first is to characterize by the set of 
+  strings leading from the state in question into accepting states. 
+  The other choice is to characterize by the set of strings leading from initial state 
+  into the state in question. For the second choice, the language recognized the automaton 
+  can be characterized by the solution of initial state, while for the second choice, 
+  the language recoginzed by the automaton can be characterized by 
+  combining solutions of all accepting states by @{text "+"}. Because of the automaton 
+  used as our intuitive guide, the $M_\Lang$, the states of which are 
+  sets of strings leading from initial state, the second choice is used in this paper.
+
+  Supposing the automaton in Fig \ref{fig_auto_part_rel} is the $M_\Lang$ for some language $\Lang$, 
+  and suppose $\Sigma = \{a, b, c, d, e\}$. Under the second choice, the equational system extracted is:
+  \begin{subequations}
+  \begin{eqnarray}
+    X_0 & = & X_1 \cdot c + X_2 \cdot d + \lambda \label{x_0_def_o} \\
+    X_1 & = & X_0 \cdot a + X_1 \cdot b + X_2 \cdot d \label{x_1_def_o} \\
+    X_2 & = & X_0 \cdot b + X_1 \cdot d + X_2 \cdot a \\
+    X_3 & = & \begin{aligned}
+                 & X_0 \cdot (c + d + e) + X_1 \cdot (a + e) + X_2 \cdot (b + e) + \\
+                 & X_3 \cdot (a + b + c + d + e)
+              \end{aligned}
+  \end{eqnarray}
+  \end{subequations} 
+
+\begin{figure}[h!]
+\centering
+\begin{minipage}{0.5\textwidth}
+\scalebox{0.8}{
+\begin{tikzpicture}[ultra thick,>=latex]
+  \node[draw,circle,initial] (start) {$X_0$};
+  \node[draw,circle,accepting] at ($(start) + (3.5cm,2cm)$) (ac1) {$X_1$};
+  \node[draw,circle,accepting] at ($(start) + (3.5cm,-2cm)$) (ac2) {$X_2$};
+  \node[draw,circle] at ($(start) + (6.5cm,0cm)$) (ab) {$X_3$};
+
+  \path[->] (start) edge [bend left] node [midway, above] {$a$} (ac1);
+  \path[->] (start) edge [bend right] node [midway, below] {$b$} (ac2);
+  \path[->] (ac1) edge [loop above] node [midway, above] {$b$} (ac1);
+  \path[->] (ac2) edge [loop below] node [midway, below] {$a$} (ac2);
+  \path[->] (ac1) edge [bend right] node [midway, left] {$c$} (ac2);
+  \path[->] (ac2) edge [bend right] node [midway, right] {$c$} (ac1);
+  \path[->] (ac1) edge node [midway, sloped, above] {$d$} (start);
+  \path[->] (ac2) edge node [midway, sloped, above] {$d$} (start);
+
+  \path [draw, rounded corners,->,dashed] (start) -- ($(start) + (0cm, 3.7cm)$)
+         -- ($(ab) + (0cm, 3.7cm)$)  node[midway,above,sloped]{$\Sigma - \{a, b\}$} -- (ab);
+  \path[->,dashed] (ac1) edge node [midway, above, sloped] {$\Sigma - \{b,c,d\}$} (ab);
+  \path[->,dashed] (ac2) edge node [midway, below, sloped] {$\Sigma - \{a,c,d\}$} (ab);
+  \path[->,dashed] (ab) edge [loop right] node [midway, right] {$\Sigma$} (ab);
+\end{tikzpicture}}
+\end{minipage}
+\caption{An example automaton}\label{fig_auto_part_rel}
+\end{figure}
+
+  Every $\cdot$-item on the right side of equations describes some state transtions, except 
+  the $\lambda$ in \eqref{x_0_def_o}, which represents empty string @{text "[]"}. 
+  The reason is that: every state is characterized by the
+  set of incoming strings leading from initial state. For non-initial state, every such
+  string can be splitted into a prefix leading into a preceding state and a single character suffix 
+  transiting into from the preceding state. The exception happens at
+  initial state, where the empty string is a incoming string which can not be splitted. The $\lambda$
+  in \eqref{x_0_def_o} is introduce to repsent this indivisible string. There is one and only one
+  $\lambda$ in every equational system such obtained, becasue $[]$ can only be contaied in one
+  equivalent class (the intial state in $M_\Lang$) and equivalent classes are disjoint. 
+  
+  Suppose all unknowns ($X_0, X_1, X_2, X_3$) are  solvable, the regular expression charactering 
+  laugnage $\Lang$ is $X_1 + X_2$. This paper gives a procedure
+  by which arbitrarily picked unknown can be solved. The basic idea to solve $X_i$ is by 
+  eliminating all variables other than $X_i$ from the equational system. If
+  $X_0$ is the one picked to be solved,  variables $X_1, X_2, X_3$ have to be removed one by 
+  one.  The order to remove does not matter as long as the remaing equations are kept valid.
+  Suppose $X_1$ is the first one to remove, the action is to replace all occurences of $X_1$ 
+  in remaining equations by the right hand side of its characterizing equation, i.e. 
+  the $ X_0 \cdot a + X_1 \cdot b + X_2 \cdot d$ in \eqref{x_1_def_o}. However, because
+  of the recursive occurence of $X_1$, this replacement does not really removed $X_1$. Arden's
+  lemma is invoked to transform recursive equations like \eqref{x_1_def_o} into non-recursive 
+  ones. For example, the recursive equation \eqref{x_1_def_o} is transformed into the follwing
+  non-recursive one:
+  \begin{equation}
+     X_1  =  (X_0 \cdot a + X_2 \cdot d) \cdot b^* = X_0 \cdot (a \cdot b^*) + X_2 \cdot (d \cdot b^*)
+  \end{equation}
+  which, by Arden's lemma, still characterizes $X_1$ correctly. By subsituting
+  $(X_0 \cdot a + X_2 \cdot d) \cdot b^*$ for all $X_1$ and removing  \eqref{x_1_def_o},
+  we get:
+  \begin{subequations}
+  \begin{eqnarray}
+    X_0 & = & \begin{aligned}
+              & (X_0 \cdot (a \cdot b^*) + X_2 \cdot (d \cdot b^*)) \cdot c + 
+                X_2 \cdot d + \lambda = \\
+              & X_0 \cdot (a \cdot b^* \cdot c) + X_2 \cdot (d \cdot b^* \cdot c) + 
+                X_2 \cdot d + \lambda = \\
+              &  X_0 \cdot (a \cdot b^* \cdot c) + X_2 \cdot (d \cdot b^* \cdot c + d) + \lambda 
+              \end{aligned} \label{x_0_def_1} \\
+    X_2 & = & \begin{aligned}
+                 & X_0 \cdot b + (X_0 \cdot (a \cdot b^*) + X_2 \cdot (d \cdot b^*)) \cdot d + X_2 \cdot a = \\
+                 & X_0 \cdot b + X_0 \cdot (a \cdot b^* \cdot d) + X_2 \cdot (d \cdot b^* \cdot d) + X_2 \cdot a = \\
+                 & X_0 \cdot (b + a \cdot b^* \cdot d) + X_2 \cdot (d \cdot b^* \cdot d + a)
+              \end{aligned} \\
+    X_3 & = & \begin{aligned}
+                 & X_0 \cdot (c + d + e) + ((X_0 \cdot a + X_2 \cdot d) \cdot b^*) \cdot (a + e)\\
+                 & + X_2 \cdot (b + e) + X_3 \cdot (a + b + c + d + e) \label{x_3_def_1}
+              \end{aligned}
+  \end{eqnarray}
+  \end{subequations}  
+Suppose $X_3$ is the one to remove next, since $X_3$ dose not appear in either $X_0$ or $X_2$, 
+the removal of equation \eqref{x_3_def_1} changes nothing in the rest equations. Therefore, we get:
+ \begin{subequations}
+  \begin{eqnarray}
+    X_0 & = & X_0 \cdot (a \cdot b^* \cdot c) + X_2 \cdot (d \cdot b^* \cdot c + d) + \lambda  \label{x_0_def_2} \\
+    X_2 & = & X_0 \cdot (b + a \cdot b^* \cdot d) + X_2 \cdot (d \cdot b^* \cdot d + a) \label{x_2_def_2}
+  \end{eqnarray}
+  \end{subequations}  
+Actually, since absorbing state like $X_3$ contributes nothing to the language $\Lang$, it could have been removed
+at the very beginning of this precedure without affecting the final result. Now, the last unknown to remove 
+is $X_2$ and the Arden's transformaton of \eqref{x_2_def_2} is:
+\begin{equation} \label{x_2_ardened}
+   X_2 ~ = ~  (X_0 \cdot (b + a \cdot b^* \cdot d)) \cdot (d \cdot b^* \cdot d + a)^* =
+              X_0 \cdot ((b + a \cdot b^* \cdot d) \cdot (d \cdot b^* \cdot d + a)^*)
+\end{equation}
+By substituting the right hand side of \eqref{x_2_ardened} into \eqref{x_0_def_2}, we get:
+\begin{equation}
+\begin{aligned}
+  X_0  & = && X_0 \cdot (a \cdot b^* \cdot c) + \\
+       &   && X_0 \cdot ((b + a \cdot b^* \cdot d) \cdot (d \cdot b^* \cdot d + a)^*) \cdot 
+           (d \cdot b^* \cdot c + d) + \lambda \\
+       & =  && X_0 \cdot ((a \cdot b^* \cdot c) + \\
+       &   && \hspace{3em} ((b + a \cdot b^* \cdot d) \cdot (d \cdot b^* \cdot d + a)^*) \cdot 
+           (d \cdot b^* \cdot c + d)) + \lambda 
+\end{aligned}
+\end{equation}
+By applying Arden's transformation to this, we get the solution of $X_0$ as:
+\begin{equation}
+\begin{aligned}
+  X_0  =  ((a \cdot b^* \cdot c) + 
+                ((b + a \cdot b^* \cdot d) \cdot (d \cdot b^* \cdot d + a)^*) \cdot 
+                        (d \cdot b^* \cdot c + d))^*
+\end{aligned}
+\end{equation}
+Using the same method, solutions for $X_1$ and $X_2$ can be obtained as well and the
+regular expressoin for $\Lang$ is just $X_1 + X_2$. The formalization of this procedure
+consititues the first direction of the {\em regular expression} verion of
+\mht. Detailed explaination are given in {\bf paper.pdf} and more details and comments
+can be found in the formal scripts.
+*}
+
+section {* Direction @{text "finite partition \<Rightarrow> regular language"} *}
+
+text {*
+  It is well known in the theory of regular languages that
+  the existence of finite  language partition amounts to the existence of 
+  a minimal automaton, i.e. the $M_\Lang$ constructed in section \ref{sec_prelim}, which recoginzes the 
+  same language $\Lang$. The standard way to prove the existence of finite language partition 
+  is to construct a automaton out of the regular expression which recoginzes the same language, from
+  which the existence of finite language partition follows immediately. As discussed in the introducton of 
+  {\bf paper.pdf} as well as in [5], the problem for this approach happens when automata 
+  of sub regular expressions are combined to form the automaton of the mother regular expression, 
+  no matter what kind of representation is used, the formalization is cubersome, as said 
+  by Nipkow in [5]: `{\em a more abstract mathod is clearly desirable}'. In this section, 
+  an {\em intrinsically abstract} method is given, which completely avoid the mentioning of automata,
+  let along any particular representations.
+  *}
+
+text {*
+  The main proof structure is a structural induction on regular expressions,
+  where base cases (cases for @{const "NULL"}, @{const "EMPTY"}, @{const "CHAR"}) are quite straightforward to
+  proof. Real difficulty lies in inductive cases.  By inductive hypothesis, languages defined by
+  sub-expressions induce finite partitiions. Under such hypothsis, we need to prove that the language 
+  defined by the composite regular expression gives rise to finite partion.  
+  The basic idea is to attach a tag @{text "tag(x)"} to every string 
+  @{text "x"}. The tagging fuction @{text "tag"} is carefully devised, which returns tags 
+  made of equivalent classes of the partitions induced by subexpressoins, and therefore has a finite
+  range. Let @{text "Lang"} be the composite language, it is proved that:
+  \begin{quote}
+  If strings with the same tag are equivalent with respect to @{text "Lang"}, expressed as:
+  \[
+  @{text "tag(x) = tag(y) \<Longrightarrow> x \<approx>Lang y"}
+  \]
+  then the partition induced by @{text "Lang"} must be finite.
+  \end{quote}
+  There are two arguments for this. The first goes as the following:
+  \begin{enumerate}
+    \item First, the tagging function @{text "tag"} induces an equivalent relation @{text "(=tag=)"} 
+          (defiintion of @{text "f_eq_rel"} and lemma @{text "equiv_f_eq_rel"}).
+    \item It is shown that: if the range of @{text "tag"} (denoted @{text "range(tag)"}) is finite, 
+           the partition given rise by @{text "(=tag=)"} is finite (lemma @{text "finite_eq_f_rel"}).
+           Since tags are made from equivalent classes from component partitions, and the inductive
+           hypothesis ensures the finiteness of these partitions, it is not difficult to prove
+           the finiteness of @{text "range(tag)"}.
+    \item It is proved that if equivalent relation @{text "R1"} is more refined than @{text "R2"}
+           (expressed as @{text "R1 \<subseteq> R2"}),
+           and the partition induced by @{text "R1"} is finite, then the partition induced by @{text "R2"}
+           is finite as well (lemma @{text "refined_partition_finite"}).
+    \item The injectivity assumption @{text "tag(x) = tag(y) \<Longrightarrow> x \<approx>Lang y"} implies that
+            @{text "(=tag=)"} is more refined than @{text "(\<approx>Lang)"}.
+    \item Combining the points above, we have: the partition induced by language @{text "Lang"}
+          is finite (lemma @{text "tag_finite_imageD"}).
+  \end{enumerate}
+
+We could have followed another approach given in  appendix II of Brzozowski's paper [?], where
+the set of derivatives of any regular expression can be proved to be finite. 
+Since it is easy to show that strings with same derivative are equivalent with respect to the 
+language, then the second direction follows. We believe that our
+apporoach is easy to formalize, with no need to do complicated derivation calculations
+and countings as in [???].
+*}
+
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/MyhillNerode.thy	Mon Jul 25 13:33:38 2011 +0000
@@ -0,0 +1,1816 @@
+theory MyhillNerode
+  imports "Main" "List_Prefix"
+begin
+
+text {* sequential composition of languages *}
+
+definition
+  lang_seq :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ; _" [100,100] 100)
+where 
+  "L1 ; L2 = {s1 @ s2 | s1 s2. s1 \<in> L1 \<and> s2 \<in> L2}"
+
+lemma lang_seq_empty:
+  shows "{[]} ; L = L"
+  and   "L ; {[]} = L"
+unfolding lang_seq_def by auto
+
+lemma lang_seq_null:
+  shows "{} ; L = {}"
+  and   "L ; {} = {}"
+unfolding lang_seq_def by auto
+
+lemma lang_seq_append:
+  assumes a: "s1 \<in> L1"
+  and     b: "s2 \<in> L2"
+  shows "s1@s2 \<in> L1 ; L2"
+unfolding lang_seq_def
+using a b by auto 
+
+lemma lang_seq_union:
+  shows "(L1 \<union> L2); L3 = (L1; L3) \<union> (L2; L3)"
+  and   "L1; (L2 \<union> L3) = (L1; L2) \<union> (L1; L3)"
+unfolding lang_seq_def by auto
+
+lemma lang_seq_assoc:
+  shows "(L1 ; L2) ; L3 = L1 ; (L2 ; L3)"
+unfolding lang_seq_def
+apply(auto)
+apply(metis)
+by (metis append_assoc)
+
+
+section {* Kleene star for languages defined as least fixed point *}
+
+inductive_set
+  Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102)
+  for L :: "string set"
+where
+  start[intro]: "[] \<in> L\<star>"
+| step[intro]:  "\<lbrakk>s1 \<in> L; s2 \<in> L\<star>\<rbrakk> \<Longrightarrow> s1@s2 \<in> L\<star>"
+
+lemma lang_star_empty:
+  shows "{}\<star> = {[]}"
+by (auto elim: Star.cases)
+
+lemma lang_star_cases:
+  shows "L\<star> =  {[]} \<union> L ; L\<star>"
+proof
+  { fix x
+    have "x \<in> L\<star> \<Longrightarrow> x \<in> {[]} \<union> L ; L\<star>"
+      unfolding lang_seq_def
+    by (induct rule: Star.induct) (auto)
+  }
+  then show "L\<star> \<subseteq> {[]} \<union> L ; L\<star>" by auto
+next
+  show "{[]} \<union> L ; L\<star> \<subseteq> L\<star>" 
+    unfolding lang_seq_def by auto
+qed
+
+lemma lang_star_cases':
+  shows "L\<star> =  {[]} \<union> L\<star> ; L"
+proof
+  { fix x
+    have "x \<in> L\<star> \<Longrightarrow> x \<in> {[]} \<union> L\<star> ; L"
+      unfolding lang_seq_def
+    apply (induct rule: Star.induct)
+    apply simp
+    apply simp
+    apply (erule disjE)
+    apply (auto)[1]
+    apply (erule exE | erule conjE)+
+    apply (rule disjI2)
+    apply (rule_tac x = "s1 @ s1a" in exI)
+    by auto
+  }
+  then show "L\<star> \<subseteq> {[]} \<union> L\<star> ; L" by auto
+next
+  show "{[]} \<union> L\<star> ; L \<subseteq> L\<star>" 
+    unfolding lang_seq_def
+    apply auto
+    apply (erule Star.induct)
+    apply auto
+    apply (drule step[of _ _ "[]"])
+    by (auto intro:start)
+qed
+
+lemma lang_star_simple:
+  shows "L \<subseteq> L\<star>"
+by (subst lang_star_cases)
+   (auto simp only: lang_seq_def)
+
+lemma lang_star_prop0_aux:
+  "s2 \<in> L\<star> \<Longrightarrow> \<forall> s1. s1 \<in> L \<longrightarrow> (\<exists> s3 s4. s3 \<in> L\<star> \<and> s4 \<in> L \<and> s1 @ s2 = s3 @ s4)" 
+apply (erule Star.induct)
+apply (clarify, rule_tac x = "[]" in exI, rule_tac x = s1 in exI, simp add:start)
+apply (clarify, drule_tac x = s1 in spec)
+apply (drule mp, simp, clarify)
+apply (rule_tac x = "s1a @ s3" in exI, rule_tac x = s4 in exI)
+by auto
+
+lemma lang_star_prop0:
+  "\<lbrakk>s1 \<in> L; s2 \<in> L\<star>\<rbrakk> \<Longrightarrow> \<exists> s3 s4. s3 \<in> L\<star> \<and> s4 \<in> L \<and> s1 @ s2 = s3 @ s4" 
+by (auto dest:lang_star_prop0_aux)
+
+lemma lang_star_prop1:
+  assumes asm: "L1; L2 \<subseteq> L2" 
+  shows "L1\<star>; L2 \<subseteq> L2"
+proof -
+  { fix s1 s2
+    assume minor: "s2 \<in> L2"
+    assume major: "s1 \<in> L1\<star>"
+    then have "s1@s2 \<in> L2"
+    proof(induct rule: Star.induct)
+      case start
+      show "[]@s2 \<in> L2" using minor by simp
+    next
+      case (step s1 s1')
+      have "s1 \<in> L1" by fact
+      moreover
+      have "s1'@s2 \<in> L2" by fact
+      ultimately have "s1@(s1'@s2) \<in> L1; L2" by (auto simp add: lang_seq_def)
+      with asm have "s1@(s1'@s2) \<in> L2" by auto
+      then show "(s1@s1')@s2 \<in> L2" by simp
+    qed
+  } 
+  then show "L1\<star>; L2 \<subseteq> L2" by (auto simp add: lang_seq_def)
+qed
+
+lemma lang_star_prop2_aux:
+  "s2 \<in> L2\<star> \<Longrightarrow> \<forall> s1. s1 \<in> L1 \<and> L1 ; L2 \<subseteq> L1 \<longrightarrow> s1 @ s2 \<in> L1"
+apply (erule Star.induct, simp)
+apply (clarify, drule_tac x = "s1a @ s1" in spec)
+by (auto simp:lang_seq_def)
+
+lemma lang_star_prop2:
+  "L1; L2 \<subseteq> L1 \<Longrightarrow> L1 ; L2\<star> \<subseteq> L1"
+by (auto dest!:lang_star_prop2_aux simp:lang_seq_def)
+
+lemma lang_star_seq_subseteq: 
+  shows "L ; L\<star> \<subseteq> L\<star>"
+using lang_star_cases by blast
+
+lemma lang_star_double:
+  shows "L\<star>; L\<star> = L\<star>"
+proof
+  show "L\<star> ; L\<star> \<subseteq> L\<star>" 
+    using lang_star_prop1 lang_star_seq_subseteq by blast
+next
+  have "L\<star> \<subseteq> L\<star> \<union> L\<star>; (L; L\<star>)" by auto
+  also have "\<dots> = L\<star>;{[]} \<union> L\<star>; (L; L\<star>)" by (simp add: lang_seq_empty)
+  also have "\<dots> = L\<star>; ({[]} \<union> L; L\<star>)" by (simp only: lang_seq_union)
+  also have "\<dots> = L\<star>; L\<star>" using lang_star_cases by simp 
+  finally show "L\<star> \<subseteq> L\<star> ; L\<star>" by simp
+qed
+
+lemma lang_star_seq_subseteq': 
+  shows "L\<star>; L \<subseteq> L\<star>"
+proof -
+  have "L \<subseteq> L\<star>" by (rule lang_star_simple)
+  then have "L\<star>; L \<subseteq> L\<star>; L\<star>" by (auto simp add: lang_seq_def)
+  then show "L\<star>; L \<subseteq> L\<star>" using lang_star_double by blast
+qed
+
+lemma
+  shows "L\<star> \<subseteq> L\<star>\<star>"
+by (rule lang_star_simple)
+
+
+section {* regular expressions *}
+
+datatype rexp =
+  NULL
+| EMPTY
+| CHAR char
+| SEQ rexp rexp
+| ALT rexp rexp
+| STAR rexp
+
+
+consts L:: "'a \<Rightarrow> string set"
+
+overloading L_rexp \<equiv> "L::  rexp \<Rightarrow> string set"
+begin
+fun
+  L_rexp :: "rexp \<Rightarrow> string set"
+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{* ************ now is the codes writen by chunhan ************************************* *}
+
+section {* Arden's Lemma revised *}
+
+lemma arden_aux1:
+  assumes a: "X \<subseteq> X ; A \<union> B"
+  and     b: "[] \<notin> A"
+  shows      "x \<in> X \<Longrightarrow> x \<in> B ; A\<star>"
+apply (induct x taking:length rule:measure_induct)
+apply (subgoal_tac "x \<in> X ; A \<union> B")
+defer
+using a
+apply (auto)[1]
+apply simp
+apply (erule disjE)
+defer
+apply (auto simp add:lang_seq_def) [1]
+apply (subgoal_tac "\<exists> x1 x2. x = x1 @ x2 \<and> x1 \<in> X \<and> x2 \<in> A")
+defer
+apply (auto simp add:lang_seq_def) [1]
+apply (erule exE | erule conjE)+
+apply simp
+apply (drule_tac x = x1 in spec)
+apply (simp)
+using b
+apply -
+apply (auto)[1]
+apply (subgoal_tac "x1 @ x2 \<in> (B ; A\<star>) ; A")
+defer
+apply (auto simp add:lang_seq_def)[1]
+by (metis Un_absorb1 lang_seq_assoc lang_seq_union(2) lang_star_double lang_star_simple mem_def sup1CI)
+
+theorem ardens_revised:
+  assumes nemp: "[] \<notin> A"
+  shows "(X = X ; A \<union> B) \<longleftrightarrow> (X = B ; A\<star>)"
+apply(rule iffI)
+defer
+apply(simp)
+apply(subst lang_star_cases')
+apply(subst lang_seq_union)
+apply(simp add: lang_seq_empty)
+apply(simp add: lang_seq_assoc)
+apply(auto)[1]
+proof -
+  assume "X = X ; A \<union> B"
+  then have as1: "X ; A \<union> B \<subseteq> X" and as2: "X \<subseteq> X ; A \<union> B" by simp_all
+  from as1 have a: "X ; A \<subseteq> X" and b: "B \<subseteq> X" by simp_all
+  from b have "B; A\<star> \<subseteq> X ; A\<star>" by (auto simp add: lang_seq_def)
+  moreover
+  from a have "X ; A\<star> \<subseteq> X" 
+
+by (rule lang_star_prop2)
+  ultimately have f1: "B ; A\<star> \<subseteq> X" by simp
+  from as2 nemp
+  have f2: "X \<subseteq> B; A\<star>" using arden_aux1 by auto
+  from f1 f2 show "X = B; A\<star>" by auto
+qed
+
+
+
+section {* equiv class' definition *}
+
+definition 
+  equiv_str :: "string \<Rightarrow> string set \<Rightarrow> string \<Rightarrow> bool" ("_ \<equiv>_\<equiv> _" [100, 100, 100] 100)
+where
+  "x \<equiv>Lang\<equiv> y \<longleftrightarrow> (\<forall>z. x @ z \<in> Lang \<longleftrightarrow> y @ z \<in> Lang)"
+
+definition
+  equiv_class :: "string \<Rightarrow> (string set) \<Rightarrow> string set" ("\<lbrakk>_\<rbrakk>_" [100, 100] 100)
+where
+  "\<lbrakk>x\<rbrakk>Lang \<equiv> {y. x \<equiv>Lang\<equiv> y}"
+
+text {* Chunhan modifies Q to Quo *}
+
+definition  
+  quot :: "string set \<Rightarrow> string set \<Rightarrow> (string set) set" ("_ Quo _" [100, 100] 100)
+where
+  "L1 Quo L2 \<equiv> { \<lbrakk>x\<rbrakk>L2 | x. x \<in> L1}" 
+
+
+lemma lang_eqs_union_of_eqcls: 
+  "Lang = \<Union> {X. X \<in> (UNIV Quo Lang) \<and> (\<forall> x \<in> X. x \<in> Lang)}"
+proof
+  show "Lang \<subseteq> \<Union>{X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}"
+  proof
+    fix x
+    assume "x \<in> Lang"
+    thus "x \<in> \<Union>{X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}"
+    proof (simp add:quot_def)
+      assume "(1)": "x \<in> Lang"
+      show "\<exists>xa. (\<exists>x. xa = \<lbrakk>x\<rbrakk>Lang) \<and> (\<forall>x\<in>xa. x \<in> Lang) \<and> x \<in> xa" (is "\<exists>xa.?P xa")
+      proof -
+        have "?P (\<lbrakk>x\<rbrakk>Lang)" using "(1)"
+          by (auto simp:equiv_class_def equiv_str_def dest: spec[where  x = "[]"])
+        thus ?thesis by blast
+      qed
+    qed
+  qed   
+next
+  show "\<Union>{X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang} \<subseteq> Lang"
+    by auto
+qed
+
+lemma empty_notin_CS: "{} \<notin> UNIV Quo Lang"
+apply (clarsimp simp:quot_def equiv_class_def)
+by (rule_tac x = x in exI, auto simp:equiv_str_def)
+
+lemma no_two_cls_inters:
+  "\<lbrakk>X \<in> UNIV Quo Lang; Y \<in> UNIV Quo Lang; X \<noteq> Y\<rbrakk> \<Longrightarrow> X \<inter> Y = {}"
+by (auto simp:quot_def equiv_class_def equiv_str_def)
+
+text {* equiv_class transition *}
+definition 
+  CT :: "string set \<Rightarrow> char \<Rightarrow> string set \<Rightarrow> bool" ("_-_\<rightarrow>_" [99,99]99)
+where
+  "X-c\<rightarrow>Y \<equiv> ((X;{[c]}) \<subseteq> Y)"
+
+types t_equa_rhs = "(string set \<times> rexp) set"
+
+types t_equa = "(string set \<times> t_equa_rhs)"
+
+types t_equas = "t_equa set"
+
+text {* 
+  "empty_rhs" generates "\<lambda>" for init-state, just like "\<lambda>" for final states 
+  in Brzozowski method. But if the init-state is "{[]}" ("\<lambda>" itself) then 
+  empty set is returned, see definition of "equation_rhs" 
+*} 
+
+definition 
+  empty_rhs :: "string set \<Rightarrow> t_equa_rhs"
+where
+  "empty_rhs X \<equiv> if ([] \<in> X) then {({[]}, EMPTY)} else {}"
+
+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"
+
+definition 
+  equation_rhs :: "(string set) set \<Rightarrow> string set \<Rightarrow> t_equa_rhs"
+where
+  "equation_rhs CS X \<equiv> if (X = {[]}) then {({[]}, EMPTY)}
+                         else {(S, folds ALT NULL {CHAR c| c. S-c\<rightarrow>X} )| S. S \<in> CS} \<union> empty_rhs X"
+
+definition 
+  equations :: "(string set) set \<Rightarrow> t_equas"
+where
+  "equations CS \<equiv> {(X, equation_rhs CS X) | X. X \<in> CS}"
+
+overloading L_rhs \<equiv> "L:: t_equa_rhs \<Rightarrow> string set"
+begin
+fun L_rhs:: "t_equa_rhs \<Rightarrow> string set"
+where
+  "L_rhs rhs = {x. \<exists> X r. (X, r) \<in> rhs \<and> x \<in> X;(L r)}"
+end
+
+definition 
+  distinct_rhs :: "t_equa_rhs \<Rightarrow> bool"
+where
+  "distinct_rhs rhs \<equiv> \<forall> X reg\<^isub>1 reg\<^isub>2. (X, reg\<^isub>1) \<in> rhs \<and> (X, reg\<^isub>2) \<in> rhs \<longrightarrow> reg\<^isub>1 = reg\<^isub>2"
+
+definition
+  distinct_equas_rhs :: "t_equas \<Rightarrow> bool"
+where
+  "distinct_equas_rhs equas \<equiv> \<forall> X rhs. (X, rhs) \<in> equas \<longrightarrow> distinct_rhs rhs"
+
+definition 
+  distinct_equas :: "t_equas \<Rightarrow> bool"
+where
+  "distinct_equas equas \<equiv> \<forall> X rhs rhs'. (X, rhs) \<in> equas \<and> (X, rhs') \<in> equas \<longrightarrow> rhs = rhs'"
+
+definition 
+  seq_rhs_r :: "t_equa_rhs \<Rightarrow> rexp \<Rightarrow> t_equa_rhs"
+where
+  "seq_rhs_r rhs r \<equiv> (\<lambda>(X, reg). (X, SEQ reg r)) ` rhs"
+
+definition 
+  del_x_paired :: "('a \<times> 'b) set \<Rightarrow> 'a \<Rightarrow> ('a \<times> 'b) set"
+where
+  "del_x_paired S x \<equiv> S - {X. X \<in> S \<and> fst X = x}"
+
+definition
+  arden_variate :: "string set \<Rightarrow> rexp \<Rightarrow> t_equa_rhs \<Rightarrow> t_equa_rhs"
+where
+  "arden_variate X r rhs \<equiv> seq_rhs_r (del_x_paired rhs X) (STAR r)"
+
+definition
+  no_EMPTY_rhs :: "t_equa_rhs \<Rightarrow> bool"
+where
+  "no_EMPTY_rhs rhs \<equiv> \<forall> X r. (X, r) \<in> rhs \<and> X \<noteq> {[]} \<longrightarrow> [] \<notin> L r"
+
+definition 
+  no_EMPTY_equas :: "t_equas \<Rightarrow> bool"
+where
+  "no_EMPTY_equas equas \<equiv> \<forall> X rhs. (X, rhs) \<in> equas \<longrightarrow> no_EMPTY_rhs rhs"
+
+lemma fold_alt_null_eqs:
+  "finite rS \<Longrightarrow> x \<in> L (folds ALT NULL rS) = (\<exists> r \<in> rS. x \<in> L r)"
+apply (simp add:folds_def)
+apply (rule someI2_ex)
+apply (erule finite_imp_fold_graph)
+apply (erule fold_graph.induct)
+by auto (*??? how do this be in Isar ?? *)
+
+lemma seq_rhs_r_prop1:
+  "L (seq_rhs_r rhs r) = (L rhs);(L r)"
+apply (auto simp:seq_rhs_r_def image_def lang_seq_def)
+apply (rule_tac x = "s1 @ s1a" in exI, rule_tac x = "s2a" in exI, simp)
+apply (rule_tac x = a in exI, rule_tac x = b in exI, simp)
+apply (rule_tac x = s1 in exI, rule_tac x = s1a in exI, simp)
+apply (rule_tac x = X in exI, rule_tac x  = "SEQ ra r" in exI, simp)
+apply (rule conjI)
+apply (rule_tac x = "(X, ra)" in bexI, simp+) 
+apply (rule_tac x = s1a in exI, rule_tac x = "s2a @ s2" in exI, simp)
+apply (simp add:lang_seq_def)
+by (rule_tac x = s2a in exI, rule_tac x = s2 in exI, simp)
+
+lemma del_x_paired_prop1:  
+  "\<lbrakk>distinct_rhs rhs; (X, r) \<in> rhs\<rbrakk> \<Longrightarrow> X ; L r \<union> L (del_x_paired rhs X) = L rhs"
+  apply (simp add:del_x_paired_def)
+  apply (simp add: distinct_rhs_def)
+  apply(auto simp add: lang_seq_def)
+  apply(metis)
+  done
+
+lemma arden_variate_prop:
+  assumes "(X, rx) \<in> rhs"
+  shows "(\<forall> Y. Y \<noteq> X \<longrightarrow> (\<exists> r. (Y, r) \<in> rhs) = (\<exists> r. (Y, r) \<in> (arden_variate X rx rhs)))"
+proof (rule allI, rule impI)
+  fix Y
+  assume "(1)": "Y \<noteq> X"
+  show "(\<exists>r. (Y, r) \<in> rhs) = (\<exists>r. (Y, r) \<in> arden_variate X rx rhs)"
+  proof
+    assume "(1_1)": "\<exists>r. (Y, r) \<in> rhs"
+    show "\<exists>r. (Y, r) \<in> arden_variate X rx rhs" (is "\<exists>r. ?P r")
+    proof -
+      from "(1_1)" obtain r where "(Y, r) \<in> rhs" ..
+      hence "?P (SEQ r (STAR rx))"
+      proof (simp add:arden_variate_def image_def)
+        have "(Y, r) \<in> rhs \<Longrightarrow> (Y, r) \<in> del_x_paired rhs X"
+          by (auto simp:del_x_paired_def "(1)")
+        thus "(Y, r) \<in> rhs \<Longrightarrow> (Y, SEQ r (STAR rx)) \<in> seq_rhs_r (del_x_paired rhs X) (STAR rx)"
+          by (auto simp:seq_rhs_r_def)
+      qed
+      thus ?thesis by blast
+    qed
+  next
+    assume "(2_1)": "\<exists>r. (Y, r) \<in> arden_variate X rx rhs"
+    thus "\<exists>r. (Y, r) \<in> rhs" (is "\<exists> r. ?P r")
+      by (auto simp:arden_variate_def del_x_paired_def seq_rhs_r_def image_def)
+  qed
+qed
+
+text {*
+  arden_variate_valid:  proves variation from 
+  
+   "X = X;r + Y;ry + \<dots>" to "X = Y;(SEQ ry (STAR r)) + \<dots>" 
+
+  holds the law of "language of left equiv language of right" 
+*}
+lemma arden_variate_valid:
+  assumes X_not_empty: "X \<noteq> {[]}"
+  and     l_eq_r:   "X = L rhs"
+  and     dist: "distinct_rhs rhs"
+  and     no_empty: "no_EMPTY_rhs rhs"
+  and     self_contained: "(X, r) \<in> rhs"
+  shows   "X = L (arden_variate X r rhs)" 
+proof -
+  have "[] \<notin> L r" using no_empty X_not_empty self_contained
+    by (auto simp:no_EMPTY_rhs_def)
+  hence ardens: "X = X;(L r) \<union> (L (del_x_paired rhs X)) \<longleftrightarrow> X = (L (del_x_paired rhs X)) ; (L r)\<star>" 
+    by (rule ardens_revised)
+  have del_x: "X = X ; L r \<union> L (del_x_paired rhs X) \<longleftrightarrow> X = L rhs" using dist l_eq_r self_contained
+    by (auto dest!:del_x_paired_prop1)
+  show ?thesis
+  proof
+    show "X \<subseteq> L (arden_variate X r rhs)"
+    proof
+      fix x
+      assume "(1_1)": "x \<in> X" with l_eq_r ardens del_x
+      show "x \<in> L (arden_variate X r rhs)" 
+        by (simp add:arden_variate_def seq_rhs_r_prop1 del:L_rhs.simps)
+    qed
+  next
+    show "L (arden_variate X r rhs) \<subseteq> X"
+    proof
+      fix x
+      assume "(2_1)": "x \<in> L (arden_variate X r rhs)" with ardens del_x l_eq_r
+      show "x \<in> X" 
+        by (simp add:arden_variate_def seq_rhs_r_prop1 del:L_rhs.simps)
+    qed
+  qed
+qed
+
+text {* 
+  merge_rhs {(x1, r1), (x2, r2}, (x4, r4), \<dots>} {(x1, r1'), (x3, r3'), \<dots>} = 
+     {(x1, ALT r1 r1'}, (x2, r2), (x3, r3'), (x4, r4), \<dots>} *}  
+definition 
+  merge_rhs :: "t_equa_rhs \<Rightarrow> t_equa_rhs \<Rightarrow> t_equa_rhs"
+where
+  "merge_rhs rhs rhs' \<equiv> {(X, r). (\<exists> r1 r2. ((X,r1) \<in> rhs \<and> (X, r2) \<in> rhs') \<and> r = ALT r1 r2)  \<or>
+                                 (\<exists> r1. (X, r1) \<in> rhs \<and> (\<not> (\<exists> r2. (X, r2) \<in> rhs')) \<and> r = r1) \<or>
+                                 (\<exists> r2. (X, r2) \<in> rhs' \<and> (\<not> (\<exists> r1. (X, r1) \<in> rhs)) \<and> r = r2)    }"                                  
+
+
+text {* rhs_subst rhs X=xrhs r: substitude all occurence of X in rhs((X,r) \<in> rhs) with xrhs *}
+definition 
+  rhs_subst :: "t_equa_rhs \<Rightarrow> string set \<Rightarrow> t_equa_rhs \<Rightarrow> rexp \<Rightarrow> t_equa_rhs"
+where
+  "rhs_subst rhs X xrhs r \<equiv> merge_rhs (del_x_paired rhs X) (seq_rhs_r xrhs r)"
+
+definition 
+  equas_subst_f :: "string set \<Rightarrow> t_equa_rhs \<Rightarrow> t_equa \<Rightarrow> t_equa"
+where
+  "equas_subst_f X xrhs equa \<equiv> let (Y, rhs) = equa in
+                                 if (\<exists> r. (X, r) \<in> rhs)
+                                 then (Y, rhs_subst rhs X xrhs (SOME r. (X, r) \<in> rhs))
+                                 else equa"
+
+definition
+  equas_subst :: "t_equas \<Rightarrow> string set \<Rightarrow> t_equa_rhs \<Rightarrow> t_equas"
+where
+  "equas_subst ES X xrhs \<equiv> del_x_paired (equas_subst_f X xrhs ` ES) X"
+
+lemma lang_seq_prop1:
+ "x \<in> X ; L r \<Longrightarrow> x \<in> X ; (L r \<union> L r')"
+by (auto simp:lang_seq_def)
+
+lemma lang_seq_prop1':
+  "x \<in> X; L r \<Longrightarrow> x \<in> X ; (L r' \<union> L r)"
+by (auto simp:lang_seq_def)
+
+lemma lang_seq_prop2:
+  "x \<in> X; (L r \<union> L r') \<Longrightarrow> x \<in> X;L r \<or> x \<in> X;L r'"
+by (auto simp:lang_seq_def)
+
+lemma merge_rhs_prop1:
+  shows "L (merge_rhs rhs rhs') = L rhs \<union> L rhs' "
+apply (auto simp add:merge_rhs_def dest!:lang_seq_prop2 intro:lang_seq_prop1)
+apply (rule_tac x = X in exI, rule_tac x = r1 in exI, simp)
+apply (case_tac "\<exists> r2. (X, r2) \<in> rhs'")
+apply (clarify, rule_tac x = X in exI, rule_tac x = "ALT r r2" in exI, simp add:lang_seq_prop1)
+apply (rule_tac x = X in exI, rule_tac x = r in exI, simp)
+apply (case_tac "\<exists> r1. (X, r1) \<in> rhs")
+apply (clarify, rule_tac x = X in exI, rule_tac x = "ALT r1 r" in exI, simp add:lang_seq_prop1')
+apply (rule_tac x = X in exI, rule_tac x = r in exI, simp)
+done
+
+lemma no_EMPTY_rhss_imp_merge_no_EMPTY:
+  "\<lbrakk>no_EMPTY_rhs rhs; no_EMPTY_rhs rhs'\<rbrakk> \<Longrightarrow> no_EMPTY_rhs (merge_rhs rhs rhs')"
+apply (simp add:no_EMPTY_rhs_def merge_rhs_def)
+apply (clarify, (erule conjE | erule exE | erule disjE)+)
+by auto
+
+lemma distinct_rhs_prop:
+  "\<lbrakk>distinct_rhs rhs; (X, r1) \<in> rhs; (X, r2) \<in> rhs\<rbrakk> \<Longrightarrow> r1 = r2"
+by (auto simp:distinct_rhs_def)
+
+lemma merge_rhs_prop2:
+  assumes dist_rhs: "distinct_rhs rhs"
+  and     dist_rhs':"distinct_rhs rhs'"
+  shows "distinct_rhs (merge_rhs rhs rhs')"
+apply (auto simp:merge_rhs_def distinct_rhs_def)
+using dist_rhs
+apply (drule distinct_rhs_prop, simp+)
+using dist_rhs'
+apply (drule distinct_rhs_prop, simp+)
+using dist_rhs
+apply (drule distinct_rhs_prop, simp+)
+using dist_rhs'
+apply (drule distinct_rhs_prop, simp+)
+done
+
+lemma seq_rhs_r_holds_distinct: 
+  "distinct_rhs rhs \<Longrightarrow> distinct_rhs (seq_rhs_r rhs r)"
+by (auto simp:distinct_rhs_def seq_rhs_r_def)
+
+lemma seq_rhs_r_prop0:
+  assumes l_eq_r: "X = L xrhs"
+  shows "L (seq_rhs_r xrhs r) = X ; L r "
+using l_eq_r
+by (auto simp only:seq_rhs_r_prop1)
+
+lemma rhs_subst_prop1:
+  assumes l_eq_r: "X = L xrhs"
+  and     dist:  "distinct_rhs rhs"
+  shows "(X, r) \<in> rhs \<Longrightarrow> L rhs = L (rhs_subst rhs X xrhs r)"
+apply (simp add:rhs_subst_def merge_rhs_prop1 del:L_rhs.simps)
+using l_eq_r 
+apply (drule_tac r = r in seq_rhs_r_prop0, simp del:L_rhs.simps)
+using dist
+by (auto dest!:del_x_paired_prop1 simp del:L_rhs.simps)
+
+lemma del_x_paired_holds_distinct_rhs:
+  "distinct_rhs rhs \<Longrightarrow> distinct_rhs (del_x_paired rhs x)"
+by (auto simp:distinct_rhs_def del_x_paired_def)
+
+lemma rhs_subst_holds_distinct_rhs:
+  "\<lbrakk>distinct_rhs rhs; distinct_rhs xrhs\<rbrakk> \<Longrightarrow> distinct_rhs (rhs_subst rhs X xrhs r)"
+apply (drule_tac r = r and rhs = xrhs in seq_rhs_r_holds_distinct)
+apply (drule_tac x = X in del_x_paired_holds_distinct_rhs)
+by (auto dest:merge_rhs_prop2[where rhs = "del_x_paired rhs X"] simp:rhs_subst_def)
+
+section {* myhill-nerode theorem *}
+
+definition left_eq_cls :: "t_equas \<Rightarrow> (string set) set"
+where
+  "left_eq_cls ES \<equiv> {X. \<exists> rhs. (X, rhs) \<in> ES} "
+
+definition right_eq_cls :: "t_equas \<Rightarrow> (string set) set"
+where
+  "right_eq_cls ES \<equiv> {Y. \<exists> X rhs r. (X, rhs) \<in> ES \<and> (Y, r) \<in> rhs }"
+
+definition rhs_eq_cls :: "t_equa_rhs \<Rightarrow> (string set) set"
+where
+  "rhs_eq_cls rhs \<equiv> {Y. \<exists> r. (Y, r) \<in> rhs}"
+
+definition ardenable :: "t_equa \<Rightarrow> bool"
+where
+  "ardenable equa \<equiv> let (X, rhs) = equa in 
+                      distinct_rhs rhs \<and> no_EMPTY_rhs rhs \<and> X = L rhs"
+
+text {*
+  Inv: Invairance of the equation-system, during the decrease of the equation-system, Inv holds.
+*}
+definition Inv :: "string set \<Rightarrow> t_equas \<Rightarrow> bool"
+where
+  "Inv X ES \<equiv> finite ES \<and> (\<exists> rhs. (X, rhs) \<in> ES) \<and> distinct_equas ES \<and> 
+            (\<forall> X xrhs. (X, xrhs) \<in> ES \<longrightarrow> ardenable (X, xrhs) \<and> X \<noteq> {} \<and> rhs_eq_cls xrhs \<subseteq> insert {[]} (left_eq_cls ES))"
+
+text {*
+  TCon: Termination Condition of the equation-system decreasion.
+*}
+definition TCon:: "'a set \<Rightarrow> bool"
+where
+  "TCon ES \<equiv> card ES = 1"
+
+
+text {* 
+  The following is a iteration principle, and is the main framework for the proof:
+   1: We can form the initial equation-system using "equations" defined above, and prove it has invariance Inv by lemma "init_ES_satisfy_Inv";
+   2: We can decrease the number of the equation-system using ardens_lemma_revised and substitution ("equas_subst", defined above), 
+        and prove it holds the property "step" in "wf_iter" by lemma "iteration_step"
+   and finally using property Inv and TCon to prove the myhill-nerode theorem
+  
+*}
+lemma wf_iter [rule_format]: 
+  fixes f
+  assumes step: "\<And> e. \<lbrakk>P e; \<not> Q e\<rbrakk> \<Longrightarrow> (\<exists> e'. P e' \<and>  (f(e'), f(e)) \<in> less_than)"
+  shows pe:     "P e \<longrightarrow> (\<exists> e'. P e' \<and>  Q e')"
+proof(induct e rule: wf_induct 
+           [OF wf_inv_image[OF wf_less_than, where f = "f"]], clarify)
+  fix x 
+  assume h [rule_format]: 
+    "\<forall>y. (y, x) \<in> inv_image less_than f \<longrightarrow> P y \<longrightarrow> (\<exists>e'. P e' \<and> Q e')"
+    and px: "P x"
+  show "\<exists>e'. P e' \<and> Q e'"
+  proof(cases "Q x")
+    assume "Q x" with px show ?thesis by blast
+  next
+    assume nq: "\<not> Q x"
+    from step [OF px nq]
+    obtain e' where pe': "P e'" and ltf: "(f e', f x) \<in> less_than" by auto
+    show ?thesis
+    proof(rule h)
+      from ltf show "(e', x) \<in> inv_image less_than f" 
+	by (simp add:inv_image_def)
+    next
+      from pe' show "P e'" .
+    qed
+  qed
+qed
+
+
+text {* ******BEGIN: proving the initial equation-system satisfies Inv ****** *}
+
+lemma distinct_rhs_equations:
+  "(X, xrhs) \<in> equations (UNIV Quo Lang) \<Longrightarrow> distinct_rhs xrhs"
+by (auto simp: equations_def equation_rhs_def distinct_rhs_def empty_rhs_def dest:no_two_cls_inters)
+
+lemma every_nonempty_eqclass_has_strings:
+  "\<lbrakk>X \<in> (UNIV Quo Lang); X \<noteq> {[]}\<rbrakk> \<Longrightarrow> \<exists> clist. clist \<in> X \<and> clist \<noteq> []"
+by (auto simp:quot_def equiv_class_def equiv_str_def)
+
+lemma every_eqclass_is_derived_from_empty:
+  assumes not_empty: "X \<noteq> {[]}"
+  shows "X \<in> (UNIV Quo Lang) \<Longrightarrow> \<exists> clist. {[]};{clist} \<subseteq> X \<and> clist \<noteq> []"
+using not_empty
+apply (drule_tac every_nonempty_eqclass_has_strings, simp)
+by (auto intro:exI[where x = clist] simp:lang_seq_def)
+
+lemma equiv_str_in_CS:
+  "\<lbrakk>clist\<rbrakk>Lang \<in> (UNIV Quo Lang)"
+by (auto simp:quot_def)
+
+lemma has_str_imp_defined_by_str:
+  "\<lbrakk>str \<in> X; X \<in> UNIV Quo Lang\<rbrakk> \<Longrightarrow> X = \<lbrakk>str\<rbrakk>Lang"
+by (auto simp:quot_def equiv_class_def equiv_str_def)
+
+lemma every_eqclass_has_ascendent:
+  assumes has_str: "clist @ [c] \<in> X"
+  and     in_CS:   "X \<in> UNIV Quo Lang"
+  shows "\<exists> Y. Y \<in> UNIV Quo Lang \<and> Y-c\<rightarrow>X \<and> clist \<in> Y" (is "\<exists> Y. ?P Y")
+proof -
+  have "?P (\<lbrakk>clist\<rbrakk>Lang)" 
+  proof -
+    have "\<lbrakk>clist\<rbrakk>Lang \<in> UNIV Quo Lang"       
+      by (simp add:quot_def, rule_tac x = clist in exI, simp)
+    moreover have "\<lbrakk>clist\<rbrakk>Lang-c\<rightarrow>X" 
+    proof -
+      have "X = \<lbrakk>(clist @ [c])\<rbrakk>Lang" using has_str in_CS
+        by (auto intro!:has_str_imp_defined_by_str)
+      moreover have "\<forall> sl. sl \<in> \<lbrakk>clist\<rbrakk>Lang \<longrightarrow> sl @ [c] \<in> \<lbrakk>(clist @ [c])\<rbrakk>Lang"
+        by (auto simp:equiv_class_def equiv_str_def)
+      ultimately show ?thesis unfolding CT_def lang_seq_def
+        by auto
+    qed
+    moreover have "clist \<in> \<lbrakk>clist\<rbrakk>Lang" 
+      by (auto simp:equiv_str_def equiv_class_def)
+    ultimately show "?P (\<lbrakk>clist\<rbrakk>Lang)" by simp
+  qed
+  thus ?thesis by blast
+qed
+
+lemma finite_charset_rS:
+  "finite {CHAR c |c. Y-c\<rightarrow>X}"
+by (rule_tac A = UNIV and f = CHAR in finite_surj, auto)
+
+lemma l_eq_r_in_equations:
+  assumes X_in_equas: "(X, xrhs) \<in> equations (UNIV Quo Lang)"
+  shows "X = L xrhs"    
+proof (cases "X = {[]}")
+  case True
+  thus ?thesis using X_in_equas 
+    by (simp add:equations_def equation_rhs_def lang_seq_def)
+next
+  case False 
+  show ?thesis
+  proof 
+    show "X \<subseteq> L xrhs"
+    proof
+      fix x
+      assume "(1)": "x \<in> X"
+      show "x \<in> L xrhs"          
+      proof (cases "x = []")
+        assume empty: "x = []"
+        hence "x \<in> L (empty_rhs X)" using "(1)"
+          by (auto simp:empty_rhs_def lang_seq_def)
+        thus ?thesis using X_in_equas False empty "(1)" 
+          unfolding equations_def equation_rhs_def by auto
+      next
+        assume not_empty: "x \<noteq> []"
+        hence "\<exists> clist c. x = clist @ [c]" by (case_tac x rule:rev_cases, auto)
+        then obtain clist c where decom: "x = clist @ [c]" by blast
+        moreover have "\<And> Y. \<lbrakk>Y \<in> UNIV Quo Lang; Y-c\<rightarrow>X; clist \<in> Y\<rbrakk>
+          \<Longrightarrow> [c] \<in> L (folds ALT NULL {CHAR c |c. Y-c\<rightarrow>X})"
+        proof -
+          fix Y
+          assume Y_is_eq_cl: "Y \<in> UNIV Quo Lang"
+            and Y_CT_X: "Y-c\<rightarrow>X"
+            and clist_in_Y: "clist \<in> Y"
+          with finite_charset_rS 
+          show "[c] \<in> L (folds ALT NULL {CHAR c |c. Y-c\<rightarrow>X})"
+            by (auto simp :fold_alt_null_eqs)
+        qed
+        hence "\<exists>Xa. Xa \<in> UNIV Quo Lang \<and> clist @ [c] \<in> Xa ; L (folds ALT NULL {CHAR c |c. Xa-c\<rightarrow>X})" 
+          using X_in_equas False not_empty "(1)" decom
+          by (auto dest!:every_eqclass_has_ascendent simp:equations_def equation_rhs_def lang_seq_def)
+        then obtain Xa where 
+          "Xa \<in> UNIV Quo Lang \<and> clist @ [c] \<in> Xa ; L (folds ALT NULL {CHAR c |c. Xa-c\<rightarrow>X})" by blast
+        hence "x \<in> L {(S, folds ALT NULL {CHAR c |c. S-c\<rightarrow>X}) |S. S \<in> UNIV Quo Lang}" 
+          using X_in_equas "(1)" decom
+          by (auto simp add:equations_def equation_rhs_def intro!:exI[where x = Xa])
+        thus "x \<in> L xrhs" using X_in_equas False not_empty unfolding equations_def equation_rhs_def
+          by auto
+      qed
+    qed
+  next
+    show "L xrhs \<subseteq> X"
+    proof
+      fix x 
+      assume "(2)": "x \<in> L xrhs"
+      have "(2_1)": "\<And> s1 s2 r Xa. \<lbrakk>s1 \<in> Xa; s2 \<in> L (folds ALT NULL {CHAR c |c. Xa-c\<rightarrow>X})\<rbrakk> \<Longrightarrow> s1 @ s2 \<in> X"
+        using finite_charset_rS
+        by (auto simp:CT_def lang_seq_def fold_alt_null_eqs)
+      have "(2_2)": "\<And> s1 s2 Xa r.\<lbrakk>s1 \<in> Xa; s2 \<in> L r; (Xa, r) \<in> empty_rhs X\<rbrakk> \<Longrightarrow> s1 @ s2 \<in> X"
+        by (simp add:empty_rhs_def split:if_splits)
+      show "x \<in> X" using X_in_equas False "(2)"         
+        by (auto intro:"(2_1)" "(2_2)" simp:equations_def equation_rhs_def lang_seq_def)
+    qed
+  qed
+qed
+
+
+
+lemma no_EMPTY_equations:
+  "(X, xrhs) \<in> equations CS \<Longrightarrow> no_EMPTY_rhs xrhs"
+apply (clarsimp simp add:equations_def equation_rhs_def)
+apply (simp add:no_EMPTY_rhs_def empty_rhs_def, auto)
+apply (subgoal_tac "finite {CHAR c |c. Xa-c\<rightarrow>X}", drule_tac x = "[]" in fold_alt_null_eqs, clarsimp, rule finite_charset_rS)+
+done
+
+lemma init_ES_satisfy_ardenable:
+  "(X, xrhs) \<in> equations (UNIV Quo Lang)  \<Longrightarrow> ardenable (X, xrhs)"  
+  unfolding ardenable_def
+  by (auto intro:distinct_rhs_equations no_EMPTY_equations simp:l_eq_r_in_equations simp del:L_rhs.simps)
+
+lemma init_ES_satisfy_Inv:
+  assumes finite_CS: "finite (UNIV Quo Lang)"
+  and X_in_eq_cls: "X \<in> UNIV Quo Lang"
+  shows "Inv X (equations (UNIV Quo Lang))"
+proof -
+  have "finite (equations (UNIV Quo Lang))" using finite_CS
+    by (auto simp:equations_def)    
+  moreover have "\<exists>rhs. (X, rhs) \<in> equations (UNIV Quo Lang)" using X_in_eq_cls 
+    by (simp add:equations_def)
+  moreover have "distinct_equas (equations (UNIV Quo Lang))" 
+    by (auto simp:distinct_equas_def equations_def)
+  moreover have "\<forall>X xrhs. (X, xrhs) \<in> equations (UNIV Quo Lang) \<longrightarrow>
+                 rhs_eq_cls xrhs \<subseteq> insert {[]} (left_eq_cls (equations (UNIV Quo Lang)))" 
+    apply (simp add:left_eq_cls_def equations_def rhs_eq_cls_def equation_rhs_def)
+    by (auto simp:empty_rhs_def split:if_splits)
+  moreover have "\<forall>X xrhs. (X, xrhs) \<in> equations (UNIV Quo Lang) \<longrightarrow> X \<noteq> {}"
+    by (clarsimp simp:equations_def empty_notin_CS intro:classical)
+  moreover have "\<forall>X xrhs. (X, xrhs) \<in> equations (UNIV Quo Lang) \<longrightarrow> ardenable (X, xrhs)"
+    by (auto intro!:init_ES_satisfy_ardenable)
+  ultimately show ?thesis by (simp add:Inv_def)
+qed
+
+
+text {* *********** END: proving the initial equation-system satisfies Inv ******* *}
+
+
+text {* ****** BEGIN: proving every equation-system's iteration step satisfies Inv ***** *}
+
+lemma not_T_aux: "\<lbrakk>\<not> TCon (insert a A); x = a\<rbrakk>
+       \<Longrightarrow> \<exists>y. x \<noteq> y \<and> y \<in> insert a A "
+apply (case_tac "insert a A = {a}")
+by (auto simp:TCon_def)
+
+lemma not_T_atleast_2[rule_format]:
+  "finite S \<Longrightarrow> \<forall> x. x \<in> S \<and> (\<not> TCon S)\<longrightarrow> (\<exists> y. x \<noteq> y \<and> y \<in> S)"
+apply (erule finite.induct, simp)
+apply (clarify, case_tac "x = a")
+by (erule not_T_aux, auto)
+
+lemma exist_another_equa: 
+  "\<lbrakk>\<not> TCon ES; finite ES; distinct_equas ES; (X, rhl) \<in> ES\<rbrakk> \<Longrightarrow> \<exists> Y yrhl. (Y, yrhl) \<in> ES \<and> X \<noteq> Y"
+apply (drule not_T_atleast_2, simp)
+apply (clarsimp simp:distinct_equas_def)
+apply (drule_tac x= X in spec, drule_tac x = rhl in spec, drule_tac x = b in spec)
+by auto
+
+lemma Inv_mono_with_lambda:
+  assumes Inv_ES: "Inv X ES"
+  and X_noteq_Y:  "X \<noteq> {[]}"
+  shows "Inv X (ES - {({[]}, yrhs)})"
+proof -
+  have "finite (ES - {({[]}, yrhs)})" using Inv_ES
+    by (simp add:Inv_def)
+  moreover have "\<exists>rhs. (X, rhs) \<in> ES - {({[]}, yrhs)}" using Inv_ES X_noteq_Y
+    by (simp add:Inv_def)
+  moreover have "distinct_equas (ES - {({[]}, yrhs)})" using Inv_ES X_noteq_Y
+    apply (clarsimp simp:Inv_def distinct_equas_def)
+    by (drule_tac x = Xa in spec, simp)    
+  moreover have "\<forall>X xrhs.(X, xrhs) \<in> ES - {({[]}, yrhs)} \<longrightarrow>
+                          ardenable (X, xrhs) \<and> X \<noteq> {}" using Inv_ES
+    by (clarify, simp add:Inv_def)
+  moreover 
+  have "insert {[]} (left_eq_cls (ES - {({[]}, yrhs)})) = insert {[]} (left_eq_cls ES)"
+    by (auto simp:left_eq_cls_def)
+  hence "\<forall>X xrhs.(X, xrhs) \<in> ES - {({[]}, yrhs)} \<longrightarrow>
+                          rhs_eq_cls xrhs \<subseteq> insert {[]} (left_eq_cls (ES - {({[]}, yrhs)}))"
+    using Inv_ES by (auto simp:Inv_def)
+  ultimately show ?thesis by (simp add:Inv_def)
+qed
+
+lemma non_empty_card_prop:
+  "finite ES \<Longrightarrow> \<forall>e. e \<in> ES \<longrightarrow> card ES - Suc 0 < card ES"
+apply (erule finite.induct, simp)
+apply (case_tac[!] "a \<in> A")
+by (auto simp:insert_absorb)
+
+lemma ardenable_prop:
+  assumes not_lambda: "Y \<noteq> {[]}"
+  and ardable: "ardenable (Y, yrhs)"
+  shows "\<exists> yrhs'. Y = L yrhs' \<and> distinct_rhs yrhs' \<and> rhs_eq_cls yrhs' = rhs_eq_cls yrhs - {Y}" (is "\<exists> yrhs'. ?P yrhs'")
+proof (cases "(\<exists> reg. (Y, reg) \<in> yrhs)")
+  case True
+  thus ?thesis 
+  proof 
+    fix reg
+    assume self_contained: "(Y, reg) \<in> yrhs"
+    show ?thesis 
+    proof -
+      have "?P (arden_variate Y reg yrhs)"
+      proof -
+        have "Y = L (arden_variate Y reg yrhs)" 
+          using self_contained not_lambda ardable
+          by (rule_tac arden_variate_valid, simp_all add:ardenable_def)
+        moreover have "distinct_rhs (arden_variate Y reg yrhs)" 
+          using ardable 
+          by (auto simp:distinct_rhs_def arden_variate_def seq_rhs_r_def del_x_paired_def ardenable_def)
+        moreover have "rhs_eq_cls (arden_variate Y reg yrhs) = rhs_eq_cls yrhs - {Y}"
+        proof -
+          have "\<And> rhs r. rhs_eq_cls (seq_rhs_r rhs r) = rhs_eq_cls rhs"
+            apply (auto simp:rhs_eq_cls_def seq_rhs_r_def image_def)
+            by (rule_tac x = "SEQ ra r" in exI, rule_tac x = "(x, ra)" in bexI, simp+)
+          moreover have "\<And> rhs X. rhs_eq_cls (del_x_paired rhs X) = rhs_eq_cls rhs - {X}"
+            by (auto simp:rhs_eq_cls_def del_x_paired_def)
+          ultimately show ?thesis by (simp add:arden_variate_def)
+        qed
+        ultimately show ?thesis by simp
+      qed
+      thus ?thesis by (rule_tac x= "arden_variate Y reg yrhs" in exI, simp)
+    qed
+  qed
+next
+  case False
+  hence "(2)": "rhs_eq_cls yrhs - {Y} = rhs_eq_cls yrhs"
+    by (auto simp:rhs_eq_cls_def)
+  show ?thesis 
+  proof -
+    have "?P yrhs" using False ardable "(2)" 
+      by (simp add:ardenable_def)      
+    thus ?thesis by blast
+  qed
+qed
+
+lemma equas_subst_f_del_no_other:
+  assumes self_contained: "(Y, rhs) \<in> ES"
+  shows "\<exists> rhs'. (Y, rhs') \<in> (equas_subst_f X xrhs ` ES)" (is "\<exists> rhs'. ?P rhs'")
+proof -
+  have "\<exists> rhs'. equas_subst_f X xrhs (Y, rhs) = (Y, rhs')"
+    by (auto simp:equas_subst_f_def)
+  then obtain rhs' where "equas_subst_f X xrhs (Y, rhs) = (Y, rhs')" by blast
+  hence "?P rhs'" unfolding image_def using self_contained
+    by (auto intro:bexI[where x = "(Y, rhs)"])
+  thus ?thesis by blast
+qed
+
+lemma del_x_paired_del_only_x: 
+  "\<lbrakk>X \<noteq> Y; (X, rhs) \<in> ES\<rbrakk> \<Longrightarrow> (X, rhs) \<in> del_x_paired ES Y"
+by (auto simp:del_x_paired_def)
+
+lemma equas_subst_del_no_other:
+ "\<lbrakk>(X, rhs) \<in> ES; X \<noteq> Y\<rbrakk> \<Longrightarrow> (\<exists>rhs. (X, rhs) \<in> equas_subst ES Y rhs')"
+unfolding equas_subst_def
+apply (drule_tac X = Y and xrhs = rhs' in equas_subst_f_del_no_other)
+by (erule exE, drule del_x_paired_del_only_x, auto)
+
+lemma equas_subst_holds_distinct:
+  "distinct_equas ES \<Longrightarrow> distinct_equas (equas_subst ES Y rhs')"
+apply (clarsimp simp add:equas_subst_def distinct_equas_def del_x_paired_def equas_subst_f_def)
+by (auto split:if_splits)
+
+lemma del_x_paired_dels: 
+  "(X, rhs) \<in> ES \<Longrightarrow> {Y. Y \<in> ES \<and> fst Y = X} \<inter> ES \<noteq> {}"
+by (auto)
+
+lemma del_x_paired_subset:
+  "(X, rhs) \<in> ES \<Longrightarrow> ES - {Y. Y \<in> ES \<and> fst Y = X} \<subset> ES"
+apply (drule del_x_paired_dels)
+by auto
+
+lemma del_x_paired_card_less: 
+  "\<lbrakk>finite ES; (X, rhs) \<in> ES\<rbrakk> \<Longrightarrow> card (del_x_paired ES X) < card ES"
+apply (simp add:del_x_paired_def)
+apply (drule del_x_paired_subset)
+by (auto intro:psubset_card_mono)
+
+lemma equas_subst_card_less:
+  "\<lbrakk>finite ES; (Y, yrhs) \<in> ES\<rbrakk> \<Longrightarrow> card (equas_subst ES Y rhs') < card ES"
+apply (simp add:equas_subst_def)
+apply (frule_tac h = "equas_subst_f Y rhs'" in finite_imageI)
+apply (drule_tac f = "equas_subst_f Y rhs'" in Finite_Set.card_image_le)
+apply (drule_tac X = Y and xrhs = rhs' in equas_subst_f_del_no_other,erule exE)
+by (drule del_x_paired_card_less, auto)
+
+lemma equas_subst_holds_distinct_rhs:
+  assumes   dist': "distinct_rhs yrhs'"
+  and     history: "\<forall>X xrhs. (X, xrhs) \<in> ES \<longrightarrow> ardenable (X, xrhs)"
+  and     X_in :  "(X, xrhs) \<in> equas_subst ES Y yrhs'"
+  shows "distinct_rhs xrhs"
+using X_in history
+apply (clarsimp simp:equas_subst_def del_x_paired_def)
+apply (drule_tac x = a in spec, drule_tac x = b in spec)
+apply (simp add:ardenable_def equas_subst_f_def)
+by (auto intro:rhs_subst_holds_distinct_rhs simp:dist' split:if_splits)
+
+lemma r_no_EMPTY_imp_seq_rhs_r_no_EMPTY:
+  "[] \<notin> L r \<Longrightarrow> no_EMPTY_rhs (seq_rhs_r rhs r)"
+by (auto simp:no_EMPTY_rhs_def seq_rhs_r_def lang_seq_def)
+
+lemma del_x_paired_holds_no_EMPTY:
+  "no_EMPTY_rhs yrhs \<Longrightarrow> no_EMPTY_rhs (del_x_paired yrhs Y)"
+by (auto simp:no_EMPTY_rhs_def del_x_paired_def)
+
+lemma rhs_subst_holds_no_EMPTY:
+  "\<lbrakk>no_EMPTY_rhs yrhs; (Y, r) \<in> yrhs; Y \<noteq> {[]}\<rbrakk> \<Longrightarrow> no_EMPTY_rhs (rhs_subst yrhs Y rhs' r)"
+apply (auto simp:rhs_subst_def intro!:no_EMPTY_rhss_imp_merge_no_EMPTY r_no_EMPTY_imp_seq_rhs_r_no_EMPTY del_x_paired_holds_no_EMPTY)
+by (auto simp:no_EMPTY_rhs_def)
+
+lemma equas_subst_holds_no_EMPTY:
+  assumes substor: "Y \<noteq> {[]}"
+  and history: "\<forall>X xrhs. (X, xrhs) \<in> ES \<longrightarrow> ardenable (X, xrhs)"
+  and X_in:"(X, xrhs) \<in> equas_subst ES Y yrhs'"
+  shows "no_EMPTY_rhs xrhs"
+proof-
+  from X_in have "\<exists> (Z, zrhs) \<in> ES. (X, xrhs) = equas_subst_f Y yrhs' (Z, zrhs)"
+    by (auto simp add:equas_subst_def del_x_paired_def)
+  then obtain Z zrhs where Z_in: "(Z, zrhs) \<in> ES"
+                       and X_Z: "(X, xrhs) = equas_subst_f Y yrhs' (Z, zrhs)" by blast
+  hence dist_zrhs: "distinct_rhs zrhs" using history
+    by (auto simp:ardenable_def)
+  show ?thesis
+  proof (cases "\<exists> r. (Y, r) \<in> zrhs")
+    case True
+    then obtain r where Y_in_zrhs: "(Y, r) \<in> zrhs" ..
+    hence some: "(SOME r. (Y, r) \<in> zrhs) = r" using Z_in dist_zrhs
+      by (auto simp:distinct_rhs_def)
+    hence "no_EMPTY_rhs (rhs_subst zrhs Y yrhs' r)"
+      using substor Y_in_zrhs history Z_in
+      by (rule_tac rhs_subst_holds_no_EMPTY, auto simp:ardenable_def)
+    thus ?thesis using X_Z True some
+      by (simp add:equas_subst_def equas_subst_f_def)
+  next
+    case False
+    hence "(X, xrhs) = (Z, zrhs)" using Z_in X_Z
+      by (simp add:equas_subst_f_def)
+    thus ?thesis using history Z_in
+      by (auto simp:ardenable_def)
+  qed
+qed
+
+lemma equas_subst_f_holds_left_eq_right:
+  assumes substor: "Y = L rhs'"
+  and     history: "\<forall>X xrhs. (X, xrhs) \<in> ES \<longrightarrow> distinct_rhs xrhs \<and> X = L xrhs"
+  and       subst: "(X, xrhs) = equas_subst_f Y rhs' (Z, zrhs)"
+  and     self_contained: "(Z, zrhs) \<in> ES"
+  shows "X = L xrhs"
+proof (cases "\<exists> r. (Y, r) \<in> zrhs")
+  case True
+  from True obtain r where "(1)":"(Y, r) \<in> zrhs" ..
+  show ?thesis
+  proof -
+    from history self_contained
+    have dist: "distinct_rhs zrhs" by auto
+    hence "(SOME r. (Y, r) \<in> zrhs) = r" using self_contained "(1)"
+      using distinct_rhs_def by (auto intro:some_equality)
+    moreover have "L zrhs = L (rhs_subst zrhs Y rhs' r)" using substor dist "(1)" self_contained
+      by (rule_tac rhs_subst_prop1, auto simp:distinct_equas_rhs_def)
+    ultimately show ?thesis using subst history self_contained
+      by (auto simp:equas_subst_f_def split:if_splits)
+  qed
+next
+  case False
+  thus ?thesis using history subst self_contained
+    by (auto simp:equas_subst_f_def)
+qed
+
+lemma equas_subst_holds_left_eq_right:
+  assumes history: "\<forall>X xrhs. (X, xrhs) \<in> ES \<longrightarrow> ardenable (X, xrhs)"
+  and     substor: "Y = L rhs'"
+  and     X_in :  "(X, xrhs) \<in> equas_subst ES Y yrhs'"
+  shows "\<forall>X xrhs. (X, xrhs) \<in> equas_subst ES Y rhs' \<longrightarrow> X = L xrhs"
+apply (clarsimp simp add:equas_subst_def del_x_paired_def)
+using substor
+apply (drule_tac equas_subst_f_holds_left_eq_right)
+using history
+by (auto simp:ardenable_def)
+
+lemma equas_subst_holds_ardenable:
+  assumes substor: "Y = L yrhs'"
+  and history: "\<forall>X xrhs. (X, xrhs) \<in> ES \<longrightarrow> ardenable (X, xrhs)"
+  and X_in:"(X, xrhs) \<in> equas_subst ES Y yrhs'"
+  and dist': "distinct_rhs yrhs'"
+  and not_lambda: "Y \<noteq> {[]}"
+  shows "ardenable (X, xrhs)"
+proof -
+  have "distinct_rhs xrhs" using history X_in dist' 
+    by (auto dest:equas_subst_holds_distinct_rhs)
+  moreover have "no_EMPTY_rhs xrhs" using history X_in not_lambda
+    by (auto intro:equas_subst_holds_no_EMPTY)
+  moreover have "X = L xrhs" using history substor X_in
+  by (auto dest: equas_subst_holds_left_eq_right simp del:L_rhs.simps)
+  ultimately show ?thesis using ardenable_def by simp
+qed
+
+lemma equas_subst_holds_cls_defined:
+  assumes         X_in: "(X, xrhs) \<in> equas_subst ES Y yrhs'"
+  and           Inv_ES: "Inv X' ES"
+  and            subst: "(Y, yrhs) \<in> ES"
+  and  cls_holds_but_Y: "rhs_eq_cls yrhs' = rhs_eq_cls yrhs - {Y}"
+  shows "rhs_eq_cls xrhs \<subseteq> insert {[]} (left_eq_cls (equas_subst ES Y yrhs'))"
+proof-
+  have tac: "\<lbrakk> A \<subseteq> B; C \<subseteq> D; E \<subseteq> A \<union> B\<rbrakk> \<Longrightarrow> E \<subseteq> B \<union> D" by auto
+  from X_in have "\<exists> (Z, zrhs) \<in> ES. (X, xrhs) = equas_subst_f Y yrhs' (Z, zrhs)"
+    by (auto simp add:equas_subst_def del_x_paired_def)
+  then obtain Z zrhs where Z_in: "(Z, zrhs) \<in> ES"
+                       and X_Z: "(X, xrhs) = equas_subst_f Y yrhs' (Z, zrhs)" by blast
+  hence "rhs_eq_cls zrhs \<subseteq> insert {[]} (left_eq_cls ES)" using Inv_ES
+    by (auto simp:Inv_def)
+  moreover have "rhs_eq_cls yrhs' \<subseteq> insert {[]} (left_eq_cls ES) - {Y}" 
+    using Inv_ES subst cls_holds_but_Y
+    by (auto simp:Inv_def)
+  moreover have "rhs_eq_cls xrhs \<subseteq> rhs_eq_cls zrhs \<union> rhs_eq_cls yrhs' - {Y}"
+    using X_Z cls_holds_but_Y
+    apply (clarsimp simp add:equas_subst_f_def rhs_subst_def split:if_splits)
+    by (auto simp:rhs_eq_cls_def merge_rhs_def del_x_paired_def seq_rhs_r_def)
+  moreover have "left_eq_cls (equas_subst ES Y yrhs') = left_eq_cls ES - {Y}" using subst
+    by (auto simp: left_eq_cls_def equas_subst_def del_x_paired_def equas_subst_f_def
+             dest: equas_subst_f_del_no_other
+             split: if_splits)
+  ultimately show ?thesis by blast
+qed
+
+lemma iteration_step: 
+  assumes Inv_ES: "Inv X ES"
+  and    not_T: "\<not> TCon ES"
+  shows "(\<exists> ES'. Inv X ES' \<and> (card ES', card ES) \<in> less_than)" 
+proof -
+  from Inv_ES not_T have another: "\<exists>Y yrhs. (Y, yrhs) \<in> ES \<and> X \<noteq> Y" unfolding Inv_def
+    by (clarify, rule_tac exist_another_equa[where X = X], auto)
+  then obtain Y yrhs where subst: "(Y, yrhs) \<in> ES" and not_X: " X \<noteq> Y" by blast
+  show ?thesis (is "\<exists> ES'. ?P ES'")
+  proof (cases "Y = {[]}") 
+    case True
+      --"in this situation, we pick a \"\<lambda>\" equation, thus directly remove it from the equation-system"
+    have "?P (ES - {(Y, yrhs)})" 
+    proof 
+      show "Inv X (ES - {(Y, yrhs)})" using True not_X
+        by (simp add:Inv_ES Inv_mono_with_lambda)
+    next 
+      show "(card (ES - {(Y, yrhs)}), card ES) \<in> less_than" using Inv_ES subst
+        by (auto elim:non_empty_card_prop[rule_format] simp:Inv_def)
+    qed
+    thus ?thesis by blast
+  next
+    case False
+      --"in this situation, we pick a equation and using ardenable to get a 
+        rhs without itself in it, then use equas_subst to form a new equation-system"
+    hence "\<exists> yrhs'. Y = L yrhs' \<and> distinct_rhs yrhs' \<and> rhs_eq_cls yrhs' = rhs_eq_cls yrhs - {Y}" 
+      using subst Inv_ES
+      by (auto intro:ardenable_prop simp add:Inv_def simp del:L_rhs.simps)
+    then obtain yrhs' where Y'_l_eq_r: "Y = L yrhs'"
+      and dist_Y': "distinct_rhs yrhs'"
+      and cls_holds_but_Y: "rhs_eq_cls yrhs' = rhs_eq_cls yrhs - {Y}" by blast
+    hence "?P (equas_subst ES Y yrhs')"
+    proof -
+      have finite_del: "\<And> S x. finite S \<Longrightarrow> finite (del_x_paired S x)" 
+        apply (rule_tac A = "del_x_paired S x" in finite_subset)
+        by (auto simp:del_x_paired_def)
+      have "finite (equas_subst ES Y yrhs')" using Inv_ES 
+        by (auto intro!:finite_del simp:equas_subst_def Inv_def)
+      moreover have "\<exists>rhs. (X, rhs) \<in> equas_subst ES Y yrhs'" using Inv_ES not_X
+        by (auto intro:equas_subst_del_no_other simp:Inv_def)
+      moreover have "distinct_equas (equas_subst ES Y yrhs')" using Inv_ES
+        by (auto intro:equas_subst_holds_distinct simp:Inv_def)
+      moreover have "\<forall>X xrhs. (X, xrhs) \<in> equas_subst ES Y yrhs' \<longrightarrow> ardenable (X, xrhs)"
+        using Inv_ES dist_Y' False Y'_l_eq_r
+        apply (clarsimp simp:Inv_def)
+        by (rule equas_subst_holds_ardenable, simp_all)
+      moreover have "\<forall>X xrhs. (X, xrhs) \<in> equas_subst ES Y yrhs' \<longrightarrow> X \<noteq> {}" using Inv_ES
+        by (clarsimp simp:equas_subst_def Inv_def del_x_paired_def equas_subst_f_def split:if_splits, auto)
+      moreover have "\<forall>X xrhs. (X, xrhs) \<in> equas_subst ES Y yrhs' \<longrightarrow>
+                               rhs_eq_cls xrhs \<subseteq> insert {[]} (left_eq_cls (equas_subst ES Y yrhs'))"
+        using Inv_ES subst cls_holds_but_Y
+        apply (rule_tac impI | rule_tac allI)+
+        by (erule equas_subst_holds_cls_defined, auto)
+      moreover have "(card (equas_subst ES Y yrhs'), card ES) \<in> less_than"using Inv_ES subst
+        by (simp add:equas_subst_card_less Inv_def)
+      ultimately show "?P (equas_subst ES Y yrhs')" by (auto simp:Inv_def)      
+    qed
+    thus ?thesis by blast
+  qed
+qed
+
+text {* ***** END: proving every equation-system's iteration step satisfies Inv ************** *}
+
+lemma iteration_conc: 
+  assumes history: "Inv X ES"
+  shows "\<exists> ES'. Inv X ES' \<and> TCon ES'" (is "\<exists> ES'. ?P ES'")
+proof (cases "TCon ES")
+  case True
+  hence "?P ES" using history by simp
+  thus ?thesis by blast
+next
+  case False  
+  thus ?thesis using history iteration_step
+    by (rule_tac f = card in wf_iter, simp_all)
+qed
+
+lemma eqset_imp_iff': "A = B \<Longrightarrow> \<forall> x. x \<in> A \<longleftrightarrow> x \<in> B"
+apply (auto simp:mem_def)
+done
+
+lemma set_cases2:
+  "\<lbrakk>(A = {} \<Longrightarrow> R A); \<And> x. (A = {x}) \<Longrightarrow> R A; \<And> x y. \<lbrakk>x \<noteq> y; x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> R A\<rbrakk> \<Longrightarrow> R A"
+apply (case_tac "A = {}", simp)
+by (case_tac "\<exists> x. A = {x}", clarsimp, blast)
+
+lemma rhs_aux:"\<lbrakk>distinct_rhs rhs; {Y. \<exists>r. (Y, r) \<in> rhs} = {X}\<rbrakk> \<Longrightarrow> (\<exists>r. rhs = {(X, r)})"
+apply (rule_tac A = rhs in set_cases2, simp)
+apply (drule_tac x = X in eqset_imp_iff, clarsimp)
+apply (drule eqset_imp_iff',clarsimp)
+apply (frule_tac x = a in spec, drule_tac x = aa in spec)
+by (auto simp:distinct_rhs_def)
+
+lemma every_eqcl_has_reg: 
+  assumes finite_CS: "finite (UNIV Quo Lang)"
+  and X_in_CS: "X \<in> (UNIV Quo Lang)"
+  shows "\<exists> (reg::rexp). L reg = X" (is "\<exists> r. ?E r")
+proof-
+  have "\<exists>ES'. Inv X ES' \<and> TCon ES'" using finite_CS X_in_CS
+    by (auto intro:init_ES_satisfy_Inv iteration_conc)
+  then obtain ES' where Inv_ES': "Inv X ES'" 
+                   and  TCon_ES': "TCon ES'" by blast
+  from Inv_ES' TCon_ES' 
+  have "\<exists> rhs. ES' = {(X, rhs)}"
+    apply (clarsimp simp:Inv_def TCon_def)
+    apply (rule_tac x = rhs in exI)
+    by (auto dest!:card_Suc_Diff1 simp:card_eq_0_iff)  
+  then obtain rhs where ES'_single_equa: "ES' = {(X, rhs)}" ..
+  hence X_ardenable: "ardenable (X, rhs)" using Inv_ES'
+    by (simp add:Inv_def)
+  
+  from X_ardenable have X_l_eq_r: "X = L rhs"
+    by (simp add:ardenable_def)
+  hence rhs_not_empty: "rhs \<noteq> {}" using Inv_ES' ES'_single_equa
+    by (auto simp:Inv_def ardenable_def)
+  have rhs_eq_cls: "rhs_eq_cls rhs \<subseteq> {X, {[]}}"
+    using Inv_ES' ES'_single_equa
+    by (auto simp:Inv_def ardenable_def left_eq_cls_def)
+  have X_not_empty: "X \<noteq> {}" using Inv_ES' ES'_single_equa
+    by (auto simp:Inv_def)    
+  show ?thesis
+  proof (cases "X = {[]}")
+    case True
+    hence "?E EMPTY" by (simp)
+    thus ?thesis by blast
+  next
+    case False with  X_ardenable
+    have "\<exists> rhs'. X = L rhs' \<and> rhs_eq_cls rhs' = rhs_eq_cls rhs - {X} \<and> distinct_rhs rhs'"
+      by (drule_tac ardenable_prop, auto)
+    then obtain rhs' where X_eq_rhs': "X = L rhs'"
+      and rhs'_eq_cls: "rhs_eq_cls rhs' = rhs_eq_cls rhs - {X}" 
+      and rhs'_dist : "distinct_rhs rhs'" by blast
+    have "rhs_eq_cls rhs' \<subseteq> {{[]}}" using rhs_eq_cls False rhs'_eq_cls rhs_not_empty 
+      by blast
+    hence "rhs_eq_cls rhs' = {{[]}}" using X_not_empty X_eq_rhs'
+      by (auto simp:rhs_eq_cls_def)
+    hence "\<exists> r. rhs' = {({[]}, r)}" using rhs'_dist
+      by (auto intro:rhs_aux simp:rhs_eq_cls_def)
+    then obtain r where "rhs' = {({[]}, r)}" ..
+    hence "?E r" using X_eq_rhs' by (auto simp add:lang_seq_def)
+    thus ?thesis by blast     
+  qed
+qed
+
+text {* definition of a regular language *}
+
+abbreviation
+  reg :: "string set => bool"
+where
+  "reg L1 \<equiv> (\<exists>r::rexp. L r = L1)"
+
+theorem myhill_nerode: 
+  assumes finite_CS: "finite (UNIV Quo Lang)"
+  shows   "\<exists> (reg::rexp). Lang = L reg" (is "\<exists> r. ?P r")
+proof -
+  have has_r_each: "\<forall>C\<in>{X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}. \<exists>(r::rexp). C = L r" using finite_CS
+    by (auto dest:every_eqcl_has_reg)  
+  have "\<exists> (rS::rexp set). finite rS \<and> 
+                          (\<forall> C \<in> {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}. \<exists> r \<in> rS. C = L r) \<and> 
+                          (\<forall> r \<in> rS. \<exists> C \<in> {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}. C = L r)" 
+       (is "\<exists> rS. ?Q rS")
+  proof-
+    have "\<And> C. C \<in> {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang} \<Longrightarrow> C = L (SOME (ra::rexp). C = L ra)"
+      using has_r_each
+      apply (erule_tac x = C in ballE, erule_tac exE)
+      by (rule_tac a = r in someI2, simp+)
+    hence "?Q ((\<lambda> C. SOME r. C = L r) ` {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang})" using has_r_each
+      using finite_CS by auto
+    thus ?thesis by blast    
+  qed
+  then obtain rS where finite_rS : "finite rS"
+    and has_r_each': "\<forall> C \<in> {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}. \<exists> r \<in> (rS::rexp set). C = L r"
+    and has_cl_each: "\<forall> r \<in> (rS::rexp set). \<exists> C \<in> {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}. C = L r" by blast
+  have "?P (folds ALT NULL rS)"
+  proof
+    show "Lang \<subseteq> L (folds ALT NULL rS)" using finite_rS lang_eqs_union_of_eqcls[of Lang] has_r_each'
+      apply (clarsimp simp:fold_alt_null_eqs) by blast
+  next 
+    show "L (folds ALT NULL rS) \<subseteq> Lang" using finite_rS lang_eqs_union_of_eqcls[of Lang] has_cl_each
+      by (clarsimp simp:fold_alt_null_eqs)
+  qed
+  thus ?thesis by blast
+qed 
+
+
+text {* tests by Christian *}
+
+(* Alternative definition for Quo *)
+definition 
+  QUOT :: "string set \<Rightarrow> (string set) set"  
+where
+  "QUOT Lang \<equiv> (\<Union>x. {\<lbrakk>x\<rbrakk>Lang})"
+
+lemma test: 
+  "UNIV Quo Lang = QUOT Lang"
+by (auto simp add: quot_def QUOT_def)
+
+lemma test1:
+  assumes finite_CS: "finite (QUOT Lang)"
+  shows "reg Lang"
+using finite_CS
+unfolding test[symmetric]
+by (auto dest: myhill_nerode)
+
+lemma cons_one: "x @ y \<in> {z} \<Longrightarrow> x @ y = z"
+by simp
+
+lemma quot_lambda: "QUOT {[]} = {{[]}, UNIV - {[]}}"
+proof 
+  show "QUOT {[]} \<subseteq> {{[]}, UNIV - {[]}}"
+  proof 
+    fix x 
+    assume in_quot: "x \<in> QUOT {[]}"
+    show "x \<in> {{[]}, UNIV - {[]}}"
+    proof -
+      from in_quot 
+      have "x = {[]} \<or> x = UNIV - {[]}"
+        unfolding QUOT_def equiv_class_def
+      proof 
+        fix xa
+        assume in_univ: "xa \<in> UNIV"
+           and in_eqiv: "x \<in> {{y. xa \<equiv>{[]}\<equiv> y}}"
+        show "x = {[]} \<or> x = UNIV - {[]}"
+        proof(cases "xa = []")
+          case True
+          hence "{y. xa \<equiv>{[]}\<equiv> y} = {[]}" using in_eqiv
+            by (auto simp add:equiv_str_def)
+          thus ?thesis using in_eqiv by (rule_tac disjI1, simp)
+        next
+          case False
+          hence "{y. xa \<equiv>{[]}\<equiv> y} = UNIV - {[]}" using in_eqiv
+            by (auto simp:equiv_str_def)
+          thus ?thesis using in_eqiv by (rule_tac disjI2, simp)
+        qed
+      qed
+      thus ?thesis by simp
+    qed
+  qed
+next
+  show "{{[]}, UNIV - {[]}} \<subseteq> QUOT {[]}"
+  proof
+    fix x
+    assume in_res: "x \<in> {{[]}, (UNIV::string set) - {[]}}"
+    show "x \<in> (QUOT {[]})"
+    proof -
+      have "x = {[]} \<Longrightarrow> x \<in> QUOT {[]}"
+        apply (simp add:QUOT_def equiv_class_def equiv_str_def)
+        by (rule_tac x = "[]" in exI, auto)
+      moreover have "x = UNIV - {[]} \<Longrightarrow> x \<in> QUOT {[]}"
+        apply (simp add:QUOT_def equiv_class_def equiv_str_def)
+        by (rule_tac x = "''a''" in exI, auto)
+      ultimately show ?thesis using in_res by blast
+    qed
+  qed
+qed
+
+lemma quot_single_aux: "\<lbrakk>x \<noteq> []; x \<noteq> [c]\<rbrakk> \<Longrightarrow> x @ z \<noteq> [c]"
+by (induct x, auto)
+
+lemma quot_single: "\<And> (c::char). QUOT {[c]} = {{[]}, {[c]}, UNIV - {[], [c]}}"
+proof - 
+  fix c::"char" 
+  have exist_another: "\<exists> a. a \<noteq> c" 
+    apply (case_tac "c = CHR ''a''")
+    apply (rule_tac x = "CHR ''b''" in exI, simp)
+    by (rule_tac x = "CHR ''a''" in exI, simp)  
+  show "QUOT {[c]} = {{[]}, {[c]}, UNIV - {[], [c]}}"
+  proof
+    show "QUOT {[c]} \<subseteq> {{[]},{[c]}, UNIV - {[], [c]}}"
+    proof 
+      fix x 
+      assume in_quot: "x \<in> QUOT {[c]}"
+      show "x \<in> {{[]}, {[c]}, UNIV - {[],[c]}}"
+      proof -
+        from in_quot 
+        have "x = {[]} \<or> x = {[c]} \<or> x = UNIV - {[],[c]}"
+          unfolding QUOT_def equiv_class_def
+        proof 
+          fix xa
+          assume in_eqiv: "x \<in> {{y. xa \<equiv>{[c]}\<equiv> y}}"
+          show "x = {[]} \<or> x = {[c]} \<or> x = UNIV - {[], [c]}"
+          proof-
+            have "xa = [] \<Longrightarrow> x = {[]}" using in_eqiv 
+              by (auto simp add:equiv_str_def)
+            moreover have "xa = [c] \<Longrightarrow> x = {[c]}"
+            proof -
+              have "xa = [c] \<Longrightarrow> {y. xa \<equiv>{[c]}\<equiv> y} = {[c]}" using in_eqiv
+                apply (simp add:equiv_str_def)
+                apply (rule set_ext, rule iffI, simp)
+                apply (drule_tac x = "[]" in spec, auto)
+                done   
+              thus "xa = [c] \<Longrightarrow> x = {[c]}" using in_eqiv by simp 
+            qed
+            moreover have "\<lbrakk>xa \<noteq> []; xa \<noteq> [c]\<rbrakk> \<Longrightarrow> x = UNIV - {[],[c]}"
+            proof -
+              have "\<lbrakk>xa \<noteq> []; xa \<noteq> [c]\<rbrakk> \<Longrightarrow> {y. xa \<equiv>{[c]}\<equiv> y} = UNIV - {[],[c]}" 
+                apply (clarsimp simp add:equiv_str_def)
+                apply (rule set_ext, rule iffI, simp)
+                apply (rule conjI)
+                apply (drule_tac x = "[c]" in spec, simp)
+                apply (drule_tac x = "[]" in spec, simp)
+                by (auto dest:quot_single_aux)
+              thus "\<lbrakk>xa \<noteq> []; xa \<noteq> [c]\<rbrakk> \<Longrightarrow> x = UNIV - {[],[c]}" using in_eqiv by simp
+            qed
+            ultimately show ?thesis by blast
+          qed
+        qed
+        thus ?thesis by simp
+      qed
+    qed
+  next
+    show "{{[]}, {[c]}, UNIV - {[],[c]}} \<subseteq> QUOT {[c]}"
+    proof
+      fix x
+      assume in_res: "x \<in> {{[]},{[c]}, (UNIV::string set) - {[],[c]}}"
+      show "x \<in> (QUOT {[c]})"
+      proof -
+        have "x = {[]} \<Longrightarrow> x \<in> QUOT {[c]}"
+          apply (simp add:QUOT_def equiv_class_def equiv_str_def)
+          by (rule_tac x = "[]" in exI, auto)
+        moreover have "x = {[c]} \<Longrightarrow> x \<in> QUOT {[c]}"
+          apply (simp add:QUOT_def equiv_class_def equiv_str_def)
+          apply (rule_tac x = "[c]" in exI, simp)
+          apply (rule set_ext, rule iffI, simp+)
+          by (drule_tac x = "[]" in spec, simp)
+        moreover have "x = UNIV - {[],[c]} \<Longrightarrow> x \<in> QUOT {[c]}"
+          using exist_another
+          apply (clarsimp simp add:QUOT_def equiv_class_def equiv_str_def)        
+          apply (rule_tac x = "[a]" in exI, simp)
+          apply (rule set_ext, rule iffI, simp)
+          apply (clarsimp simp:quot_single_aux, simp)
+          apply (rule conjI)
+          apply (drule_tac x = "[c]" in spec, simp)
+          by (drule_tac x = "[]" in spec, simp)     
+        ultimately show ?thesis using in_res by blast
+      qed
+    qed
+  qed
+qed
+
+lemma eq_class_imp_eq_str:
+  "\<lbrakk>x\<rbrakk>lang = \<lbrakk>y\<rbrakk>lang \<Longrightarrow> x \<equiv>lang\<equiv> y"
+by (auto simp:equiv_class_def equiv_str_def)
+
+lemma finite_tag_image: 
+  "finite (range tag) \<Longrightarrow> finite (((op `) tag) ` S)"
+apply (rule_tac B = "Pow (tag ` UNIV)" in finite_subset)
+by (auto simp add:image_def Pow_def)
+
+lemma str_inj_imps:
+  assumes str_inj: "\<And> m n. tag m = tag (n::string) \<Longrightarrow> m \<equiv>lang\<equiv> n"
+  shows "inj_on ((op `) tag) (QUOT lang)"
+proof (clarsimp simp add:inj_on_def QUOT_def)
+  fix x y
+  assume eq_tag: "tag ` \<lbrakk>x\<rbrakk>lang = tag ` \<lbrakk>y\<rbrakk>lang"
+  show "\<lbrakk>x\<rbrakk>lang = \<lbrakk>y\<rbrakk>lang"
+  proof -
+    have aux1:"\<And>a b. a \<in> (\<lbrakk>b\<rbrakk>lang) \<Longrightarrow> (a \<equiv>lang\<equiv> b)"
+      by (simp add:equiv_class_def equiv_str_def)
+    have aux2: "\<And> A B f. \<lbrakk>f ` A = f ` B; A \<noteq> {}\<rbrakk> \<Longrightarrow> \<exists> a b. a \<in> A \<and> b \<in> B \<and> f a = f b"
+      by auto
+    have aux3: "\<And> a l. \<lbrakk>a\<rbrakk>l \<noteq> {}" 
+      by (auto simp:equiv_class_def equiv_str_def)
+    show ?thesis using eq_tag
+      apply (drule_tac aux2, simp add:aux3, clarsimp)
+      apply (drule_tac str_inj, (drule_tac aux1)+)
+      by (simp add:equiv_str_def equiv_class_def)
+  qed
+qed
+
+definition tag_str_ALT :: "string set \<Rightarrow> string set \<Rightarrow> string \<Rightarrow> (string set \<times> string set)"
+where
+  "tag_str_ALT L\<^isub>1 L\<^isub>2 x \<equiv> (\<lbrakk>x\<rbrakk>L\<^isub>1, \<lbrakk>x\<rbrakk>L\<^isub>2)"
+
+lemma tag_str_alt_range_finite:
+  assumes finite1: "finite (QUOT L\<^isub>1)"
+  and finite2: "finite (QUOT L\<^isub>2)"
+  shows "finite (range (tag_str_ALT L\<^isub>1 L\<^isub>2))"
+proof -
+  have "{y. \<exists>x. y = (\<lbrakk>x\<rbrakk>L\<^isub>1, \<lbrakk>x\<rbrakk>L\<^isub>2)} \<subseteq> (QUOT L\<^isub>1) \<times> (QUOT L\<^isub>2)"
+    by (auto simp:QUOT_def)
+  thus ?thesis using finite1 finite2
+    by (auto simp: image_def tag_str_ALT_def UNION_def 
+            intro: finite_subset[where B = "(QUOT L\<^isub>1) \<times> (QUOT L\<^isub>2)"])
+qed
+
+lemma tag_str_alt_inj:
+  "tag_str_ALT L\<^isub>1 L\<^isub>2 x = tag_str_ALT L\<^isub>1 L\<^isub>2 y \<Longrightarrow> x \<equiv>(L\<^isub>1 \<union> L\<^isub>2)\<equiv> y"
+apply (simp add:tag_str_ALT_def equiv_class_def equiv_str_def)
+by blast
+  
+lemma quot_alt:
+  assumes finite1: "finite (QUOT L\<^isub>1)"
+  and finite2: "finite (QUOT L\<^isub>2)"
+  shows "finite (QUOT (L\<^isub>1 \<union> L\<^isub>2))"
+proof (rule_tac f = "(op `) (tag_str_ALT L\<^isub>1 L\<^isub>2)" in finite_imageD)
+  show "finite (op ` (tag_str_ALT L\<^isub>1 L\<^isub>2) ` QUOT (L\<^isub>1 \<union> L\<^isub>2))"
+    using finite_tag_image tag_str_alt_range_finite finite1 finite2
+    by auto
+next
+  show "inj_on (op ` (tag_str_ALT L\<^isub>1 L\<^isub>2)) (QUOT (L\<^isub>1 \<union> L\<^isub>2))"
+    apply (rule_tac str_inj_imps)
+    by (erule_tac tag_str_alt_inj)
+qed
+
+(* list_diff:: list substract, once different return tailer *)
+fun list_diff :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infix "-" 51)
+where
+  "list_diff []  xs = []" |
+  "list_diff (x#xs) [] = x#xs" |
+  "list_diff (x#xs) (y#ys) = (if x = y then list_diff xs ys else (x#xs))"
+
+lemma [simp]: "(x @ y) - x = y"
+apply (induct x)
+by (case_tac y, simp+)
+
+lemma [simp]: "x - x = []"
+by (induct x, auto)
+
+lemma [simp]: "x = xa @ y \<Longrightarrow> x - xa = y "
+by (induct x, auto)
+
+lemma [simp]: "x - [] = x"
+by (induct x, auto)
+
+lemma [simp]: "xa \<le> x \<Longrightarrow> (x @ y) - xa = (x - xa) @ y"
+by (auto elim:prefixE)
+
+definition tag_str_SEQ:: "string set \<Rightarrow> string set \<Rightarrow> string \<Rightarrow> (string set \<times> string set set)"
+where
+  "tag_str_SEQ L\<^isub>1 L\<^isub>2 x \<equiv> if (\<exists> xa \<le> x. xa \<in> L\<^isub>1)
+                         then (\<lbrakk>x\<rbrakk>L\<^isub>1, {\<lbrakk>(x - xa)\<rbrakk>L\<^isub>2 | xa.  xa \<le> x \<and> xa \<in> L\<^isub>1})
+                         else (\<lbrakk>x\<rbrakk>L\<^isub>1, {})"
+
+lemma tag_seq_eq_E:
+  "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y \<Longrightarrow>
+   ((\<exists> xa \<le> x. xa \<in> L\<^isub>1) \<and> \<lbrakk>x\<rbrakk>L\<^isub>1 = \<lbrakk>y\<rbrakk>L\<^isub>1 \<and> 
+    {\<lbrakk>(x - xa)\<rbrakk>L\<^isub>2 | xa. xa \<le> x \<and> xa \<in> L\<^isub>1} = {\<lbrakk>(y - ya)\<rbrakk>L\<^isub>2 | ya. ya \<le> y \<and> ya \<in> L\<^isub>1} ) \<or>
+   ((\<forall> xa \<le> x. xa \<notin> L\<^isub>1) \<and> \<lbrakk>x\<rbrakk>L\<^isub>1 = \<lbrakk>y\<rbrakk>L\<^isub>1)"
+by (simp add:tag_str_SEQ_def split:if_splits, blast)
+
+lemma tag_str_seq_range_finite:
+  assumes finite1: "finite (QUOT L\<^isub>1)"
+  and finite2: "finite (QUOT L\<^isub>2)"
+  shows "finite (range (tag_str_SEQ L\<^isub>1 L\<^isub>2))"
+proof -
+  have "(range (tag_str_SEQ L\<^isub>1 L\<^isub>2)) \<subseteq> (QUOT L\<^isub>1) \<times> (Pow (QUOT L\<^isub>2))"
+    by (auto simp:image_def tag_str_SEQ_def QUOT_def)
+  thus ?thesis using finite1 finite2 
+    by (rule_tac B = "(QUOT L\<^isub>1) \<times> (Pow (QUOT L\<^isub>2))" in finite_subset, auto)
+qed
+  
+lemma app_in_seq_decom[rule_format]:
+  "\<forall> x. x @ z \<in> L\<^isub>1 ; L\<^isub>2 \<longrightarrow> (\<exists> xa \<le> x. xa \<in> L\<^isub>1 \<and> (x - xa) @ z \<in> L\<^isub>2) \<or> 
+                            (\<exists> za \<le> z. (x @ za) \<in> L\<^isub>1 \<and> (z - za) \<in> L\<^isub>2)"
+apply (induct z)
+apply (simp, rule allI, rule impI, rule disjI1)
+apply (clarsimp simp add:lang_seq_def)
+apply (rule_tac x = s1 in exI, simp)
+apply (rule allI | rule impI)+
+apply (drule_tac x = "x @ [a]" in spec, simp)
+apply (erule exE | erule conjE | erule disjE)+
+apply (rule disjI2, rule_tac x = "[a]" in exI, simp)
+apply (rule disjI1, rule_tac x = xa in exI, simp) 
+apply (erule exE | erule conjE)+
+apply (rule disjI2, rule_tac x = "a # za" in exI, simp)
+done
+
+lemma tag_str_seq_inj:
+  assumes tag_eq: "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y"
+  shows "(x::string) \<equiv>(L\<^isub>1 ; L\<^isub>2)\<equiv> y"
+proof -
+  have aux: "\<And> x y z. \<lbrakk>tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y; x @ z \<in> L\<^isub>1 ; L\<^isub>2\<rbrakk> 
+                       \<Longrightarrow> y @ z \<in> L\<^isub>1 ; L\<^isub>2"
+  proof (drule app_in_seq_decom, erule disjE)
+    fix x y z
+    assume tag_eq: "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y"
+      and x_gets_l2: "\<exists>xa\<le>x. xa \<in> L\<^isub>1 \<and> (x - xa) @ z \<in> L\<^isub>2"
+    from x_gets_l2 have "\<exists> xa \<le> x. xa \<in> L\<^isub>1" by blast
+    hence xy_l2:"{\<lbrakk>(x - xa)\<rbrakk>L\<^isub>2 | xa. xa \<le> x \<and> xa \<in> L\<^isub>1} = {\<lbrakk>(y - ya)\<rbrakk>L\<^isub>2 | ya. ya \<le> y \<and> ya \<in> L\<^isub>1}"
+      using tag_eq tag_seq_eq_E by blast
+    from x_gets_l2 obtain xa where xa_le_x: "xa \<le> x"
+                               and xa_in_l1: "xa \<in> L\<^isub>1"
+                               and rest_in_l2: "(x - xa) @ z \<in> L\<^isub>2" by blast
+    hence "\<exists> ya. \<lbrakk>(x - xa)\<rbrakk>L\<^isub>2 = \<lbrakk>(y - ya)\<rbrakk>L\<^isub>2 \<and> ya \<le> y \<and> ya \<in> L\<^isub>1" using xy_l2 by auto
+    then obtain ya where ya_le_x: "ya \<le> y"
+                     and ya_in_l1: "ya \<in> L\<^isub>1"
+                     and rest_eq: "\<lbrakk>(x - xa)\<rbrakk>L\<^isub>2 = \<lbrakk>(y - ya)\<rbrakk>L\<^isub>2" by blast
+    from rest_eq rest_in_l2 have "(y - ya) @ z \<in> L\<^isub>2" 
+      by (auto simp:equiv_class_def equiv_str_def)
+    hence "ya @ ((y - ya) @ z) \<in> L\<^isub>1 ; L\<^isub>2" using ya_in_l1
+      by (auto simp:lang_seq_def)
+    thus "y @ z \<in> L\<^isub>1 ; L\<^isub>2" using ya_le_x 
+      by (erule_tac prefixE, simp)
+  next
+    fix x y z
+    assume tag_eq: "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y"
+      and x_gets_l1: "\<exists>za\<le>z. x @ za \<in> L\<^isub>1 \<and> z - za \<in> L\<^isub>2"
+    from tag_eq tag_seq_eq_E have x_y_eq: "\<lbrakk>x\<rbrakk>L\<^isub>1 = \<lbrakk>y\<rbrakk>L\<^isub>1" by blast
+    from x_gets_l1 obtain za where za_le_z: "za \<le> z"
+                               and x_za_in_l1: "(x @ za) \<in> L\<^isub>1"
+                               and rest_in_l2: "z - za \<in> L\<^isub>2" by blast
+    from x_y_eq x_za_in_l1 have y_za_in_l1: "y @ za \<in> L\<^isub>1"
+      by (auto simp:equiv_class_def equiv_str_def)
+    hence "(y @ za) @ (z - za) \<in> L\<^isub>1 ; L\<^isub>2" using rest_in_l2
+      apply (simp add:lang_seq_def)
+      by (rule_tac x = "y @ za" in exI, rule_tac x = "z - za" in exI, simp)
+    thus "y @ z \<in> L\<^isub>1 ; L\<^isub>2" using za_le_z
+      by (erule_tac prefixE, simp)
+  qed
+  show ?thesis using tag_eq
+    apply (simp add:equiv_str_def)
+    by (auto intro:aux)
+qed
+
+lemma quot_seq: 
+  assumes finite1: "finite (QUOT L\<^isub>1)"
+  and finite2: "finite (QUOT L\<^isub>2)"
+  shows "finite (QUOT (L\<^isub>1;L\<^isub>2))"
+proof (rule_tac f = "(op `) (tag_str_SEQ L\<^isub>1 L\<^isub>2)" in finite_imageD)
+  show "finite (op ` (tag_str_SEQ L\<^isub>1 L\<^isub>2) ` QUOT (L\<^isub>1 ; L\<^isub>2))"
+    using finite_tag_image tag_str_seq_range_finite finite1 finite2
+    by auto
+next
+  show "inj_on (op ` (tag_str_SEQ L\<^isub>1 L\<^isub>2)) (QUOT (L\<^isub>1 ; L\<^isub>2))"
+    apply (rule_tac str_inj_imps)
+    by (erule_tac tag_str_seq_inj)
+qed
+
+(****************** the STAR case ************************)
+
+lemma app_eq_elim[rule_format]:
+  "\<And> a. \<forall> b x y. a @ b = x @ y \<longrightarrow> (\<exists> aa ab. a = aa @ ab \<and> x = aa \<and> y = ab @ b) \<or>
+                                   (\<exists> ba bb. b = ba @ bb \<and> x = a @ ba \<and> y = bb \<and> ba \<noteq> [])"
+  apply (induct_tac a rule:List.induct, simp)
+  apply (rule allI | rule impI)+
+  by (case_tac x, auto)
+
+definition tag_str_STAR:: "string set \<Rightarrow> string \<Rightarrow> string set set"
+where
+  "tag_str_STAR L\<^isub>1 x \<equiv> if (x = []) 
+                       then {}
+                       else {\<lbrakk>x\<^isub>2\<rbrakk>L\<^isub>1 | x\<^isub>1 x\<^isub>2. x =  x\<^isub>1 @ x\<^isub>2 \<and> x\<^isub>1 \<in> L\<^isub>1\<star> \<and> x\<^isub>2 \<noteq> []}"
+
+lemma tag_str_star_range_finite:
+  assumes finite1: "finite (QUOT L\<^isub>1)"
+  shows "finite (range (tag_str_STAR L\<^isub>1))"
+proof -
+  have "range (tag_str_STAR L\<^isub>1) \<subseteq> Pow (QUOT L\<^isub>1)"
+    by (auto simp:image_def tag_str_STAR_def QUOT_def)
+  thus ?thesis using finite1
+    by (rule_tac B = "Pow (QUOT L\<^isub>1)" in finite_subset, auto)
+qed
+
+lemma star_prop[rule_format]: "x \<in> lang\<star> \<Longrightarrow> \<forall> y. y \<in> lang\<star> \<longrightarrow> x @ y \<in> lang\<star>"
+by (erule Star.induct, auto)
+
+lemma star_prop2: "y \<in> lang \<Longrightarrow> y \<in> lang\<star>"
+by (drule step[of y lang "[]"], auto simp:start)
+
+lemma star_prop3[rule_format]: "x \<in> lang\<star> \<Longrightarrow> \<forall>y . y \<in> lang \<longrightarrow> x @ y \<in> lang\<star>"
+by (erule Star.induct, auto intro:star_prop2)
+
+lemma postfix_prop: "y >>= (x @ y) \<Longrightarrow> x = []"
+by (erule postfixE, induct x arbitrary:y, auto)
+
+lemma inj_aux:
+  "\<lbrakk>(m @ z) \<in> L\<^isub>1\<star>; m \<equiv>L\<^isub>1\<equiv> yb; xa @ m = x; xa \<in> L\<^isub>1\<star>; m \<noteq> [];
+    \<forall> xa xb. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star> \<longrightarrow> xb >>= m\<rbrakk> 
+  \<Longrightarrow> (yb @ z) \<in> L\<^isub>1\<star>"
+proof- 
+  have "\<And>s. s \<in> L\<^isub>1\<star> \<Longrightarrow> \<forall> m z yb. (s = m @ z \<and> m \<equiv>L\<^isub>1\<equiv> yb \<and> x = xa @ m \<and> xa \<in> L\<^isub>1\<star> \<and> m \<noteq> [] \<and>  
+    (\<forall> xa xb. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star> \<longrightarrow> xb >>= m) \<longrightarrow> (yb @ z) \<in> L\<^isub>1\<star>)"    
+    apply (erule Star.induct, simp)
+    apply (rule allI | rule impI | erule conjE)+
+    apply (drule app_eq_elim)
+    apply (erule disjE | erule exE | erule conjE)+
+    apply simp
+    apply (simp (no_asm) only:append_assoc[THEN sym])
+    apply (rule step) 
+    apply (simp add:equiv_str_def)
+    apply simp
+
+    apply (erule exE | erule conjE)+    
+    apply (rotate_tac 3)
+    apply (frule_tac x = "xa @ s1" in spec)
+    apply (rotate_tac 12)
+    apply (drule_tac x = ba in spec)
+    apply (erule impE)
+    apply ( simp add:star_prop3)
+    apply (simp)
+    apply (drule postfix_prop)
+    apply simp
+    done
+  thus "\<lbrakk>(m @ z) \<in> L\<^isub>1\<star>; m \<equiv>L\<^isub>1\<equiv> yb; xa @ m = x; xa \<in> L\<^isub>1\<star>; m \<noteq> [];
+         \<forall> xa xb. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star> \<longrightarrow> xb >>= m\<rbrakk> 
+        \<Longrightarrow> (yb @ z) \<in> L\<^isub>1\<star>" by blast
+qed
+
+
+lemma min_postfix_exists[rule_format]:
+  "finite A \<Longrightarrow> A \<noteq> {} \<and> (\<forall> a \<in> A. \<forall> b \<in> A. ((b >>= a) \<or> (a >>= b))) 
+                \<longrightarrow> (\<exists> min. (min \<in> A \<and> (\<forall> a \<in> A. a >>= min)))"
+apply (erule finite.induct)
+apply simp
+apply simp
+apply (case_tac "A = {}")
+apply simp
+apply clarsimp
+apply (case_tac "a >>= min")
+apply (rule_tac x = min in exI, simp)
+apply (rule_tac x = a in exI, simp)
+apply clarify
+apply (rotate_tac 5)
+apply (erule_tac x = aa in ballE) defer apply simp
+apply (erule_tac ys = min in postfix_trans)
+apply (erule_tac x = min in ballE) 
+by simp+
+
+lemma tag_str_star_inj:
+  "tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 (y::string) \<Longrightarrow> x \<equiv>L\<^isub>1\<star>\<equiv> y"
+proof -
+  have aux: "\<And> x y z. \<lbrakk>tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 y; x @ z \<in> L\<^isub>1\<star>\<rbrakk> \<Longrightarrow> y @ z \<in> L\<^isub>1\<star>"
+  proof-
+    fix x y z
+    assume tag_eq: "tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 y"
+      and x_z: "x @ z \<in> L\<^isub>1\<star>"
+    show "y @ z \<in> L\<^isub>1\<star>"
+    proof (cases "x = []")
+      case True
+      with tag_eq have "y = []" by (simp add:tag_str_STAR_def split:if_splits, blast)
+      thus ?thesis using x_z True by simp
+    next
+      case False
+      hence not_empty: "{xb. \<exists> xa. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star>} \<noteq> {}" using x_z
+        by (simp, rule_tac x = x in exI, rule_tac x = "[]" in exI, simp add:start)
+      have finite_set: "finite {xb. \<exists> xa. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star>}"
+        apply (rule_tac B = "{xb. \<exists> xa. x = xa @ xb}" in finite_subset)
+        apply auto
+        apply (induct x, simp)
+        apply (subgoal_tac "{xb. \<exists>xa. a # x = xa @ xb} = insert (a # x) {xb. \<exists>xa. x = xa @ xb}")
+        apply auto
+        by (case_tac xaa, simp+)
+      have comparable: "\<forall> a \<in> {xb. \<exists> xa. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star>}. 
+                        \<forall> b \<in> {xb. \<exists> xa. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star>}.
+                          ((b >>= a) \<or> (a >>= b))"
+        by (auto simp:postfix_def, drule app_eq_elim, blast)
+      hence "\<exists> min. min \<in> {xb. \<exists> xa. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star>} 
+                \<and> (\<forall> xa xb. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star> \<longrightarrow> xb >>= min)"
+        using finite_set not_empty comparable
+        apply (drule_tac min_postfix_exists, simp)
+        by (erule exE, rule_tac x = min in exI, auto)
+      then obtain min xa where x_decom: "x = xa @ min \<and> xa \<in> L\<^isub>1\<star>"
+        and min_not_empty: "min \<noteq> []"
+        and min_z_in_star: "min @ z \<in> L\<^isub>1\<star>"
+        and is_min: "\<forall> xa xb. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star> \<longrightarrow> xb >>= min"  by blast
+      from x_decom min_not_empty have "\<lbrakk>min\<rbrakk>L\<^isub>1 \<in> tag_str_STAR L\<^isub>1 x"  by (auto simp:tag_str_STAR_def)
+      hence "\<exists> yb. \<lbrakk>yb\<rbrakk>L\<^isub>1 \<in> tag_str_STAR L\<^isub>1 y \<and> \<lbrakk>min\<rbrakk>L\<^isub>1 = \<lbrakk>yb\<rbrakk>L\<^isub>1" using tag_eq by auto
+      hence "\<exists> ya yb. y = ya @ yb \<and> ya \<in> L\<^isub>1\<star> \<and> min \<equiv>L\<^isub>1\<equiv> yb \<and> yb \<noteq> [] " 
+        by (simp add:tag_str_STAR_def equiv_class_def equiv_str_def split:if_splits, blast)        
+      then obtain ya yb where y_decom: "y = ya @ yb"
+                          and ya_in_star: "ya \<in> L\<^isub>1\<star>"
+                          and yb_not_empty: "yb \<noteq> []"
+                          and min_yb_eq: "min \<equiv>L\<^isub>1\<equiv> yb"  by blast
+      from min_z_in_star min_yb_eq min_not_empty is_min x_decom
+      have "yb @ z \<in> L\<^isub>1\<star>"
+        by (rule_tac x = x and xa = xa in inj_aux, simp+)
+      thus ?thesis using ya_in_star y_decom
+        by (auto dest:star_prop)        
+    qed
+  qed
+  show "tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 (y::string) \<Longrightarrow> x \<equiv>L\<^isub>1\<star>\<equiv> y"
+    by (auto intro:aux simp:equiv_str_def)
+qed
+
+lemma quot_star:  
+  assumes finite1: "finite (QUOT L\<^isub>1)"
+  shows "finite (QUOT (L\<^isub>1\<star>))"
+proof (rule_tac f = "(op `) (tag_str_STAR L\<^isub>1)" in finite_imageD)
+  show "finite (op ` (tag_str_STAR L\<^isub>1) ` QUOT (L\<^isub>1\<star>))"
+    using finite_tag_image tag_str_star_range_finite finite1
+    by auto
+next
+  show "inj_on (op ` (tag_str_STAR L\<^isub>1)) (QUOT (L\<^isub>1\<star>))"
+    apply (rule_tac str_inj_imps)
+    by (erule_tac tag_str_star_inj)
+qed
+
+lemma other_direction:
+  "Lang = L (r::rexp) \<Longrightarrow> finite (QUOT Lang)"
+apply (induct arbitrary:Lang rule:rexp.induct)
+apply (simp add:QUOT_def equiv_class_def equiv_str_def)
+by (simp_all add:quot_lambda quot_single quot_seq quot_alt quot_star)  
+
+end 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Prelude.thy	Mon Jul 25 13:33:38 2011 +0000
@@ -0,0 +1,19 @@
+theory Prelude
+imports Main
+begin
+
+(*
+To make the theory work under Isabelle 2009 and 2011
+
+Isabelle 2009: set_ext
+Isabelle 2011: set_eqI
+
+*)
+
+
+lemma set_eq_intro:
+  "(\<And>x. (x \<in> A) = (x \<in> B)) \<Longrightarrow> A = B"
+by blast
+
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/old/Closure.thy	Mon Jul 25 13:33:38 2011 +0000
@@ -0,0 +1,158 @@
+(* Author: Christian Urban, Xingyuan Zhang, Chunhan Wu *)
+theory Closure
+imports Derivs
+begin
+
+section {* Closure properties of regular languages *}
+
+abbreviation
+  regular :: "lang \<Rightarrow> bool"
+where
+  "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_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 \<cdot> B)"
+proof -
+  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_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)"
+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
+
+subsection {* 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[simp]:
+  shows "rev ` (B \<cdot> A) = (rev ` A) \<cdot> (rev ` B)"
+unfolding Seq_def image_def
+by (auto) (metis rev_append)+
+
+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[simp]:
+  shows " rev ` (A\<star>) = (rev ` A)\<star>"
+using rev_star1 rev_star2 by auto
+
+lemma rev_lang:
+  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 "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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/old/Derivs.thy	Mon Jul 25 13:33:38 2011 +0000
@@ -0,0 +1,492 @@
+theory Derivs
+imports Myhill_2
+begin
+
+section {* Left-Quotients and Derivatives *}
+
+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 \<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) \<cdot> A\<star>"
+proof -
+  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 \<cdot> A\<star>)"
+    by (simp only: star_cases[symmetric])
+  also have "... = Der c (A \<cdot> A\<star>)"
+    by (simp only: Der_union Der_empty) (simp)
+  also have "... = (Der c A) \<cdot> A\<star> \<union> (Delta A \<cdot> Der c (A\<star>))"
+    by simp
+  also have "... =  (Der c A) \<cdot> A\<star>"
+    using incl by auto
+  finally show "Der c (A\<star>) = (Der c A) \<cdot> A\<star>" . 
+qed
+
+
+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_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_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_rexp r) = L_rexp (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_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 \<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_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_rexp r) = \<Union> L_rexp ` (pders s r)"
+proof (induct s rule: rev_induct)
+  case (snoc c s)
+  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_rexp ` (pders s r))" using ih
+    by (simp add: Ders_singleton)
+  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_rexp `  (pder c r)))"
+    by (simp add: Der_pder)
+  also have "\<dots> = (\<Union>L_rexp ` (pder_set c (pders s r)))"
+    by (simp add: pder_set_lang)
+  also have "\<dots> = (\<Union>L_rexp ` (pders (s @ [c]) r))"
+    by simp
+  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_rexp r) = (\<Union> L_rexp ` (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) \<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 \<cdot> {[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) (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))"
+    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 (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))"
+      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 finite_pders2:
+  shows "finite {pders s r | s. s \<in> A}"
+proof -
+  have "{pders s r | s. s \<in> A} \<subseteq> Pow (pders_set A r)" by auto
+  moreover
+  have "finite (Pow (pders_set A r))"
+    using finite_pders_set by simp
+  ultimately 
+  show "finite {pders s r | s. s \<in> A}"
+    by(rule finite_subset)
+qed
+
+
+lemma Myhill_Nerode3:
+  fixes r::"rexp"
+  shows "finite (UNIV // \<approx>(L_rexp r))"
+proof -
+  have "finite (UNIV // =(\<lambda>x. pders x r)=)"
+  proof - 
+    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 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_rexp r)"
+    unfolding tag_eq_rel_def
+    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_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_rexp r))" 
+    by (rule refined_partition_finite)
+qed
+
+
+section {* Relating derivatives and partial derivatives *}
+
+lemma
+  shows "(\<Union> L_rexp ` (pder c r)) = L_rexp (der c r)"
+unfolding Der_der[symmetric] Der_pder by simp
+
+lemma
+  shows "(\<Union> L_rexp ` (pders s r)) = L_rexp (ders s r)"
+unfolding Ders_ders[symmetric] Ders_pders by simp
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/old/Folds.thy	Mon Jul 25 13:33:38 2011 +0000
@@ -0,0 +1,22 @@
+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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/old/My.thy	Mon Jul 25 13:33:38 2011 +0000
@@ -0,0 +1,389 @@
+theory My
+imports Main Infinite_Set
+begin
+
+
+definition
+  Seq :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ;; _" [100,100] 100)
+where 
+  "L1 ;; L2 = {s1 @ s2 | s1 s2. s1 \<in> L1 \<and> s2 \<in> L2}"
+
+inductive_set
+  Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102)
+  for L :: "string set"
+where
+  start[intro]: "[] \<in> L\<star>"
+| step[intro]:  "\<lbrakk>s1 \<in> L; s2 \<in> L\<star>\<rbrakk> \<Longrightarrow> s1@s2 \<in> L\<star>"
+
+lemma lang_star_cases:
+  shows "L\<star> =  {[]} \<union> L ;; L\<star>"
+unfolding Seq_def
+by (auto) (metis Star.simps)
+
+lemma lang_star_cases2:
+  shows "L ;; L\<star>  = L\<star> ;; L"
+sorry
+
+
+theorem ardens_revised:
+  assumes nemp: "[] \<notin> A"
+  shows "(X = X ;; A \<union> B) \<longleftrightarrow> (X = B ;; A\<star>)"
+proof
+  assume eq: "X = B ;; A\<star>"
+  have "A\<star> =  {[]} \<union> A\<star> ;; A" sorry
+  then have "B ;; A\<star> = B ;; ({[]} \<union> A\<star> ;; A)" unfolding Seq_def by simp
+  also have "\<dots> = B \<union> B ;; (A\<star> ;; A)"  unfolding Seq_def by auto
+  also have "\<dots> = B \<union> (B ;; A\<star>) ;; A"  unfolding Seq_def
+    by (auto) (metis append_assoc)+
+  finally show "X = X ;; A \<union> B" using eq by auto
+next
+  assume "X = X ;; A \<union> B"
+  then have "B \<subseteq> X" "X ;; A \<subseteq> X" by auto
+  show "X = B ;; A\<star>" sorry
+qed
+
+datatype rexp =
+  NULL
+| EMPTY
+| CHAR char
+| SEQ rexp rexp
+| ALT rexp rexp
+| STAR rexp
+
+fun
+  Sem :: "rexp \<Rightarrow> string set" ("\<lparr>_\<rparr>" [0] 1000)
+where
+    "\<lparr>NULL\<rparr> = {}"
+  | "\<lparr>EMPTY\<rparr> = {[]}"
+  | "\<lparr>CHAR c\<rparr> = {[c]}"
+  | "\<lparr>SEQ r1 r2\<rparr> = \<lparr>r1\<rparr> ;; \<lparr>r2\<rparr>"
+  | "\<lparr>ALT r1 r2\<rparr> = \<lparr>r1\<rparr> \<union> \<lparr>r2\<rparr>"
+  | "\<lparr>STAR r\<rparr> = \<lparr>r\<rparr>\<star>"
+
+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"
+
+lemma folds_simp_null [simp]:
+  "finite rs \<Longrightarrow> x \<in> \<lparr>folds ALT NULL rs\<rparr> \<longleftrightarrow> (\<exists>r \<in> rs. x \<in> \<lparr>r\<rparr>)"
+apply (simp add: folds_def)
+apply (rule someI2_ex)
+apply (erule finite_imp_fold_graph)
+apply (erule fold_graph.induct)
+apply (auto)
+done
+
+lemma folds_simp_empty [simp]:
+  "finite rs \<Longrightarrow> x \<in> \<lparr>folds ALT EMPTY rs\<rparr> \<longleftrightarrow> (\<exists>r \<in> rs. x \<in> \<lparr>r\<rparr>) \<or> x = []"
+apply (simp add: folds_def)
+apply (rule someI2_ex)
+apply (erule finite_imp_fold_graph)
+apply (erule fold_graph.induct)
+apply (auto)
+done
+
+lemma [simp]:
+  shows "(x, y) \<in> {(x, y). P x y} \<longleftrightarrow> P x y"
+by simp
+
+definition
+  str_eq ("_ \<approx>_ _")
+where
+  "x \<approx>Lang y \<equiv> (\<forall>z. x @ z \<in> Lang \<longleftrightarrow> y @ z \<in> Lang)"
+
+definition
+  str_eq_rel ("\<approx>_")
+where
+  "\<approx>Lang \<equiv> {(x, y). x \<approx>Lang y}"
+
+definition
+  final :: "string set \<Rightarrow> string set \<Rightarrow> bool"
+where
+  "final X Lang \<equiv> (X \<in> UNIV // \<approx>Lang) \<and> (\<forall>s \<in> X. s \<in> Lang)"
+
+lemma lang_is_union_of_finals: 
+  "Lang = \<Union> {X. final X Lang}"
+proof -
+  have  "Lang \<subseteq> \<Union> {X. final X Lang}"
+    unfolding final_def
+    unfolding quotient_def Image_def
+    unfolding str_eq_rel_def
+    apply(simp)
+    apply(auto)
+    apply(rule_tac x="(\<approx>Lang) `` {x}" in exI)
+    unfolding Image_def
+    unfolding str_eq_rel_def
+    apply(auto)
+    unfolding str_eq_def
+    apply(auto)
+    apply(drule_tac x="[]" in spec)
+    apply(simp)
+    done
+  moreover
+  have "\<Union>{X. final X Lang} \<subseteq> Lang" 
+    unfolding final_def by auto
+  ultimately 
+  show "Lang = \<Union> {X. final X Lang}"
+    by blast
+qed
+
+lemma all_rexp:
+  "\<lbrakk>finite (UNIV // \<approx>Lang); X \<in> (UNIV // \<approx>Lang)\<rbrakk> \<Longrightarrow> \<exists>r. X = \<lparr>r\<rparr>"
+sorry
+
+lemma final_rexp:
+  "\<lbrakk>finite (UNIV // (\<approx>Lang)); final X Lang\<rbrakk> \<Longrightarrow> \<exists>r. X = \<lparr>r\<rparr>"
+unfolding final_def
+using all_rexp by blast
+
+lemma finite_f_one_to_one:
+  assumes "finite B"
+  and "\<forall>x \<in> B. \<exists>y. f y = x"
+  shows "\<exists>rs. finite rs \<and> (B = {f y | y . y \<in> rs})"
+using assms
+by (induct) (auto)
+
+lemma finite_final:
+  assumes "finite (UNIV // (\<approx>Lang))"
+  shows "finite {X. final X Lang}"
+using assms
+proof -
+  have "{X. final X Lang} \<subseteq> (UNIV // (\<approx>Lang))"
+    unfolding final_def by auto
+  with assms show "finite {X. final X Lang}" 
+    using finite_subset by auto
+qed
+
+lemma finite_regular_aux:
+  fixes Lang :: "string set"
+  assumes "finite (UNIV // (\<approx>Lang))"
+  shows "\<exists>rs. Lang = \<lparr>folds ALT NULL rs\<rparr>"
+apply(subst lang_is_union_of_finals)
+using assms
+apply -
+apply(drule finite_final)
+apply(drule_tac f="Sem" in finite_f_one_to_one)
+apply(clarify)
+apply(drule final_rexp[OF assms])
+apply(auto)[1]
+apply(clarify)
+apply(rule_tac x="rs" in exI)
+apply(simp)
+apply(rule set_eqI)
+apply(auto)
+done
+
+lemma finite_regular:
+  fixes Lang :: "string set"
+  assumes "finite (UNIV // (\<approx>Lang))"
+  shows "\<exists>r. Lang =  \<lparr>r\<rparr>"
+using assms finite_regular_aux
+by auto
+
+
+
+section {* other direction *}
+
+
+lemma inj_image_lang:
+  fixes f::"string \<Rightarrow> 'a"
+  assumes str_inj: "\<And>x y. f x = f y \<Longrightarrow> x \<approx>Lang y"
+  shows "inj_on (image f) (UNIV // (\<approx>Lang))"
+proof - 
+  { fix x y::string
+    assume eq_tag: "f ` {z. x \<approx>Lang z} = f ` {z. y \<approx>Lang z}"
+    moreover
+    have "{z. x \<approx>Lang z} \<noteq> {}" unfolding str_eq_def by auto
+    ultimately obtain a b where "x \<approx>Lang a" "y \<approx>Lang b" "f a = f b" by blast
+    then have "x \<approx>Lang a" "y \<approx>Lang b" "a \<approx>Lang b" using str_inj by auto
+    then have "x \<approx>Lang y" unfolding str_eq_def by simp 
+    then have "{z. x \<approx>Lang z} = {z. y \<approx>Lang z}" unfolding str_eq_def by simp
+  }
+  then have "\<forall>x\<in>UNIV // \<approx>Lang. \<forall>y\<in>UNIV // \<approx>Lang. f ` x = f ` y \<longrightarrow> x = y"
+    unfolding quotient_def Image_def str_eq_rel_def by simp
+  then show "inj_on (image f) (UNIV // (\<approx>Lang))"
+    unfolding inj_on_def by simp
+qed
+
+
+lemma finite_range_image: 
+  assumes fin: "finite (range f)"
+  shows "finite ((image f) ` X)"
+proof -
+  from fin have "finite (Pow (f ` UNIV))" by auto
+  moreover
+  have "(image f) ` X \<subseteq> Pow (f ` UNIV)" by auto
+  ultimately show "finite ((image f) ` X)" using finite_subset by auto
+qed
+
+definition 
+  tag1 :: "string set \<Rightarrow> string set \<Rightarrow> string \<Rightarrow> (string set \<times> string set)"
+where
+  "tag1 L\<^isub>1 L\<^isub>2 \<equiv> \<lambda>x. ((\<approx>L\<^isub>1) `` {x}, (\<approx>L\<^isub>2) `` {x})"
+
+lemma tag1_range_finite:
+  assumes finite1: "finite (UNIV // \<approx>L\<^isub>1)"
+  and finite2: "finite (UNIV // \<approx>L\<^isub>2)"
+  shows "finite (range (tag1 L\<^isub>1 L\<^isub>2))"
+proof -
+  have "finite (UNIV // \<approx>L\<^isub>1 \<times> UNIV // \<approx>L\<^isub>2)" using finite1 finite2 by auto
+  moreover
+  have "range (tag1 L\<^isub>1 L\<^isub>2) \<subseteq> (UNIV // \<approx>L\<^isub>1) \<times> (UNIV // \<approx>L\<^isub>2)"
+    unfolding tag1_def quotient_def by auto
+  ultimately show "finite (range (tag1 L\<^isub>1 L\<^isub>2))" 
+    using finite_subset by blast
+qed
+
+lemma tag1_inj:
+  "tag1 L\<^isub>1 L\<^isub>2 x = tag1 L\<^isub>1 L\<^isub>2 y \<Longrightarrow> x \<approx>(L\<^isub>1 \<union> L\<^isub>2) y"
+unfolding tag1_def Image_def str_eq_rel_def str_eq_def
+by auto
+
+lemma quot_alt_cu:
+  fixes L\<^isub>1 L\<^isub>2::"string set"
+  assumes fin1: "finite (UNIV // \<approx>L\<^isub>1)"
+  and fin2: "finite (UNIV // \<approx>L\<^isub>2)"
+  shows "finite (UNIV // \<approx>(L\<^isub>1 \<union> L\<^isub>2))"
+proof -
+  have "finite (range (tag1 L\<^isub>1 L\<^isub>2))" 
+    using fin1 fin2 tag1_range_finite by simp
+  then have "finite (image (tag1 L\<^isub>1 L\<^isub>2) ` (UNIV // \<approx>(L\<^isub>1 \<union> L\<^isub>2)))" 
+    using finite_range_image by blast
+  moreover 
+  have "\<And>x y. tag1 L\<^isub>1 L\<^isub>2 x = tag1 L\<^isub>1 L\<^isub>2 y \<Longrightarrow> x \<approx>(L\<^isub>1 \<union> L\<^isub>2) y" 
+    using tag1_inj by simp
+  then have "inj_on (image (tag1 L\<^isub>1 L\<^isub>2)) (UNIV // \<approx>(L\<^isub>1 \<union> L\<^isub>2))" 
+    using inj_image_lang by blast
+  ultimately 
+  show "finite (UNIV // \<approx>(L\<^isub>1 \<union> L\<^isub>2))" by (rule finite_imageD)
+qed
+
+
+section {* finite \<Rightarrow> regular *}
+
+definition
+  transitions :: "string set \<Rightarrow> string set \<Rightarrow> rexp set" ("_\<Turnstile>\<Rightarrow>_")
+where
+  "Y \<Turnstile>\<Rightarrow> X \<equiv> {CHAR c | c. Y ;; {[c]} \<subseteq> X}"
+
+definition
+  transitions_rexp ("_ \<turnstile>\<rightarrow> _")
+where
+  "Y \<turnstile>\<rightarrow> X \<equiv> if [] \<in> X then folds ALT EMPTY (Y \<Turnstile>\<Rightarrow>X) else folds ALT NULL (Y \<Turnstile>\<Rightarrow>X)"
+
+definition
+  "rhs CS X \<equiv> if X = {[]} then {({[]}, EMPTY)} else {(Y, Y \<turnstile>\<rightarrow>X) | Y. Y \<in> CS}"
+
+definition
+  "rhs_sem CS X \<equiv> \<Union> {(Y;; \<lparr>r\<rparr>) | Y r . (Y, r) \<in> rhs CS X}"
+
+definition
+  "eqs CS \<equiv> (\<Union>X \<in> CS. {(X, rhs CS X)})"
+
+definition
+  "eqs_sem CS \<equiv> (\<Union>X \<in> CS. {(X, rhs_sem CS X)})"
+
+lemma [simp]:
+  shows "finite (Y \<Turnstile>\<Rightarrow> X)"
+unfolding transitions_def
+by auto
+
+
+lemma defined_by_str:
+  assumes "s \<in> X" 
+  and "X \<in> UNIV // (\<approx>Lang)"
+  shows "X = (\<approx>Lang) `` {s}"
+using assms
+unfolding quotient_def Image_def
+unfolding str_eq_rel_def str_eq_def
+by auto
+
+lemma every_eqclass_has_transition:
+  assumes has_str: "s @ [c] \<in> X"
+  and     in_CS:   "X \<in> UNIV // (\<approx>Lang)"
+  obtains Y where "Y \<in> UNIV // (\<approx>Lang)" and "Y ;; {[c]} \<subseteq> X" and "s \<in> Y"
+proof -
+  def Y \<equiv> "(\<approx>Lang) `` {s}"
+  have "Y \<in> UNIV // (\<approx>Lang)" 
+    unfolding Y_def quotient_def by auto
+  moreover
+  have "X = (\<approx>Lang) `` {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 (auto) (simp add: str_eq_def)
+  moreover
+  have "s \<in> Y" unfolding Y_def 
+    unfolding Image_def str_eq_rel_def str_eq_def by simp
+  (*moreover 
+  have "True" by simp FIXME *) 
+  ultimately show thesis by (blast intro: that)
+qed
+
+lemma test:
+  assumes "[] \<in> X"
+  shows "[] \<in> \<lparr>Y \<turnstile>\<rightarrow> X\<rparr>"
+using assms
+by (simp add: transitions_rexp_def)
+
+lemma rhs_sem:
+  assumes "X \<in> (UNIV // (\<approx>Lang))"
+  shows "X \<subseteq> rhs_sem (UNIV // (\<approx>Lang)) X"
+apply(case_tac "X = {[]}")
+apply(simp)
+apply(simp add: rhs_sem_def rhs_def Seq_def)
+apply(rule subsetI)
+apply(case_tac "x = []")
+apply(simp add: rhs_sem_def rhs_def)
+apply(rule_tac x = "X" in exI)
+apply(simp)
+apply(rule_tac x = "X" in exI)
+apply(simp add: assms)
+apply(simp add: transitions_rexp_def)
+oops
+
+
+(*
+fun
+  power :: "string \<Rightarrow> nat \<Rightarrow> string" (infixr "\<Up>" 100)
+where
+  "s \<Up> 0 = s"
+| "s \<Up> (Suc n) = s @ (s \<Up> n)"
+
+definition 
+ "Lone = {(''0'' \<Up> n) @ (''1'' \<Up> n) | n. True }"
+
+lemma
+  "infinite (UNIV // (\<approx>Lone))"
+unfolding infinite_iff_countable_subset
+apply(rule_tac x="\<lambda>n. {(''0'' \<Up> n) @ (''1'' \<Up> i) | i. i \<in> {..n} }" in exI)
+apply(auto)
+prefer 2
+unfolding Lone_def
+unfolding quotient_def
+unfolding Image_def
+apply(simp)
+unfolding str_eq_rel_def
+unfolding str_eq_def
+apply(auto)
+apply(rule_tac x="''0'' \<Up> n" in exI)
+apply(auto)
+unfolding infinite_nat_iff_unbounded
+unfolding Lone_def
+*)
+
+
+
+text {* Derivatives *}
+
+definition
+  DERS :: "string \<Rightarrow> string set \<Rightarrow> string set"
+where
+  "DERS s L \<equiv> {s'. s @ s' \<in> L}"
+
+lemma
+  shows "x \<approx>L y \<longleftrightarrow> DERS x L = DERS y L"
+unfolding DERS_def str_eq_def
+by auto
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/old/Myhill_1.thy	Mon Jul 25 13:33:38 2011 +0000
@@ -0,0 +1,776 @@
+theory Myhill_1
+imports 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 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"
+
+fun 
+  L_rhs::"trm set \<Rightarrow> lang"
+where 
+  "L_rhs rhs = \<Union> (L_trm ` rhs)"
+
+lemma L_rhs_set:
+  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::"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 \<cdot> {[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> trm \<Rightarrow> 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> 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}"
+
+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 rhs"
+
+definition 
+  "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"
+
+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 trm_soundness:
+  assumes finite:"finite 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_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_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_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 {* Intial Equational System *}
+
+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 \<cdot> {[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 \<cdot> {[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 rhs"
+proof 
+  show "X \<subseteq> L_rhs rhs"
+  proof
+    fix x
+    assume in_X: "x \<in> X"
+    { assume empty: "x = []"
+      then have "x \<in> L_rhs 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 \<cdot> {[c]} \<subseteq> X" "s \<in> Y"
+        using decom in_X every_eqclass_has_transition by blast
+      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 rhs" using X_in_eqs in_X
+	unfolding Init_def Init_rhs_def by simp
+    }
+    ultimately show "x \<in> L_rhs rhs" by blast
+  qed
+next
+  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_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 \<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 \<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 \<cdot> {[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 rhs"
+  and not_empty: "ardenable rhs"
+  and finite: "finite rhs"
+  shows "X = L_rhs (Arden X rhs)"
+proof -
+  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 (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 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 trm_soundness[OF finite]
+    unfolding A_def
+    by blast
+  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_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_rhs (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_rhs xrhs"
+  and finite: "finite rhs"
+  shows "L_rhs (Subst rhs X xrhs) = L_rhs rhs" (is "?Left = ?Right")
+proof-
+  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_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 
+      unfolding A_def
+      unfolding L_rhs_union_distrib
+      by simp
+  qed
+  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
+
+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_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)
+  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_rhs (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. X = L_rexp 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_rhs xrhs" using Inv_ES unfolding invariant_def soundness_def
+    by simp
+  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_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:
+  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. 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. X = L_rexp r" 
+    using finite_CS every_eqcl_has_reg by blast
+  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_rexp ` rs)" "finite rs"
+    using fin by (auto dest: bchoice_finite_set)
+  then have "A = L_rexp (\<Uplus>rs)" 
+    unfolding lang_is_union_of_finals[symmetric] by simp
+  then show "\<exists>r. A = L_rexp r" by blast
+qed 
+
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/old/Myhill_2.thy	Mon Jul 25 13:33:38 2011 +0000
@@ -0,0 +1,473 @@
+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)"
+
+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). 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]:
+  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 \<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
+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 \<cdot> B) y"
+proof -
+  { fix x y z
+    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 \<cdot> 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 \<cdot> 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 \<cdot> B"
+	  unfolding prefix_def Seq_def
+	  by (auto) (metis append_assoc)
+      }
+      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 \<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 \<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 \<cdot> 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>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>A))" 
+    using finite1 by auto
+  show "finite (range (tag_str_STAR A))"
+    unfolding tag_str_STAR_def
+    apply(rule finite_subset[OF _ *])
+    unfolding quotient_def
+    by auto
+qed
+
+subsubsection{* The conclusion *}
+
+lemma Myhill_Nerode2:
+  shows "finite (UNIV // \<approx>(L_rexp r))"
+by (induct r) (auto)
+
+
+theorem Myhill_Nerode:
+  shows "(\<exists>r. A = L_rexp r) \<longleftrightarrow> finite (UNIV // \<approx>A)"
+using Myhill_Nerode1 Myhill_Nerode2 by auto
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/old/Prefix_subtract.thy	Mon Jul 25 13:33:38 2011 +0000
@@ -0,0 +1,60 @@
+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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/old/Regular.thy	Mon Jul 25 13:33:38 2011 +0000
@@ -0,0 +1,302 @@
+(* Author: Christian Urban, Xingyuan Zhang, Chunhan Wu *)
+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 "\<cdot>" 100)
+where 
+  "A \<cdot> 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 "\<cdot>"}. *}
+
+lemma seq_add_left:
+  assumes a: "A = B"
+  shows "C \<cdot> A = C \<cdot> B"
+using a by simp
+
+lemma seq_union_distrib_right:
+  shows "(A \<union> B) \<cdot> C = (A \<cdot> C) \<union> (B \<cdot> C)"
+unfolding Seq_def by auto
+
+lemma seq_union_distrib_left:
+  shows "C \<cdot> (A \<union> B) = (C \<cdot> A) \<union> (C \<cdot> B)"
+unfolding Seq_def by  auto
+
+lemma seq_intro:
+  assumes a: "x \<in> A" "y \<in> B"
+  shows "x @ y \<in> A \<cdot> B "
+using a by (auto simp: Seq_def)
+
+lemma seq_assoc:
+  shows "(A \<cdot> B) \<cdot> C = A \<cdot> (B \<cdot> C)"
+unfolding Seq_def
+apply(auto)
+apply(blast)
+by (metis append_assoc)
+
+lemma seq_empty [simp]:
+  shows "A \<cdot> {[]} = A"
+  and   "{[]} \<cdot> A = A"
+by (simp_all add: Seq_def)
+
+lemma seq_null [simp]:
+  shows "A \<cdot> {} = {}"
+  and   "{} \<cdot> 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 \<cdot> (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 \<cdot> A\<star>"
+proof
+  { fix x
+    have "x \<in> A\<star> \<Longrightarrow> x \<in> {[]} \<union> A \<cdot> A\<star>"
+      unfolding Seq_def
+    by (induct rule: star_induct) (auto)
+  }
+  then show "A\<star> \<subseteq> {[]} \<union> A \<cdot> A\<star>" by auto
+next
+  show "{[]} \<union> A \<cdot> 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 seq_Union_left: 
+  shows "B \<cdot> (\<Union>n. A \<up> n) = (\<Union>n. B \<cdot> (A \<up> n))"
+unfolding Seq_def by auto
+
+lemma seq_Union_right: 
+  shows "(\<Union>n. A \<up> n) \<cdot> B = (\<Union>n. (A \<up> n) \<cdot> B)"
+unfolding Seq_def by auto
+
+lemma seq_pow_comm:
+  shows "A \<cdot> (A \<up> n) = (A \<up> n) \<cdot> A"
+by (induct n) (simp_all add: seq_assoc[symmetric])
+
+lemma seq_star_comm:
+  shows "A \<cdot> A\<star> = A\<star> \<cdot> 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 \<cdot> (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 \<cdot> A \<union> B"
+  shows "X = X \<cdot> (A \<up> Suc n) \<union> (\<Union>m\<in>{0..n}. B \<cdot> (A \<up> m))"
+proof (induct n)
+  case 0 
+  show "X = X \<cdot> (A \<up> Suc 0) \<union> (\<Union>(m::nat)\<in>{0..0}. B \<cdot> (A \<up> m))"
+    using eq by simp
+next
+  case (Suc n)
+  have ih: "X = X \<cdot> (A \<up> Suc n) \<union> (\<Union>m\<in>{0..n}. B \<cdot> (A \<up> m))" by fact
+  also have "\<dots> = (X \<cdot> A \<union> B) \<cdot> (A \<up> Suc n) \<union> (\<Union>m\<in>{0..n}. B \<cdot> (A \<up> m))" using eq by simp
+  also have "\<dots> = X \<cdot> (A \<up> Suc (Suc n)) \<union> (B \<cdot> (A \<up> Suc n)) \<union> (\<Union>m\<in>{0..n}. B \<cdot> (A \<up> m))"
+    by (simp add: seq_union_distrib_right seq_assoc)
+  also have "\<dots> = X \<cdot> (A \<up> Suc (Suc n)) \<union> (\<Union>m\<in>{0..Suc n}. B \<cdot> (A \<up> m))"
+    by (auto simp add: le_Suc_eq)
+  finally show "X = X \<cdot> (A \<up> Suc (Suc n)) \<union> (\<Union>m\<in>{0..Suc n}. B \<cdot> (A \<up> m))" .
+qed
+
+theorem arden:
+  assumes nemp: "[] \<notin> A"
+  shows "X = X \<cdot> A \<union> B \<longleftrightarrow> X = B \<cdot> A\<star>"
+proof
+  assume eq: "X = B \<cdot> A\<star>"
+  have "A\<star> = {[]} \<union> A\<star> \<cdot> A" 
+    unfolding seq_star_comm[symmetric]
+    by (rule star_cases)
+  then have "B \<cdot> A\<star> = B \<cdot> ({[]} \<union> A\<star> \<cdot> A)"
+    by (rule seq_add_left)
+  also have "\<dots> = B \<union> B \<cdot> (A\<star> \<cdot> A)"
+    unfolding seq_union_distrib_left by simp
+  also have "\<dots> = B \<union> (B \<cdot> A\<star>) \<cdot> A" 
+    by (simp only: seq_assoc)
+  finally show "X = X \<cdot> A \<union> B" 
+    using eq by blast 
+next
+  assume eq: "X = X \<cdot> A \<union> B"
+  { fix n::nat
+    have "B \<cdot> (A \<up> n) \<subseteq> X" using arden_helper[OF eq, of "n"] by auto }
+  then have "B \<cdot> 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 \<cdot> (A \<up> Suc k)" 
+      using seq_pow_length[OF nemp] by blast
+    assume "s \<in> X"
+    then have "s \<in> X \<cdot> (A \<up> Suc k) \<union> (\<Union>m\<in>{0..k}. B \<cdot> (A \<up> m))"
+      using arden_helper[OF eq, of "k"] by auto
+    then have "s \<in> (\<Union>m\<in>{0..k}. B \<cdot> (A \<up> m))" using not_in by auto
+    moreover
+    have "(\<Union>m\<in>{0..k}. B \<cdot> (A \<up> m)) \<subseteq> (\<Union>n. B \<cdot> (A \<up> n))" by auto
+    ultimately 
+    have "s \<in> B \<cdot> A\<star>" 
+      unfolding seq_Union_left Star_def by auto }
+  then have "X \<subseteq> B \<cdot> A\<star>" by auto
+  ultimately 
+  show "X = B \<cdot> A\<star>" by simp
+qed
+
+
+section {* Regular Expressions *}
+
+datatype rexp =
+  NULL
+| EMPTY
+| CHAR char
+| SEQ rexp rexp
+| ALT rexp rexp
+| STAR rexp
+
+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) \<cdot> (L_rexp r2)"
+  | "L_rexp (ALT r1 r2) = (L_rexp r1) \<union> (L_rexp r2)"
+  | "L_rexp (STAR r) = (L_rexp r)\<star>"
+
+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_exp}.
+*}
+
+lemma folds_alt_simp [simp]:
+  fixes rs::"rexp set"
+  assumes a: "finite rs"
+  shows "L_rexp (\<Uplus>rs) = \<Union> (L_rexp ` 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
--- a/Closure.thy	Fri Jun 03 13:59:21 2011 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,158 +0,0 @@
-(* Author: Christian Urban, Xingyuan Zhang, Chunhan Wu *)
-theory Closure
-imports Derivs
-begin
-
-section {* Closure properties of regular languages *}
-
-abbreviation
-  regular :: "lang \<Rightarrow> bool"
-where
-  "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_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 \<cdot> B)"
-proof -
-  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_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)"
-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
-
-subsection {* 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[simp]:
-  shows "rev ` (B \<cdot> A) = (rev ` A) \<cdot> (rev ` B)"
-unfolding Seq_def image_def
-by (auto) (metis rev_append)+
-
-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[simp]:
-  shows " rev ` (A\<star>) = (rev ` A)\<star>"
-using rev_star1 rev_star2 by auto
-
-lemma rev_lang:
-  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 "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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Closures.thy	Mon Jul 25 13:33:38 2011 +0000
@@ -0,0 +1,161 @@
+(* Author: Christian Urban, Xingyuan Zhang, Chunhan Wu *)
+theory Closures
+imports Derivatives
+begin
+
+section {* Closure properties of regular languages *}
+
+abbreviation
+  regular :: "'a lang \<Rightarrow> bool"
+where
+  "regular A \<equiv> \<exists>r. A = lang 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::"'a rexp" where "lang r1 = A" "lang r2 = B" by auto
+  then have "A \<union> B = lang (Plus 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 \<cdot> B)"
+proof -
+  from assms obtain r1 r2::"'a rexp" where "lang r1 = A" "lang r2 = B" by auto
+  then have "A \<cdot> B = lang (Times 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::"'a rexp" where "lang r = A" by auto
+  then have "A\<star> = lang (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]:
+  fixes A::"('a::finite) lang"
+  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]:
+  fixes A::"('a::finite) lang"
+  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]:
+  fixes A::"('a::finite) lang"
+  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
+
+subsection {* Closure under string reversal *}
+
+fun
+  Rev :: "'a rexp \<Rightarrow> 'a rexp"
+where
+  "Rev Zero = Zero"
+| "Rev One = One"
+| "Rev (Atom c) = Atom c"
+| "Rev (Plus r1 r2) = Plus (Rev r1) (Rev r2)"
+| "Rev (Times r1 r2) = Times (Rev r2) (Rev r1)"
+| "Rev (Star r) = Star (Rev r)"
+
+lemma rev_seq[simp]:
+  shows "rev ` (B \<cdot> A) = (rev ` A) \<cdot> (rev ` B)"
+unfolding conc_def image_def
+by (auto) (metis rev_append)+
+
+lemma rev_star1:
+  assumes a: "s \<in> (rev ` A)\<star>"
+  shows "s \<in> rev ` (A\<star>)"
+using a
+proof(induct rule: star_induct)
+  case (append s1 s2)
+  have inj: "inj (rev::'a list \<Rightarrow> 'a list)" 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)
+  then have "x2 @ x1 \<in> A\<star>" by (auto)
+  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 (append s1 s2)
+  have inj: "inj (rev::'a list \<Rightarrow> 'a list)" 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)
+  moreover
+  have "rev s2 \<in> (rev ` A)\<star>" by fact
+  ultimately show "rev (s1 @ s2) \<in>  (rev ` A)\<star>" by (auto)
+qed (auto)
+
+lemma rev_star [simp]:
+  shows " rev ` (A\<star>) = (rev ` A)\<star>"
+using rev_star1 rev_star2 by auto
+
+lemma rev_lang:
+  shows "rev ` (lang r) = lang (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::"'a rexp" where "A = lang r" by auto
+  then have "lang (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::"'a rexp" where eq: "lang r = A" by auto
+  have fin: "finite (pders_set B r)" by (rule finite_pders_set)
+  
+  have "Ders_set B (lang r) = (\<Union> lang ` (pders_set B r))"
+    by (simp add: Ders_set_pders_set)
+  also have "\<dots> = lang (\<Uplus>(pders_set B r))" using fin by simp
+  finally have "Ders_set B A = lang (\<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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Derivatives.thy	Mon Jul 25 13:33:38 2011 +0000
@@ -0,0 +1,490 @@
+theory Derivatives
+imports Myhill_2
+begin
+
+section {* Left-Quotients and Derivatives *}
+
+subsection {* Left-Quotients *}
+
+definition
+  Delta :: "'a lang \<Rightarrow> 'a lang"
+where
+  "Delta A = (if [] \<in> A then {[]} else {})"
+
+definition
+  Der :: "'a \<Rightarrow> 'a lang \<Rightarrow> 'a lang"
+where
+  "Der c A \<equiv> {s. [c] @ s \<in> A}"
+
+definition
+  Ders :: "'a list \<Rightarrow> 'a lang \<Rightarrow> 'a lang"
+where
+  "Ders s A \<equiv> {s'. s @ s' \<in> A}"
+
+definition
+  Ders_set :: "'a lang \<Rightarrow> 'a lang \<Rightarrow> 'a 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_zero [simp]:
+  shows "Der c {} = {}"
+unfolding Der_def
+by auto
+
+lemma Der_one [simp]:
+  shows "Der c {[]} = {}"
+unfolding Der_def
+by auto
+
+lemma Der_atom [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_conc [simp]:
+  shows "Der c (A \<cdot> B) = (Der c A) \<cdot> B \<union> (Delta A \<cdot> Der c B)"
+unfolding Der_def Delta_def conc_def
+by (auto simp add: Cons_eq_append_conv)
+
+lemma Der_star [simp]:
+  shows "Der c (A\<star>) = (Der c A) \<cdot> A\<star>"
+proof -
+  have incl: "Delta A \<cdot> Der c (A\<star>) \<subseteq> (Der c A) \<cdot> A\<star>"
+    unfolding Der_def Delta_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 \<cdot> A\<star>)"
+    by (simp only: star_cases[symmetric])
+  also have "... = Der c (A \<cdot> A\<star>)"
+    by (simp only: Der_union Der_one) (simp)
+  also have "... = (Der c A) \<cdot> A\<star> \<union> (Delta A \<cdot> Der c (A\<star>))"
+    by simp
+  also have "... =  (Der c A) \<cdot> A\<star>"
+    using incl by auto
+  finally show "Der c (A\<star>) = (Der c A) \<cdot> A\<star>" . 
+qed
+
+
+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 
+
+
+text {* Relating the Myhill-Nerode relation with left-quotients. *}
+
+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 :: "'a rexp \<Rightarrow> bool"
+where
+  "nullable (Zero) = False"
+| "nullable (One) = True"
+| "nullable (Atom c) = False"
+| "nullable (Plus r1 r2) = (nullable r1 \<or> nullable r2)"
+| "nullable (Times r1 r2) = (nullable r1 \<and> nullable r2)"
+| "nullable (Star r) = True"
+
+fun
+  der :: "'a \<Rightarrow> 'a rexp \<Rightarrow> 'a rexp"
+where
+  "der c (Zero) = Zero"
+| "der c (One) = Zero"
+| "der c (Atom c') = (if c = c' then One else Zero)"
+| "der c (Plus r1 r2) = Plus (der c r1) (der c r2)"
+| "der c (Times r1 r2) = Plus (Times (der c r1) r2) (if nullable r1 then der c r2 else Zero)"
+| "der c (Star r) = Times (der c r) (Star r)"
+
+function 
+  ders :: "'a list \<Rightarrow> 'a rexp \<Rightarrow> 'a 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 (lang r) = (if nullable r then {[]} else {})"
+unfolding Delta_def
+by (induct r) (auto simp add: conc_def split: if_splits)
+
+lemma Der_der:
+  shows "Der c (lang r) = lang (der c r)"
+by (induct r) (simp_all add: Delta_nullable)
+
+lemma Ders_ders:
+  shows "Ders s (lang r) = lang (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
+  "Times_set rs r \<equiv> {Times r' r | r'. r' \<in> rs}"
+
+fun
+  pder :: "'a \<Rightarrow> 'a rexp \<Rightarrow> ('a rexp) set"
+where
+  "pder c Zero = {Zero}"
+| "pder c One = {Zero}"
+| "pder c (Atom c') = (if c = c' then {One} else {Zero})"
+| "pder c (Plus r1 r2) = (pder c r1) \<union> (pder c r2)"
+| "pder c (Times r1 r2) = Times_set (pder c r1) r2 \<union> (if nullable r1 then pder c r2 else {})"
+| "pder c (Star r) = Times_set (pder c r) (Star r)"
+
+abbreviation
+  "pder_set c rs \<equiv> \<Union>r \<in> rs. pder c r"
+
+function 
+  pders :: "'a list \<Rightarrow> 'a rexp \<Rightarrow> ('a 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 pders_set_lang:
+  shows "(\<Union> (lang ` pder_set c rs)) = (\<Union>r \<in> rs. (\<Union>lang ` (pder c r)))"
+unfolding image_def 
+by auto
+
+lemma pders_Zero [simp]:
+  shows "pders s Zero = {Zero}"
+by (induct s rule: rev_induct) (simp_all)
+
+lemma pders_One [simp]:
+  shows "pders s One = (if s = [] then {One} else {Zero})"
+by (induct s rule: rev_induct) (auto)
+
+lemma pders_Atom [simp]:
+  shows "pders s (Atom c) = (if s = [] then {Atom c} else (if s = [c] then {One} else {Zero}))"
+by (induct s rule: rev_induct) (auto)
+
+lemma pders_Plus [simp]:
+  shows "pders s (Plus r1 r2) = (if s = [] then {Plus r1 r2} else (pders s r1) \<union> (pders s r2))"
+by (induct s rule: rev_induct) (auto)
+
+text {* Non-empty suffixes of a string *}
+
+definition
+  "Suf s \<equiv> {v. v \<noteq> [] \<and> (\<exists>u. u @ v = s)}"
+
+lemma Suf:
+  shows "Suf (s @ [c]) = (Suf s) \<cdot> {[c]} \<union> {[c]}"
+unfolding Suf_def conc_def
+by (auto simp add: append_eq_append_conv2 append_eq_Cons_conv)
+
+lemma Suf_Union:
+  shows "(\<Union>v \<in> Suf s \<cdot> {[c]}. P v) = (\<Union>v \<in> Suf s. P (v @ [c]))"
+by (auto simp add: conc_def)
+
+lemma pders_Times:
+  shows "pders s (Times r1 r2) \<subseteq> Times_set (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 (Times r1 r2) \<subseteq> Times_set (pders s r1) r2 \<union> (\<Union>v \<in> Suf s. pders v r2)" 
+    by fact
+  have "pders (s @ [c]) (Times r1 r2) = pder_set c (pders s (Times r1 r2))" by simp
+  also have "\<dots> \<subseteq> pder_set c (Times_set (pders s r1) r2 \<union> (\<Union>v \<in> Suf s. pders v r2))"
+    using ih by (auto) (blast)
+  also have "\<dots> = pder_set c (Times_set (pders s r1) r2) \<union> pder_set c (\<Union>v \<in> Suf s. pders v r2)"
+    by (simp)
+  also have "\<dots> = pder_set c (Times_set (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 (Times_set (pders s r1) r2) \<union> (pder c r2) \<union> (\<Union>v \<in> Suf s. pders (v @ [c]) r2)"
+    by (auto)
+  also have "\<dots> \<subseteq> Times_set (pder_set c (pders s r1)) r2 \<union> (pder c r2) \<union> (\<Union>v \<in> Suf s. pders (v @ [c]) r2)"
+    by (auto simp add: if_splits) (blast)
+  also have "\<dots> = Times_set (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. Times_set (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. Times_set (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. Times_set (pders v r) (Star r)))"
+      using ih[OF asm] by blast
+    also have "\<dots> = (\<Union>v\<in>Suf s. pder_set c (Times_set (pders v r) (Star r)))"
+      by simp
+    also have "\<dots> \<subseteq> (\<Union>v\<in>Suf s. (Times_set (pder_set c (pders v r)) (Star r) \<union> pder c (Star r)))"
+      by (auto split: if_splits) 
+    also have "\<dots> = (\<Union>v\<in>Suf s. (Times_set (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. (Times_set (pders (v @ [c]) r) (Star r))) \<union> (Times_set (pder c r) (Star r))"
+      by simp
+    also have "\<dots> = (\<Union>v\<in>Suf (s @ [c]). (Times_set (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_Zero:
+  shows "pders_set UNIV1 Zero = {Zero}"
+by auto
+
+lemma pders_set_One:
+  shows "pders_set UNIV1 One = {Zero}"
+by (auto split: if_splits)
+
+lemma pders_set_Atom:
+  shows "pders_set UNIV1 (Atom c) \<subseteq> {One, Zero}"
+by (auto split: if_splits)
+
+lemma pders_set_Plus:
+  shows "pders_set UNIV1 (Plus r1 r2) = pders_set UNIV1 r1 \<union> pders_set UNIV1 r2"
+by auto
+
+lemma pders_set_Times_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_Times:
+  shows "pders_set UNIV1 (Times r1 r2) \<subseteq> Times_set (pders_set UNIV1 r1) r2 \<union> pders_set UNIV1 r2"
+apply(rule UN_least)
+apply(rule subset_trans)
+apply(rule pders_Times)
+apply(simp)
+apply(rule conjI) 
+apply(auto)[1]
+apply(rule subset_trans)
+apply(rule pders_set_Times_aux)
+apply(auto)
+done
+
+lemma pders_set_Star:
+  shows "pders_set UNIV1 (Star r) \<subseteq> Times_set (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_Times_set:
+  assumes a: "finite A"
+  shows "finite (Times_set 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_One)
+apply(simp)
+apply(rule finite_subset)
+apply(rule pders_set_Atom)
+apply(simp)
+apply(simp only: pders_set_Plus)
+apply(simp)
+apply(rule finite_subset)
+apply(rule pders_set_Times)
+apply(simp only: finite_Times_set finite_Un)
+apply(simp)
+apply(rule finite_subset)
+apply(rule pders_set_Star)
+apply(simp only: finite_Times_set)
+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 finite_pders2:
+  shows "finite {pders s r | s. s \<in> A}"
+proof -
+  have "{pders s r | s. s \<in> A} \<subseteq> Pow (pders_set A r)" by auto
+  moreover
+  have "finite (Pow (pders_set A r))"
+    using finite_pders_set by simp
+  ultimately 
+  show "finite {pders s r | s. s \<in> A}"
+    by(rule finite_subset)
+qed
+
+
+subsection {* Relating left-quotients and partial derivatives *}
+
+lemma Der_pder:
+  shows "Der c (lang r) = \<Union> lang ` (pder c r)"
+by (induct r) (auto simp add: Delta_nullable conc_UNION_distrib)
+
+lemma Ders_pders:
+  shows "Ders s (lang r) = \<Union> lang ` (pders s r)"
+proof (induct s rule: rev_induct)
+  case (snoc c s)
+  have ih: "Ders s (lang r) = \<Union> lang ` (pders s r)" by fact
+  have "Ders (s @ [c]) (lang r) = Ders [c] (Ders s (lang r))"
+    by (simp add: Ders_append)
+  also have "\<dots> = Der c (\<Union> lang ` (pders s r))" using ih
+    by (simp add: Ders_singleton)
+  also have "\<dots> = (\<Union>r\<in>pders s r. Der c (lang r))" 
+    unfolding Der_def image_def by auto
+  also have "\<dots> = (\<Union>r\<in>pders s r. (\<Union> lang `  (pder c r)))"
+    by (simp add: Der_pder)
+  also have "\<dots> = (\<Union>lang ` (pder_set c (pders s r)))"
+    by (simp add: pders_set_lang)
+  also have "\<dots> = (\<Union>lang ` (pders (s @ [c]) r))"
+    by simp
+  finally show "Ders (s @ [c]) (lang r) = \<Union> lang ` pders (s @ [c]) r" .
+qed (simp add: Ders_def)
+
+lemma Ders_set_pders_set:
+  shows "Ders_set A (lang r) = (\<Union> lang ` (pders_set A r))"
+by (simp add: Ders_set_Ders Ders_pders)
+
+
+subsection {* Relating derivatives and partial derivatives *}
+
+lemma
+  shows "(\<Union> lang ` (pder c r)) = lang (der c r)"
+unfolding Der_der[symmetric] Der_pder by simp
+
+lemma
+  shows "(\<Union> lang ` (pders s r)) = lang (ders s r)"
+unfolding Ders_ders[symmetric] Ders_pders by simp
+
+
+
+subsection {*
+  The second direction of the Myhill-Nerode theorem using
+  partial derivatives.
+*}
+
+lemma Myhill_Nerode3:
+  fixes r::"'a rexp"
+  shows "finite (UNIV // \<approx>(lang r))"
+proof -
+  have "finite (UNIV // =(\<lambda>x. pders x r)=)"
+  proof - 
+    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 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>(lang r)"
+    unfolding tag_eq_rel_def
+    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>(lang r))"
+    unfolding equiv_def refl_on_def sym_def trans_def
+    unfolding str_eq_rel_def
+    by auto
+  ultimately show "finite (UNIV // \<approx>(lang r))" 
+    by (rule refined_partition_finite)
+qed
+
+end
\ No newline at end of file
--- a/Derivs.thy	Fri Jun 03 13:59:21 2011 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,492 +0,0 @@
-theory Derivs
-imports Myhill_2
-begin
-
-section {* Left-Quotients and Derivatives *}
-
-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 \<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) \<cdot> A\<star>"
-proof -
-  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 \<cdot> A\<star>)"
-    by (simp only: star_cases[symmetric])
-  also have "... = Der c (A \<cdot> A\<star>)"
-    by (simp only: Der_union Der_empty) (simp)
-  also have "... = (Der c A) \<cdot> A\<star> \<union> (Delta A \<cdot> Der c (A\<star>))"
-    by simp
-  also have "... =  (Der c A) \<cdot> A\<star>"
-    using incl by auto
-  finally show "Der c (A\<star>) = (Der c A) \<cdot> A\<star>" . 
-qed
-
-
-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_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_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_rexp r) = L_rexp (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_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 \<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_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_rexp r) = \<Union> L_rexp ` (pders s r)"
-proof (induct s rule: rev_induct)
-  case (snoc c s)
-  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_rexp ` (pders s r))" using ih
-    by (simp add: Ders_singleton)
-  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_rexp `  (pder c r)))"
-    by (simp add: Der_pder)
-  also have "\<dots> = (\<Union>L_rexp ` (pder_set c (pders s r)))"
-    by (simp add: pder_set_lang)
-  also have "\<dots> = (\<Union>L_rexp ` (pders (s @ [c]) r))"
-    by simp
-  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_rexp r) = (\<Union> L_rexp ` (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) \<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 \<cdot> {[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) (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))"
-    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 (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))"
-      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 finite_pders2:
-  shows "finite {pders s r | s. s \<in> A}"
-proof -
-  have "{pders s r | s. s \<in> A} \<subseteq> Pow (pders_set A r)" by auto
-  moreover
-  have "finite (Pow (pders_set A r))"
-    using finite_pders_set by simp
-  ultimately 
-  show "finite {pders s r | s. s \<in> A}"
-    by(rule finite_subset)
-qed
-
-
-lemma Myhill_Nerode3:
-  fixes r::"rexp"
-  shows "finite (UNIV // \<approx>(L_rexp r))"
-proof -
-  have "finite (UNIV // =(\<lambda>x. pders x r)=)"
-  proof - 
-    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 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_rexp r)"
-    unfolding tag_eq_rel_def
-    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_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_rexp r))" 
-    by (rule refined_partition_finite)
-qed
-
-
-section {* Relating derivatives and partial derivatives *}
-
-lemma
-  shows "(\<Union> L_rexp ` (pder c r)) = L_rexp (der c r)"
-unfolding Der_der[symmetric] Der_pder by simp
-
-lemma
-  shows "(\<Union> L_rexp ` (pders s r)) = L_rexp (ders s r)"
-unfolding Ders_ders[symmetric] Ders_pders by simp
-
-end
\ No newline at end of file
--- a/Journal/Paper.thy	Fri Jun 03 13:59:21 2011 +0000
+++ b/Journal/Paper.thy	Mon Jul 25 13:33:38 2011 +0000
@@ -93,6 +93,7 @@
   and contain very detailed `pencil-and-paper' proofs
   (e.g.~\cite{Kozen97}). It seems natural to exercise theorem provers by
   formalising the theorems and by verifying formally the algorithms.
+  Some of the popular theorem provers are based on Higher-Order Logic (HOL).
 
   There is however a problem: the typical approach to regular languages is to
   introduce finite automata and then define everything in terms of them.  For
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/More_Regular_Set.thy	Mon Jul 25 13:33:38 2011 +0000
@@ -0,0 +1,170 @@
+(* Author: Christian Urban, Xingyuan Zhang, Chunhan Wu *)
+theory More_Regular_Set
+imports "Regular_Exp" "Folds"
+begin
+
+text {* Some properties of operator @{text "@@"}. *}
+
+notation 
+  conc (infixr "\<cdot>" 100) and
+  star ("_\<star>" [101] 102)
+
+lemma conc_add_left:
+  assumes a: "A = B"
+  shows "C \<cdot> A = C \<cdot> B"
+using a by simp
+
+lemma star_cases:
+  shows "A\<star> =  {[]} \<union> A \<cdot> A\<star>"
+proof
+  { fix x
+    have "x \<in> A\<star> \<Longrightarrow> x \<in> {[]} \<union> A \<cdot> A\<star>"
+      unfolding conc_def
+    by (induct rule: star_induct) (auto)
+  }
+  then show "A\<star> \<subseteq> {[]} \<union> A \<cdot> A\<star>" by auto
+next
+  show "{[]} \<union> A \<cdot> A\<star> \<subseteq> A\<star>"
+    unfolding conc_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 seq_pow_comm:
+  shows "A \<cdot> (A ^^ n) = (A ^^ n) \<cdot> A"
+by (induct n) (simp_all add: conc_assoc[symmetric])
+
+lemma seq_star_comm:
+  shows "A \<cdot> A\<star> = A\<star> \<cdot> A"
+unfolding star_def seq_pow_comm conc_UNION_distrib
+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 ^^ Suc n"
+  shows "n < length s"
+using b
+proof (induct n arbitrary: s)
+  case 0
+  have "s \<in> A ^^ 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 ^^ Suc n \<Longrightarrow> n < length s" by fact
+  have "s \<in> A ^^ Suc (Suc n)" by fact
+  then obtain s1 s2 where eq: "s = s1 @ s2" and *: "s1 \<in> A" and **: "s2 \<in> A ^^ Suc n"
+    by (auto simp add: conc_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 \<cdot> (A ^^ Suc n)"
+  shows "n < length s"
+proof -
+  from b obtain s1 s2 where eq: "s = s1 @ s2" and *: "s2 \<in> A ^^ 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 \<cdot> A \<union> B"
+  shows "X = X \<cdot> (A ^^ Suc n) \<union> (\<Union>m\<in>{0..n}. B \<cdot> (A ^^ m))"
+proof (induct n)
+  case 0 
+  show "X = X \<cdot> (A ^^ Suc 0) \<union> (\<Union>(m::nat)\<in>{0..0}. B \<cdot> (A ^^ m))"
+    using eq by simp
+next
+  case (Suc n)
+  have ih: "X = X \<cdot> (A ^^ Suc n) \<union> (\<Union>m\<in>{0..n}. B \<cdot> (A ^^ m))" by fact
+  also have "\<dots> = (X \<cdot> A \<union> B) \<cdot> (A ^^ Suc n) \<union> (\<Union>m\<in>{0..n}. B \<cdot> (A ^^ m))" using eq by simp
+  also have "\<dots> = X \<cdot> (A ^^ Suc (Suc n)) \<union> (B \<cdot> (A ^^ Suc n)) \<union> (\<Union>m\<in>{0..n}. B \<cdot> (A ^^ m))"
+    by (simp add: conc_Un_distrib conc_assoc)
+  also have "\<dots> = X \<cdot> (A ^^ Suc (Suc n)) \<union> (\<Union>m\<in>{0..Suc n}. B \<cdot> (A ^^ m))"
+    by (auto simp add: le_Suc_eq)
+  finally show "X = X \<cdot> (A ^^ Suc (Suc n)) \<union> (\<Union>m\<in>{0..Suc n}. B \<cdot> (A ^^ m))" .
+qed
+
+theorem arden:
+  assumes nemp: "[] \<notin> A"
+  shows "X = X \<cdot> A \<union> B \<longleftrightarrow> X = B \<cdot> A\<star>"
+proof
+  assume eq: "X = B \<cdot> A\<star>"
+  have "A\<star> = {[]} \<union> A\<star> \<cdot> A" 
+    unfolding seq_star_comm[symmetric]
+    by (rule star_cases)
+  then have "B \<cdot> A\<star> = B \<cdot> ({[]} \<union> A\<star> \<cdot> A)"
+    by (rule conc_add_left)
+  also have "\<dots> = B \<union> B \<cdot> (A\<star> \<cdot> A)"
+    unfolding conc_Un_distrib by simp
+  also have "\<dots> = B \<union> (B \<cdot> A\<star>) \<cdot> A" 
+    by (simp only: conc_assoc)
+  finally show "X = X \<cdot> A \<union> B" 
+    using eq by blast 
+next
+  assume eq: "X = X \<cdot> A \<union> B"
+  { fix n::nat
+    have "B \<cdot> (A ^^ n) \<subseteq> X" using arden_helper[OF eq, of "n"] by auto }
+  then have "B \<cdot> A\<star> \<subseteq> X" 
+    unfolding conc_def star_def UNION_def by auto
+  moreover
+  { fix s::"'a list"
+    obtain k where "k = length s" by auto
+    then have not_in: "s \<notin> X \<cdot> (A ^^ Suc k)" 
+      using seq_pow_length[OF nemp] by blast
+    assume "s \<in> X"
+    then have "s \<in> X \<cdot> (A ^^ Suc k) \<union> (\<Union>m\<in>{0..k}. B \<cdot> (A ^^ m))"
+      using arden_helper[OF eq, of "k"] by auto
+    then have "s \<in> (\<Union>m\<in>{0..k}. B \<cdot> (A ^^ m))" using not_in by auto
+    moreover
+    have "(\<Union>m\<in>{0..k}. B \<cdot> (A ^^ m)) \<subseteq> (\<Union>n. B \<cdot> (A ^^ n))" by auto
+    ultimately 
+    have "s \<in> B \<cdot> A\<star>" 
+      unfolding conc_Un_distrib star_def by auto }
+  then have "X \<subseteq> B \<cdot> A\<star>" by auto
+  ultimately 
+  show "X = B \<cdot> A\<star>" by simp
+qed
+
+
+text {* Plus-combination for a set of regular expressions *}
+
+abbreviation
+  Setalt  ("\<Uplus>_" [1000] 999) 
+where
+  "\<Uplus>A \<equiv> folds Plus Zero A"
+
+text {* 
+  For finite sets, @{term Setalt} is preserved under @{term lang}.
+*}
+
+lemma folds_alt_simp [simp]:
+  fixes rs::"('a rexp) set"
+  assumes a: "finite rs"
+  shows "lang (\<Uplus>rs) = \<Union> (lang ` 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
--- a/My.thy	Fri Jun 03 13:59:21 2011 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,389 +0,0 @@
-theory My
-imports Main Infinite_Set
-begin
-
-
-definition
-  Seq :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ;; _" [100,100] 100)
-where 
-  "L1 ;; L2 = {s1 @ s2 | s1 s2. s1 \<in> L1 \<and> s2 \<in> L2}"
-
-inductive_set
-  Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102)
-  for L :: "string set"
-where
-  start[intro]: "[] \<in> L\<star>"
-| step[intro]:  "\<lbrakk>s1 \<in> L; s2 \<in> L\<star>\<rbrakk> \<Longrightarrow> s1@s2 \<in> L\<star>"
-
-lemma lang_star_cases:
-  shows "L\<star> =  {[]} \<union> L ;; L\<star>"
-unfolding Seq_def
-by (auto) (metis Star.simps)
-
-lemma lang_star_cases2:
-  shows "L ;; L\<star>  = L\<star> ;; L"
-sorry
-
-
-theorem ardens_revised:
-  assumes nemp: "[] \<notin> A"
-  shows "(X = X ;; A \<union> B) \<longleftrightarrow> (X = B ;; A\<star>)"
-proof
-  assume eq: "X = B ;; A\<star>"
-  have "A\<star> =  {[]} \<union> A\<star> ;; A" sorry
-  then have "B ;; A\<star> = B ;; ({[]} \<union> A\<star> ;; A)" unfolding Seq_def by simp
-  also have "\<dots> = B \<union> B ;; (A\<star> ;; A)"  unfolding Seq_def by auto
-  also have "\<dots> = B \<union> (B ;; A\<star>) ;; A"  unfolding Seq_def
-    by (auto) (metis append_assoc)+
-  finally show "X = X ;; A \<union> B" using eq by auto
-next
-  assume "X = X ;; A \<union> B"
-  then have "B \<subseteq> X" "X ;; A \<subseteq> X" by auto
-  show "X = B ;; A\<star>" sorry
-qed
-
-datatype rexp =
-  NULL
-| EMPTY
-| CHAR char
-| SEQ rexp rexp
-| ALT rexp rexp
-| STAR rexp
-
-fun
-  Sem :: "rexp \<Rightarrow> string set" ("\<lparr>_\<rparr>" [0] 1000)
-where
-    "\<lparr>NULL\<rparr> = {}"
-  | "\<lparr>EMPTY\<rparr> = {[]}"
-  | "\<lparr>CHAR c\<rparr> = {[c]}"
-  | "\<lparr>SEQ r1 r2\<rparr> = \<lparr>r1\<rparr> ;; \<lparr>r2\<rparr>"
-  | "\<lparr>ALT r1 r2\<rparr> = \<lparr>r1\<rparr> \<union> \<lparr>r2\<rparr>"
-  | "\<lparr>STAR r\<rparr> = \<lparr>r\<rparr>\<star>"
-
-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"
-
-lemma folds_simp_null [simp]:
-  "finite rs \<Longrightarrow> x \<in> \<lparr>folds ALT NULL rs\<rparr> \<longleftrightarrow> (\<exists>r \<in> rs. x \<in> \<lparr>r\<rparr>)"
-apply (simp add: folds_def)
-apply (rule someI2_ex)
-apply (erule finite_imp_fold_graph)
-apply (erule fold_graph.induct)
-apply (auto)
-done
-
-lemma folds_simp_empty [simp]:
-  "finite rs \<Longrightarrow> x \<in> \<lparr>folds ALT EMPTY rs\<rparr> \<longleftrightarrow> (\<exists>r \<in> rs. x \<in> \<lparr>r\<rparr>) \<or> x = []"
-apply (simp add: folds_def)
-apply (rule someI2_ex)
-apply (erule finite_imp_fold_graph)
-apply (erule fold_graph.induct)
-apply (auto)
-done
-
-lemma [simp]:
-  shows "(x, y) \<in> {(x, y). P x y} \<longleftrightarrow> P x y"
-by simp
-
-definition
-  str_eq ("_ \<approx>_ _")
-where
-  "x \<approx>Lang y \<equiv> (\<forall>z. x @ z \<in> Lang \<longleftrightarrow> y @ z \<in> Lang)"
-
-definition
-  str_eq_rel ("\<approx>_")
-where
-  "\<approx>Lang \<equiv> {(x, y). x \<approx>Lang y}"
-
-definition
-  final :: "string set \<Rightarrow> string set \<Rightarrow> bool"
-where
-  "final X Lang \<equiv> (X \<in> UNIV // \<approx>Lang) \<and> (\<forall>s \<in> X. s \<in> Lang)"
-
-lemma lang_is_union_of_finals: 
-  "Lang = \<Union> {X. final X Lang}"
-proof -
-  have  "Lang \<subseteq> \<Union> {X. final X Lang}"
-    unfolding final_def
-    unfolding quotient_def Image_def
-    unfolding str_eq_rel_def
-    apply(simp)
-    apply(auto)
-    apply(rule_tac x="(\<approx>Lang) `` {x}" in exI)
-    unfolding Image_def
-    unfolding str_eq_rel_def
-    apply(auto)
-    unfolding str_eq_def
-    apply(auto)
-    apply(drule_tac x="[]" in spec)
-    apply(simp)
-    done
-  moreover
-  have "\<Union>{X. final X Lang} \<subseteq> Lang" 
-    unfolding final_def by auto
-  ultimately 
-  show "Lang = \<Union> {X. final X Lang}"
-    by blast
-qed
-
-lemma all_rexp:
-  "\<lbrakk>finite (UNIV // \<approx>Lang); X \<in> (UNIV // \<approx>Lang)\<rbrakk> \<Longrightarrow> \<exists>r. X = \<lparr>r\<rparr>"
-sorry
-
-lemma final_rexp:
-  "\<lbrakk>finite (UNIV // (\<approx>Lang)); final X Lang\<rbrakk> \<Longrightarrow> \<exists>r. X = \<lparr>r\<rparr>"
-unfolding final_def
-using all_rexp by blast
-
-lemma finite_f_one_to_one:
-  assumes "finite B"
-  and "\<forall>x \<in> B. \<exists>y. f y = x"
-  shows "\<exists>rs. finite rs \<and> (B = {f y | y . y \<in> rs})"
-using assms
-by (induct) (auto)
-
-lemma finite_final:
-  assumes "finite (UNIV // (\<approx>Lang))"
-  shows "finite {X. final X Lang}"
-using assms
-proof -
-  have "{X. final X Lang} \<subseteq> (UNIV // (\<approx>Lang))"
-    unfolding final_def by auto
-  with assms show "finite {X. final X Lang}" 
-    using finite_subset by auto
-qed
-
-lemma finite_regular_aux:
-  fixes Lang :: "string set"
-  assumes "finite (UNIV // (\<approx>Lang))"
-  shows "\<exists>rs. Lang = \<lparr>folds ALT NULL rs\<rparr>"
-apply(subst lang_is_union_of_finals)
-using assms
-apply -
-apply(drule finite_final)
-apply(drule_tac f="Sem" in finite_f_one_to_one)
-apply(clarify)
-apply(drule final_rexp[OF assms])
-apply(auto)[1]
-apply(clarify)
-apply(rule_tac x="rs" in exI)
-apply(simp)
-apply(rule set_eqI)
-apply(auto)
-done
-
-lemma finite_regular:
-  fixes Lang :: "string set"
-  assumes "finite (UNIV // (\<approx>Lang))"
-  shows "\<exists>r. Lang =  \<lparr>r\<rparr>"
-using assms finite_regular_aux
-by auto
-
-
-
-section {* other direction *}
-
-
-lemma inj_image_lang:
-  fixes f::"string \<Rightarrow> 'a"
-  assumes str_inj: "\<And>x y. f x = f y \<Longrightarrow> x \<approx>Lang y"
-  shows "inj_on (image f) (UNIV // (\<approx>Lang))"
-proof - 
-  { fix x y::string
-    assume eq_tag: "f ` {z. x \<approx>Lang z} = f ` {z. y \<approx>Lang z}"
-    moreover
-    have "{z. x \<approx>Lang z} \<noteq> {}" unfolding str_eq_def by auto
-    ultimately obtain a b where "x \<approx>Lang a" "y \<approx>Lang b" "f a = f b" by blast
-    then have "x \<approx>Lang a" "y \<approx>Lang b" "a \<approx>Lang b" using str_inj by auto
-    then have "x \<approx>Lang y" unfolding str_eq_def by simp 
-    then have "{z. x \<approx>Lang z} = {z. y \<approx>Lang z}" unfolding str_eq_def by simp
-  }
-  then have "\<forall>x\<in>UNIV // \<approx>Lang. \<forall>y\<in>UNIV // \<approx>Lang. f ` x = f ` y \<longrightarrow> x = y"
-    unfolding quotient_def Image_def str_eq_rel_def by simp
-  then show "inj_on (image f) (UNIV // (\<approx>Lang))"
-    unfolding inj_on_def by simp
-qed
-
-
-lemma finite_range_image: 
-  assumes fin: "finite (range f)"
-  shows "finite ((image f) ` X)"
-proof -
-  from fin have "finite (Pow (f ` UNIV))" by auto
-  moreover
-  have "(image f) ` X \<subseteq> Pow (f ` UNIV)" by auto
-  ultimately show "finite ((image f) ` X)" using finite_subset by auto
-qed
-
-definition 
-  tag1 :: "string set \<Rightarrow> string set \<Rightarrow> string \<Rightarrow> (string set \<times> string set)"
-where
-  "tag1 L\<^isub>1 L\<^isub>2 \<equiv> \<lambda>x. ((\<approx>L\<^isub>1) `` {x}, (\<approx>L\<^isub>2) `` {x})"
-
-lemma tag1_range_finite:
-  assumes finite1: "finite (UNIV // \<approx>L\<^isub>1)"
-  and finite2: "finite (UNIV // \<approx>L\<^isub>2)"
-  shows "finite (range (tag1 L\<^isub>1 L\<^isub>2))"
-proof -
-  have "finite (UNIV // \<approx>L\<^isub>1 \<times> UNIV // \<approx>L\<^isub>2)" using finite1 finite2 by auto
-  moreover
-  have "range (tag1 L\<^isub>1 L\<^isub>2) \<subseteq> (UNIV // \<approx>L\<^isub>1) \<times> (UNIV // \<approx>L\<^isub>2)"
-    unfolding tag1_def quotient_def by auto
-  ultimately show "finite (range (tag1 L\<^isub>1 L\<^isub>2))" 
-    using finite_subset by blast
-qed
-
-lemma tag1_inj:
-  "tag1 L\<^isub>1 L\<^isub>2 x = tag1 L\<^isub>1 L\<^isub>2 y \<Longrightarrow> x \<approx>(L\<^isub>1 \<union> L\<^isub>2) y"
-unfolding tag1_def Image_def str_eq_rel_def str_eq_def
-by auto
-
-lemma quot_alt_cu:
-  fixes L\<^isub>1 L\<^isub>2::"string set"
-  assumes fin1: "finite (UNIV // \<approx>L\<^isub>1)"
-  and fin2: "finite (UNIV // \<approx>L\<^isub>2)"
-  shows "finite (UNIV // \<approx>(L\<^isub>1 \<union> L\<^isub>2))"
-proof -
-  have "finite (range (tag1 L\<^isub>1 L\<^isub>2))" 
-    using fin1 fin2 tag1_range_finite by simp
-  then have "finite (image (tag1 L\<^isub>1 L\<^isub>2) ` (UNIV // \<approx>(L\<^isub>1 \<union> L\<^isub>2)))" 
-    using finite_range_image by blast
-  moreover 
-  have "\<And>x y. tag1 L\<^isub>1 L\<^isub>2 x = tag1 L\<^isub>1 L\<^isub>2 y \<Longrightarrow> x \<approx>(L\<^isub>1 \<union> L\<^isub>2) y" 
-    using tag1_inj by simp
-  then have "inj_on (image (tag1 L\<^isub>1 L\<^isub>2)) (UNIV // \<approx>(L\<^isub>1 \<union> L\<^isub>2))" 
-    using inj_image_lang by blast
-  ultimately 
-  show "finite (UNIV // \<approx>(L\<^isub>1 \<union> L\<^isub>2))" by (rule finite_imageD)
-qed
-
-
-section {* finite \<Rightarrow> regular *}
-
-definition
-  transitions :: "string set \<Rightarrow> string set \<Rightarrow> rexp set" ("_\<Turnstile>\<Rightarrow>_")
-where
-  "Y \<Turnstile>\<Rightarrow> X \<equiv> {CHAR c | c. Y ;; {[c]} \<subseteq> X}"
-
-definition
-  transitions_rexp ("_ \<turnstile>\<rightarrow> _")
-where
-  "Y \<turnstile>\<rightarrow> X \<equiv> if [] \<in> X then folds ALT EMPTY (Y \<Turnstile>\<Rightarrow>X) else folds ALT NULL (Y \<Turnstile>\<Rightarrow>X)"
-
-definition
-  "rhs CS X \<equiv> if X = {[]} then {({[]}, EMPTY)} else {(Y, Y \<turnstile>\<rightarrow>X) | Y. Y \<in> CS}"
-
-definition
-  "rhs_sem CS X \<equiv> \<Union> {(Y;; \<lparr>r\<rparr>) | Y r . (Y, r) \<in> rhs CS X}"
-
-definition
-  "eqs CS \<equiv> (\<Union>X \<in> CS. {(X, rhs CS X)})"
-
-definition
-  "eqs_sem CS \<equiv> (\<Union>X \<in> CS. {(X, rhs_sem CS X)})"
-
-lemma [simp]:
-  shows "finite (Y \<Turnstile>\<Rightarrow> X)"
-unfolding transitions_def
-by auto
-
-
-lemma defined_by_str:
-  assumes "s \<in> X" 
-  and "X \<in> UNIV // (\<approx>Lang)"
-  shows "X = (\<approx>Lang) `` {s}"
-using assms
-unfolding quotient_def Image_def
-unfolding str_eq_rel_def str_eq_def
-by auto
-
-lemma every_eqclass_has_transition:
-  assumes has_str: "s @ [c] \<in> X"
-  and     in_CS:   "X \<in> UNIV // (\<approx>Lang)"
-  obtains Y where "Y \<in> UNIV // (\<approx>Lang)" and "Y ;; {[c]} \<subseteq> X" and "s \<in> Y"
-proof -
-  def Y \<equiv> "(\<approx>Lang) `` {s}"
-  have "Y \<in> UNIV // (\<approx>Lang)" 
-    unfolding Y_def quotient_def by auto
-  moreover
-  have "X = (\<approx>Lang) `` {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 (auto) (simp add: str_eq_def)
-  moreover
-  have "s \<in> Y" unfolding Y_def 
-    unfolding Image_def str_eq_rel_def str_eq_def by simp
-  (*moreover 
-  have "True" by simp FIXME *) 
-  ultimately show thesis by (blast intro: that)
-qed
-
-lemma test:
-  assumes "[] \<in> X"
-  shows "[] \<in> \<lparr>Y \<turnstile>\<rightarrow> X\<rparr>"
-using assms
-by (simp add: transitions_rexp_def)
-
-lemma rhs_sem:
-  assumes "X \<in> (UNIV // (\<approx>Lang))"
-  shows "X \<subseteq> rhs_sem (UNIV // (\<approx>Lang)) X"
-apply(case_tac "X = {[]}")
-apply(simp)
-apply(simp add: rhs_sem_def rhs_def Seq_def)
-apply(rule subsetI)
-apply(case_tac "x = []")
-apply(simp add: rhs_sem_def rhs_def)
-apply(rule_tac x = "X" in exI)
-apply(simp)
-apply(rule_tac x = "X" in exI)
-apply(simp add: assms)
-apply(simp add: transitions_rexp_def)
-oops
-
-
-(*
-fun
-  power :: "string \<Rightarrow> nat \<Rightarrow> string" (infixr "\<Up>" 100)
-where
-  "s \<Up> 0 = s"
-| "s \<Up> (Suc n) = s @ (s \<Up> n)"
-
-definition 
- "Lone = {(''0'' \<Up> n) @ (''1'' \<Up> n) | n. True }"
-
-lemma
-  "infinite (UNIV // (\<approx>Lone))"
-unfolding infinite_iff_countable_subset
-apply(rule_tac x="\<lambda>n. {(''0'' \<Up> n) @ (''1'' \<Up> i) | i. i \<in> {..n} }" in exI)
-apply(auto)
-prefer 2
-unfolding Lone_def
-unfolding quotient_def
-unfolding Image_def
-apply(simp)
-unfolding str_eq_rel_def
-unfolding str_eq_def
-apply(auto)
-apply(rule_tac x="''0'' \<Up> n" in exI)
-apply(auto)
-unfolding infinite_nat_iff_unbounded
-unfolding Lone_def
-*)
-
-
-
-text {* Derivatives *}
-
-definition
-  DERS :: "string \<Rightarrow> string set \<Rightarrow> string set"
-where
-  "DERS s L \<equiv> {s'. s @ s' \<in> L}"
-
-lemma
-  shows "x \<approx>L y \<longleftrightarrow> DERS x L = DERS y L"
-unfolding DERS_def str_eq_def
-by auto
\ No newline at end of file
--- a/Myhill.thy	Fri Jun 03 13:59:21 2011 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,331 +0,0 @@
-theory Myhill
-  imports Myhill_2
-begin
-
-section {* Preliminaries \label{sec_prelim}*}
-
-subsection {* Finite automata and \mht  \label{sec_fa_mh} *}
-
-text {*
-
-A {\em determinisitc finite automata (DFA)} $M$ is a 5-tuple 
-$(Q, \Sigma, \delta, s, F)$, where:
-\begin{enumerate}
-  \item $Q$ is a finite set of {\em states}, also denoted $Q_M$.
-  \item $\Sigma$ is a finite set of {\em alphabets}, also denoted $\Sigma_M$.
-  \item $\delta$ is a {\em transition function} of type @{text "Q \<times> \<Sigma> \<Rightarrow> Q"} (a total function),
-           also denoted $\delta_M$.
-  \item @{text "s \<in> Q"} is a state called {\em initial state}, also denoted $s_M$.
-  \item @{text "F \<subseteq> Q"} is a set of states named {\em accepting states}, also denoted $F_M$.
-\end{enumerate}
-Therefore, we have $M = (Q_M, \Sigma_M, \delta_M, s_M, F_M)$. Every DFA $M$ can be interpreted as 
-a function assigning states to strings, denoted $\dfa{M}$, the  definition of which is as the following:
-\begin{equation}
-\begin{aligned}
-  \dfa{M}([]) &\equiv s_M \\
-   \dfa{M}(xa) &\equiv  \delta_M(\dfa{M}(x), a)
-\end{aligned}
-\end{equation}
-A string @{text "x"} is said to be {\em accepted} (or {\em recognized}) by a DFA $M$ if 
-$\dfa{M}(x) \in F_M$. The language recoginzed by DFA $M$, denoted
-$L(M)$, is defined as:
-\begin{equation}
-  L(M) \equiv \{x~|~\dfa{M}(x) \in F_M\}
-\end{equation}
-The standard way of specifying a laugage $\Lang$ as {\em regular} is by stipulating that:
-$\Lang = L(M)$ for some DFA $M$.
-
-For any DFA $M$, the DFA obtained by changing initial state to another $p \in Q_M$ is denoted $M_p$, 
-which is defined as:
-\begin{equation}
-    M_p \ \equiv\ (Q_M, \Sigma_M, \delta_M, p, F_M) 
-\end{equation}
-Two states $p, q \in Q_M$ are said to be {\em equivalent}, denoted $p \approx_M q$, iff.
-\begin{equation}\label{m_eq_def}
-  L(M_p) = L(M_q)
-\end{equation}
-It is obvious that $\approx_M$ is an equivalent relation over $Q_M$. and 
-the partition induced by $\approx_M$ has $|Q_M|$ equivalent classes.
-By overloading $\approx_M$,  an equivalent relation over strings can be defined:
-\begin{equation}
-  x \approx_M y ~ ~ \equiv ~ ~ \dfa{M}(x) \approx_M \dfa{M}(y)
-\end{equation}
-It can be proved that the the partition induced by $\approx_M$ also has $|Q_M|$ equivalent classes.
-It is also easy to show that: if $x \approx_M y$, then $x \approx_{L(M)} y$, and this means
-$\approx_M$ is a more refined equivalent relation than $\approx_{L(M)}$. Since partition induced by
-$\approx_M$ is finite, the one induced by $\approx_{L(M)}$ must also be finite, and this is 
-one of the two directions of \mht:
-\begin{Lem}[\mht , Direction two]
-  If a language $\Lang$ is regular (i.e. $\Lang = L(M)$ for some DFA $M$), then 
-  the partition induced by $\approx_\Lang$ is finite.
-\end{Lem}
-
-The other direction is:
-\begin{Lem}[\mht , Direction one]\label{auto_mh_d1}
-  If  the partition induced by $\approx_\Lang$ is finite, then
-  $\Lang$ is regular (i.e. $\Lang = L(M)$ for some DFA $M$).
-\end{Lem}
-The $M$ we are seeking when prove lemma \ref{auto_mh_d2} can be constructed out of $\approx_\Lang$,
-denoted $M_\Lang$ and defined as the following:
-\begin{subequations}
-\begin{eqnarray}
-  Q_{M_\Lang} ~ & \equiv & ~ \{ \cls{x}{\approx_\Lang}~|~ x \in \Sigma^* \}\\
-  \Sigma_{M_\Lang} ~ & \equiv & ~ \Sigma_M \\
-  \delta_{M_\Lang} ~ & \equiv & ~ \left (\lambda (\cls{x}{\approx_\Lang}, a).  \cls{xa}{\approx_\Lang} \right) \\
-  s_{M_\Lang} ~ & \equiv & ~ \cls{[]}{\approx_\Lang} \\
-  F_{M_\Lang} ~ & \equiv & ~ \{ \cls{x}{\approx_\Lang}~|~ x \in \Lang \}
-\end{eqnarray}
-\end{subequations}
-It can be proved that $Q_{M_\Lang}$ is indeed finite and $\Lang = L(M_\Lang)$, so lemma \ref{auto_mh_d1} holds.
-It can also be proved that $M_\Lang$ is the minimal DFA (therefore unique) which recoginzes $\Lang$.
-
-
-
-*}
-
-subsection {* The objective and the underlying intuition *}
-
-text {*
-  It is now obvious from section \ref{sec_fa_mh} that \mht\ can be established easily when
-  {\em reglar languages} are defined as ones recognized by finite automata. 
-  Under the context where the use of finite automata is forbiden, the situation is quite different.
-  The theorem now has to be expressed as:
-  \begin{Thm}[\mht , Regular expression version]
-      A language $\Lang$ is regular (i.e. $\Lang = L(\re)$ for some {\em regular expression} $\re$)
-      iff. the partition induced by $\approx_\Lang$ is finite.
-  \end{Thm}
-  The proof of this version consists of two directions (if the use of automata are not allowed):
-  \begin{description}
-    \item [Direction one:] 
-      generating a regular expression $\re$ out of the finite partition induced by $\approx_\Lang$,
-      such that $\Lang = L(\re)$.
-    \item [Direction two:]
-          showing the finiteness of the partition induced by $\approx_\Lang$, under the assmption
-          that $\Lang$ is recognized by some regular expression $\re$ (i.e. $\Lang = L(\re)$).
-  \end{description}
-  The development of these two directions consititutes the body of this paper.
-
-*}
-
-section {* Direction @{text "regular language \<Rightarrow>finite partition"} *}
-
-text {*
-  Although not used explicitly, the notion of finite autotmata and its relationship with 
-  language partition, as outlined in section \ref{sec_fa_mh}, still servers as important intuitive 
-  guides in the development of this paper.
-  For example, {\em Direction one} follows the {\em Brzozowski algebraic method}
-  used to convert finite autotmata to regular expressions, under the intuition that every 
-  partition member $\cls{x}{\approx_\Lang}$ is a state in the DFA $M_\Lang$  constructed to prove 
-  lemma \ref{auto_mh_d1} of section \ref{sec_fa_mh}.
-
-  The basic idea of Brzozowski method is to extract an equational system out of the 
-  transition relationship of the automaton in question. In the equational system, every
-  automaton state is represented by an unknown, the solution of which is expected to be 
-  a regular expresion characterizing the state in a certain sense. There are two choices of 
-  how a automaton state can be characterized.  The first is to characterize by the set of 
-  strings leading from the state in question into accepting states. 
-  The other choice is to characterize by the set of strings leading from initial state 
-  into the state in question. For the second choice, the language recognized the automaton 
-  can be characterized by the solution of initial state, while for the second choice, 
-  the language recoginzed by the automaton can be characterized by 
-  combining solutions of all accepting states by @{text "+"}. Because of the automaton 
-  used as our intuitive guide, the $M_\Lang$, the states of which are 
-  sets of strings leading from initial state, the second choice is used in this paper.
-
-  Supposing the automaton in Fig \ref{fig_auto_part_rel} is the $M_\Lang$ for some language $\Lang$, 
-  and suppose $\Sigma = \{a, b, c, d, e\}$. Under the second choice, the equational system extracted is:
-  \begin{subequations}
-  \begin{eqnarray}
-    X_0 & = & X_1 \cdot c + X_2 \cdot d + \lambda \label{x_0_def_o} \\
-    X_1 & = & X_0 \cdot a + X_1 \cdot b + X_2 \cdot d \label{x_1_def_o} \\
-    X_2 & = & X_0 \cdot b + X_1 \cdot d + X_2 \cdot a \\
-    X_3 & = & \begin{aligned}
-                 & X_0 \cdot (c + d + e) + X_1 \cdot (a + e) + X_2 \cdot (b + e) + \\
-                 & X_3 \cdot (a + b + c + d + e)
-              \end{aligned}
-  \end{eqnarray}
-  \end{subequations} 
-
-\begin{figure}[h!]
-\centering
-\begin{minipage}{0.5\textwidth}
-\scalebox{0.8}{
-\begin{tikzpicture}[ultra thick,>=latex]
-  \node[draw,circle,initial] (start) {$X_0$};
-  \node[draw,circle,accepting] at ($(start) + (3.5cm,2cm)$) (ac1) {$X_1$};
-  \node[draw,circle,accepting] at ($(start) + (3.5cm,-2cm)$) (ac2) {$X_2$};
-  \node[draw,circle] at ($(start) + (6.5cm,0cm)$) (ab) {$X_3$};
-
-  \path[->] (start) edge [bend left] node [midway, above] {$a$} (ac1);
-  \path[->] (start) edge [bend right] node [midway, below] {$b$} (ac2);
-  \path[->] (ac1) edge [loop above] node [midway, above] {$b$} (ac1);
-  \path[->] (ac2) edge [loop below] node [midway, below] {$a$} (ac2);
-  \path[->] (ac1) edge [bend right] node [midway, left] {$c$} (ac2);
-  \path[->] (ac2) edge [bend right] node [midway, right] {$c$} (ac1);
-  \path[->] (ac1) edge node [midway, sloped, above] {$d$} (start);
-  \path[->] (ac2) edge node [midway, sloped, above] {$d$} (start);
-
-  \path [draw, rounded corners,->,dashed] (start) -- ($(start) + (0cm, 3.7cm)$)
-         -- ($(ab) + (0cm, 3.7cm)$)  node[midway,above,sloped]{$\Sigma - \{a, b\}$} -- (ab);
-  \path[->,dashed] (ac1) edge node [midway, above, sloped] {$\Sigma - \{b,c,d\}$} (ab);
-  \path[->,dashed] (ac2) edge node [midway, below, sloped] {$\Sigma - \{a,c,d\}$} (ab);
-  \path[->,dashed] (ab) edge [loop right] node [midway, right] {$\Sigma$} (ab);
-\end{tikzpicture}}
-\end{minipage}
-\caption{An example automaton}\label{fig_auto_part_rel}
-\end{figure}
-
-  Every $\cdot$-item on the right side of equations describes some state transtions, except 
-  the $\lambda$ in \eqref{x_0_def_o}, which represents empty string @{text "[]"}. 
-  The reason is that: every state is characterized by the
-  set of incoming strings leading from initial state. For non-initial state, every such
-  string can be splitted into a prefix leading into a preceding state and a single character suffix 
-  transiting into from the preceding state. The exception happens at
-  initial state, where the empty string is a incoming string which can not be splitted. The $\lambda$
-  in \eqref{x_0_def_o} is introduce to repsent this indivisible string. There is one and only one
-  $\lambda$ in every equational system such obtained, becasue $[]$ can only be contaied in one
-  equivalent class (the intial state in $M_\Lang$) and equivalent classes are disjoint. 
-  
-  Suppose all unknowns ($X_0, X_1, X_2, X_3$) are  solvable, the regular expression charactering 
-  laugnage $\Lang$ is $X_1 + X_2$. This paper gives a procedure
-  by which arbitrarily picked unknown can be solved. The basic idea to solve $X_i$ is by 
-  eliminating all variables other than $X_i$ from the equational system. If
-  $X_0$ is the one picked to be solved,  variables $X_1, X_2, X_3$ have to be removed one by 
-  one.  The order to remove does not matter as long as the remaing equations are kept valid.
-  Suppose $X_1$ is the first one to remove, the action is to replace all occurences of $X_1$ 
-  in remaining equations by the right hand side of its characterizing equation, i.e. 
-  the $ X_0 \cdot a + X_1 \cdot b + X_2 \cdot d$ in \eqref{x_1_def_o}. However, because
-  of the recursive occurence of $X_1$, this replacement does not really removed $X_1$. Arden's
-  lemma is invoked to transform recursive equations like \eqref{x_1_def_o} into non-recursive 
-  ones. For example, the recursive equation \eqref{x_1_def_o} is transformed into the follwing
-  non-recursive one:
-  \begin{equation}
-     X_1  =  (X_0 \cdot a + X_2 \cdot d) \cdot b^* = X_0 \cdot (a \cdot b^*) + X_2 \cdot (d \cdot b^*)
-  \end{equation}
-  which, by Arden's lemma, still characterizes $X_1$ correctly. By subsituting
-  $(X_0 \cdot a + X_2 \cdot d) \cdot b^*$ for all $X_1$ and removing  \eqref{x_1_def_o},
-  we get:
-  \begin{subequations}
-  \begin{eqnarray}
-    X_0 & = & \begin{aligned}
-              & (X_0 \cdot (a \cdot b^*) + X_2 \cdot (d \cdot b^*)) \cdot c + 
-                X_2 \cdot d + \lambda = \\
-              & X_0 \cdot (a \cdot b^* \cdot c) + X_2 \cdot (d \cdot b^* \cdot c) + 
-                X_2 \cdot d + \lambda = \\
-              &  X_0 \cdot (a \cdot b^* \cdot c) + X_2 \cdot (d \cdot b^* \cdot c + d) + \lambda 
-              \end{aligned} \label{x_0_def_1} \\
-    X_2 & = & \begin{aligned}
-                 & X_0 \cdot b + (X_0 \cdot (a \cdot b^*) + X_2 \cdot (d \cdot b^*)) \cdot d + X_2 \cdot a = \\
-                 & X_0 \cdot b + X_0 \cdot (a \cdot b^* \cdot d) + X_2 \cdot (d \cdot b^* \cdot d) + X_2 \cdot a = \\
-                 & X_0 \cdot (b + a \cdot b^* \cdot d) + X_2 \cdot (d \cdot b^* \cdot d + a)
-              \end{aligned} \\
-    X_3 & = & \begin{aligned}
-                 & X_0 \cdot (c + d + e) + ((X_0 \cdot a + X_2 \cdot d) \cdot b^*) \cdot (a + e)\\
-                 & + X_2 \cdot (b + e) + X_3 \cdot (a + b + c + d + e) \label{x_3_def_1}
-              \end{aligned}
-  \end{eqnarray}
-  \end{subequations}  
-Suppose $X_3$ is the one to remove next, since $X_3$ dose not appear in either $X_0$ or $X_2$, 
-the removal of equation \eqref{x_3_def_1} changes nothing in the rest equations. Therefore, we get:
- \begin{subequations}
-  \begin{eqnarray}
-    X_0 & = & X_0 \cdot (a \cdot b^* \cdot c) + X_2 \cdot (d \cdot b^* \cdot c + d) + \lambda  \label{x_0_def_2} \\
-    X_2 & = & X_0 \cdot (b + a \cdot b^* \cdot d) + X_2 \cdot (d \cdot b^* \cdot d + a) \label{x_2_def_2}
-  \end{eqnarray}
-  \end{subequations}  
-Actually, since absorbing state like $X_3$ contributes nothing to the language $\Lang$, it could have been removed
-at the very beginning of this precedure without affecting the final result. Now, the last unknown to remove 
-is $X_2$ and the Arden's transformaton of \eqref{x_2_def_2} is:
-\begin{equation} \label{x_2_ardened}
-   X_2 ~ = ~  (X_0 \cdot (b + a \cdot b^* \cdot d)) \cdot (d \cdot b^* \cdot d + a)^* =
-              X_0 \cdot ((b + a \cdot b^* \cdot d) \cdot (d \cdot b^* \cdot d + a)^*)
-\end{equation}
-By substituting the right hand side of \eqref{x_2_ardened} into \eqref{x_0_def_2}, we get:
-\begin{equation}
-\begin{aligned}
-  X_0  & = && X_0 \cdot (a \cdot b^* \cdot c) + \\
-       &   && X_0 \cdot ((b + a \cdot b^* \cdot d) \cdot (d \cdot b^* \cdot d + a)^*) \cdot 
-           (d \cdot b^* \cdot c + d) + \lambda \\
-       & =  && X_0 \cdot ((a \cdot b^* \cdot c) + \\
-       &   && \hspace{3em} ((b + a \cdot b^* \cdot d) \cdot (d \cdot b^* \cdot d + a)^*) \cdot 
-           (d \cdot b^* \cdot c + d)) + \lambda 
-\end{aligned}
-\end{equation}
-By applying Arden's transformation to this, we get the solution of $X_0$ as:
-\begin{equation}
-\begin{aligned}
-  X_0  =  ((a \cdot b^* \cdot c) + 
-                ((b + a \cdot b^* \cdot d) \cdot (d \cdot b^* \cdot d + a)^*) \cdot 
-                        (d \cdot b^* \cdot c + d))^*
-\end{aligned}
-\end{equation}
-Using the same method, solutions for $X_1$ and $X_2$ can be obtained as well and the
-regular expressoin for $\Lang$ is just $X_1 + X_2$. The formalization of this procedure
-consititues the first direction of the {\em regular expression} verion of
-\mht. Detailed explaination are given in {\bf paper.pdf} and more details and comments
-can be found in the formal scripts.
-*}
-
-section {* Direction @{text "finite partition \<Rightarrow> regular language"} *}
-
-text {*
-  It is well known in the theory of regular languages that
-  the existence of finite  language partition amounts to the existence of 
-  a minimal automaton, i.e. the $M_\Lang$ constructed in section \ref{sec_prelim}, which recoginzes the 
-  same language $\Lang$. The standard way to prove the existence of finite language partition 
-  is to construct a automaton out of the regular expression which recoginzes the same language, from
-  which the existence of finite language partition follows immediately. As discussed in the introducton of 
-  {\bf paper.pdf} as well as in [5], the problem for this approach happens when automata 
-  of sub regular expressions are combined to form the automaton of the mother regular expression, 
-  no matter what kind of representation is used, the formalization is cubersome, as said 
-  by Nipkow in [5]: `{\em a more abstract mathod is clearly desirable}'. In this section, 
-  an {\em intrinsically abstract} method is given, which completely avoid the mentioning of automata,
-  let along any particular representations.
-  *}
-
-text {*
-  The main proof structure is a structural induction on regular expressions,
-  where base cases (cases for @{const "NULL"}, @{const "EMPTY"}, @{const "CHAR"}) are quite straightforward to
-  proof. Real difficulty lies in inductive cases.  By inductive hypothesis, languages defined by
-  sub-expressions induce finite partitiions. Under such hypothsis, we need to prove that the language 
-  defined by the composite regular expression gives rise to finite partion.  
-  The basic idea is to attach a tag @{text "tag(x)"} to every string 
-  @{text "x"}. The tagging fuction @{text "tag"} is carefully devised, which returns tags 
-  made of equivalent classes of the partitions induced by subexpressoins, and therefore has a finite
-  range. Let @{text "Lang"} be the composite language, it is proved that:
-  \begin{quote}
-  If strings with the same tag are equivalent with respect to @{text "Lang"}, expressed as:
-  \[
-  @{text "tag(x) = tag(y) \<Longrightarrow> x \<approx>Lang y"}
-  \]
-  then the partition induced by @{text "Lang"} must be finite.
-  \end{quote}
-  There are two arguments for this. The first goes as the following:
-  \begin{enumerate}
-    \item First, the tagging function @{text "tag"} induces an equivalent relation @{text "(=tag=)"} 
-          (defiintion of @{text "f_eq_rel"} and lemma @{text "equiv_f_eq_rel"}).
-    \item It is shown that: if the range of @{text "tag"} (denoted @{text "range(tag)"}) is finite, 
-           the partition given rise by @{text "(=tag=)"} is finite (lemma @{text "finite_eq_f_rel"}).
-           Since tags are made from equivalent classes from component partitions, and the inductive
-           hypothesis ensures the finiteness of these partitions, it is not difficult to prove
-           the finiteness of @{text "range(tag)"}.
-    \item It is proved that if equivalent relation @{text "R1"} is more refined than @{text "R2"}
-           (expressed as @{text "R1 \<subseteq> R2"}),
-           and the partition induced by @{text "R1"} is finite, then the partition induced by @{text "R2"}
-           is finite as well (lemma @{text "refined_partition_finite"}).
-    \item The injectivity assumption @{text "tag(x) = tag(y) \<Longrightarrow> x \<approx>Lang y"} implies that
-            @{text "(=tag=)"} is more refined than @{text "(\<approx>Lang)"}.
-    \item Combining the points above, we have: the partition induced by language @{text "Lang"}
-          is finite (lemma @{text "tag_finite_imageD"}).
-  \end{enumerate}
-
-We could have followed another approach given in  appendix II of Brzozowski's paper [?], where
-the set of derivatives of any regular expression can be proved to be finite. 
-Since it is easy to show that strings with same derivative are equivalent with respect to the 
-language, then the second direction follows. We believe that our
-apporoach is easy to formalize, with no need to do complicated derivation calculations
-and countings as in [???].
-*}
-
-
-end
--- a/MyhillNerode.thy	Fri Jun 03 13:59:21 2011 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1816 +0,0 @@
-theory MyhillNerode
-  imports "Main" "List_Prefix"
-begin
-
-text {* sequential composition of languages *}
-
-definition
-  lang_seq :: "string set \<Rightarrow> string set \<Rightarrow> string set" ("_ ; _" [100,100] 100)
-where 
-  "L1 ; L2 = {s1 @ s2 | s1 s2. s1 \<in> L1 \<and> s2 \<in> L2}"
-
-lemma lang_seq_empty:
-  shows "{[]} ; L = L"
-  and   "L ; {[]} = L"
-unfolding lang_seq_def by auto
-
-lemma lang_seq_null:
-  shows "{} ; L = {}"
-  and   "L ; {} = {}"
-unfolding lang_seq_def by auto
-
-lemma lang_seq_append:
-  assumes a: "s1 \<in> L1"
-  and     b: "s2 \<in> L2"
-  shows "s1@s2 \<in> L1 ; L2"
-unfolding lang_seq_def
-using a b by auto 
-
-lemma lang_seq_union:
-  shows "(L1 \<union> L2); L3 = (L1; L3) \<union> (L2; L3)"
-  and   "L1; (L2 \<union> L3) = (L1; L2) \<union> (L1; L3)"
-unfolding lang_seq_def by auto
-
-lemma lang_seq_assoc:
-  shows "(L1 ; L2) ; L3 = L1 ; (L2 ; L3)"
-unfolding lang_seq_def
-apply(auto)
-apply(metis)
-by (metis append_assoc)
-
-
-section {* Kleene star for languages defined as least fixed point *}
-
-inductive_set
-  Star :: "string set \<Rightarrow> string set" ("_\<star>" [101] 102)
-  for L :: "string set"
-where
-  start[intro]: "[] \<in> L\<star>"
-| step[intro]:  "\<lbrakk>s1 \<in> L; s2 \<in> L\<star>\<rbrakk> \<Longrightarrow> s1@s2 \<in> L\<star>"
-
-lemma lang_star_empty:
-  shows "{}\<star> = {[]}"
-by (auto elim: Star.cases)
-
-lemma lang_star_cases:
-  shows "L\<star> =  {[]} \<union> L ; L\<star>"
-proof
-  { fix x
-    have "x \<in> L\<star> \<Longrightarrow> x \<in> {[]} \<union> L ; L\<star>"
-      unfolding lang_seq_def
-    by (induct rule: Star.induct) (auto)
-  }
-  then show "L\<star> \<subseteq> {[]} \<union> L ; L\<star>" by auto
-next
-  show "{[]} \<union> L ; L\<star> \<subseteq> L\<star>" 
-    unfolding lang_seq_def by auto
-qed
-
-lemma lang_star_cases':
-  shows "L\<star> =  {[]} \<union> L\<star> ; L"
-proof
-  { fix x
-    have "x \<in> L\<star> \<Longrightarrow> x \<in> {[]} \<union> L\<star> ; L"
-      unfolding lang_seq_def
-    apply (induct rule: Star.induct)
-    apply simp
-    apply simp
-    apply (erule disjE)
-    apply (auto)[1]
-    apply (erule exE | erule conjE)+
-    apply (rule disjI2)
-    apply (rule_tac x = "s1 @ s1a" in exI)
-    by auto
-  }
-  then show "L\<star> \<subseteq> {[]} \<union> L\<star> ; L" by auto
-next
-  show "{[]} \<union> L\<star> ; L \<subseteq> L\<star>" 
-    unfolding lang_seq_def
-    apply auto
-    apply (erule Star.induct)
-    apply auto
-    apply (drule step[of _ _ "[]"])
-    by (auto intro:start)
-qed
-
-lemma lang_star_simple:
-  shows "L \<subseteq> L\<star>"
-by (subst lang_star_cases)
-   (auto simp only: lang_seq_def)
-
-lemma lang_star_prop0_aux:
-  "s2 \<in> L\<star> \<Longrightarrow> \<forall> s1. s1 \<in> L \<longrightarrow> (\<exists> s3 s4. s3 \<in> L\<star> \<and> s4 \<in> L \<and> s1 @ s2 = s3 @ s4)" 
-apply (erule Star.induct)
-apply (clarify, rule_tac x = "[]" in exI, rule_tac x = s1 in exI, simp add:start)
-apply (clarify, drule_tac x = s1 in spec)
-apply (drule mp, simp, clarify)
-apply (rule_tac x = "s1a @ s3" in exI, rule_tac x = s4 in exI)
-by auto
-
-lemma lang_star_prop0:
-  "\<lbrakk>s1 \<in> L; s2 \<in> L\<star>\<rbrakk> \<Longrightarrow> \<exists> s3 s4. s3 \<in> L\<star> \<and> s4 \<in> L \<and> s1 @ s2 = s3 @ s4" 
-by (auto dest:lang_star_prop0_aux)
-
-lemma lang_star_prop1:
-  assumes asm: "L1; L2 \<subseteq> L2" 
-  shows "L1\<star>; L2 \<subseteq> L2"
-proof -
-  { fix s1 s2
-    assume minor: "s2 \<in> L2"
-    assume major: "s1 \<in> L1\<star>"
-    then have "s1@s2 \<in> L2"
-    proof(induct rule: Star.induct)
-      case start
-      show "[]@s2 \<in> L2" using minor by simp
-    next
-      case (step s1 s1')
-      have "s1 \<in> L1" by fact
-      moreover
-      have "s1'@s2 \<in> L2" by fact
-      ultimately have "s1@(s1'@s2) \<in> L1; L2" by (auto simp add: lang_seq_def)
-      with asm have "s1@(s1'@s2) \<in> L2" by auto
-      then show "(s1@s1')@s2 \<in> L2" by simp
-    qed
-  } 
-  then show "L1\<star>; L2 \<subseteq> L2" by (auto simp add: lang_seq_def)
-qed
-
-lemma lang_star_prop2_aux:
-  "s2 \<in> L2\<star> \<Longrightarrow> \<forall> s1. s1 \<in> L1 \<and> L1 ; L2 \<subseteq> L1 \<longrightarrow> s1 @ s2 \<in> L1"
-apply (erule Star.induct, simp)
-apply (clarify, drule_tac x = "s1a @ s1" in spec)
-by (auto simp:lang_seq_def)
-
-lemma lang_star_prop2:
-  "L1; L2 \<subseteq> L1 \<Longrightarrow> L1 ; L2\<star> \<subseteq> L1"
-by (auto dest!:lang_star_prop2_aux simp:lang_seq_def)
-
-lemma lang_star_seq_subseteq: 
-  shows "L ; L\<star> \<subseteq> L\<star>"
-using lang_star_cases by blast
-
-lemma lang_star_double:
-  shows "L\<star>; L\<star> = L\<star>"
-proof
-  show "L\<star> ; L\<star> \<subseteq> L\<star>" 
-    using lang_star_prop1 lang_star_seq_subseteq by blast
-next
-  have "L\<star> \<subseteq> L\<star> \<union> L\<star>; (L; L\<star>)" by auto
-  also have "\<dots> = L\<star>;{[]} \<union> L\<star>; (L; L\<star>)" by (simp add: lang_seq_empty)
-  also have "\<dots> = L\<star>; ({[]} \<union> L; L\<star>)" by (simp only: lang_seq_union)
-  also have "\<dots> = L\<star>; L\<star>" using lang_star_cases by simp 
-  finally show "L\<star> \<subseteq> L\<star> ; L\<star>" by simp
-qed
-
-lemma lang_star_seq_subseteq': 
-  shows "L\<star>; L \<subseteq> L\<star>"
-proof -
-  have "L \<subseteq> L\<star>" by (rule lang_star_simple)
-  then have "L\<star>; L \<subseteq> L\<star>; L\<star>" by (auto simp add: lang_seq_def)
-  then show "L\<star>; L \<subseteq> L\<star>" using lang_star_double by blast
-qed
-
-lemma
-  shows "L\<star> \<subseteq> L\<star>\<star>"
-by (rule lang_star_simple)
-
-
-section {* regular expressions *}
-
-datatype rexp =
-  NULL
-| EMPTY
-| CHAR char
-| SEQ rexp rexp
-| ALT rexp rexp
-| STAR rexp
-
-
-consts L:: "'a \<Rightarrow> string set"
-
-overloading L_rexp \<equiv> "L::  rexp \<Rightarrow> string set"
-begin
-fun
-  L_rexp :: "rexp \<Rightarrow> string set"
-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{* ************ now is the codes writen by chunhan ************************************* *}
-
-section {* Arden's Lemma revised *}
-
-lemma arden_aux1:
-  assumes a: "X \<subseteq> X ; A \<union> B"
-  and     b: "[] \<notin> A"
-  shows      "x \<in> X \<Longrightarrow> x \<in> B ; A\<star>"
-apply (induct x taking:length rule:measure_induct)
-apply (subgoal_tac "x \<in> X ; A \<union> B")
-defer
-using a
-apply (auto)[1]
-apply simp
-apply (erule disjE)
-defer
-apply (auto simp add:lang_seq_def) [1]
-apply (subgoal_tac "\<exists> x1 x2. x = x1 @ x2 \<and> x1 \<in> X \<and> x2 \<in> A")
-defer
-apply (auto simp add:lang_seq_def) [1]
-apply (erule exE | erule conjE)+
-apply simp
-apply (drule_tac x = x1 in spec)
-apply (simp)
-using b
-apply -
-apply (auto)[1]
-apply (subgoal_tac "x1 @ x2 \<in> (B ; A\<star>) ; A")
-defer
-apply (auto simp add:lang_seq_def)[1]
-by (metis Un_absorb1 lang_seq_assoc lang_seq_union(2) lang_star_double lang_star_simple mem_def sup1CI)
-
-theorem ardens_revised:
-  assumes nemp: "[] \<notin> A"
-  shows "(X = X ; A \<union> B) \<longleftrightarrow> (X = B ; A\<star>)"
-apply(rule iffI)
-defer
-apply(simp)
-apply(subst lang_star_cases')
-apply(subst lang_seq_union)
-apply(simp add: lang_seq_empty)
-apply(simp add: lang_seq_assoc)
-apply(auto)[1]
-proof -
-  assume "X = X ; A \<union> B"
-  then have as1: "X ; A \<union> B \<subseteq> X" and as2: "X \<subseteq> X ; A \<union> B" by simp_all
-  from as1 have a: "X ; A \<subseteq> X" and b: "B \<subseteq> X" by simp_all
-  from b have "B; A\<star> \<subseteq> X ; A\<star>" by (auto simp add: lang_seq_def)
-  moreover
-  from a have "X ; A\<star> \<subseteq> X" 
-
-by (rule lang_star_prop2)
-  ultimately have f1: "B ; A\<star> \<subseteq> X" by simp
-  from as2 nemp
-  have f2: "X \<subseteq> B; A\<star>" using arden_aux1 by auto
-  from f1 f2 show "X = B; A\<star>" by auto
-qed
-
-
-
-section {* equiv class' definition *}
-
-definition 
-  equiv_str :: "string \<Rightarrow> string set \<Rightarrow> string \<Rightarrow> bool" ("_ \<equiv>_\<equiv> _" [100, 100, 100] 100)
-where
-  "x \<equiv>Lang\<equiv> y \<longleftrightarrow> (\<forall>z. x @ z \<in> Lang \<longleftrightarrow> y @ z \<in> Lang)"
-
-definition
-  equiv_class :: "string \<Rightarrow> (string set) \<Rightarrow> string set" ("\<lbrakk>_\<rbrakk>_" [100, 100] 100)
-where
-  "\<lbrakk>x\<rbrakk>Lang \<equiv> {y. x \<equiv>Lang\<equiv> y}"
-
-text {* Chunhan modifies Q to Quo *}
-
-definition  
-  quot :: "string set \<Rightarrow> string set \<Rightarrow> (string set) set" ("_ Quo _" [100, 100] 100)
-where
-  "L1 Quo L2 \<equiv> { \<lbrakk>x\<rbrakk>L2 | x. x \<in> L1}" 
-
-
-lemma lang_eqs_union_of_eqcls: 
-  "Lang = \<Union> {X. X \<in> (UNIV Quo Lang) \<and> (\<forall> x \<in> X. x \<in> Lang)}"
-proof
-  show "Lang \<subseteq> \<Union>{X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}"
-  proof
-    fix x
-    assume "x \<in> Lang"
-    thus "x \<in> \<Union>{X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}"
-    proof (simp add:quot_def)
-      assume "(1)": "x \<in> Lang"
-      show "\<exists>xa. (\<exists>x. xa = \<lbrakk>x\<rbrakk>Lang) \<and> (\<forall>x\<in>xa. x \<in> Lang) \<and> x \<in> xa" (is "\<exists>xa.?P xa")
-      proof -
-        have "?P (\<lbrakk>x\<rbrakk>Lang)" using "(1)"
-          by (auto simp:equiv_class_def equiv_str_def dest: spec[where  x = "[]"])
-        thus ?thesis by blast
-      qed
-    qed
-  qed   
-next
-  show "\<Union>{X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang} \<subseteq> Lang"
-    by auto
-qed
-
-lemma empty_notin_CS: "{} \<notin> UNIV Quo Lang"
-apply (clarsimp simp:quot_def equiv_class_def)
-by (rule_tac x = x in exI, auto simp:equiv_str_def)
-
-lemma no_two_cls_inters:
-  "\<lbrakk>X \<in> UNIV Quo Lang; Y \<in> UNIV Quo Lang; X \<noteq> Y\<rbrakk> \<Longrightarrow> X \<inter> Y = {}"
-by (auto simp:quot_def equiv_class_def equiv_str_def)
-
-text {* equiv_class transition *}
-definition 
-  CT :: "string set \<Rightarrow> char \<Rightarrow> string set \<Rightarrow> bool" ("_-_\<rightarrow>_" [99,99]99)
-where
-  "X-c\<rightarrow>Y \<equiv> ((X;{[c]}) \<subseteq> Y)"
-
-types t_equa_rhs = "(string set \<times> rexp) set"
-
-types t_equa = "(string set \<times> t_equa_rhs)"
-
-types t_equas = "t_equa set"
-
-text {* 
-  "empty_rhs" generates "\<lambda>" for init-state, just like "\<lambda>" for final states 
-  in Brzozowski method. But if the init-state is "{[]}" ("\<lambda>" itself) then 
-  empty set is returned, see definition of "equation_rhs" 
-*} 
-
-definition 
-  empty_rhs :: "string set \<Rightarrow> t_equa_rhs"
-where
-  "empty_rhs X \<equiv> if ([] \<in> X) then {({[]}, EMPTY)} else {}"
-
-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"
-
-definition 
-  equation_rhs :: "(string set) set \<Rightarrow> string set \<Rightarrow> t_equa_rhs"
-where
-  "equation_rhs CS X \<equiv> if (X = {[]}) then {({[]}, EMPTY)}
-                         else {(S, folds ALT NULL {CHAR c| c. S-c\<rightarrow>X} )| S. S \<in> CS} \<union> empty_rhs X"
-
-definition 
-  equations :: "(string set) set \<Rightarrow> t_equas"
-where
-  "equations CS \<equiv> {(X, equation_rhs CS X) | X. X \<in> CS}"
-
-overloading L_rhs \<equiv> "L:: t_equa_rhs \<Rightarrow> string set"
-begin
-fun L_rhs:: "t_equa_rhs \<Rightarrow> string set"
-where
-  "L_rhs rhs = {x. \<exists> X r. (X, r) \<in> rhs \<and> x \<in> X;(L r)}"
-end
-
-definition 
-  distinct_rhs :: "t_equa_rhs \<Rightarrow> bool"
-where
-  "distinct_rhs rhs \<equiv> \<forall> X reg\<^isub>1 reg\<^isub>2. (X, reg\<^isub>1) \<in> rhs \<and> (X, reg\<^isub>2) \<in> rhs \<longrightarrow> reg\<^isub>1 = reg\<^isub>2"
-
-definition
-  distinct_equas_rhs :: "t_equas \<Rightarrow> bool"
-where
-  "distinct_equas_rhs equas \<equiv> \<forall> X rhs. (X, rhs) \<in> equas \<longrightarrow> distinct_rhs rhs"
-
-definition 
-  distinct_equas :: "t_equas \<Rightarrow> bool"
-where
-  "distinct_equas equas \<equiv> \<forall> X rhs rhs'. (X, rhs) \<in> equas \<and> (X, rhs') \<in> equas \<longrightarrow> rhs = rhs'"
-
-definition 
-  seq_rhs_r :: "t_equa_rhs \<Rightarrow> rexp \<Rightarrow> t_equa_rhs"
-where
-  "seq_rhs_r rhs r \<equiv> (\<lambda>(X, reg). (X, SEQ reg r)) ` rhs"
-
-definition 
-  del_x_paired :: "('a \<times> 'b) set \<Rightarrow> 'a \<Rightarrow> ('a \<times> 'b) set"
-where
-  "del_x_paired S x \<equiv> S - {X. X \<in> S \<and> fst X = x}"
-
-definition
-  arden_variate :: "string set \<Rightarrow> rexp \<Rightarrow> t_equa_rhs \<Rightarrow> t_equa_rhs"
-where
-  "arden_variate X r rhs \<equiv> seq_rhs_r (del_x_paired rhs X) (STAR r)"
-
-definition
-  no_EMPTY_rhs :: "t_equa_rhs \<Rightarrow> bool"
-where
-  "no_EMPTY_rhs rhs \<equiv> \<forall> X r. (X, r) \<in> rhs \<and> X \<noteq> {[]} \<longrightarrow> [] \<notin> L r"
-
-definition 
-  no_EMPTY_equas :: "t_equas \<Rightarrow> bool"
-where
-  "no_EMPTY_equas equas \<equiv> \<forall> X rhs. (X, rhs) \<in> equas \<longrightarrow> no_EMPTY_rhs rhs"
-
-lemma fold_alt_null_eqs:
-  "finite rS \<Longrightarrow> x \<in> L (folds ALT NULL rS) = (\<exists> r \<in> rS. x \<in> L r)"
-apply (simp add:folds_def)
-apply (rule someI2_ex)
-apply (erule finite_imp_fold_graph)
-apply (erule fold_graph.induct)
-by auto (*??? how do this be in Isar ?? *)
-
-lemma seq_rhs_r_prop1:
-  "L (seq_rhs_r rhs r) = (L rhs);(L r)"
-apply (auto simp:seq_rhs_r_def image_def lang_seq_def)
-apply (rule_tac x = "s1 @ s1a" in exI, rule_tac x = "s2a" in exI, simp)
-apply (rule_tac x = a in exI, rule_tac x = b in exI, simp)
-apply (rule_tac x = s1 in exI, rule_tac x = s1a in exI, simp)
-apply (rule_tac x = X in exI, rule_tac x  = "SEQ ra r" in exI, simp)
-apply (rule conjI)
-apply (rule_tac x = "(X, ra)" in bexI, simp+) 
-apply (rule_tac x = s1a in exI, rule_tac x = "s2a @ s2" in exI, simp)
-apply (simp add:lang_seq_def)
-by (rule_tac x = s2a in exI, rule_tac x = s2 in exI, simp)
-
-lemma del_x_paired_prop1:  
-  "\<lbrakk>distinct_rhs rhs; (X, r) \<in> rhs\<rbrakk> \<Longrightarrow> X ; L r \<union> L (del_x_paired rhs X) = L rhs"
-  apply (simp add:del_x_paired_def)
-  apply (simp add: distinct_rhs_def)
-  apply(auto simp add: lang_seq_def)
-  apply(metis)
-  done
-
-lemma arden_variate_prop:
-  assumes "(X, rx) \<in> rhs"
-  shows "(\<forall> Y. Y \<noteq> X \<longrightarrow> (\<exists> r. (Y, r) \<in> rhs) = (\<exists> r. (Y, r) \<in> (arden_variate X rx rhs)))"
-proof (rule allI, rule impI)
-  fix Y
-  assume "(1)": "Y \<noteq> X"
-  show "(\<exists>r. (Y, r) \<in> rhs) = (\<exists>r. (Y, r) \<in> arden_variate X rx rhs)"
-  proof
-    assume "(1_1)": "\<exists>r. (Y, r) \<in> rhs"
-    show "\<exists>r. (Y, r) \<in> arden_variate X rx rhs" (is "\<exists>r. ?P r")
-    proof -
-      from "(1_1)" obtain r where "(Y, r) \<in> rhs" ..
-      hence "?P (SEQ r (STAR rx))"
-      proof (simp add:arden_variate_def image_def)
-        have "(Y, r) \<in> rhs \<Longrightarrow> (Y, r) \<in> del_x_paired rhs X"
-          by (auto simp:del_x_paired_def "(1)")
-        thus "(Y, r) \<in> rhs \<Longrightarrow> (Y, SEQ r (STAR rx)) \<in> seq_rhs_r (del_x_paired rhs X) (STAR rx)"
-          by (auto simp:seq_rhs_r_def)
-      qed
-      thus ?thesis by blast
-    qed
-  next
-    assume "(2_1)": "\<exists>r. (Y, r) \<in> arden_variate X rx rhs"
-    thus "\<exists>r. (Y, r) \<in> rhs" (is "\<exists> r. ?P r")
-      by (auto simp:arden_variate_def del_x_paired_def seq_rhs_r_def image_def)
-  qed
-qed
-
-text {*
-  arden_variate_valid:  proves variation from 
-  
-   "X = X;r + Y;ry + \<dots>" to "X = Y;(SEQ ry (STAR r)) + \<dots>" 
-
-  holds the law of "language of left equiv language of right" 
-*}
-lemma arden_variate_valid:
-  assumes X_not_empty: "X \<noteq> {[]}"
-  and     l_eq_r:   "X = L rhs"
-  and     dist: "distinct_rhs rhs"
-  and     no_empty: "no_EMPTY_rhs rhs"
-  and     self_contained: "(X, r) \<in> rhs"
-  shows   "X = L (arden_variate X r rhs)" 
-proof -
-  have "[] \<notin> L r" using no_empty X_not_empty self_contained
-    by (auto simp:no_EMPTY_rhs_def)
-  hence ardens: "X = X;(L r) \<union> (L (del_x_paired rhs X)) \<longleftrightarrow> X = (L (del_x_paired rhs X)) ; (L r)\<star>" 
-    by (rule ardens_revised)
-  have del_x: "X = X ; L r \<union> L (del_x_paired rhs X) \<longleftrightarrow> X = L rhs" using dist l_eq_r self_contained
-    by (auto dest!:del_x_paired_prop1)
-  show ?thesis
-  proof
-    show "X \<subseteq> L (arden_variate X r rhs)"
-    proof
-      fix x
-      assume "(1_1)": "x \<in> X" with l_eq_r ardens del_x
-      show "x \<in> L (arden_variate X r rhs)" 
-        by (simp add:arden_variate_def seq_rhs_r_prop1 del:L_rhs.simps)
-    qed
-  next
-    show "L (arden_variate X r rhs) \<subseteq> X"
-    proof
-      fix x
-      assume "(2_1)": "x \<in> L (arden_variate X r rhs)" with ardens del_x l_eq_r
-      show "x \<in> X" 
-        by (simp add:arden_variate_def seq_rhs_r_prop1 del:L_rhs.simps)
-    qed
-  qed
-qed
-
-text {* 
-  merge_rhs {(x1, r1), (x2, r2}, (x4, r4), \<dots>} {(x1, r1'), (x3, r3'), \<dots>} = 
-     {(x1, ALT r1 r1'}, (x2, r2), (x3, r3'), (x4, r4), \<dots>} *}  
-definition 
-  merge_rhs :: "t_equa_rhs \<Rightarrow> t_equa_rhs \<Rightarrow> t_equa_rhs"
-where
-  "merge_rhs rhs rhs' \<equiv> {(X, r). (\<exists> r1 r2. ((X,r1) \<in> rhs \<and> (X, r2) \<in> rhs') \<and> r = ALT r1 r2)  \<or>
-                                 (\<exists> r1. (X, r1) \<in> rhs \<and> (\<not> (\<exists> r2. (X, r2) \<in> rhs')) \<and> r = r1) \<or>
-                                 (\<exists> r2. (X, r2) \<in> rhs' \<and> (\<not> (\<exists> r1. (X, r1) \<in> rhs)) \<and> r = r2)    }"                                  
-
-
-text {* rhs_subst rhs X=xrhs r: substitude all occurence of X in rhs((X,r) \<in> rhs) with xrhs *}
-definition 
-  rhs_subst :: "t_equa_rhs \<Rightarrow> string set \<Rightarrow> t_equa_rhs \<Rightarrow> rexp \<Rightarrow> t_equa_rhs"
-where
-  "rhs_subst rhs X xrhs r \<equiv> merge_rhs (del_x_paired rhs X) (seq_rhs_r xrhs r)"
-
-definition 
-  equas_subst_f :: "string set \<Rightarrow> t_equa_rhs \<Rightarrow> t_equa \<Rightarrow> t_equa"
-where
-  "equas_subst_f X xrhs equa \<equiv> let (Y, rhs) = equa in
-                                 if (\<exists> r. (X, r) \<in> rhs)
-                                 then (Y, rhs_subst rhs X xrhs (SOME r. (X, r) \<in> rhs))
-                                 else equa"
-
-definition
-  equas_subst :: "t_equas \<Rightarrow> string set \<Rightarrow> t_equa_rhs \<Rightarrow> t_equas"
-where
-  "equas_subst ES X xrhs \<equiv> del_x_paired (equas_subst_f X xrhs ` ES) X"
-
-lemma lang_seq_prop1:
- "x \<in> X ; L r \<Longrightarrow> x \<in> X ; (L r \<union> L r')"
-by (auto simp:lang_seq_def)
-
-lemma lang_seq_prop1':
-  "x \<in> X; L r \<Longrightarrow> x \<in> X ; (L r' \<union> L r)"
-by (auto simp:lang_seq_def)
-
-lemma lang_seq_prop2:
-  "x \<in> X; (L r \<union> L r') \<Longrightarrow> x \<in> X;L r \<or> x \<in> X;L r'"
-by (auto simp:lang_seq_def)
-
-lemma merge_rhs_prop1:
-  shows "L (merge_rhs rhs rhs') = L rhs \<union> L rhs' "
-apply (auto simp add:merge_rhs_def dest!:lang_seq_prop2 intro:lang_seq_prop1)
-apply (rule_tac x = X in exI, rule_tac x = r1 in exI, simp)
-apply (case_tac "\<exists> r2. (X, r2) \<in> rhs'")
-apply (clarify, rule_tac x = X in exI, rule_tac x = "ALT r r2" in exI, simp add:lang_seq_prop1)
-apply (rule_tac x = X in exI, rule_tac x = r in exI, simp)
-apply (case_tac "\<exists> r1. (X, r1) \<in> rhs")
-apply (clarify, rule_tac x = X in exI, rule_tac x = "ALT r1 r" in exI, simp add:lang_seq_prop1')
-apply (rule_tac x = X in exI, rule_tac x = r in exI, simp)
-done
-
-lemma no_EMPTY_rhss_imp_merge_no_EMPTY:
-  "\<lbrakk>no_EMPTY_rhs rhs; no_EMPTY_rhs rhs'\<rbrakk> \<Longrightarrow> no_EMPTY_rhs (merge_rhs rhs rhs')"
-apply (simp add:no_EMPTY_rhs_def merge_rhs_def)
-apply (clarify, (erule conjE | erule exE | erule disjE)+)
-by auto
-
-lemma distinct_rhs_prop:
-  "\<lbrakk>distinct_rhs rhs; (X, r1) \<in> rhs; (X, r2) \<in> rhs\<rbrakk> \<Longrightarrow> r1 = r2"
-by (auto simp:distinct_rhs_def)
-
-lemma merge_rhs_prop2:
-  assumes dist_rhs: "distinct_rhs rhs"
-  and     dist_rhs':"distinct_rhs rhs'"
-  shows "distinct_rhs (merge_rhs rhs rhs')"
-apply (auto simp:merge_rhs_def distinct_rhs_def)
-using dist_rhs
-apply (drule distinct_rhs_prop, simp+)
-using dist_rhs'
-apply (drule distinct_rhs_prop, simp+)
-using dist_rhs
-apply (drule distinct_rhs_prop, simp+)
-using dist_rhs'
-apply (drule distinct_rhs_prop, simp+)
-done
-
-lemma seq_rhs_r_holds_distinct: 
-  "distinct_rhs rhs \<Longrightarrow> distinct_rhs (seq_rhs_r rhs r)"
-by (auto simp:distinct_rhs_def seq_rhs_r_def)
-
-lemma seq_rhs_r_prop0:
-  assumes l_eq_r: "X = L xrhs"
-  shows "L (seq_rhs_r xrhs r) = X ; L r "
-using l_eq_r
-by (auto simp only:seq_rhs_r_prop1)
-
-lemma rhs_subst_prop1:
-  assumes l_eq_r: "X = L xrhs"
-  and     dist:  "distinct_rhs rhs"
-  shows "(X, r) \<in> rhs \<Longrightarrow> L rhs = L (rhs_subst rhs X xrhs r)"
-apply (simp add:rhs_subst_def merge_rhs_prop1 del:L_rhs.simps)
-using l_eq_r 
-apply (drule_tac r = r in seq_rhs_r_prop0, simp del:L_rhs.simps)
-using dist
-by (auto dest!:del_x_paired_prop1 simp del:L_rhs.simps)
-
-lemma del_x_paired_holds_distinct_rhs:
-  "distinct_rhs rhs \<Longrightarrow> distinct_rhs (del_x_paired rhs x)"
-by (auto simp:distinct_rhs_def del_x_paired_def)
-
-lemma rhs_subst_holds_distinct_rhs:
-  "\<lbrakk>distinct_rhs rhs; distinct_rhs xrhs\<rbrakk> \<Longrightarrow> distinct_rhs (rhs_subst rhs X xrhs r)"
-apply (drule_tac r = r and rhs = xrhs in seq_rhs_r_holds_distinct)
-apply (drule_tac x = X in del_x_paired_holds_distinct_rhs)
-by (auto dest:merge_rhs_prop2[where rhs = "del_x_paired rhs X"] simp:rhs_subst_def)
-
-section {* myhill-nerode theorem *}
-
-definition left_eq_cls :: "t_equas \<Rightarrow> (string set) set"
-where
-  "left_eq_cls ES \<equiv> {X. \<exists> rhs. (X, rhs) \<in> ES} "
-
-definition right_eq_cls :: "t_equas \<Rightarrow> (string set) set"
-where
-  "right_eq_cls ES \<equiv> {Y. \<exists> X rhs r. (X, rhs) \<in> ES \<and> (Y, r) \<in> rhs }"
-
-definition rhs_eq_cls :: "t_equa_rhs \<Rightarrow> (string set) set"
-where
-  "rhs_eq_cls rhs \<equiv> {Y. \<exists> r. (Y, r) \<in> rhs}"
-
-definition ardenable :: "t_equa \<Rightarrow> bool"
-where
-  "ardenable equa \<equiv> let (X, rhs) = equa in 
-                      distinct_rhs rhs \<and> no_EMPTY_rhs rhs \<and> X = L rhs"
-
-text {*
-  Inv: Invairance of the equation-system, during the decrease of the equation-system, Inv holds.
-*}
-definition Inv :: "string set \<Rightarrow> t_equas \<Rightarrow> bool"
-where
-  "Inv X ES \<equiv> finite ES \<and> (\<exists> rhs. (X, rhs) \<in> ES) \<and> distinct_equas ES \<and> 
-            (\<forall> X xrhs. (X, xrhs) \<in> ES \<longrightarrow> ardenable (X, xrhs) \<and> X \<noteq> {} \<and> rhs_eq_cls xrhs \<subseteq> insert {[]} (left_eq_cls ES))"
-
-text {*
-  TCon: Termination Condition of the equation-system decreasion.
-*}
-definition TCon:: "'a set \<Rightarrow> bool"
-where
-  "TCon ES \<equiv> card ES = 1"
-
-
-text {* 
-  The following is a iteration principle, and is the main framework for the proof:
-   1: We can form the initial equation-system using "equations" defined above, and prove it has invariance Inv by lemma "init_ES_satisfy_Inv";
-   2: We can decrease the number of the equation-system using ardens_lemma_revised and substitution ("equas_subst", defined above), 
-        and prove it holds the property "step" in "wf_iter" by lemma "iteration_step"
-   and finally using property Inv and TCon to prove the myhill-nerode theorem
-  
-*}
-lemma wf_iter [rule_format]: 
-  fixes f
-  assumes step: "\<And> e. \<lbrakk>P e; \<not> Q e\<rbrakk> \<Longrightarrow> (\<exists> e'. P e' \<and>  (f(e'), f(e)) \<in> less_than)"
-  shows pe:     "P e \<longrightarrow> (\<exists> e'. P e' \<and>  Q e')"
-proof(induct e rule: wf_induct 
-           [OF wf_inv_image[OF wf_less_than, where f = "f"]], clarify)
-  fix x 
-  assume h [rule_format]: 
-    "\<forall>y. (y, x) \<in> inv_image less_than f \<longrightarrow> P y \<longrightarrow> (\<exists>e'. P e' \<and> Q e')"
-    and px: "P x"
-  show "\<exists>e'. P e' \<and> Q e'"
-  proof(cases "Q x")
-    assume "Q x" with px show ?thesis by blast
-  next
-    assume nq: "\<not> Q x"
-    from step [OF px nq]
-    obtain e' where pe': "P e'" and ltf: "(f e', f x) \<in> less_than" by auto
-    show ?thesis
-    proof(rule h)
-      from ltf show "(e', x) \<in> inv_image less_than f" 
-	by (simp add:inv_image_def)
-    next
-      from pe' show "P e'" .
-    qed
-  qed
-qed
-
-
-text {* ******BEGIN: proving the initial equation-system satisfies Inv ****** *}
-
-lemma distinct_rhs_equations:
-  "(X, xrhs) \<in> equations (UNIV Quo Lang) \<Longrightarrow> distinct_rhs xrhs"
-by (auto simp: equations_def equation_rhs_def distinct_rhs_def empty_rhs_def dest:no_two_cls_inters)
-
-lemma every_nonempty_eqclass_has_strings:
-  "\<lbrakk>X \<in> (UNIV Quo Lang); X \<noteq> {[]}\<rbrakk> \<Longrightarrow> \<exists> clist. clist \<in> X \<and> clist \<noteq> []"
-by (auto simp:quot_def equiv_class_def equiv_str_def)
-
-lemma every_eqclass_is_derived_from_empty:
-  assumes not_empty: "X \<noteq> {[]}"
-  shows "X \<in> (UNIV Quo Lang) \<Longrightarrow> \<exists> clist. {[]};{clist} \<subseteq> X \<and> clist \<noteq> []"
-using not_empty
-apply (drule_tac every_nonempty_eqclass_has_strings, simp)
-by (auto intro:exI[where x = clist] simp:lang_seq_def)
-
-lemma equiv_str_in_CS:
-  "\<lbrakk>clist\<rbrakk>Lang \<in> (UNIV Quo Lang)"
-by (auto simp:quot_def)
-
-lemma has_str_imp_defined_by_str:
-  "\<lbrakk>str \<in> X; X \<in> UNIV Quo Lang\<rbrakk> \<Longrightarrow> X = \<lbrakk>str\<rbrakk>Lang"
-by (auto simp:quot_def equiv_class_def equiv_str_def)
-
-lemma every_eqclass_has_ascendent:
-  assumes has_str: "clist @ [c] \<in> X"
-  and     in_CS:   "X \<in> UNIV Quo Lang"
-  shows "\<exists> Y. Y \<in> UNIV Quo Lang \<and> Y-c\<rightarrow>X \<and> clist \<in> Y" (is "\<exists> Y. ?P Y")
-proof -
-  have "?P (\<lbrakk>clist\<rbrakk>Lang)" 
-  proof -
-    have "\<lbrakk>clist\<rbrakk>Lang \<in> UNIV Quo Lang"       
-      by (simp add:quot_def, rule_tac x = clist in exI, simp)
-    moreover have "\<lbrakk>clist\<rbrakk>Lang-c\<rightarrow>X" 
-    proof -
-      have "X = \<lbrakk>(clist @ [c])\<rbrakk>Lang" using has_str in_CS
-        by (auto intro!:has_str_imp_defined_by_str)
-      moreover have "\<forall> sl. sl \<in> \<lbrakk>clist\<rbrakk>Lang \<longrightarrow> sl @ [c] \<in> \<lbrakk>(clist @ [c])\<rbrakk>Lang"
-        by (auto simp:equiv_class_def equiv_str_def)
-      ultimately show ?thesis unfolding CT_def lang_seq_def
-        by auto
-    qed
-    moreover have "clist \<in> \<lbrakk>clist\<rbrakk>Lang" 
-      by (auto simp:equiv_str_def equiv_class_def)
-    ultimately show "?P (\<lbrakk>clist\<rbrakk>Lang)" by simp
-  qed
-  thus ?thesis by blast
-qed
-
-lemma finite_charset_rS:
-  "finite {CHAR c |c. Y-c\<rightarrow>X}"
-by (rule_tac A = UNIV and f = CHAR in finite_surj, auto)
-
-lemma l_eq_r_in_equations:
-  assumes X_in_equas: "(X, xrhs) \<in> equations (UNIV Quo Lang)"
-  shows "X = L xrhs"    
-proof (cases "X = {[]}")
-  case True
-  thus ?thesis using X_in_equas 
-    by (simp add:equations_def equation_rhs_def lang_seq_def)
-next
-  case False 
-  show ?thesis
-  proof 
-    show "X \<subseteq> L xrhs"
-    proof
-      fix x
-      assume "(1)": "x \<in> X"
-      show "x \<in> L xrhs"          
-      proof (cases "x = []")
-        assume empty: "x = []"
-        hence "x \<in> L (empty_rhs X)" using "(1)"
-          by (auto simp:empty_rhs_def lang_seq_def)
-        thus ?thesis using X_in_equas False empty "(1)" 
-          unfolding equations_def equation_rhs_def by auto
-      next
-        assume not_empty: "x \<noteq> []"
-        hence "\<exists> clist c. x = clist @ [c]" by (case_tac x rule:rev_cases, auto)
-        then obtain clist c where decom: "x = clist @ [c]" by blast
-        moreover have "\<And> Y. \<lbrakk>Y \<in> UNIV Quo Lang; Y-c\<rightarrow>X; clist \<in> Y\<rbrakk>
-          \<Longrightarrow> [c] \<in> L (folds ALT NULL {CHAR c |c. Y-c\<rightarrow>X})"
-        proof -
-          fix Y
-          assume Y_is_eq_cl: "Y \<in> UNIV Quo Lang"
-            and Y_CT_X: "Y-c\<rightarrow>X"
-            and clist_in_Y: "clist \<in> Y"
-          with finite_charset_rS 
-          show "[c] \<in> L (folds ALT NULL {CHAR c |c. Y-c\<rightarrow>X})"
-            by (auto simp :fold_alt_null_eqs)
-        qed
-        hence "\<exists>Xa. Xa \<in> UNIV Quo Lang \<and> clist @ [c] \<in> Xa ; L (folds ALT NULL {CHAR c |c. Xa-c\<rightarrow>X})" 
-          using X_in_equas False not_empty "(1)" decom
-          by (auto dest!:every_eqclass_has_ascendent simp:equations_def equation_rhs_def lang_seq_def)
-        then obtain Xa where 
-          "Xa \<in> UNIV Quo Lang \<and> clist @ [c] \<in> Xa ; L (folds ALT NULL {CHAR c |c. Xa-c\<rightarrow>X})" by blast
-        hence "x \<in> L {(S, folds ALT NULL {CHAR c |c. S-c\<rightarrow>X}) |S. S \<in> UNIV Quo Lang}" 
-          using X_in_equas "(1)" decom
-          by (auto simp add:equations_def equation_rhs_def intro!:exI[where x = Xa])
-        thus "x \<in> L xrhs" using X_in_equas False not_empty unfolding equations_def equation_rhs_def
-          by auto
-      qed
-    qed
-  next
-    show "L xrhs \<subseteq> X"
-    proof
-      fix x 
-      assume "(2)": "x \<in> L xrhs"
-      have "(2_1)": "\<And> s1 s2 r Xa. \<lbrakk>s1 \<in> Xa; s2 \<in> L (folds ALT NULL {CHAR c |c. Xa-c\<rightarrow>X})\<rbrakk> \<Longrightarrow> s1 @ s2 \<in> X"
-        using finite_charset_rS
-        by (auto simp:CT_def lang_seq_def fold_alt_null_eqs)
-      have "(2_2)": "\<And> s1 s2 Xa r.\<lbrakk>s1 \<in> Xa; s2 \<in> L r; (Xa, r) \<in> empty_rhs X\<rbrakk> \<Longrightarrow> s1 @ s2 \<in> X"
-        by (simp add:empty_rhs_def split:if_splits)
-      show "x \<in> X" using X_in_equas False "(2)"         
-        by (auto intro:"(2_1)" "(2_2)" simp:equations_def equation_rhs_def lang_seq_def)
-    qed
-  qed
-qed
-
-
-
-lemma no_EMPTY_equations:
-  "(X, xrhs) \<in> equations CS \<Longrightarrow> no_EMPTY_rhs xrhs"
-apply (clarsimp simp add:equations_def equation_rhs_def)
-apply (simp add:no_EMPTY_rhs_def empty_rhs_def, auto)
-apply (subgoal_tac "finite {CHAR c |c. Xa-c\<rightarrow>X}", drule_tac x = "[]" in fold_alt_null_eqs, clarsimp, rule finite_charset_rS)+
-done
-
-lemma init_ES_satisfy_ardenable:
-  "(X, xrhs) \<in> equations (UNIV Quo Lang)  \<Longrightarrow> ardenable (X, xrhs)"  
-  unfolding ardenable_def
-  by (auto intro:distinct_rhs_equations no_EMPTY_equations simp:l_eq_r_in_equations simp del:L_rhs.simps)
-
-lemma init_ES_satisfy_Inv:
-  assumes finite_CS: "finite (UNIV Quo Lang)"
-  and X_in_eq_cls: "X \<in> UNIV Quo Lang"
-  shows "Inv X (equations (UNIV Quo Lang))"
-proof -
-  have "finite (equations (UNIV Quo Lang))" using finite_CS
-    by (auto simp:equations_def)    
-  moreover have "\<exists>rhs. (X, rhs) \<in> equations (UNIV Quo Lang)" using X_in_eq_cls 
-    by (simp add:equations_def)
-  moreover have "distinct_equas (equations (UNIV Quo Lang))" 
-    by (auto simp:distinct_equas_def equations_def)
-  moreover have "\<forall>X xrhs. (X, xrhs) \<in> equations (UNIV Quo Lang) \<longrightarrow>
-                 rhs_eq_cls xrhs \<subseteq> insert {[]} (left_eq_cls (equations (UNIV Quo Lang)))" 
-    apply (simp add:left_eq_cls_def equations_def rhs_eq_cls_def equation_rhs_def)
-    by (auto simp:empty_rhs_def split:if_splits)
-  moreover have "\<forall>X xrhs. (X, xrhs) \<in> equations (UNIV Quo Lang) \<longrightarrow> X \<noteq> {}"
-    by (clarsimp simp:equations_def empty_notin_CS intro:classical)
-  moreover have "\<forall>X xrhs. (X, xrhs) \<in> equations (UNIV Quo Lang) \<longrightarrow> ardenable (X, xrhs)"
-    by (auto intro!:init_ES_satisfy_ardenable)
-  ultimately show ?thesis by (simp add:Inv_def)
-qed
-
-
-text {* *********** END: proving the initial equation-system satisfies Inv ******* *}
-
-
-text {* ****** BEGIN: proving every equation-system's iteration step satisfies Inv ***** *}
-
-lemma not_T_aux: "\<lbrakk>\<not> TCon (insert a A); x = a\<rbrakk>
-       \<Longrightarrow> \<exists>y. x \<noteq> y \<and> y \<in> insert a A "
-apply (case_tac "insert a A = {a}")
-by (auto simp:TCon_def)
-
-lemma not_T_atleast_2[rule_format]:
-  "finite S \<Longrightarrow> \<forall> x. x \<in> S \<and> (\<not> TCon S)\<longrightarrow> (\<exists> y. x \<noteq> y \<and> y \<in> S)"
-apply (erule finite.induct, simp)
-apply (clarify, case_tac "x = a")
-by (erule not_T_aux, auto)
-
-lemma exist_another_equa: 
-  "\<lbrakk>\<not> TCon ES; finite ES; distinct_equas ES; (X, rhl) \<in> ES\<rbrakk> \<Longrightarrow> \<exists> Y yrhl. (Y, yrhl) \<in> ES \<and> X \<noteq> Y"
-apply (drule not_T_atleast_2, simp)
-apply (clarsimp simp:distinct_equas_def)
-apply (drule_tac x= X in spec, drule_tac x = rhl in spec, drule_tac x = b in spec)
-by auto
-
-lemma Inv_mono_with_lambda:
-  assumes Inv_ES: "Inv X ES"
-  and X_noteq_Y:  "X \<noteq> {[]}"
-  shows "Inv X (ES - {({[]}, yrhs)})"
-proof -
-  have "finite (ES - {({[]}, yrhs)})" using Inv_ES
-    by (simp add:Inv_def)
-  moreover have "\<exists>rhs. (X, rhs) \<in> ES - {({[]}, yrhs)}" using Inv_ES X_noteq_Y
-    by (simp add:Inv_def)
-  moreover have "distinct_equas (ES - {({[]}, yrhs)})" using Inv_ES X_noteq_Y
-    apply (clarsimp simp:Inv_def distinct_equas_def)
-    by (drule_tac x = Xa in spec, simp)    
-  moreover have "\<forall>X xrhs.(X, xrhs) \<in> ES - {({[]}, yrhs)} \<longrightarrow>
-                          ardenable (X, xrhs) \<and> X \<noteq> {}" using Inv_ES
-    by (clarify, simp add:Inv_def)
-  moreover 
-  have "insert {[]} (left_eq_cls (ES - {({[]}, yrhs)})) = insert {[]} (left_eq_cls ES)"
-    by (auto simp:left_eq_cls_def)
-  hence "\<forall>X xrhs.(X, xrhs) \<in> ES - {({[]}, yrhs)} \<longrightarrow>
-                          rhs_eq_cls xrhs \<subseteq> insert {[]} (left_eq_cls (ES - {({[]}, yrhs)}))"
-    using Inv_ES by (auto simp:Inv_def)
-  ultimately show ?thesis by (simp add:Inv_def)
-qed
-
-lemma non_empty_card_prop:
-  "finite ES \<Longrightarrow> \<forall>e. e \<in> ES \<longrightarrow> card ES - Suc 0 < card ES"
-apply (erule finite.induct, simp)
-apply (case_tac[!] "a \<in> A")
-by (auto simp:insert_absorb)
-
-lemma ardenable_prop:
-  assumes not_lambda: "Y \<noteq> {[]}"
-  and ardable: "ardenable (Y, yrhs)"
-  shows "\<exists> yrhs'. Y = L yrhs' \<and> distinct_rhs yrhs' \<and> rhs_eq_cls yrhs' = rhs_eq_cls yrhs - {Y}" (is "\<exists> yrhs'. ?P yrhs'")
-proof (cases "(\<exists> reg. (Y, reg) \<in> yrhs)")
-  case True
-  thus ?thesis 
-  proof 
-    fix reg
-    assume self_contained: "(Y, reg) \<in> yrhs"
-    show ?thesis 
-    proof -
-      have "?P (arden_variate Y reg yrhs)"
-      proof -
-        have "Y = L (arden_variate Y reg yrhs)" 
-          using self_contained not_lambda ardable
-          by (rule_tac arden_variate_valid, simp_all add:ardenable_def)
-        moreover have "distinct_rhs (arden_variate Y reg yrhs)" 
-          using ardable 
-          by (auto simp:distinct_rhs_def arden_variate_def seq_rhs_r_def del_x_paired_def ardenable_def)
-        moreover have "rhs_eq_cls (arden_variate Y reg yrhs) = rhs_eq_cls yrhs - {Y}"
-        proof -
-          have "\<And> rhs r. rhs_eq_cls (seq_rhs_r rhs r) = rhs_eq_cls rhs"
-            apply (auto simp:rhs_eq_cls_def seq_rhs_r_def image_def)
-            by (rule_tac x = "SEQ ra r" in exI, rule_tac x = "(x, ra)" in bexI, simp+)
-          moreover have "\<And> rhs X. rhs_eq_cls (del_x_paired rhs X) = rhs_eq_cls rhs - {X}"
-            by (auto simp:rhs_eq_cls_def del_x_paired_def)
-          ultimately show ?thesis by (simp add:arden_variate_def)
-        qed
-        ultimately show ?thesis by simp
-      qed
-      thus ?thesis by (rule_tac x= "arden_variate Y reg yrhs" in exI, simp)
-    qed
-  qed
-next
-  case False
-  hence "(2)": "rhs_eq_cls yrhs - {Y} = rhs_eq_cls yrhs"
-    by (auto simp:rhs_eq_cls_def)
-  show ?thesis 
-  proof -
-    have "?P yrhs" using False ardable "(2)" 
-      by (simp add:ardenable_def)      
-    thus ?thesis by blast
-  qed
-qed
-
-lemma equas_subst_f_del_no_other:
-  assumes self_contained: "(Y, rhs) \<in> ES"
-  shows "\<exists> rhs'. (Y, rhs') \<in> (equas_subst_f X xrhs ` ES)" (is "\<exists> rhs'. ?P rhs'")
-proof -
-  have "\<exists> rhs'. equas_subst_f X xrhs (Y, rhs) = (Y, rhs')"
-    by (auto simp:equas_subst_f_def)
-  then obtain rhs' where "equas_subst_f X xrhs (Y, rhs) = (Y, rhs')" by blast
-  hence "?P rhs'" unfolding image_def using self_contained
-    by (auto intro:bexI[where x = "(Y, rhs)"])
-  thus ?thesis by blast
-qed
-
-lemma del_x_paired_del_only_x: 
-  "\<lbrakk>X \<noteq> Y; (X, rhs) \<in> ES\<rbrakk> \<Longrightarrow> (X, rhs) \<in> del_x_paired ES Y"
-by (auto simp:del_x_paired_def)
-
-lemma equas_subst_del_no_other:
- "\<lbrakk>(X, rhs) \<in> ES; X \<noteq> Y\<rbrakk> \<Longrightarrow> (\<exists>rhs. (X, rhs) \<in> equas_subst ES Y rhs')"
-unfolding equas_subst_def
-apply (drule_tac X = Y and xrhs = rhs' in equas_subst_f_del_no_other)
-by (erule exE, drule del_x_paired_del_only_x, auto)
-
-lemma equas_subst_holds_distinct:
-  "distinct_equas ES \<Longrightarrow> distinct_equas (equas_subst ES Y rhs')"
-apply (clarsimp simp add:equas_subst_def distinct_equas_def del_x_paired_def equas_subst_f_def)
-by (auto split:if_splits)
-
-lemma del_x_paired_dels: 
-  "(X, rhs) \<in> ES \<Longrightarrow> {Y. Y \<in> ES \<and> fst Y = X} \<inter> ES \<noteq> {}"
-by (auto)
-
-lemma del_x_paired_subset:
-  "(X, rhs) \<in> ES \<Longrightarrow> ES - {Y. Y \<in> ES \<and> fst Y = X} \<subset> ES"
-apply (drule del_x_paired_dels)
-by auto
-
-lemma del_x_paired_card_less: 
-  "\<lbrakk>finite ES; (X, rhs) \<in> ES\<rbrakk> \<Longrightarrow> card (del_x_paired ES X) < card ES"
-apply (simp add:del_x_paired_def)
-apply (drule del_x_paired_subset)
-by (auto intro:psubset_card_mono)
-
-lemma equas_subst_card_less:
-  "\<lbrakk>finite ES; (Y, yrhs) \<in> ES\<rbrakk> \<Longrightarrow> card (equas_subst ES Y rhs') < card ES"
-apply (simp add:equas_subst_def)
-apply (frule_tac h = "equas_subst_f Y rhs'" in finite_imageI)
-apply (drule_tac f = "equas_subst_f Y rhs'" in Finite_Set.card_image_le)
-apply (drule_tac X = Y and xrhs = rhs' in equas_subst_f_del_no_other,erule exE)
-by (drule del_x_paired_card_less, auto)
-
-lemma equas_subst_holds_distinct_rhs:
-  assumes   dist': "distinct_rhs yrhs'"
-  and     history: "\<forall>X xrhs. (X, xrhs) \<in> ES \<longrightarrow> ardenable (X, xrhs)"
-  and     X_in :  "(X, xrhs) \<in> equas_subst ES Y yrhs'"
-  shows "distinct_rhs xrhs"
-using X_in history
-apply (clarsimp simp:equas_subst_def del_x_paired_def)
-apply (drule_tac x = a in spec, drule_tac x = b in spec)
-apply (simp add:ardenable_def equas_subst_f_def)
-by (auto intro:rhs_subst_holds_distinct_rhs simp:dist' split:if_splits)
-
-lemma r_no_EMPTY_imp_seq_rhs_r_no_EMPTY:
-  "[] \<notin> L r \<Longrightarrow> no_EMPTY_rhs (seq_rhs_r rhs r)"
-by (auto simp:no_EMPTY_rhs_def seq_rhs_r_def lang_seq_def)
-
-lemma del_x_paired_holds_no_EMPTY:
-  "no_EMPTY_rhs yrhs \<Longrightarrow> no_EMPTY_rhs (del_x_paired yrhs Y)"
-by (auto simp:no_EMPTY_rhs_def del_x_paired_def)
-
-lemma rhs_subst_holds_no_EMPTY:
-  "\<lbrakk>no_EMPTY_rhs yrhs; (Y, r) \<in> yrhs; Y \<noteq> {[]}\<rbrakk> \<Longrightarrow> no_EMPTY_rhs (rhs_subst yrhs Y rhs' r)"
-apply (auto simp:rhs_subst_def intro!:no_EMPTY_rhss_imp_merge_no_EMPTY r_no_EMPTY_imp_seq_rhs_r_no_EMPTY del_x_paired_holds_no_EMPTY)
-by (auto simp:no_EMPTY_rhs_def)
-
-lemma equas_subst_holds_no_EMPTY:
-  assumes substor: "Y \<noteq> {[]}"
-  and history: "\<forall>X xrhs. (X, xrhs) \<in> ES \<longrightarrow> ardenable (X, xrhs)"
-  and X_in:"(X, xrhs) \<in> equas_subst ES Y yrhs'"
-  shows "no_EMPTY_rhs xrhs"
-proof-
-  from X_in have "\<exists> (Z, zrhs) \<in> ES. (X, xrhs) = equas_subst_f Y yrhs' (Z, zrhs)"
-    by (auto simp add:equas_subst_def del_x_paired_def)
-  then obtain Z zrhs where Z_in: "(Z, zrhs) \<in> ES"
-                       and X_Z: "(X, xrhs) = equas_subst_f Y yrhs' (Z, zrhs)" by blast
-  hence dist_zrhs: "distinct_rhs zrhs" using history
-    by (auto simp:ardenable_def)
-  show ?thesis
-  proof (cases "\<exists> r. (Y, r) \<in> zrhs")
-    case True
-    then obtain r where Y_in_zrhs: "(Y, r) \<in> zrhs" ..
-    hence some: "(SOME r. (Y, r) \<in> zrhs) = r" using Z_in dist_zrhs
-      by (auto simp:distinct_rhs_def)
-    hence "no_EMPTY_rhs (rhs_subst zrhs Y yrhs' r)"
-      using substor Y_in_zrhs history Z_in
-      by (rule_tac rhs_subst_holds_no_EMPTY, auto simp:ardenable_def)
-    thus ?thesis using X_Z True some
-      by (simp add:equas_subst_def equas_subst_f_def)
-  next
-    case False
-    hence "(X, xrhs) = (Z, zrhs)" using Z_in X_Z
-      by (simp add:equas_subst_f_def)
-    thus ?thesis using history Z_in
-      by (auto simp:ardenable_def)
-  qed
-qed
-
-lemma equas_subst_f_holds_left_eq_right:
-  assumes substor: "Y = L rhs'"
-  and     history: "\<forall>X xrhs. (X, xrhs) \<in> ES \<longrightarrow> distinct_rhs xrhs \<and> X = L xrhs"
-  and       subst: "(X, xrhs) = equas_subst_f Y rhs' (Z, zrhs)"
-  and     self_contained: "(Z, zrhs) \<in> ES"
-  shows "X = L xrhs"
-proof (cases "\<exists> r. (Y, r) \<in> zrhs")
-  case True
-  from True obtain r where "(1)":"(Y, r) \<in> zrhs" ..
-  show ?thesis
-  proof -
-    from history self_contained
-    have dist: "distinct_rhs zrhs" by auto
-    hence "(SOME r. (Y, r) \<in> zrhs) = r" using self_contained "(1)"
-      using distinct_rhs_def by (auto intro:some_equality)
-    moreover have "L zrhs = L (rhs_subst zrhs Y rhs' r)" using substor dist "(1)" self_contained
-      by (rule_tac rhs_subst_prop1, auto simp:distinct_equas_rhs_def)
-    ultimately show ?thesis using subst history self_contained
-      by (auto simp:equas_subst_f_def split:if_splits)
-  qed
-next
-  case False
-  thus ?thesis using history subst self_contained
-    by (auto simp:equas_subst_f_def)
-qed
-
-lemma equas_subst_holds_left_eq_right:
-  assumes history: "\<forall>X xrhs. (X, xrhs) \<in> ES \<longrightarrow> ardenable (X, xrhs)"
-  and     substor: "Y = L rhs'"
-  and     X_in :  "(X, xrhs) \<in> equas_subst ES Y yrhs'"
-  shows "\<forall>X xrhs. (X, xrhs) \<in> equas_subst ES Y rhs' \<longrightarrow> X = L xrhs"
-apply (clarsimp simp add:equas_subst_def del_x_paired_def)
-using substor
-apply (drule_tac equas_subst_f_holds_left_eq_right)
-using history
-by (auto simp:ardenable_def)
-
-lemma equas_subst_holds_ardenable:
-  assumes substor: "Y = L yrhs'"
-  and history: "\<forall>X xrhs. (X, xrhs) \<in> ES \<longrightarrow> ardenable (X, xrhs)"
-  and X_in:"(X, xrhs) \<in> equas_subst ES Y yrhs'"
-  and dist': "distinct_rhs yrhs'"
-  and not_lambda: "Y \<noteq> {[]}"
-  shows "ardenable (X, xrhs)"
-proof -
-  have "distinct_rhs xrhs" using history X_in dist' 
-    by (auto dest:equas_subst_holds_distinct_rhs)
-  moreover have "no_EMPTY_rhs xrhs" using history X_in not_lambda
-    by (auto intro:equas_subst_holds_no_EMPTY)
-  moreover have "X = L xrhs" using history substor X_in
-  by (auto dest: equas_subst_holds_left_eq_right simp del:L_rhs.simps)
-  ultimately show ?thesis using ardenable_def by simp
-qed
-
-lemma equas_subst_holds_cls_defined:
-  assumes         X_in: "(X, xrhs) \<in> equas_subst ES Y yrhs'"
-  and           Inv_ES: "Inv X' ES"
-  and            subst: "(Y, yrhs) \<in> ES"
-  and  cls_holds_but_Y: "rhs_eq_cls yrhs' = rhs_eq_cls yrhs - {Y}"
-  shows "rhs_eq_cls xrhs \<subseteq> insert {[]} (left_eq_cls (equas_subst ES Y yrhs'))"
-proof-
-  have tac: "\<lbrakk> A \<subseteq> B; C \<subseteq> D; E \<subseteq> A \<union> B\<rbrakk> \<Longrightarrow> E \<subseteq> B \<union> D" by auto
-  from X_in have "\<exists> (Z, zrhs) \<in> ES. (X, xrhs) = equas_subst_f Y yrhs' (Z, zrhs)"
-    by (auto simp add:equas_subst_def del_x_paired_def)
-  then obtain Z zrhs where Z_in: "(Z, zrhs) \<in> ES"
-                       and X_Z: "(X, xrhs) = equas_subst_f Y yrhs' (Z, zrhs)" by blast
-  hence "rhs_eq_cls zrhs \<subseteq> insert {[]} (left_eq_cls ES)" using Inv_ES
-    by (auto simp:Inv_def)
-  moreover have "rhs_eq_cls yrhs' \<subseteq> insert {[]} (left_eq_cls ES) - {Y}" 
-    using Inv_ES subst cls_holds_but_Y
-    by (auto simp:Inv_def)
-  moreover have "rhs_eq_cls xrhs \<subseteq> rhs_eq_cls zrhs \<union> rhs_eq_cls yrhs' - {Y}"
-    using X_Z cls_holds_but_Y
-    apply (clarsimp simp add:equas_subst_f_def rhs_subst_def split:if_splits)
-    by (auto simp:rhs_eq_cls_def merge_rhs_def del_x_paired_def seq_rhs_r_def)
-  moreover have "left_eq_cls (equas_subst ES Y yrhs') = left_eq_cls ES - {Y}" using subst
-    by (auto simp: left_eq_cls_def equas_subst_def del_x_paired_def equas_subst_f_def
-             dest: equas_subst_f_del_no_other
-             split: if_splits)
-  ultimately show ?thesis by blast
-qed
-
-lemma iteration_step: 
-  assumes Inv_ES: "Inv X ES"
-  and    not_T: "\<not> TCon ES"
-  shows "(\<exists> ES'. Inv X ES' \<and> (card ES', card ES) \<in> less_than)" 
-proof -
-  from Inv_ES not_T have another: "\<exists>Y yrhs. (Y, yrhs) \<in> ES \<and> X \<noteq> Y" unfolding Inv_def
-    by (clarify, rule_tac exist_another_equa[where X = X], auto)
-  then obtain Y yrhs where subst: "(Y, yrhs) \<in> ES" and not_X: " X \<noteq> Y" by blast
-  show ?thesis (is "\<exists> ES'. ?P ES'")
-  proof (cases "Y = {[]}") 
-    case True
-      --"in this situation, we pick a \"\<lambda>\" equation, thus directly remove it from the equation-system"
-    have "?P (ES - {(Y, yrhs)})" 
-    proof 
-      show "Inv X (ES - {(Y, yrhs)})" using True not_X
-        by (simp add:Inv_ES Inv_mono_with_lambda)
-    next 
-      show "(card (ES - {(Y, yrhs)}), card ES) \<in> less_than" using Inv_ES subst
-        by (auto elim:non_empty_card_prop[rule_format] simp:Inv_def)
-    qed
-    thus ?thesis by blast
-  next
-    case False
-      --"in this situation, we pick a equation and using ardenable to get a 
-        rhs without itself in it, then use equas_subst to form a new equation-system"
-    hence "\<exists> yrhs'. Y = L yrhs' \<and> distinct_rhs yrhs' \<and> rhs_eq_cls yrhs' = rhs_eq_cls yrhs - {Y}" 
-      using subst Inv_ES
-      by (auto intro:ardenable_prop simp add:Inv_def simp del:L_rhs.simps)
-    then obtain yrhs' where Y'_l_eq_r: "Y = L yrhs'"
-      and dist_Y': "distinct_rhs yrhs'"
-      and cls_holds_but_Y: "rhs_eq_cls yrhs' = rhs_eq_cls yrhs - {Y}" by blast
-    hence "?P (equas_subst ES Y yrhs')"
-    proof -
-      have finite_del: "\<And> S x. finite S \<Longrightarrow> finite (del_x_paired S x)" 
-        apply (rule_tac A = "del_x_paired S x" in finite_subset)
-        by (auto simp:del_x_paired_def)
-      have "finite (equas_subst ES Y yrhs')" using Inv_ES 
-        by (auto intro!:finite_del simp:equas_subst_def Inv_def)
-      moreover have "\<exists>rhs. (X, rhs) \<in> equas_subst ES Y yrhs'" using Inv_ES not_X
-        by (auto intro:equas_subst_del_no_other simp:Inv_def)
-      moreover have "distinct_equas (equas_subst ES Y yrhs')" using Inv_ES
-        by (auto intro:equas_subst_holds_distinct simp:Inv_def)
-      moreover have "\<forall>X xrhs. (X, xrhs) \<in> equas_subst ES Y yrhs' \<longrightarrow> ardenable (X, xrhs)"
-        using Inv_ES dist_Y' False Y'_l_eq_r
-        apply (clarsimp simp:Inv_def)
-        by (rule equas_subst_holds_ardenable, simp_all)
-      moreover have "\<forall>X xrhs. (X, xrhs) \<in> equas_subst ES Y yrhs' \<longrightarrow> X \<noteq> {}" using Inv_ES
-        by (clarsimp simp:equas_subst_def Inv_def del_x_paired_def equas_subst_f_def split:if_splits, auto)
-      moreover have "\<forall>X xrhs. (X, xrhs) \<in> equas_subst ES Y yrhs' \<longrightarrow>
-                               rhs_eq_cls xrhs \<subseteq> insert {[]} (left_eq_cls (equas_subst ES Y yrhs'))"
-        using Inv_ES subst cls_holds_but_Y
-        apply (rule_tac impI | rule_tac allI)+
-        by (erule equas_subst_holds_cls_defined, auto)
-      moreover have "(card (equas_subst ES Y yrhs'), card ES) \<in> less_than"using Inv_ES subst
-        by (simp add:equas_subst_card_less Inv_def)
-      ultimately show "?P (equas_subst ES Y yrhs')" by (auto simp:Inv_def)      
-    qed
-    thus ?thesis by blast
-  qed
-qed
-
-text {* ***** END: proving every equation-system's iteration step satisfies Inv ************** *}
-
-lemma iteration_conc: 
-  assumes history: "Inv X ES"
-  shows "\<exists> ES'. Inv X ES' \<and> TCon ES'" (is "\<exists> ES'. ?P ES'")
-proof (cases "TCon ES")
-  case True
-  hence "?P ES" using history by simp
-  thus ?thesis by blast
-next
-  case False  
-  thus ?thesis using history iteration_step
-    by (rule_tac f = card in wf_iter, simp_all)
-qed
-
-lemma eqset_imp_iff': "A = B \<Longrightarrow> \<forall> x. x \<in> A \<longleftrightarrow> x \<in> B"
-apply (auto simp:mem_def)
-done
-
-lemma set_cases2:
-  "\<lbrakk>(A = {} \<Longrightarrow> R A); \<And> x. (A = {x}) \<Longrightarrow> R A; \<And> x y. \<lbrakk>x \<noteq> y; x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> R A\<rbrakk> \<Longrightarrow> R A"
-apply (case_tac "A = {}", simp)
-by (case_tac "\<exists> x. A = {x}", clarsimp, blast)
-
-lemma rhs_aux:"\<lbrakk>distinct_rhs rhs; {Y. \<exists>r. (Y, r) \<in> rhs} = {X}\<rbrakk> \<Longrightarrow> (\<exists>r. rhs = {(X, r)})"
-apply (rule_tac A = rhs in set_cases2, simp)
-apply (drule_tac x = X in eqset_imp_iff, clarsimp)
-apply (drule eqset_imp_iff',clarsimp)
-apply (frule_tac x = a in spec, drule_tac x = aa in spec)
-by (auto simp:distinct_rhs_def)
-
-lemma every_eqcl_has_reg: 
-  assumes finite_CS: "finite (UNIV Quo Lang)"
-  and X_in_CS: "X \<in> (UNIV Quo Lang)"
-  shows "\<exists> (reg::rexp). L reg = X" (is "\<exists> r. ?E r")
-proof-
-  have "\<exists>ES'. Inv X ES' \<and> TCon ES'" using finite_CS X_in_CS
-    by (auto intro:init_ES_satisfy_Inv iteration_conc)
-  then obtain ES' where Inv_ES': "Inv X ES'" 
-                   and  TCon_ES': "TCon ES'" by blast
-  from Inv_ES' TCon_ES' 
-  have "\<exists> rhs. ES' = {(X, rhs)}"
-    apply (clarsimp simp:Inv_def TCon_def)
-    apply (rule_tac x = rhs in exI)
-    by (auto dest!:card_Suc_Diff1 simp:card_eq_0_iff)  
-  then obtain rhs where ES'_single_equa: "ES' = {(X, rhs)}" ..
-  hence X_ardenable: "ardenable (X, rhs)" using Inv_ES'
-    by (simp add:Inv_def)
-  
-  from X_ardenable have X_l_eq_r: "X = L rhs"
-    by (simp add:ardenable_def)
-  hence rhs_not_empty: "rhs \<noteq> {}" using Inv_ES' ES'_single_equa
-    by (auto simp:Inv_def ardenable_def)
-  have rhs_eq_cls: "rhs_eq_cls rhs \<subseteq> {X, {[]}}"
-    using Inv_ES' ES'_single_equa
-    by (auto simp:Inv_def ardenable_def left_eq_cls_def)
-  have X_not_empty: "X \<noteq> {}" using Inv_ES' ES'_single_equa
-    by (auto simp:Inv_def)    
-  show ?thesis
-  proof (cases "X = {[]}")
-    case True
-    hence "?E EMPTY" by (simp)
-    thus ?thesis by blast
-  next
-    case False with  X_ardenable
-    have "\<exists> rhs'. X = L rhs' \<and> rhs_eq_cls rhs' = rhs_eq_cls rhs - {X} \<and> distinct_rhs rhs'"
-      by (drule_tac ardenable_prop, auto)
-    then obtain rhs' where X_eq_rhs': "X = L rhs'"
-      and rhs'_eq_cls: "rhs_eq_cls rhs' = rhs_eq_cls rhs - {X}" 
-      and rhs'_dist : "distinct_rhs rhs'" by blast
-    have "rhs_eq_cls rhs' \<subseteq> {{[]}}" using rhs_eq_cls False rhs'_eq_cls rhs_not_empty 
-      by blast
-    hence "rhs_eq_cls rhs' = {{[]}}" using X_not_empty X_eq_rhs'
-      by (auto simp:rhs_eq_cls_def)
-    hence "\<exists> r. rhs' = {({[]}, r)}" using rhs'_dist
-      by (auto intro:rhs_aux simp:rhs_eq_cls_def)
-    then obtain r where "rhs' = {({[]}, r)}" ..
-    hence "?E r" using X_eq_rhs' by (auto simp add:lang_seq_def)
-    thus ?thesis by blast     
-  qed
-qed
-
-text {* definition of a regular language *}
-
-abbreviation
-  reg :: "string set => bool"
-where
-  "reg L1 \<equiv> (\<exists>r::rexp. L r = L1)"
-
-theorem myhill_nerode: 
-  assumes finite_CS: "finite (UNIV Quo Lang)"
-  shows   "\<exists> (reg::rexp). Lang = L reg" (is "\<exists> r. ?P r")
-proof -
-  have has_r_each: "\<forall>C\<in>{X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}. \<exists>(r::rexp). C = L r" using finite_CS
-    by (auto dest:every_eqcl_has_reg)  
-  have "\<exists> (rS::rexp set). finite rS \<and> 
-                          (\<forall> C \<in> {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}. \<exists> r \<in> rS. C = L r) \<and> 
-                          (\<forall> r \<in> rS. \<exists> C \<in> {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}. C = L r)" 
-       (is "\<exists> rS. ?Q rS")
-  proof-
-    have "\<And> C. C \<in> {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang} \<Longrightarrow> C = L (SOME (ra::rexp). C = L ra)"
-      using has_r_each
-      apply (erule_tac x = C in ballE, erule_tac exE)
-      by (rule_tac a = r in someI2, simp+)
-    hence "?Q ((\<lambda> C. SOME r. C = L r) ` {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang})" using has_r_each
-      using finite_CS by auto
-    thus ?thesis by blast    
-  qed
-  then obtain rS where finite_rS : "finite rS"
-    and has_r_each': "\<forall> C \<in> {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}. \<exists> r \<in> (rS::rexp set). C = L r"
-    and has_cl_each: "\<forall> r \<in> (rS::rexp set). \<exists> C \<in> {X \<in> UNIV Quo Lang. \<forall>x\<in>X. x \<in> Lang}. C = L r" by blast
-  have "?P (folds ALT NULL rS)"
-  proof
-    show "Lang \<subseteq> L (folds ALT NULL rS)" using finite_rS lang_eqs_union_of_eqcls[of Lang] has_r_each'
-      apply (clarsimp simp:fold_alt_null_eqs) by blast
-  next 
-    show "L (folds ALT NULL rS) \<subseteq> Lang" using finite_rS lang_eqs_union_of_eqcls[of Lang] has_cl_each
-      by (clarsimp simp:fold_alt_null_eqs)
-  qed
-  thus ?thesis by blast
-qed 
-
-
-text {* tests by Christian *}
-
-(* Alternative definition for Quo *)
-definition 
-  QUOT :: "string set \<Rightarrow> (string set) set"  
-where
-  "QUOT Lang \<equiv> (\<Union>x. {\<lbrakk>x\<rbrakk>Lang})"
-
-lemma test: 
-  "UNIV Quo Lang = QUOT Lang"
-by (auto simp add: quot_def QUOT_def)
-
-lemma test1:
-  assumes finite_CS: "finite (QUOT Lang)"
-  shows "reg Lang"
-using finite_CS
-unfolding test[symmetric]
-by (auto dest: myhill_nerode)
-
-lemma cons_one: "x @ y \<in> {z} \<Longrightarrow> x @ y = z"
-by simp
-
-lemma quot_lambda: "QUOT {[]} = {{[]}, UNIV - {[]}}"
-proof 
-  show "QUOT {[]} \<subseteq> {{[]}, UNIV - {[]}}"
-  proof 
-    fix x 
-    assume in_quot: "x \<in> QUOT {[]}"
-    show "x \<in> {{[]}, UNIV - {[]}}"
-    proof -
-      from in_quot 
-      have "x = {[]} \<or> x = UNIV - {[]}"
-        unfolding QUOT_def equiv_class_def
-      proof 
-        fix xa
-        assume in_univ: "xa \<in> UNIV"
-           and in_eqiv: "x \<in> {{y. xa \<equiv>{[]}\<equiv> y}}"
-        show "x = {[]} \<or> x = UNIV - {[]}"
-        proof(cases "xa = []")
-          case True
-          hence "{y. xa \<equiv>{[]}\<equiv> y} = {[]}" using in_eqiv
-            by (auto simp add:equiv_str_def)
-          thus ?thesis using in_eqiv by (rule_tac disjI1, simp)
-        next
-          case False
-          hence "{y. xa \<equiv>{[]}\<equiv> y} = UNIV - {[]}" using in_eqiv
-            by (auto simp:equiv_str_def)
-          thus ?thesis using in_eqiv by (rule_tac disjI2, simp)
-        qed
-      qed
-      thus ?thesis by simp
-    qed
-  qed
-next
-  show "{{[]}, UNIV - {[]}} \<subseteq> QUOT {[]}"
-  proof
-    fix x
-    assume in_res: "x \<in> {{[]}, (UNIV::string set) - {[]}}"
-    show "x \<in> (QUOT {[]})"
-    proof -
-      have "x = {[]} \<Longrightarrow> x \<in> QUOT {[]}"
-        apply (simp add:QUOT_def equiv_class_def equiv_str_def)
-        by (rule_tac x = "[]" in exI, auto)
-      moreover have "x = UNIV - {[]} \<Longrightarrow> x \<in> QUOT {[]}"
-        apply (simp add:QUOT_def equiv_class_def equiv_str_def)
-        by (rule_tac x = "''a''" in exI, auto)
-      ultimately show ?thesis using in_res by blast
-    qed
-  qed
-qed
-
-lemma quot_single_aux: "\<lbrakk>x \<noteq> []; x \<noteq> [c]\<rbrakk> \<Longrightarrow> x @ z \<noteq> [c]"
-by (induct x, auto)
-
-lemma quot_single: "\<And> (c::char). QUOT {[c]} = {{[]}, {[c]}, UNIV - {[], [c]}}"
-proof - 
-  fix c::"char" 
-  have exist_another: "\<exists> a. a \<noteq> c" 
-    apply (case_tac "c = CHR ''a''")
-    apply (rule_tac x = "CHR ''b''" in exI, simp)
-    by (rule_tac x = "CHR ''a''" in exI, simp)  
-  show "QUOT {[c]} = {{[]}, {[c]}, UNIV - {[], [c]}}"
-  proof
-    show "QUOT {[c]} \<subseteq> {{[]},{[c]}, UNIV - {[], [c]}}"
-    proof 
-      fix x 
-      assume in_quot: "x \<in> QUOT {[c]}"
-      show "x \<in> {{[]}, {[c]}, UNIV - {[],[c]}}"
-      proof -
-        from in_quot 
-        have "x = {[]} \<or> x = {[c]} \<or> x = UNIV - {[],[c]}"
-          unfolding QUOT_def equiv_class_def
-        proof 
-          fix xa
-          assume in_eqiv: "x \<in> {{y. xa \<equiv>{[c]}\<equiv> y}}"
-          show "x = {[]} \<or> x = {[c]} \<or> x = UNIV - {[], [c]}"
-          proof-
-            have "xa = [] \<Longrightarrow> x = {[]}" using in_eqiv 
-              by (auto simp add:equiv_str_def)
-            moreover have "xa = [c] \<Longrightarrow> x = {[c]}"
-            proof -
-              have "xa = [c] \<Longrightarrow> {y. xa \<equiv>{[c]}\<equiv> y} = {[c]}" using in_eqiv
-                apply (simp add:equiv_str_def)
-                apply (rule set_ext, rule iffI, simp)
-                apply (drule_tac x = "[]" in spec, auto)
-                done   
-              thus "xa = [c] \<Longrightarrow> x = {[c]}" using in_eqiv by simp 
-            qed
-            moreover have "\<lbrakk>xa \<noteq> []; xa \<noteq> [c]\<rbrakk> \<Longrightarrow> x = UNIV - {[],[c]}"
-            proof -
-              have "\<lbrakk>xa \<noteq> []; xa \<noteq> [c]\<rbrakk> \<Longrightarrow> {y. xa \<equiv>{[c]}\<equiv> y} = UNIV - {[],[c]}" 
-                apply (clarsimp simp add:equiv_str_def)
-                apply (rule set_ext, rule iffI, simp)
-                apply (rule conjI)
-                apply (drule_tac x = "[c]" in spec, simp)
-                apply (drule_tac x = "[]" in spec, simp)
-                by (auto dest:quot_single_aux)
-              thus "\<lbrakk>xa \<noteq> []; xa \<noteq> [c]\<rbrakk> \<Longrightarrow> x = UNIV - {[],[c]}" using in_eqiv by simp
-            qed
-            ultimately show ?thesis by blast
-          qed
-        qed
-        thus ?thesis by simp
-      qed
-    qed
-  next
-    show "{{[]}, {[c]}, UNIV - {[],[c]}} \<subseteq> QUOT {[c]}"
-    proof
-      fix x
-      assume in_res: "x \<in> {{[]},{[c]}, (UNIV::string set) - {[],[c]}}"
-      show "x \<in> (QUOT {[c]})"
-      proof -
-        have "x = {[]} \<Longrightarrow> x \<in> QUOT {[c]}"
-          apply (simp add:QUOT_def equiv_class_def equiv_str_def)
-          by (rule_tac x = "[]" in exI, auto)
-        moreover have "x = {[c]} \<Longrightarrow> x \<in> QUOT {[c]}"
-          apply (simp add:QUOT_def equiv_class_def equiv_str_def)
-          apply (rule_tac x = "[c]" in exI, simp)
-          apply (rule set_ext, rule iffI, simp+)
-          by (drule_tac x = "[]" in spec, simp)
-        moreover have "x = UNIV - {[],[c]} \<Longrightarrow> x \<in> QUOT {[c]}"
-          using exist_another
-          apply (clarsimp simp add:QUOT_def equiv_class_def equiv_str_def)        
-          apply (rule_tac x = "[a]" in exI, simp)
-          apply (rule set_ext, rule iffI, simp)
-          apply (clarsimp simp:quot_single_aux, simp)
-          apply (rule conjI)
-          apply (drule_tac x = "[c]" in spec, simp)
-          by (drule_tac x = "[]" in spec, simp)     
-        ultimately show ?thesis using in_res by blast
-      qed
-    qed
-  qed
-qed
-
-lemma eq_class_imp_eq_str:
-  "\<lbrakk>x\<rbrakk>lang = \<lbrakk>y\<rbrakk>lang \<Longrightarrow> x \<equiv>lang\<equiv> y"
-by (auto simp:equiv_class_def equiv_str_def)
-
-lemma finite_tag_image: 
-  "finite (range tag) \<Longrightarrow> finite (((op `) tag) ` S)"
-apply (rule_tac B = "Pow (tag ` UNIV)" in finite_subset)
-by (auto simp add:image_def Pow_def)
-
-lemma str_inj_imps:
-  assumes str_inj: "\<And> m n. tag m = tag (n::string) \<Longrightarrow> m \<equiv>lang\<equiv> n"
-  shows "inj_on ((op `) tag) (QUOT lang)"
-proof (clarsimp simp add:inj_on_def QUOT_def)
-  fix x y
-  assume eq_tag: "tag ` \<lbrakk>x\<rbrakk>lang = tag ` \<lbrakk>y\<rbrakk>lang"
-  show "\<lbrakk>x\<rbrakk>lang = \<lbrakk>y\<rbrakk>lang"
-  proof -
-    have aux1:"\<And>a b. a \<in> (\<lbrakk>b\<rbrakk>lang) \<Longrightarrow> (a \<equiv>lang\<equiv> b)"
-      by (simp add:equiv_class_def equiv_str_def)
-    have aux2: "\<And> A B f. \<lbrakk>f ` A = f ` B; A \<noteq> {}\<rbrakk> \<Longrightarrow> \<exists> a b. a \<in> A \<and> b \<in> B \<and> f a = f b"
-      by auto
-    have aux3: "\<And> a l. \<lbrakk>a\<rbrakk>l \<noteq> {}" 
-      by (auto simp:equiv_class_def equiv_str_def)
-    show ?thesis using eq_tag
-      apply (drule_tac aux2, simp add:aux3, clarsimp)
-      apply (drule_tac str_inj, (drule_tac aux1)+)
-      by (simp add:equiv_str_def equiv_class_def)
-  qed
-qed
-
-definition tag_str_ALT :: "string set \<Rightarrow> string set \<Rightarrow> string \<Rightarrow> (string set \<times> string set)"
-where
-  "tag_str_ALT L\<^isub>1 L\<^isub>2 x \<equiv> (\<lbrakk>x\<rbrakk>L\<^isub>1, \<lbrakk>x\<rbrakk>L\<^isub>2)"
-
-lemma tag_str_alt_range_finite:
-  assumes finite1: "finite (QUOT L\<^isub>1)"
-  and finite2: "finite (QUOT L\<^isub>2)"
-  shows "finite (range (tag_str_ALT L\<^isub>1 L\<^isub>2))"
-proof -
-  have "{y. \<exists>x. y = (\<lbrakk>x\<rbrakk>L\<^isub>1, \<lbrakk>x\<rbrakk>L\<^isub>2)} \<subseteq> (QUOT L\<^isub>1) \<times> (QUOT L\<^isub>2)"
-    by (auto simp:QUOT_def)
-  thus ?thesis using finite1 finite2
-    by (auto simp: image_def tag_str_ALT_def UNION_def 
-            intro: finite_subset[where B = "(QUOT L\<^isub>1) \<times> (QUOT L\<^isub>2)"])
-qed
-
-lemma tag_str_alt_inj:
-  "tag_str_ALT L\<^isub>1 L\<^isub>2 x = tag_str_ALT L\<^isub>1 L\<^isub>2 y \<Longrightarrow> x \<equiv>(L\<^isub>1 \<union> L\<^isub>2)\<equiv> y"
-apply (simp add:tag_str_ALT_def equiv_class_def equiv_str_def)
-by blast
-  
-lemma quot_alt:
-  assumes finite1: "finite (QUOT L\<^isub>1)"
-  and finite2: "finite (QUOT L\<^isub>2)"
-  shows "finite (QUOT (L\<^isub>1 \<union> L\<^isub>2))"
-proof (rule_tac f = "(op `) (tag_str_ALT L\<^isub>1 L\<^isub>2)" in finite_imageD)
-  show "finite (op ` (tag_str_ALT L\<^isub>1 L\<^isub>2) ` QUOT (L\<^isub>1 \<union> L\<^isub>2))"
-    using finite_tag_image tag_str_alt_range_finite finite1 finite2
-    by auto
-next
-  show "inj_on (op ` (tag_str_ALT L\<^isub>1 L\<^isub>2)) (QUOT (L\<^isub>1 \<union> L\<^isub>2))"
-    apply (rule_tac str_inj_imps)
-    by (erule_tac tag_str_alt_inj)
-qed
-
-(* list_diff:: list substract, once different return tailer *)
-fun list_diff :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infix "-" 51)
-where
-  "list_diff []  xs = []" |
-  "list_diff (x#xs) [] = x#xs" |
-  "list_diff (x#xs) (y#ys) = (if x = y then list_diff xs ys else (x#xs))"
-
-lemma [simp]: "(x @ y) - x = y"
-apply (induct x)
-by (case_tac y, simp+)
-
-lemma [simp]: "x - x = []"
-by (induct x, auto)
-
-lemma [simp]: "x = xa @ y \<Longrightarrow> x - xa = y "
-by (induct x, auto)
-
-lemma [simp]: "x - [] = x"
-by (induct x, auto)
-
-lemma [simp]: "xa \<le> x \<Longrightarrow> (x @ y) - xa = (x - xa) @ y"
-by (auto elim:prefixE)
-
-definition tag_str_SEQ:: "string set \<Rightarrow> string set \<Rightarrow> string \<Rightarrow> (string set \<times> string set set)"
-where
-  "tag_str_SEQ L\<^isub>1 L\<^isub>2 x \<equiv> if (\<exists> xa \<le> x. xa \<in> L\<^isub>1)
-                         then (\<lbrakk>x\<rbrakk>L\<^isub>1, {\<lbrakk>(x - xa)\<rbrakk>L\<^isub>2 | xa.  xa \<le> x \<and> xa \<in> L\<^isub>1})
-                         else (\<lbrakk>x\<rbrakk>L\<^isub>1, {})"
-
-lemma tag_seq_eq_E:
-  "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y \<Longrightarrow>
-   ((\<exists> xa \<le> x. xa \<in> L\<^isub>1) \<and> \<lbrakk>x\<rbrakk>L\<^isub>1 = \<lbrakk>y\<rbrakk>L\<^isub>1 \<and> 
-    {\<lbrakk>(x - xa)\<rbrakk>L\<^isub>2 | xa. xa \<le> x \<and> xa \<in> L\<^isub>1} = {\<lbrakk>(y - ya)\<rbrakk>L\<^isub>2 | ya. ya \<le> y \<and> ya \<in> L\<^isub>1} ) \<or>
-   ((\<forall> xa \<le> x. xa \<notin> L\<^isub>1) \<and> \<lbrakk>x\<rbrakk>L\<^isub>1 = \<lbrakk>y\<rbrakk>L\<^isub>1)"
-by (simp add:tag_str_SEQ_def split:if_splits, blast)
-
-lemma tag_str_seq_range_finite:
-  assumes finite1: "finite (QUOT L\<^isub>1)"
-  and finite2: "finite (QUOT L\<^isub>2)"
-  shows "finite (range (tag_str_SEQ L\<^isub>1 L\<^isub>2))"
-proof -
-  have "(range (tag_str_SEQ L\<^isub>1 L\<^isub>2)) \<subseteq> (QUOT L\<^isub>1) \<times> (Pow (QUOT L\<^isub>2))"
-    by (auto simp:image_def tag_str_SEQ_def QUOT_def)
-  thus ?thesis using finite1 finite2 
-    by (rule_tac B = "(QUOT L\<^isub>1) \<times> (Pow (QUOT L\<^isub>2))" in finite_subset, auto)
-qed
-  
-lemma app_in_seq_decom[rule_format]:
-  "\<forall> x. x @ z \<in> L\<^isub>1 ; L\<^isub>2 \<longrightarrow> (\<exists> xa \<le> x. xa \<in> L\<^isub>1 \<and> (x - xa) @ z \<in> L\<^isub>2) \<or> 
-                            (\<exists> za \<le> z. (x @ za) \<in> L\<^isub>1 \<and> (z - za) \<in> L\<^isub>2)"
-apply (induct z)
-apply (simp, rule allI, rule impI, rule disjI1)
-apply (clarsimp simp add:lang_seq_def)
-apply (rule_tac x = s1 in exI, simp)
-apply (rule allI | rule impI)+
-apply (drule_tac x = "x @ [a]" in spec, simp)
-apply (erule exE | erule conjE | erule disjE)+
-apply (rule disjI2, rule_tac x = "[a]" in exI, simp)
-apply (rule disjI1, rule_tac x = xa in exI, simp) 
-apply (erule exE | erule conjE)+
-apply (rule disjI2, rule_tac x = "a # za" in exI, simp)
-done
-
-lemma tag_str_seq_inj:
-  assumes tag_eq: "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y"
-  shows "(x::string) \<equiv>(L\<^isub>1 ; L\<^isub>2)\<equiv> y"
-proof -
-  have aux: "\<And> x y z. \<lbrakk>tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y; x @ z \<in> L\<^isub>1 ; L\<^isub>2\<rbrakk> 
-                       \<Longrightarrow> y @ z \<in> L\<^isub>1 ; L\<^isub>2"
-  proof (drule app_in_seq_decom, erule disjE)
-    fix x y z
-    assume tag_eq: "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y"
-      and x_gets_l2: "\<exists>xa\<le>x. xa \<in> L\<^isub>1 \<and> (x - xa) @ z \<in> L\<^isub>2"
-    from x_gets_l2 have "\<exists> xa \<le> x. xa \<in> L\<^isub>1" by blast
-    hence xy_l2:"{\<lbrakk>(x - xa)\<rbrakk>L\<^isub>2 | xa. xa \<le> x \<and> xa \<in> L\<^isub>1} = {\<lbrakk>(y - ya)\<rbrakk>L\<^isub>2 | ya. ya \<le> y \<and> ya \<in> L\<^isub>1}"
-      using tag_eq tag_seq_eq_E by blast
-    from x_gets_l2 obtain xa where xa_le_x: "xa \<le> x"
-                               and xa_in_l1: "xa \<in> L\<^isub>1"
-                               and rest_in_l2: "(x - xa) @ z \<in> L\<^isub>2" by blast
-    hence "\<exists> ya. \<lbrakk>(x - xa)\<rbrakk>L\<^isub>2 = \<lbrakk>(y - ya)\<rbrakk>L\<^isub>2 \<and> ya \<le> y \<and> ya \<in> L\<^isub>1" using xy_l2 by auto
-    then obtain ya where ya_le_x: "ya \<le> y"
-                     and ya_in_l1: "ya \<in> L\<^isub>1"
-                     and rest_eq: "\<lbrakk>(x - xa)\<rbrakk>L\<^isub>2 = \<lbrakk>(y - ya)\<rbrakk>L\<^isub>2" by blast
-    from rest_eq rest_in_l2 have "(y - ya) @ z \<in> L\<^isub>2" 
-      by (auto simp:equiv_class_def equiv_str_def)
-    hence "ya @ ((y - ya) @ z) \<in> L\<^isub>1 ; L\<^isub>2" using ya_in_l1
-      by (auto simp:lang_seq_def)
-    thus "y @ z \<in> L\<^isub>1 ; L\<^isub>2" using ya_le_x 
-      by (erule_tac prefixE, simp)
-  next
-    fix x y z
-    assume tag_eq: "tag_str_SEQ L\<^isub>1 L\<^isub>2 x = tag_str_SEQ L\<^isub>1 L\<^isub>2 y"
-      and x_gets_l1: "\<exists>za\<le>z. x @ za \<in> L\<^isub>1 \<and> z - za \<in> L\<^isub>2"
-    from tag_eq tag_seq_eq_E have x_y_eq: "\<lbrakk>x\<rbrakk>L\<^isub>1 = \<lbrakk>y\<rbrakk>L\<^isub>1" by blast
-    from x_gets_l1 obtain za where za_le_z: "za \<le> z"
-                               and x_za_in_l1: "(x @ za) \<in> L\<^isub>1"
-                               and rest_in_l2: "z - za \<in> L\<^isub>2" by blast
-    from x_y_eq x_za_in_l1 have y_za_in_l1: "y @ za \<in> L\<^isub>1"
-      by (auto simp:equiv_class_def equiv_str_def)
-    hence "(y @ za) @ (z - za) \<in> L\<^isub>1 ; L\<^isub>2" using rest_in_l2
-      apply (simp add:lang_seq_def)
-      by (rule_tac x = "y @ za" in exI, rule_tac x = "z - za" in exI, simp)
-    thus "y @ z \<in> L\<^isub>1 ; L\<^isub>2" using za_le_z
-      by (erule_tac prefixE, simp)
-  qed
-  show ?thesis using tag_eq
-    apply (simp add:equiv_str_def)
-    by (auto intro:aux)
-qed
-
-lemma quot_seq: 
-  assumes finite1: "finite (QUOT L\<^isub>1)"
-  and finite2: "finite (QUOT L\<^isub>2)"
-  shows "finite (QUOT (L\<^isub>1;L\<^isub>2))"
-proof (rule_tac f = "(op `) (tag_str_SEQ L\<^isub>1 L\<^isub>2)" in finite_imageD)
-  show "finite (op ` (tag_str_SEQ L\<^isub>1 L\<^isub>2) ` QUOT (L\<^isub>1 ; L\<^isub>2))"
-    using finite_tag_image tag_str_seq_range_finite finite1 finite2
-    by auto
-next
-  show "inj_on (op ` (tag_str_SEQ L\<^isub>1 L\<^isub>2)) (QUOT (L\<^isub>1 ; L\<^isub>2))"
-    apply (rule_tac str_inj_imps)
-    by (erule_tac tag_str_seq_inj)
-qed
-
-(****************** the STAR case ************************)
-
-lemma app_eq_elim[rule_format]:
-  "\<And> a. \<forall> b x y. a @ b = x @ y \<longrightarrow> (\<exists> aa ab. a = aa @ ab \<and> x = aa \<and> y = ab @ b) \<or>
-                                   (\<exists> ba bb. b = ba @ bb \<and> x = a @ ba \<and> y = bb \<and> ba \<noteq> [])"
-  apply (induct_tac a rule:List.induct, simp)
-  apply (rule allI | rule impI)+
-  by (case_tac x, auto)
-
-definition tag_str_STAR:: "string set \<Rightarrow> string \<Rightarrow> string set set"
-where
-  "tag_str_STAR L\<^isub>1 x \<equiv> if (x = []) 
-                       then {}
-                       else {\<lbrakk>x\<^isub>2\<rbrakk>L\<^isub>1 | x\<^isub>1 x\<^isub>2. x =  x\<^isub>1 @ x\<^isub>2 \<and> x\<^isub>1 \<in> L\<^isub>1\<star> \<and> x\<^isub>2 \<noteq> []}"
-
-lemma tag_str_star_range_finite:
-  assumes finite1: "finite (QUOT L\<^isub>1)"
-  shows "finite (range (tag_str_STAR L\<^isub>1))"
-proof -
-  have "range (tag_str_STAR L\<^isub>1) \<subseteq> Pow (QUOT L\<^isub>1)"
-    by (auto simp:image_def tag_str_STAR_def QUOT_def)
-  thus ?thesis using finite1
-    by (rule_tac B = "Pow (QUOT L\<^isub>1)" in finite_subset, auto)
-qed
-
-lemma star_prop[rule_format]: "x \<in> lang\<star> \<Longrightarrow> \<forall> y. y \<in> lang\<star> \<longrightarrow> x @ y \<in> lang\<star>"
-by (erule Star.induct, auto)
-
-lemma star_prop2: "y \<in> lang \<Longrightarrow> y \<in> lang\<star>"
-by (drule step[of y lang "[]"], auto simp:start)
-
-lemma star_prop3[rule_format]: "x \<in> lang\<star> \<Longrightarrow> \<forall>y . y \<in> lang \<longrightarrow> x @ y \<in> lang\<star>"
-by (erule Star.induct, auto intro:star_prop2)
-
-lemma postfix_prop: "y >>= (x @ y) \<Longrightarrow> x = []"
-by (erule postfixE, induct x arbitrary:y, auto)
-
-lemma inj_aux:
-  "\<lbrakk>(m @ z) \<in> L\<^isub>1\<star>; m \<equiv>L\<^isub>1\<equiv> yb; xa @ m = x; xa \<in> L\<^isub>1\<star>; m \<noteq> [];
-    \<forall> xa xb. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star> \<longrightarrow> xb >>= m\<rbrakk> 
-  \<Longrightarrow> (yb @ z) \<in> L\<^isub>1\<star>"
-proof- 
-  have "\<And>s. s \<in> L\<^isub>1\<star> \<Longrightarrow> \<forall> m z yb. (s = m @ z \<and> m \<equiv>L\<^isub>1\<equiv> yb \<and> x = xa @ m \<and> xa \<in> L\<^isub>1\<star> \<and> m \<noteq> [] \<and>  
-    (\<forall> xa xb. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star> \<longrightarrow> xb >>= m) \<longrightarrow> (yb @ z) \<in> L\<^isub>1\<star>)"    
-    apply (erule Star.induct, simp)
-    apply (rule allI | rule impI | erule conjE)+
-    apply (drule app_eq_elim)
-    apply (erule disjE | erule exE | erule conjE)+
-    apply simp
-    apply (simp (no_asm) only:append_assoc[THEN sym])
-    apply (rule step) 
-    apply (simp add:equiv_str_def)
-    apply simp
-
-    apply (erule exE | erule conjE)+    
-    apply (rotate_tac 3)
-    apply (frule_tac x = "xa @ s1" in spec)
-    apply (rotate_tac 12)
-    apply (drule_tac x = ba in spec)
-    apply (erule impE)
-    apply ( simp add:star_prop3)
-    apply (simp)
-    apply (drule postfix_prop)
-    apply simp
-    done
-  thus "\<lbrakk>(m @ z) \<in> L\<^isub>1\<star>; m \<equiv>L\<^isub>1\<equiv> yb; xa @ m = x; xa \<in> L\<^isub>1\<star>; m \<noteq> [];
-         \<forall> xa xb. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star> \<longrightarrow> xb >>= m\<rbrakk> 
-        \<Longrightarrow> (yb @ z) \<in> L\<^isub>1\<star>" by blast
-qed
-
-
-lemma min_postfix_exists[rule_format]:
-  "finite A \<Longrightarrow> A \<noteq> {} \<and> (\<forall> a \<in> A. \<forall> b \<in> A. ((b >>= a) \<or> (a >>= b))) 
-                \<longrightarrow> (\<exists> min. (min \<in> A \<and> (\<forall> a \<in> A. a >>= min)))"
-apply (erule finite.induct)
-apply simp
-apply simp
-apply (case_tac "A = {}")
-apply simp
-apply clarsimp
-apply (case_tac "a >>= min")
-apply (rule_tac x = min in exI, simp)
-apply (rule_tac x = a in exI, simp)
-apply clarify
-apply (rotate_tac 5)
-apply (erule_tac x = aa in ballE) defer apply simp
-apply (erule_tac ys = min in postfix_trans)
-apply (erule_tac x = min in ballE) 
-by simp+
-
-lemma tag_str_star_inj:
-  "tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 (y::string) \<Longrightarrow> x \<equiv>L\<^isub>1\<star>\<equiv> y"
-proof -
-  have aux: "\<And> x y z. \<lbrakk>tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 y; x @ z \<in> L\<^isub>1\<star>\<rbrakk> \<Longrightarrow> y @ z \<in> L\<^isub>1\<star>"
-  proof-
-    fix x y z
-    assume tag_eq: "tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 y"
-      and x_z: "x @ z \<in> L\<^isub>1\<star>"
-    show "y @ z \<in> L\<^isub>1\<star>"
-    proof (cases "x = []")
-      case True
-      with tag_eq have "y = []" by (simp add:tag_str_STAR_def split:if_splits, blast)
-      thus ?thesis using x_z True by simp
-    next
-      case False
-      hence not_empty: "{xb. \<exists> xa. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star>} \<noteq> {}" using x_z
-        by (simp, rule_tac x = x in exI, rule_tac x = "[]" in exI, simp add:start)
-      have finite_set: "finite {xb. \<exists> xa. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star>}"
-        apply (rule_tac B = "{xb. \<exists> xa. x = xa @ xb}" in finite_subset)
-        apply auto
-        apply (induct x, simp)
-        apply (subgoal_tac "{xb. \<exists>xa. a # x = xa @ xb} = insert (a # x) {xb. \<exists>xa. x = xa @ xb}")
-        apply auto
-        by (case_tac xaa, simp+)
-      have comparable: "\<forall> a \<in> {xb. \<exists> xa. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star>}. 
-                        \<forall> b \<in> {xb. \<exists> xa. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star>}.
-                          ((b >>= a) \<or> (a >>= b))"
-        by (auto simp:postfix_def, drule app_eq_elim, blast)
-      hence "\<exists> min. min \<in> {xb. \<exists> xa. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star>} 
-                \<and> (\<forall> xa xb. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star> \<longrightarrow> xb >>= min)"
-        using finite_set not_empty comparable
-        apply (drule_tac min_postfix_exists, simp)
-        by (erule exE, rule_tac x = min in exI, auto)
-      then obtain min xa where x_decom: "x = xa @ min \<and> xa \<in> L\<^isub>1\<star>"
-        and min_not_empty: "min \<noteq> []"
-        and min_z_in_star: "min @ z \<in> L\<^isub>1\<star>"
-        and is_min: "\<forall> xa xb. x = xa @ xb \<and> xa \<in> L\<^isub>1\<star> \<and> xb \<noteq> [] \<and> xb @ z \<in> L\<^isub>1\<star> \<longrightarrow> xb >>= min"  by blast
-      from x_decom min_not_empty have "\<lbrakk>min\<rbrakk>L\<^isub>1 \<in> tag_str_STAR L\<^isub>1 x"  by (auto simp:tag_str_STAR_def)
-      hence "\<exists> yb. \<lbrakk>yb\<rbrakk>L\<^isub>1 \<in> tag_str_STAR L\<^isub>1 y \<and> \<lbrakk>min\<rbrakk>L\<^isub>1 = \<lbrakk>yb\<rbrakk>L\<^isub>1" using tag_eq by auto
-      hence "\<exists> ya yb. y = ya @ yb \<and> ya \<in> L\<^isub>1\<star> \<and> min \<equiv>L\<^isub>1\<equiv> yb \<and> yb \<noteq> [] " 
-        by (simp add:tag_str_STAR_def equiv_class_def equiv_str_def split:if_splits, blast)        
-      then obtain ya yb where y_decom: "y = ya @ yb"
-                          and ya_in_star: "ya \<in> L\<^isub>1\<star>"
-                          and yb_not_empty: "yb \<noteq> []"
-                          and min_yb_eq: "min \<equiv>L\<^isub>1\<equiv> yb"  by blast
-      from min_z_in_star min_yb_eq min_not_empty is_min x_decom
-      have "yb @ z \<in> L\<^isub>1\<star>"
-        by (rule_tac x = x and xa = xa in inj_aux, simp+)
-      thus ?thesis using ya_in_star y_decom
-        by (auto dest:star_prop)        
-    qed
-  qed
-  show "tag_str_STAR L\<^isub>1 x = tag_str_STAR L\<^isub>1 (y::string) \<Longrightarrow> x \<equiv>L\<^isub>1\<star>\<equiv> y"
-    by (auto intro:aux simp:equiv_str_def)
-qed
-
-lemma quot_star:  
-  assumes finite1: "finite (QUOT L\<^isub>1)"
-  shows "finite (QUOT (L\<^isub>1\<star>))"
-proof (rule_tac f = "(op `) (tag_str_STAR L\<^isub>1)" in finite_imageD)
-  show "finite (op ` (tag_str_STAR L\<^isub>1) ` QUOT (L\<^isub>1\<star>))"
-    using finite_tag_image tag_str_star_range_finite finite1
-    by auto
-next
-  show "inj_on (op ` (tag_str_STAR L\<^isub>1)) (QUOT (L\<^isub>1\<star>))"
-    apply (rule_tac str_inj_imps)
-    by (erule_tac tag_str_star_inj)
-qed
-
-lemma other_direction:
-  "Lang = L (r::rexp) \<Longrightarrow> finite (QUOT Lang)"
-apply (induct arbitrary:Lang rule:rexp.induct)
-apply (simp add:QUOT_def equiv_class_def equiv_str_def)
-by (simp_all add:quot_lambda quot_single quot_seq quot_alt quot_star)  
-
-end 
--- a/Myhill_1.thy	Fri Jun 03 13:59:21 2011 +0000
+++ b/Myhill_1.thy	Mon Jul 25 13:33:38 2011 +0000
@@ -1,5 +1,5 @@
 theory Myhill_1
-imports Regular
+imports More_Regular_Set
         "~~/src/HOL/Library/While_Combinator" 
 begin
 
@@ -12,12 +12,12 @@
 text {* Myhill-Nerode relation *}
 
 definition
-  str_eq_rel :: "lang \<Rightarrow> (string \<times> string) set" ("\<approx>_" [100] 100)
+  str_eq_rel :: "'a lang \<Rightarrow> ('a list \<times> 'a list) 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"
+  finals :: "'a lang \<Rightarrow> 'a lang set"
 where
   "finals A \<equiv> {\<approx>A `` {s} | s . s \<in> A}"
 
@@ -37,35 +37,35 @@
 
 text {* The two kinds of terms in the rhs of equations. *}
 
-datatype trm = 
-   Lam "rexp"            (* Lambda-marker *)
- | Trn "lang" "rexp"     (* Transition *)
+datatype 'a trm = 
+   Lam "'a rexp"            (* Lambda-marker *)
+ | Trn "'a lang" "'a rexp"     (* Transition *)
 
 fun 
-  L_trm::"trm \<Rightarrow> lang"
+  lang_trm::"'a trm \<Rightarrow> 'a lang"
 where
-  "L_trm (Lam r) = L_rexp r" 
-| "L_trm (Trn X r) = X \<cdot> L_rexp r"
+  "lang_trm (Lam r) = lang r" 
+| "lang_trm (Trn X r) = X \<cdot> lang r"
 
 fun 
-  L_rhs::"trm set \<Rightarrow> lang"
+  lang_rhs::"('a trm) set \<Rightarrow> 'a lang"
 where 
-  "L_rhs rhs = \<Union> (L_trm ` rhs)"
+  "lang_rhs rhs = \<Union> (lang_trm ` rhs)"
 
-lemma L_rhs_set:
-  shows "L_rhs {Trn X r | r. P r} = \<Union>{L_trm (Trn X r) | r. P r}"
+lemma lang_rhs_set:
+  shows "lang_rhs {Trn X r | r. P r} = \<Union>{lang_trm (Trn X r) | r. P r}"
 by (auto)
 
-lemma L_rhs_union_distrib:
-  fixes A B::"trm set"
-  shows "L_rhs A \<union> L_rhs B = L_rhs (A \<union> B)"
+lemma lang_rhs_union_distrib:
+  fixes A B::"('a trm) set"
+  shows "lang_rhs A \<union> lang_rhs B = lang_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)
+  transition :: "'a lang \<Rightarrow> 'a \<Rightarrow> 'a lang \<Rightarrow> bool" ("_ \<Turnstile>_\<Rightarrow>_" [100,100,100] 100)
 where
   "Y \<Turnstile>c\<Rightarrow> X \<equiv> Y \<cdot> {[c]} \<subseteq> X"
 
@@ -74,9 +74,9 @@
 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}
+          {Lam One} \<union> {Trn Y (Atom 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}"
+          {Trn Y (Atom 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}"
@@ -85,10 +85,10 @@
 section {* Arden Operation on equations *}
 
 fun 
-  Append_rexp :: "rexp \<Rightarrow> trm \<Rightarrow> trm"
+  Append_rexp :: "'a rexp \<Rightarrow> 'a trm \<Rightarrow> 'a trm"
 where
-  "Append_rexp r (Lam rexp)   = Lam (SEQ rexp r)"
-| "Append_rexp r (Trn X rexp) = Trn X (SEQ rexp r)"
+  "Append_rexp r (Lam rexp)   = Lam (Times rexp r)"
+| "Append_rexp r (Trn X rexp) = Trn X (Times rexp r)"
 
 
 definition
@@ -96,7 +96,7 @@
 
 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}))"
+     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 *}
@@ -106,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> trm set) set \<Rightarrow> lang \<Rightarrow> trm set \<Rightarrow> (lang \<times> trm set) set"
+  Subst_all :: "('a lang \<times> ('a trm) set) set \<Rightarrow> 'a lang \<Rightarrow> ('a trm) set \<Rightarrow> ('a lang \<times> ('a trm) set) set"
 where
   "Subst_all ES X xrhs \<equiv> {(Y, Subst yrhs X xrhs) | Y yrhs. (Y, yrhs) \<in> ES}"
 
@@ -143,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 rhs"
+  "soundness ES \<equiv> \<forall>(X, rhs) \<in> ES. X = lang_rhs rhs"
 
 definition 
-  "ardenable rhs \<equiv> (\<forall> Y r. Trn Y r \<in> rhs \<longrightarrow> [] \<notin> L_rexp r)"
+  "ardenable rhs \<equiv> (\<forall> Y r. Trn Y r \<in> rhs \<longrightarrow> [] \<notin> lang r)"
 
 definition 
   "ardenable_all ES \<equiv> \<forall>(X, rhs) \<in> ES. ardenable rhs"
@@ -223,23 +223,23 @@
 
 lemma trm_soundness:
   assumes finite:"finite rhs"
-  shows "L_rhs ({Trn X r| r. Trn X r \<in> rhs}) = X \<cdot> (L_rexp (\<Uplus>{r. Trn X r \<in> rhs}))"
+  shows "lang_rhs ({Trn X r| r. Trn X r \<in> rhs}) = X \<cdot> (lang (\<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_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)
+  then show "lang_rhs ({Trn X r| r. Trn X r \<in> rhs}) = X \<cdot> (lang (\<Uplus>{r. Trn X r \<in> rhs}))"
+    by (simp only: lang_rhs_set lang_trm.simps) (auto simp add: conc_def)
 qed
 
 lemma lang_of_append_rexp:
-  "L_trm (Append_rexp r trm) = L_trm trm \<cdot> L_rexp r"
+  "lang_trm (Append_rexp r trm) = lang_trm trm \<cdot> lang r"
 by (induct rule: Append_rexp.induct)
-   (auto simp add: seq_assoc)
+   (auto simp add: conc_assoc)
 
 lemma lang_of_append_rexp_rhs:
-  "L_rhs (Append_rexp_rhs rhs r) = L_rhs rhs \<cdot> L_rexp r"
+  "lang_rhs (Append_rexp_rhs rhs r) = lang_rhs rhs \<cdot> lang r"
 unfolding Append_rexp_rhs_def
-by (auto simp add: Seq_def lang_of_append_rexp)
+by (auto simp add: conc_def lang_of_append_rexp)
 
 
 subsubsection {* Intial Equational System *}
@@ -263,7 +263,7 @@
   have "X = \<approx>A `` {s @ [c]}" 
     using has_str in_CS defined_by_str by blast
   then have "Y \<cdot> {[c]} \<subseteq> X" 
-    unfolding Y_def Image_def Seq_def
+    unfolding Y_def Image_def conc_def
     unfolding str_eq_rel_def
     by clarsimp
   moreover
@@ -274,14 +274,14 @@
 
 lemma l_eq_r_in_eqs:
   assumes X_in_eqs: "(X, rhs) \<in> Init (UNIV // \<approx>A)"
-  shows "X = L_rhs rhs"
+  shows "X = lang_rhs rhs"
 proof 
-  show "X \<subseteq> L_rhs rhs"
+  show "X \<subseteq> lang_rhs rhs"
   proof
     fix x
     assume in_X: "x \<in> X"
     { assume empty: "x = []"
-      then have "x \<in> L_rhs rhs" using X_in_eqs in_X
+      then have "x \<in> lang_rhs rhs" using X_in_eqs in_X
 	unfolding Init_def Init_rhs_def
         by auto
     }
@@ -291,43 +291,42 @@
 	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 \<cdot> {[c]} \<subseteq> X" "s \<in> Y"
-        using decom in_X every_eqclass_has_transition by blast
-      then have "x \<in> L_rhs {Trn Y (CHAR c)| Y c. Y \<in> UNIV // \<approx>A \<and> Y \<Turnstile>c\<Rightarrow> X}"
+        using decom in_X every_eqclass_has_transition by metis
+      then have "x \<in> lang_rhs {Trn Y (Atom 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 rhs" using X_in_eqs in_X
+	using decom by (force simp add: conc_def)
+      then have "x \<in> lang_rhs rhs" using X_in_eqs in_X
 	unfolding Init_def Init_rhs_def by simp
     }
-    ultimately show "x \<in> L_rhs rhs" by blast
+    ultimately show "x \<in> lang_rhs rhs" by blast
   qed
 next
-  show "L_rhs rhs \<subseteq> X" using X_in_eqs
+  show "lang_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_trm `  rhs)"
-using assms l_eq_r_in_eqs by (simp)
 
 lemma finite_Init_rhs: 
+  fixes CS::"(('a::finite) lang) set"
   assumes finite: "finite CS"
   shows "finite (Init_rhs CS X)"
 proof-
-  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
+  def S \<equiv> "{(Y, c)| Y c::'a. Y \<in> CS \<and> Y \<cdot> {[c]} \<subseteq> X}" 
+  def h \<equiv> "\<lambda> (Y, c::'a). Trn Y (Atom c)"
+  have "finite (CS \<times> (UNIV::('a::finite) 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 \<cdot> {[c]} \<subseteq> X} = h ` S"
+  moreover have "{Trn Y (Atom c) |Y c::'a. 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 \<cdot> {[c]} \<subseteq> X}" by auto
+  have "finite {Trn Y (Atom 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
 
+
 lemma Init_ES_satisfies_invariant:
+  fixes A::"(('a::finite) lang)"
   assumes finite_CS: "finite (UNIV // \<approx>A)"
   shows "invariant (Init (UNIV // \<approx>A))"
 proof (rule invariantI)
@@ -352,20 +351,20 @@
 subsubsection {* Interation step *}
 
 lemma Arden_keeps_eq:
-  assumes l_eq_r: "X = L_rhs rhs"
+  assumes l_eq_r: "X = lang_rhs rhs"
   and not_empty: "ardenable rhs"
   and finite: "finite rhs"
-  shows "X = L_rhs (Arden X rhs)"
+  shows "X = lang_rhs (Arden X rhs)"
 proof -
-  def A \<equiv> "L_rexp (\<Uplus>{r. Trn X r \<in> rhs})"
+  def A \<equiv> "lang (\<Uplus>{r. Trn X r \<in> rhs})"
   def b \<equiv> "{Trn X r | r. Trn X r \<in> rhs}"
-  def B \<equiv> "L_rhs (rhs - b)"
+  def B \<equiv> "lang_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 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)
+  have "X = lang_rhs rhs" using l_eq_r by simp
+  also have "\<dots> = lang_rhs (b \<union> (rhs - b))" unfolding b_def by auto
+  also have "\<dots> = lang_rhs b \<union> B" unfolding B_def by (simp only: lang_rhs_union_distrib)
   also have "\<dots> = X \<cdot> A \<union> B"
     unfolding b_def
     unfolding trm_soundness[OF finite]
@@ -374,24 +373,24 @@
   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_rhs (Arden X rhs)"
+  also have "\<dots> = lang_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_rhs (Arden X rhs)" by simp
+    by (simp only: lang_of_append_rexp_rhs lang.simps)
+  finally show "X = lang_rhs (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)
+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)
+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)
+apply (auto simp: ardenable_def Append_rexp_rhs_def)
+by (case_tac x, auto simp: conc_def)
 
 lemma nonempty_set_sub:
   "ardenable rhs \<Longrightarrow> ardenable (rhs - A)"
@@ -411,24 +410,25 @@
 by (simp only: Subst_def Append_keeps_nonempty nonempty_set_union nonempty_set_sub)
 
 lemma Subst_keeps_eq:
-  assumes substor: "X = L_rhs xrhs"
+  assumes substor: "X = lang_rhs xrhs"
   and finite: "finite rhs"
-  shows "L_rhs (Subst rhs X xrhs) = L_rhs rhs" (is "?Left = ?Right")
+  shows "lang_rhs (Subst rhs X xrhs) = lang_rhs rhs" (is "?Left = ?Right")
 proof-
-  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}))"
+  def A \<equiv> "lang_rhs (rhs - {Trn X r | r. Trn X r \<in> rhs})"
+  have "?Left = A \<union> lang_rhs (Append_rexp_rhs xrhs (\<Uplus>{r. Trn X r \<in> rhs}))"
     unfolding Subst_def
-    unfolding L_rhs_union_distrib[symmetric]
+    unfolding lang_rhs_union_distrib[symmetric]
     by (simp add: A_def)
-  moreover have "?Right = A \<union> L_rhs {Trn X r | r. Trn X r \<in> rhs}"
+  moreover have "?Right = A \<union> lang_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 
       unfolding A_def
-      unfolding L_rhs_union_distrib
+      unfolding lang_rhs_union_distrib
       by simp
   qed
-  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}" 
+  moreover 
+  have "lang_rhs (Append_rexp_rhs xrhs (\<Uplus>{r. Trn X r \<in> rhs})) = lang_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
@@ -441,8 +441,8 @@
   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)"
+  def eqns \<equiv> "{(X::'a lang, rhs) |X rhs. (X, rhs) \<in> ES}"
+  def h \<equiv> "\<lambda>(X::'a 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
@@ -456,24 +456,24 @@
 
 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+)
+apply (auto simp: rhss_def Append_rexp_rhs_def)
+apply (case_tac xa, auto simp: image_def)
+by (rule_tac x = "Times 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)
+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)
+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)
+by (auto simp: rhss_def)
 
 lemma Subst_all_keeps_validity:
   assumes sc: "validity (ES \<union> {(Y, yrhs)})"        (is "validity ?A")
@@ -490,17 +490,17 @@
       moreover have "rhss xrhs' \<subseteq> lhss ES"
       proof-
         have "rhss xrhs' \<subseteq>  rhss xrhs \<union> rhss (Arden Y yrhs) - {Y}"
-        proof-
+        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)
+            using Arden_removes_cl by auto
+          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)
+          by (auto simp add: Arden_removes_cl validity_def lhss_def)
         ultimately show ?thesis by auto
       qed
       ultimately show ?thesis by simp
@@ -512,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_rhs yrhs" 
+  have Y_eq_yrhs: "Y = lang_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)
@@ -520,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_rhs (Arden Y yrhs)" 
+    have "Y = lang_rhs (Arden Y yrhs)" 
       using Y_eq_yrhs invariant_ES finite_yrhs
       using finite_Trn[OF finite_yrhs]
       apply(rule_tac Arden_keeps_eq)
@@ -530,7 +530,7 @@
       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)
+      by (auto simp add: Subst_keeps_eq simp del: lang_rhs.simps)
   qed
   show "finite (Subst_all ES Y (Arden Y yrhs))" 
     using invariant_ES by (simp add:invariant_def Subst_all_keeps_finite)
@@ -557,13 +557,13 @@
     proof -
       have "finite yrhs" using invariant_ES 
         by (auto simp:invariant_def finite_rhs_def)
-      thus ?thesis using Arden_keeps_finite by simp
+      thus ?thesis using Arden_keeps_finite by auto
     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)
+    using invariant_ES Subst_all_keeps_validity by (auto simp add: invariant_def)
 qed
 
 lemma Remove_in_card_measure:
@@ -571,7 +571,7 @@
   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 f \<equiv> "\<lambda> x. ((fst x)::'a 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)
@@ -674,6 +674,7 @@
 subsubsection {* Conclusion of the proof *}
 
 lemma Solve:
+  fixes A::"('a::finite) lang"
   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)}"
@@ -714,9 +715,10 @@
 qed
 
 lemma every_eqcl_has_reg:
+  fixes A::"('a::finite) lang"
   assumes finite_CS: "finite (UNIV // \<approx>A)"
   and X_in_CS: "X \<in> (UNIV // \<approx>A)"
-  shows "\<exists>r. X = L_rexp r" 
+  shows "\<exists>r. X = lang r" 
 proof -
   from finite_CS X_in_CS 
   obtain xrhs where Inv_ES: "invariant {(X, xrhs)}"
@@ -735,14 +737,14 @@
     using Arden_keeps_finite by auto
   then have fin: "finite {r. Lam r \<in> A}" by (rule finite_Lam)
 
-  have "X = L_rhs xrhs" using Inv_ES unfolding invariant_def soundness_def
+  have "X = lang_rhs xrhs" using Inv_ES unfolding invariant_def soundness_def
     by simp
-  then have "X = L_rhs A" using Inv_ES 
+  then have "X = lang_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_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
+  then have "X = lang_rhs {Lam r | r. Lam r \<in> A}" using eq by simp
+  then have "X = lang (\<Uplus>{r. Lam r \<in> A})" using fin by auto
+  then show "\<exists>r. X = lang r" by blast
 qed
 
 lemma bchoice_finite_set:
@@ -756,20 +758,21 @@
 done
 
 theorem Myhill_Nerode1:
+  fixes A::"('a::finite) lang"
   assumes finite_CS: "finite (UNIV // \<approx>A)"
-  shows   "\<exists>r. A = L_rexp r"
+  shows   "\<exists>r. A = lang 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. X = L_rexp r" 
+  have "\<forall>X \<in> (UNIV // \<approx>A). \<exists>r. X = lang r" 
     using finite_CS every_eqcl_has_reg by blast
-  then have a: "\<forall>X \<in> finals A. \<exists>r. X = L_rexp r"
+  then have a: "\<forall>X \<in> finals A. \<exists>r. X = lang r"
     using finals_in_partitions by auto
-  then obtain rs::"rexp set" where "\<Union> (finals A) = \<Union>(L_rexp ` rs)" "finite rs"
+  then obtain rs::"('a rexp) set" where "\<Union> (finals A) = \<Union>(lang ` rs)" "finite rs"
     using fin by (auto dest: bchoice_finite_set)
-  then have "A = L_rexp (\<Uplus>rs)" 
+  then have "A = lang (\<Uplus>rs)" 
     unfolding lang_is_union_of_finals[symmetric] by simp
-  then show "\<exists>r. A = L_rexp r" by blast
+  then show "\<exists>r. A = lang r" by blast
 qed 
 
 
--- a/Myhill_2.thy	Fri Jun 03 13:59:21 2011 +0000
+++ b/Myhill_2.thy	Mon Jul 25 13:33:38 2011 +0000
@@ -6,7 +6,7 @@
 section {* Direction @{text "regular language \<Rightarrow> finite partition"} *}
 
 definition
-  str_eq :: "string \<Rightarrow> lang \<Rightarrow> string \<Rightarrow> bool" ("_ \<approx>_ _")
+  str_eq :: "'a list \<Rightarrow> 'a lang \<Rightarrow> 'a list \<Rightarrow> bool" ("_ \<approx>_ _")
 where
   "x \<approx>A y \<equiv> (x, y) \<in> (\<approx>A)"
 
@@ -16,7 +16,7 @@
 by simp
 
 definition 
-   tag_eq_rel :: "(string \<Rightarrow> 'b) \<Rightarrow> (string \<times> string) set" ("=_=")
+   tag_eq_rel :: "('a list \<Rightarrow> 'b) \<Rightarrow> ('a list \<times> 'a list) set" ("=_=")
 where
    "=tag= \<equiv> {(x, y). tag x = tag y}"
 
@@ -116,20 +116,20 @@
 
 subsection {* The proof *}
 
-subsubsection {* The base case for @{const "NULL"} *}
+subsubsection {* The base case for @{const "Zero"} *}
 
-lemma quot_null_eq:
+lemma quot_zero_eq:
   shows "UNIV // \<approx>{} = {UNIV}"
 unfolding quotient_def Image_def str_eq_rel_def by auto
 
-lemma quot_null_finiteI [intro]:
+lemma quot_zero_finiteI [intro]:
   shows "finite (UNIV // \<approx>{})"
-unfolding quot_null_eq by simp
+unfolding quot_zero_eq by simp
 
 
-subsubsection {* The base case for @{const "EMPTY"} *}
+subsubsection {* The base case for @{const "One"} *}
 
-lemma quot_empty_subset:
+lemma quot_one_subset:
   shows "UNIV // \<approx>{[]} \<subseteq> {{[]}, UNIV - {[]}}"
 proof
   fix x
@@ -148,14 +148,14 @@
   qed
 qed
 
-lemma quot_empty_finiteI [intro]:
+lemma quot_one_finiteI [intro]:
   shows "finite (UNIV // \<approx>{[]})"
-by (rule finite_subset[OF quot_empty_subset]) (simp)
+by (rule finite_subset[OF quot_one_subset]) (simp)
 
 
-subsubsection {* The base case for @{const "CHAR"} *}
+subsubsection {* The base case for @{const "Atom"} *}
 
-lemma quot_char_subset:
+lemma quot_atom_subset:
   "UNIV // (\<approx>{[c]}) \<subseteq> {{[]},{[c]}, UNIV - {[], [c]}}"
 proof 
   fix x 
@@ -181,43 +181,43 @@
   qed
 qed
 
-lemma quot_char_finiteI [intro]:
+lemma quot_atom_finiteI [intro]:
   shows "finite (UNIV // \<approx>{[c]})"
-by (rule finite_subset[OF quot_char_subset]) (simp)
+by (rule finite_subset[OF quot_atom_subset]) (simp)
 
 
-subsubsection {* The inductive case for @{const ALT} *}
+subsubsection {* The inductive case for @{const Plus} *}
 
 definition 
-  tag_str_ALT :: "lang \<Rightarrow> lang \<Rightarrow> string \<Rightarrow> (lang \<times> lang)"
+  tag_str_Plus :: "'a lang \<Rightarrow> 'a lang \<Rightarrow> 'a list \<Rightarrow> ('a lang \<times> 'a lang)"
 where
-  "tag_str_ALT A B \<equiv> (\<lambda>x. (\<approx>A `` {x}, \<approx>B `` {x}))"
+  "tag_str_Plus A B \<equiv> (\<lambda>x. (\<approx>A `` {x}, \<approx>B `` {x}))"
 
 lemma quot_union_finiteI [intro]:
   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)
+proof (rule_tac tag = "tag_str_Plus 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
+  then show "finite (range (tag_str_Plus A B))"
+    unfolding tag_str_Plus_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
+  show "\<And>x y. tag_str_Plus A B x = tag_str_Plus A B y \<Longrightarrow> x \<approx>(A \<union> B) y"
+    unfolding tag_str_Plus_def
     unfolding str_eq_def
     unfolding str_eq_rel_def
     by auto
 qed
 
 
-subsubsection {* The inductive case for @{text "SEQ"}*}
+subsubsection {* The inductive case for @{text "Times"}*}
 
 definition 
-  tag_str_SEQ :: "lang \<Rightarrow> lang \<Rightarrow> string \<Rightarrow> (lang \<times> lang set)"
+  tag_str_Times :: "'a lang \<Rightarrow> 'a lang \<Rightarrow> 'a list \<Rightarrow> ('a lang \<times> 'a lang set)"
 where
-  "tag_str_SEQ L1 L2 \<equiv>
+  "tag_str_Times L1 L2 \<equiv>
      (\<lambda>x. (\<approx>L1 `` {x}, {(\<approx>L2 `` {x - xa}) | xa.  xa \<le> x \<and> xa \<in> L1}))"
 
 lemma Seq_in_cases:
@@ -225,16 +225,16 @@
   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
+unfolding conc_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" 
+lemma tag_str_Times_injI:
+  assumes eq_tag: "tag_str_Times A B x = tag_str_Times A B y" 
   shows "x \<approx>(A \<cdot> B) y"
 proof -
   { fix x y z
     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"
+    and tag_xy: "tag_str_Times A B x = tag_str_Times A B y"
     have"y @ z \<in> A \<cdot> B" 
     proof -
       { (* first case with x' in A and (x - x') @ z in B *)
@@ -247,7 +247,7 @@
         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
+            using tag_xy unfolding tag_str_Times_def by simp
           moreover 
 	  have "\<approx>B `` {x - x'} \<in> ?Left" using h1 h2 by auto
           ultimately 
@@ -271,11 +271,11 @@
         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
+           using tag_xy unfolding tag_str_Times_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 \<cdot> B"
-	  unfolding prefix_def Seq_def
+	  unfolding prefix_def conc_def
 	  by (auto) (metis append_assoc)
       }
       ultimately show "y @ z \<in> A \<cdot> B" 
@@ -287,30 +287,30 @@
 qed 
 
 lemma quot_seq_finiteI [intro]:
-  fixes L1 L2::"lang"
+  fixes L1 L2::"'a lang"
   assumes fin1: "finite (UNIV // \<approx>L1)" 
   and     fin2: "finite (UNIV // \<approx>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 \<cdot> L2) y"
-    by (rule tag_str_SEQ_injI)
+proof (rule_tac tag = "tag_str_Times L1 L2" in tag_finite_imageD)
+  show "\<And>x y. tag_str_Times L1 L2 x = tag_str_Times L1 L2 y \<Longrightarrow> x \<approx>(L1 \<cdot> L2) y"
+    by (rule tag_str_Times_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
+  show "finite (range (tag_str_Times L1 L2))" 
+    unfolding tag_str_Times_def
     apply(rule finite_subset[OF _ *])
     unfolding quotient_def
     by auto
 qed
 
 
-subsubsection {* The inductive case for @{const "STAR"} *}
+subsubsection {* The inductive case for @{const "Star"} *}
 
 definition 
-  tag_str_STAR :: "lang \<Rightarrow> string \<Rightarrow> lang set"
+  tag_str_Star :: "'a lang \<Rightarrow> 'a list \<Rightarrow> ('a lang) set"
 where
-  "tag_str_STAR L1 \<equiv> (\<lambda>x. {\<approx>L1 `` {x - xa} | xa. xa < x \<and> xa \<in> L1\<star>})"
+  "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> 
@@ -342,31 +342,33 @@
 
 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)}"
+lemma finite_strict_prefix_set: 
+  shows "finite {xa. xa < (x::'a list)}"
 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"
+lemma tag_str_Star_injI:
+  fixes L\<^isub>1::"('a::finite) lang"
+  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"
+      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)
+        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>}"
+      let ?S = "{xa::('a::finite) list. xa < x \<and> xa \<in> L\<^isub>1\<star> \<and> (x - xa) @ z \<in> L\<^isub>1\<star>}"
       have "finite ?S"
-        by (rule_tac B = "{xa. xa < x}" in finite_subset, 
-          auto simp:finite_strict_prefix_set)
+	by (rule_tac B = "{xa. xa < x}" in finite_subset)
+	   (auto simp: finite_strict_prefix_set)
       moreover have "?S \<noteq> {}" using False xz_in_star
         by (simp, rule_tac x = "[]" in exI, auto simp:strict_prefix_def)
       ultimately have "\<exists> xa_max \<in> ?S. \<forall> xa \<in> ?S. length xa \<le> length xa_max" 
@@ -384,7 +386,7 @@
       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)
+          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 
@@ -417,7 +419,7 @@
                 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) 
+                  using a_in h2 by (auto)
                 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)" 
@@ -432,11 +434,12 @@
            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
+        have "((y - ya) @ za) @ zb \<in> L\<^isub>1\<star>" using  l_za ls_zb 
+	  by (rule_tac append_in_starI) (auto)
         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)
+      with h5 h6 show ?thesis 	
+        by (auto simp:strict_prefix_def elim: prefixE)
     qed
   } 
   from this [OF _ eq_tag] and this [OF _ eq_tag [THEN sym]]
@@ -444,16 +447,17 @@
 qed
 
 lemma quot_star_finiteI [intro]:
+   fixes A::"('a::finite) lang"
   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)
+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>A))" 
     using finite1 by auto
-  show "finite (range (tag_str_STAR A))"
-    unfolding tag_str_STAR_def
+  show "finite (range (tag_str_Star A))"
+    unfolding tag_str_Star_def
     apply(rule finite_subset[OF _ *])
     unfolding quotient_def
     by auto
@@ -462,12 +466,13 @@
 subsubsection{* The conclusion *}
 
 lemma Myhill_Nerode2:
-  shows "finite (UNIV // \<approx>(L_rexp r))"
+  fixes r::"('a::finite) rexp"
+  shows "finite (UNIV // \<approx>(lang r))"
 by (induct r) (auto)
 
-
 theorem Myhill_Nerode:
-  shows "(\<exists>r. A = L_rexp r) \<longleftrightarrow> finite (UNIV // \<approx>A)"
+  fixes A::"('a::finite) lang"
+  shows "(\<exists>r. A = lang r) \<longleftrightarrow> finite (UNIV // \<approx>A)"
 using Myhill_Nerode1 Myhill_Nerode2 by auto
 
 end
--- a/Paper/Paper.thy	Fri Jun 03 13:59:21 2011 +0000
+++ b/Paper/Paper.thy	Mon Jul 25 13:33:38 2011 +0000
@@ -15,17 +15,33 @@
 abbreviation
   "Append_rexp2 r_itm r == Append_rexp r r_itm"
 
+abbreviation
+  "pow" (infixl "\<up>" 100)
+where
+  "A \<up> n \<equiv> A ^^ n"
+
+
+abbreviation "NULL \<equiv> Zero"
+abbreviation "EMPTY \<equiv> One"
+abbreviation "CHAR \<equiv> Atom"
+abbreviation "ALT \<equiv> Plus"
+abbreviation "SEQ \<equiv> Times"
+abbreviation "STAR \<equiv> Star"
+
+
+ML {* @{term "op ^^"} *}
+
 notation (latex output)
   str_eq_rel ("\<approx>\<^bsub>_\<^esub>") and
   str_eq ("_ \<approx>\<^bsub>_\<^esub> _") and
-  Seq (infixr "\<cdot>" 100) and
-  Star ("_\<^bsup>\<star>\<^esup>") and
+  conc (infixr "\<cdot>" 100) and
+  star ("_\<^bsup>\<star>\<^esup>") and
   pow ("_\<^bsup>_\<^esup>" [100, 100] 100) and
   Suc ("_+1" [100] 100) and
   quotient ("_ \<^raw:\ensuremath{\!\sslash\!}> _" [90, 90] 90) and
   REL ("\<approx>") and
   UPLUS ("_ \<^raw:\ensuremath{\uplus}> _" [90, 90] 90) and
-  L_rexp ("\<^raw:\ensuremath{\cal{L}}>'(_')" [0] 101) and
+  lang ("\<^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
@@ -35,18 +51,27 @@
   Append_rexp_rhs ("_ \<^raw:\ensuremath{\triangleleft}> _" [100, 100] 50) and
   
   uminus ("\<^raw:\ensuremath{\overline{>_\<^raw:}}>" [100] 100) and
-  tag_str_ALT ("tag\<^isub>A\<^isub>L\<^isub>T _ _" [100, 100] 100) and
-  tag_str_ALT ("tag\<^isub>A\<^isub>L\<^isub>T _ _ _" [100, 100, 100] 100) and
-  tag_str_SEQ ("tag\<^isub>S\<^isub>E\<^isub>Q _ _" [100, 100] 100) and
-  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)
+  tag_str_Plus ("tag\<^isub>A\<^isub>L\<^isub>T _ _" [100, 100] 100) and
+  tag_str_Plus ("tag\<^isub>A\<^isub>L\<^isub>T _ _ _" [100, 100, 100] 100) and
+  tag_str_Times ("tag\<^isub>S\<^isub>E\<^isub>Q _ _" [100, 100] 100) and
+  tag_str_Times ("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
 
+lemma conc_def':
+  "A \<cdot> B = {s\<^isub>1 @ s\<^isub>2 | s\<^isub>1 s\<^isub>2. s\<^isub>1 \<in> A \<and> s\<^isub>2 \<in> B}"
+unfolding conc_def by simp
+
 (* THEOREMS *)
+
+lemma conc_Union_left: 
+  shows "B \<cdot> (\<Union>n. A \<up> n) = (\<Union>n. B \<cdot> (A \<up> n))"
+unfolding conc_def by auto
+
 notation (Rule output)
   "==>"  ("\<^raw:\mbox{}\inferrule{\mbox{>_\<^raw:}}>\<^raw:{\mbox{>_\<^raw:}}>")
 
@@ -277,23 +302,23 @@
   @{term "A \<up> n"}. They are defined as usual
 
   \begin{center}
-  @{thm Seq_def[THEN eq_reflection, where A1="A" and B1="B"]}
+  @{thm conc_def'[THEN eq_reflection, where A1="A" and B1="B"]}
   \hspace{7mm}
-  @{thm pow.simps(1)[THEN eq_reflection, where A1="A"]}
+  @{thm lang_pow.simps(1)[THEN eq_reflection, where A1="A"]}
   \hspace{7mm}
-  @{thm pow.simps(2)[THEN eq_reflection, where A1="A" and n1="n"]}
+  @{thm lang_pow.simps(2)[THEN eq_reflection, where A1="A" and n1="n"]}
   \end{center}
 
   \noindent
   where @{text "@"} is the list-append operation. The Kleene-star of a language @{text A}
-  is defined as the union over all powers, namely @{thm Star_def}. In the paper
+  is defined as the union over all powers, namely @{thm star_def}. In the paper
   we will make use of the following properties of these constructions.
   
   \begin{proposition}\label{langprops}\mbox{}\\
   \begin{tabular}{@ {}ll}
   (i)   & @{thm star_cases}     \\ 
   (ii)  & @{thm[mode=IfThen] pow_length}\\
-  (iii) & @{thm seq_Union_left} \\ 
+  (iii) & @{thm conc_Union_left} \\ 
   \end{tabular}
   \end{proposition}
 
@@ -372,18 +397,18 @@
   \begin{center}
   \begin{tabular}{c@ {\hspace{10mm}}c}
   \begin{tabular}{rcl}
-  @{thm (lhs) L_rexp.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) L_rexp.simps(1)}\\
-  @{thm (lhs) L_rexp.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) L_rexp.simps(2)}\\
-  @{thm (lhs) L_rexp.simps(3)[where c="c"]} & @{text "\<equiv>"} & @{thm (rhs) L_rexp.simps(3)[where c="c"]}\\
+  @{thm (lhs) lang.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) lang.simps(1)}\\
+  @{thm (lhs) lang.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) lang.simps(2)}\\
+  @{thm (lhs) lang.simps(3)[where a="c"]} & @{text "\<equiv>"} & @{thm (rhs) lang.simps(3)[where a="c"]}\\
   \end{tabular}
   &
   \begin{tabular}{rcl}
-  @{thm (lhs) L_rexp.simps(4)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]} & @{text "\<equiv>"} &
-       @{thm (rhs) L_rexp.simps(4)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]}\\
-  @{thm (lhs) L_rexp.simps(5)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]} & @{text "\<equiv>"} &
-       @{thm (rhs) L_rexp.simps(5)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]}\\
-  @{thm (lhs) L_rexp.simps(6)[where r="r"]} & @{text "\<equiv>"} &
-      @{thm (rhs) L_rexp.simps(6)[where r="r"]}\\
+  @{thm (lhs) lang.simps(4)[where ?r="r\<^isub>1" and ?s="r\<^isub>2"]} & @{text "\<equiv>"} &
+       @{thm (rhs) lang.simps(4)[where ?r="r\<^isub>1" and ?s="r\<^isub>2"]}\\
+  @{thm (lhs) lang.simps(5)[where ?r="r\<^isub>1" and ?s="r\<^isub>2"]} & @{text "\<equiv>"} &
+       @{thm (rhs) lang.simps(5)[where ?r="r\<^isub>1" and ?s="r\<^isub>2"]}\\
+  @{thm (lhs) lang.simps(6)[where r="r"]} & @{text "\<equiv>"} &
+      @{thm (rhs) lang.simps(6)[where r="r"]}\\
   \end{tabular}
   \end{tabular}
   \end{center}
@@ -528,8 +553,8 @@
   
   \begin{center}
   @{text "\<calL>(Y, r) \<equiv>"} %
-  @{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]}
+  @{thm (rhs) lang_trm.simps(2)[where X="Y" and r="r", THEN eq_reflection]}\hspace{10mm}
+  @{thm lang_trm.simps(1)[where r="r", THEN eq_reflection]}
   \end{center}
 
   \noindent
@@ -568,9 +593,14 @@
   @{thm (lhs) Init_def}     & @{text "\<equiv>"} & @{thm (rhs) Init_def}
   \end{tabular}}
   \end{equation}
+*}(*<*)
 
-  
+lemma test:
+  assumes X_in_eqs: "(X, rhs) \<in> Init (UNIV // \<approx>A)"
+  shows "X = \<Union> (lang_trm `  rhs)"
+using assms l_eq_r_in_eqs by (simp)
 
+(*>*)text {*
   \noindent
   Because we use sets of terms 
   for representing the right-hand sides of equations, we can 
@@ -926,9 +956,9 @@
 
   \begin{center}
   \begin{tabular}{l}
-  @{thm quot_null_eq}\\
-  @{thm quot_empty_subset}\\
-  @{thm quot_char_subset}
+  @{thm quot_zero_eq}\\
+  @{thm quot_one_subset}\\
+  @{thm quot_atom_subset}
   \end{tabular}
   \end{center}
 
@@ -1021,7 +1051,7 @@
   We take as tagging-function 
   %
   \begin{center}
-  @{thm tag_str_ALT_def[where A="A" and B="B", THEN meta_eq_app]}
+  @{thm tag_str_Plus_def[where A="A" and B="B", THEN meta_eq_app]}
   \end{center}
 
   \noindent
@@ -1154,7 +1184,7 @@
   following tagging-function
   %
   \begin{center}
-  @{thm tag_str_SEQ_def[where ?L1.0="A" and ?L2.0="B", THEN meta_eq_app]}
+  @{thm tag_str_Times_def[where ?L1.0="A" and ?L2.0="B", THEN meta_eq_app]}
   \end{center}
 
   \noindent
@@ -1268,7 +1298,7 @@
   the following tagging-function:
   %
   \begin{center}
-  @{thm tag_str_STAR_def[where ?L1.0="A", THEN meta_eq_app]}\smallskip
+  @{thm tag_str_Star_def[where ?L1.0="A", THEN meta_eq_app]}\smallskip
   \end{center}
 
   \begin{proof}[@{const STAR}-Case]
--- a/Paper/ROOT.ML	Fri Jun 03 13:59:21 2011 +0000
+++ b/Paper/ROOT.ML	Mon Jul 25 13:33:38 2011 +0000
@@ -1,3 +1,3 @@
-no_document use_thy "../Myhill";
+no_document use_thy "../Myhill_2";
 
 use_thy "Paper"
\ No newline at end of file
--- a/Prelude.thy	Fri Jun 03 13:59:21 2011 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,19 +0,0 @@
-theory Prelude
-imports Main
-begin
-
-(*
-To make the theory work under Isabelle 2009 and 2011
-
-Isabelle 2009: set_ext
-Isabelle 2011: set_eqI
-
-*)
-
-
-lemma set_eq_intro:
-  "(\<And>x. (x \<in> A) = (x \<in> B)) \<Longrightarrow> A = B"
-by blast
-
-
-end
\ No newline at end of file
--- a/Regular.thy	Fri Jun 03 13:59:21 2011 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,302 +0,0 @@
-(* Author: Christian Urban, Xingyuan Zhang, Chunhan Wu *)
-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 "\<cdot>" 100)
-where 
-  "A \<cdot> 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 "\<cdot>"}. *}
-
-lemma seq_add_left:
-  assumes a: "A = B"
-  shows "C \<cdot> A = C \<cdot> B"
-using a by simp
-
-lemma seq_union_distrib_right:
-  shows "(A \<union> B) \<cdot> C = (A \<cdot> C) \<union> (B \<cdot> C)"
-unfolding Seq_def by auto
-
-lemma seq_union_distrib_left:
-  shows "C \<cdot> (A \<union> B) = (C \<cdot> A) \<union> (C \<cdot> B)"
-unfolding Seq_def by  auto
-
-lemma seq_intro:
-  assumes a: "x \<in> A" "y \<in> B"
-  shows "x @ y \<in> A \<cdot> B "
-using a by (auto simp: Seq_def)
-
-lemma seq_assoc:
-  shows "(A \<cdot> B) \<cdot> C = A \<cdot> (B \<cdot> C)"
-unfolding Seq_def
-apply(auto)
-apply(blast)
-by (metis append_assoc)
-
-lemma seq_empty [simp]:
-  shows "A \<cdot> {[]} = A"
-  and   "{[]} \<cdot> A = A"
-by (simp_all add: Seq_def)
-
-lemma seq_null [simp]:
-  shows "A \<cdot> {} = {}"
-  and   "{} \<cdot> 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 \<cdot> (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 \<cdot> A\<star>"
-proof
-  { fix x
-    have "x \<in> A\<star> \<Longrightarrow> x \<in> {[]} \<union> A \<cdot> A\<star>"
-      unfolding Seq_def
-    by (induct rule: star_induct) (auto)
-  }
-  then show "A\<star> \<subseteq> {[]} \<union> A \<cdot> A\<star>" by auto
-next
-  show "{[]} \<union> A \<cdot> 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 seq_Union_left: 
-  shows "B \<cdot> (\<Union>n. A \<up> n) = (\<Union>n. B \<cdot> (A \<up> n))"
-unfolding Seq_def by auto
-
-lemma seq_Union_right: 
-  shows "(\<Union>n. A \<up> n) \<cdot> B = (\<Union>n. (A \<up> n) \<cdot> B)"
-unfolding Seq_def by auto
-
-lemma seq_pow_comm:
-  shows "A \<cdot> (A \<up> n) = (A \<up> n) \<cdot> A"
-by (induct n) (simp_all add: seq_assoc[symmetric])
-
-lemma seq_star_comm:
-  shows "A \<cdot> A\<star> = A\<star> \<cdot> 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 \<cdot> (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 \<cdot> A \<union> B"
-  shows "X = X \<cdot> (A \<up> Suc n) \<union> (\<Union>m\<in>{0..n}. B \<cdot> (A \<up> m))"
-proof (induct n)
-  case 0 
-  show "X = X \<cdot> (A \<up> Suc 0) \<union> (\<Union>(m::nat)\<in>{0..0}. B \<cdot> (A \<up> m))"
-    using eq by simp
-next
-  case (Suc n)
-  have ih: "X = X \<cdot> (A \<up> Suc n) \<union> (\<Union>m\<in>{0..n}. B \<cdot> (A \<up> m))" by fact
-  also have "\<dots> = (X \<cdot> A \<union> B) \<cdot> (A \<up> Suc n) \<union> (\<Union>m\<in>{0..n}. B \<cdot> (A \<up> m))" using eq by simp
-  also have "\<dots> = X \<cdot> (A \<up> Suc (Suc n)) \<union> (B \<cdot> (A \<up> Suc n)) \<union> (\<Union>m\<in>{0..n}. B \<cdot> (A \<up> m))"
-    by (simp add: seq_union_distrib_right seq_assoc)
-  also have "\<dots> = X \<cdot> (A \<up> Suc (Suc n)) \<union> (\<Union>m\<in>{0..Suc n}. B \<cdot> (A \<up> m))"
-    by (auto simp add: le_Suc_eq)
-  finally show "X = X \<cdot> (A \<up> Suc (Suc n)) \<union> (\<Union>m\<in>{0..Suc n}. B \<cdot> (A \<up> m))" .
-qed
-
-theorem arden:
-  assumes nemp: "[] \<notin> A"
-  shows "X = X \<cdot> A \<union> B \<longleftrightarrow> X = B \<cdot> A\<star>"
-proof
-  assume eq: "X = B \<cdot> A\<star>"
-  have "A\<star> = {[]} \<union> A\<star> \<cdot> A" 
-    unfolding seq_star_comm[symmetric]
-    by (rule star_cases)
-  then have "B \<cdot> A\<star> = B \<cdot> ({[]} \<union> A\<star> \<cdot> A)"
-    by (rule seq_add_left)
-  also have "\<dots> = B \<union> B \<cdot> (A\<star> \<cdot> A)"
-    unfolding seq_union_distrib_left by simp
-  also have "\<dots> = B \<union> (B \<cdot> A\<star>) \<cdot> A" 
-    by (simp only: seq_assoc)
-  finally show "X = X \<cdot> A \<union> B" 
-    using eq by blast 
-next
-  assume eq: "X = X \<cdot> A \<union> B"
-  { fix n::nat
-    have "B \<cdot> (A \<up> n) \<subseteq> X" using arden_helper[OF eq, of "n"] by auto }
-  then have "B \<cdot> 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 \<cdot> (A \<up> Suc k)" 
-      using seq_pow_length[OF nemp] by blast
-    assume "s \<in> X"
-    then have "s \<in> X \<cdot> (A \<up> Suc k) \<union> (\<Union>m\<in>{0..k}. B \<cdot> (A \<up> m))"
-      using arden_helper[OF eq, of "k"] by auto
-    then have "s \<in> (\<Union>m\<in>{0..k}. B \<cdot> (A \<up> m))" using not_in by auto
-    moreover
-    have "(\<Union>m\<in>{0..k}. B \<cdot> (A \<up> m)) \<subseteq> (\<Union>n. B \<cdot> (A \<up> n))" by auto
-    ultimately 
-    have "s \<in> B \<cdot> A\<star>" 
-      unfolding seq_Union_left Star_def by auto }
-  then have "X \<subseteq> B \<cdot> A\<star>" by auto
-  ultimately 
-  show "X = B \<cdot> A\<star>" by simp
-qed
-
-
-section {* Regular Expressions *}
-
-datatype rexp =
-  NULL
-| EMPTY
-| CHAR char
-| SEQ rexp rexp
-| ALT rexp rexp
-| STAR rexp
-
-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) \<cdot> (L_rexp r2)"
-  | "L_rexp (ALT r1 r2) = (L_rexp r1) \<union> (L_rexp r2)"
-  | "L_rexp (STAR r) = (L_rexp r)\<star>"
-
-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_exp}.
-*}
-
-lemma folds_alt_simp [simp]:
-  fixes rs::"rexp set"
-  assumes a: "finite rs"
-  shows "L_rexp (\<Uplus>rs) = \<Union> (L_rexp ` 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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Regular_Exp.thy	Mon Jul 25 13:33:38 2011 +0000
@@ -0,0 +1,36 @@
+(*  Author:     Tobias Nipkow
+    Copyright   1998 TUM
+*)
+
+header "Regular expressions"
+
+theory Regular_Exp
+imports Regular_Set
+begin
+
+datatype 'a rexp =
+  Zero |
+  One |
+  Atom 'a |
+  Plus "('a rexp)" "('a rexp)" |
+  Times "('a rexp)" "('a rexp)" |
+  Star "('a rexp)"
+
+primrec lang :: "'a rexp => 'a list set" where
+"lang Zero = {}" |
+"lang One = {[]}" |
+"lang (Atom a) = {[a]}" |
+"lang (Plus r s) = (lang r) Un (lang s)" |
+"lang (Times r s) = conc (lang r) (lang s)" |
+"lang (Star r) = star(lang r)"
+
+primrec atoms :: "'a rexp \<Rightarrow> 'a set"
+where
+"atoms Zero = {}" |
+"atoms One = {}" |
+"atoms (Atom a) = {a}" |
+"atoms (Plus r s) = atoms r \<union> atoms s" |
+"atoms (Times r s) = atoms r \<union> atoms s" |
+"atoms (Star r) = atoms r"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Regular_Set.thy	Mon Jul 25 13:33:38 2011 +0000
@@ -0,0 +1,348 @@
+(*  Author: Tobias Nipkow, Alex Krauss  *)
+
+header "Regular sets"
+
+theory Regular_Set
+imports Main
+begin
+
+type_synonym 'a lang = "'a list set"
+
+definition conc :: "'a lang \<Rightarrow> 'a lang \<Rightarrow> 'a lang" (infixr "@@" 75) where
+"A @@ B = {xs@ys | xs ys. xs:A & ys:B}"
+
+overloading lang_pow == "compow :: nat \<Rightarrow> 'a lang \<Rightarrow> 'a lang"
+begin
+  primrec lang_pow :: "nat \<Rightarrow> 'a lang \<Rightarrow> 'a lang" where
+  "lang_pow 0 A = {[]}" |
+  "lang_pow (Suc n) A = A @@ (lang_pow n A)"
+end
+
+definition star :: "'a lang \<Rightarrow> 'a lang" where
+"star A = (\<Union>n. A ^^ n)"
+
+
+
+definition deriv :: "'a \<Rightarrow> 'a lang \<Rightarrow> 'a lang"
+where "deriv x L = { xs. x#xs \<in> L }"
+
+
+coinductive bisimilar :: "'a list set \<Rightarrow> 'a list set \<Rightarrow> bool" where
+"([] \<in> K \<longleftrightarrow> [] \<in> L) 
+ \<Longrightarrow> (\<And>x. bisimilar (deriv x K) (deriv x L))
+ \<Longrightarrow> bisimilar K L"
+
+
+subsection{* @{term "op @@"} *}
+
+lemma concI[simp,intro]: "u : A \<Longrightarrow> v : B \<Longrightarrow> u@v : A @@ B"
+by (auto simp add: conc_def)
+
+lemma concE[elim]: 
+assumes "w \<in> A @@ B"
+obtains u v where "u \<in> A" "v \<in> B" "w = u@v"
+using assms by (auto simp: conc_def)
+
+lemma conc_mono: "A \<subseteq> C \<Longrightarrow> B \<subseteq> D \<Longrightarrow> A @@ B \<subseteq> C @@ D"
+by (auto simp: conc_def) 
+
+lemma conc_empty[simp]: shows "{} @@ A = {}" and "A @@ {} = {}"
+by auto
+
+lemma conc_epsilon[simp]: shows "{[]} @@ A = A" and "A @@ {[]} = A"
+by (simp_all add:conc_def)
+
+lemma conc_assoc: "(A @@ B) @@ C = A @@ (B @@ C)"
+by (auto elim!: concE) (simp only: append_assoc[symmetric] concI)
+
+lemma conc_Un_distrib:
+shows "A @@ (B \<union> C) = A @@ B \<union> A @@ C"
+and   "(A \<union> B) @@ C = A @@ C \<union> B @@ C"
+by auto
+
+lemma conc_UNION_distrib:
+shows "A @@ UNION I M = UNION I (%i. A @@ M i)"
+and   "UNION I M @@ A = UNION I (%i. M i @@ A)"
+by auto
+
+
+subsection{* @{term "A ^^ n"} *}
+
+lemma lang_pow_add: "A ^^ (n + m) = A ^^ n @@ A ^^ m"
+by (induct n) (auto simp: conc_assoc)
+
+lemma lang_pow_empty: "{} ^^ n = (if n = 0 then {[]} else {})"
+by (induct n) auto
+
+lemma lang_pow_empty_Suc[simp]: "({}::'a lang) ^^ Suc n = {}"
+by (simp add: lang_pow_empty)
+
+
+lemma length_lang_pow_ub:
+  "ALL w : A. length w \<le> k \<Longrightarrow> w : A^^n \<Longrightarrow> length w \<le> k*n"
+by(induct n arbitrary: w) (fastsimp simp: conc_def)+
+
+lemma length_lang_pow_lb:
+  "ALL w : A. length w \<ge> k \<Longrightarrow> w : A^^n \<Longrightarrow> length w \<ge> k*n"
+by(induct n arbitrary: w) (fastsimp simp: conc_def)+
+
+
+subsection{* @{const star} *}
+
+lemma star_if_lang_pow[simp]: "w : A ^^ n \<Longrightarrow> w : star A"
+by (auto simp: star_def)
+
+lemma Nil_in_star[iff]: "[] : star A"
+proof (rule star_if_lang_pow)
+  show "[] : A ^^ 0" by simp
+qed
+
+lemma star_if_lang[simp]: assumes "w : A" shows "w : star A"
+proof (rule star_if_lang_pow)
+  show "w : A ^^ 1" using `w : A` by simp
+qed
+
+lemma append_in_starI[simp]:
+assumes "u : star A" and "v : star A" shows "u@v : star A"
+proof -
+  from `u : star A` obtain m where "u : A ^^ m" by (auto simp: star_def)
+  moreover
+  from `v : star A` obtain n where "v : A ^^ n" by (auto simp: star_def)
+  ultimately have "u@v : A ^^ (m+n)" by (simp add: lang_pow_add)
+  thus ?thesis by simp
+qed
+
+lemma conc_star_star: "star A @@ star A = star A"
+by (auto simp: conc_def)
+
+lemma star_induct[consumes 1, case_names Nil append, induct set: star]:
+assumes "w : star A"
+  and "P []"
+  and step: "!!u v. u : A \<Longrightarrow> v : star A \<Longrightarrow> P v \<Longrightarrow> P (u@v)"
+shows "P w"
+proof -
+  { fix n have "w : A ^^ n \<Longrightarrow> P w"
+    by (induct n arbitrary: w) (auto intro: `P []` step star_if_lang_pow) }
+  with `w : star A` show "P w" by (auto simp: star_def)
+qed
+
+lemma star_empty[simp]: "star {} = {[]}"
+by (auto elim: star_induct)
+
+lemma star_epsilon[simp]: "star {[]} = {[]}"
+by (auto elim: star_induct)
+
+lemma star_idemp[simp]: "star (star A) = star A"
+by (auto elim: star_induct)
+
+lemma star_unfold_left: "star A = A @@ star A \<union> {[]}" (is "?L = ?R")
+proof
+  show "?L \<subseteq> ?R" by (rule, erule star_induct) auto
+qed auto
+
+lemma concat_in_star: "set ws \<subseteq> A \<Longrightarrow> concat ws : star A"
+by (induct ws) simp_all
+
+lemma in_star_iff_concat:
+  "w : star A = (EX ws. set ws \<subseteq> A & w = concat ws)"
+  (is "_ = (EX ws. ?R w ws)")
+proof
+  assume "w : star A" thus "EX ws. ?R w ws"
+  proof induct
+    case Nil have "?R [] []" by simp
+    thus ?case ..
+  next
+    case (append u v)
+    moreover
+    then obtain ws where "set ws \<subseteq> A \<and> v = concat ws" by blast
+    ultimately have "?R (u@v) (u#ws)" by auto
+    thus ?case ..
+  qed
+next
+  assume "EX us. ?R w us" thus "w : star A"
+  by (auto simp: concat_in_star)
+qed
+
+lemma star_conv_concat: "star A = {concat ws|ws. set ws \<subseteq> A}"
+by (fastsimp simp: in_star_iff_concat)
+
+lemma star_insert_eps[simp]: "star (insert [] A) = star(A)"
+proof-
+  { fix us
+    have "set us \<subseteq> insert [] A \<Longrightarrow> EX vs. concat us = concat vs \<and> set vs \<subseteq> A"
+      (is "?P \<Longrightarrow> EX vs. ?Q vs")
+    proof
+      let ?vs = "filter (%u. u \<noteq> []) us"
+      show "?P \<Longrightarrow> ?Q ?vs" by (induct us) auto
+    qed
+  } thus ?thesis by (auto simp: star_conv_concat)
+qed
+
+lemma Arden:
+assumes "[] \<notin> A" and "X = A @@ X \<union> B"
+shows "X = star A @@ B"
+proof -
+  { fix n have "X = A^^(n+1)@@X \<union> (\<Union>i\<le>n. A^^i@@B)"
+    proof(induct n)
+      case 0 show ?case using `X = A @@ X \<union> B` by simp
+    next
+      case (Suc n)
+      have "X = A@@X Un B" by(rule assms(2))
+      also have "\<dots> = A@@(A^^(n+1)@@X \<union> (\<Union>i\<le>n. A^^i@@B)) Un B"
+        by(subst Suc)(rule refl)
+      also have "\<dots> =  A^^(n+2)@@X \<union> (\<Union>i\<le>n. A^^(i+1)@@B) Un B"
+        by(simp add:conc_UNION_distrib conc_assoc conc_Un_distrib)
+      also have "\<dots> =  A^^(n+2)@@X \<union> (UN i : {1..n+1}. A^^i@@B) \<union> B"
+        by(subst UN_le_add_shift)(rule refl)
+      also have "\<dots> =  A^^(n+2)@@X \<union> (UN i : {1..n+1}. A^^i@@B) \<union> A^^0@@B"
+        by(simp)
+      also have "\<dots> =  A^^(n+2)@@X \<union> (\<Union>i\<le>n+1. A^^i@@B)"
+        by(auto simp: UN_le_eq_Un0)
+      finally show ?case by simp
+    qed
+  } note 1 = this
+  { fix w assume "w : X"
+    let ?n = "size w"
+    from `[] \<notin> A` have "ALL u : A. length u \<ge> 1"
+      by (metis Suc_eq_plus1 add_leD2 le_0_eq length_0_conv not_less_eq_eq)
+    hence "ALL u : A^^(?n+1). length u \<ge> ?n+1"
+      by (metis length_lang_pow_lb nat_mult_1)
+    hence "ALL u : A^^(?n+1)@@X. length u \<ge> ?n+1"
+      by(auto simp only: conc_def length_append)
+    hence "w \<notin> A^^(?n+1)@@X" by auto
+    hence "w : star A @@ B" using `w : X` 1[of ?n]
+      by(auto simp add: star_def conc_UNION_distrib)
+  } moreover
+  { fix w assume "w : star A @@ B"
+    hence "EX n. w : A^^n @@ B" by(auto simp: conc_def star_def)
+    hence "w : X" using 1 by blast
+  } ultimately show ?thesis by blast
+qed
+
+subsection{* @{const deriv} *}
+
+lemma deriv_empty[simp]: "deriv a {} = {}"
+and deriv_epsilon[simp]: "deriv a {[]} = {}"
+and deriv_char[simp]: "deriv a {[b]} = (if a = b then {[]} else {})"
+and deriv_union[simp]: "deriv a (A \<union> B) = deriv a A \<union> deriv a B"
+by (auto simp: deriv_def)
+
+lemma deriv_conc_subset:
+"deriv a A @@ B \<subseteq> deriv a (A @@ B)" (is "?L \<subseteq> ?R")
+proof 
+  fix w assume "w \<in> ?L"
+  then obtain u v where "w = u @ v" "a # u \<in> A" "v \<in> B"
+    by (auto simp: deriv_def)
+  then have "a # w \<in> A @@ B"
+    by (auto intro: concI[of "a # u", simplified])
+  thus "w \<in> ?R" by (auto simp: deriv_def)
+qed
+
+lemma deriv_conc1:
+assumes "[] \<notin> A"
+shows "deriv a (A @@ B) = deriv a A @@ B" (is "?L = ?R")
+proof
+  show "?L \<subseteq> ?R"
+  proof
+    fix w assume "w \<in> ?L"
+    then have "a # w \<in> A @@ B" by (simp add: deriv_def)
+    then obtain x y 
+      where aw: "a # w = x @ y" "x \<in> A" "y \<in> B" by auto
+    with `[] \<notin> A` obtain x' where "x = a # x'"
+      by (cases x) auto
+    with aw have "w = x' @ y" "x' \<in> deriv a A"
+      by (auto simp: deriv_def)
+    with `y \<in> B` show "w \<in> ?R" by simp
+  qed
+  show "?R \<subseteq> ?L" by (fact deriv_conc_subset)
+qed
+
+lemma deriv_conc2:
+assumes "[] \<in> A"
+shows "deriv a (A @@ B) = deriv a A @@ B \<union> deriv a B"
+(is "?L = ?R")
+proof
+  show "?L \<subseteq> ?R"
+  proof
+    fix w assume "w \<in> ?L"
+    then have "a # w \<in> A @@ B" by (simp add: deriv_def)
+    then obtain x y 
+      where aw: "a # w = x @ y" "x \<in> A" "y \<in> B" by auto
+    show "w \<in> ?R"
+    proof (cases x)
+      case Nil
+      with aw have "w \<in> deriv a B" by (auto simp: deriv_def)
+      thus ?thesis ..
+    next
+      case (Cons b x')
+      with aw have "w = x' @ y" "x' \<in> deriv a A"
+        by (auto simp: deriv_def)
+      with `y \<in> B` show "w \<in> ?R" by simp
+    qed
+  qed
+
+  from concI[OF `[] \<in> A`, simplified]
+  have "B \<subseteq> A @@ B" ..
+  then have "deriv a B \<subseteq> deriv a (A @@ B)" by (auto simp: deriv_def)
+  with deriv_conc_subset[of a A B]
+  show "?R \<subseteq> ?L" by auto
+qed
+
+lemma wlog_no_eps: 
+assumes PB: "\<And>B. [] \<notin> B \<Longrightarrow> P B"
+assumes preserved: "\<And>A. P A \<Longrightarrow> P (insert [] A)"
+shows "P A"
+proof -
+  let ?B = "A - {[]}"
+  have "P ?B" by (rule PB) auto
+  thus "P A"
+  proof cases
+    assume "[] \<in> A"
+    then have "(insert [] ?B) = A" by auto
+    with preserved[OF `P ?B`]
+    show ?thesis by simp
+  qed auto
+qed
+
+lemma deriv_insert_eps[simp]: 
+"deriv a (insert [] A) = deriv a A"
+by (auto simp: deriv_def)
+
+lemma deriv_star[simp]: "deriv a (star A) = deriv a A @@ star A"
+proof (induct A rule: wlog_no_eps)
+  fix B :: "'a list set" assume "[] \<notin> B"
+  thus "deriv a (star B) = deriv a B @@ star B" 
+    by (subst star_unfold_left) (simp add: deriv_conc1)
+qed auto
+
+
+subsection{* @{const bisimilar} *}
+
+lemma equal_if_bisimilar:
+assumes "bisimilar K L" shows "K = L"
+proof (rule set_eqI)
+  fix w
+  from `bisimilar K L` show "w \<in> K \<longleftrightarrow> w \<in> L"
+  proof (induct w arbitrary: K L)
+    case Nil thus ?case by (auto elim: bisimilar.cases)
+  next
+    case (Cons a w K L)
+    from `bisimilar K L` have "bisimilar (deriv a K) (deriv a L)"
+      by (auto elim: bisimilar.cases)
+    then have "w \<in> deriv a K \<longleftrightarrow> w \<in> deriv a L" by (rule Cons(1))
+    thus ?case by (auto simp: deriv_def)
+  qed
+qed
+
+lemma language_coinduct:
+fixes R (infixl "\<sim>" 50)
+assumes "K \<sim> L"
+assumes "\<And>K L. K \<sim> L \<Longrightarrow> ([] \<in> K \<longleftrightarrow> [] \<in> L)"
+assumes "\<And>K L x. K \<sim> L \<Longrightarrow> deriv x K \<sim> deriv x L"
+shows "K = L"
+apply (rule equal_if_bisimilar)
+apply (rule bisimilar.coinduct[of R, OF `K \<sim> L`])
+apply (auto simp: assms)
+done
+
+end