Myhill.thy
author zhang
Sun, 06 Feb 2011 10:28:29 +0000
changeset 69 ecf6c61a4541
parent 68 dceb84acdee8
child 73 79de7de104c8
permissions -rw-r--r--
A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
     1
theory Myhill
62
d94209ad2880 Copy Myhill.thy into Myhill_2.thyand put Myhill_2.thy into logic ListP. Myhill.thy is now almost empty ready to be extended by new manuscripts.
zhang
parents: 57
diff changeset
     2
  imports Myhill_2
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
     3
begin
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
     4
48
61d9684a557a Myhill.thy and Myhill_1.thy changed.
zhang
parents: 47
diff changeset
     5
section {* Direction @{text "regular language \<Rightarrow>finite partition"} *}
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
     6
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
     7
text {*
64
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
     8
69
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
     9
A {\em determinisitc finite automata (DFA)} $M$ is a 5-tuple 
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    10
$(Q, \Sigma, \delta, s, F)$, where:
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    11
\begin{enumerate}
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    12
  \item $Q$ is a finite set of {\em states}, also denoted $Q_M$.
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    13
  \item $\Sigma$ is a finite set of {\em alphabets}, also denoted $\Sigma_M$.
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    14
  \item $\delta$ is a {\em transition function} of type @{text "Q \<times> \<Sigma> \<Rightarrow> Q"} (a total function),
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    15
           also denoted $\delta_M$.
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    16
  \item @{text "s \<in> Q"} is a state called {\em initial state}, also denoted $s_M$.
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    17
  \item @{text "F \<subseteq> Q"} is a set of states named {\em accepting states}, also denoted $F_M$.
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    18
\end{enumerate}
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    19
Therefore, we have $M = (Q_M, \Sigma_M, \delta_M, s_M, F_M)$. Every DFA $M$ can be interpreted as 
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    20
a function assigning states to strings, denoted $\dfa{M}$, the  definition of which is as the following:
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    21
\begin{equation}
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    22
\begin{aligned}
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    23
  \dfa{M}([]) &\equiv s_M \\
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    24
   \dfa{M}(xa) &\equiv  \delta_M(\dfa{M}(x), a)
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    25
\end{aligned}
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    26
\end{equation}
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    27
A string @{text "x"} is said to be {\em accepted} (or {\em recognized}) by a DFA $M$ if 
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    28
$\dfa{M}(x) \in F_M$. The language recoginzed by DFA $M$, denoted
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    29
$L(M)$, is defined as:
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    30
\begin{equation}
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    31
  L(M) \equiv \{x~|~\dfa{M}(x) \in F_M\}
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    32
\end{equation}
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    33
The standard way of specifying a laugage $\Lang$ as {\em regular} is by stipulating that:
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    34
$\Lang = L(M)$ for some DFA $M$.
64
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
    35
69
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    36
For any DFA $M$, the DFA obtained by changing initial state to another $p \in Q_M$ is denoted $M_p$, 
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    37
which is defined as:
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    38
\begin{equation}
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    39
    M_p \ \equiv\ (Q_M, \Sigma_M, \delta_M, p, F_M) 
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    40
\end{equation}
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    41
Two states $p, q \in Q_M$ are said to be {\em equivalent}, denoted $p \approx_M q$, iff.
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    42
\begin{equation}\label{m_eq_def}
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    43
  L(M_p) = L(M_q)
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    44
\end{equation}
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    45
It is obvious that $\approx_M$ is an equivalent relation over $Q_M$. and 
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    46
the partition induced by $\approx_M$ has $|Q_M|$ equivalent classes.
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    47
By overloading $\approx_M$,  an equivalent relation over strings can be defined:
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    48
\begin{equation}
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    49
  x \approx_M y ~ ~ \equiv ~ ~ \dfa{M}(x) \approx_M \dfa{M}(y)
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    50
\end{equation}
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    51
It can be proved that the the partition induced by $\approx_M$ also has $|Q_M|$ equivalent classes.
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    52
It is also easy to show that: if $x \approx_M y$, then $x \approx_{L(M)} y$, and this means
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    53
$\approx_M$ is a more refined equivalent relation than $\approx_{L(M)}$. Since partition induced by
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    54
$\approx_M$ is finite, the one induced by $\approx_{L(M)}$ must also be finite. Now, we get one direction
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    55
of \mht:
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    56
\begin{Lem}[\mht , Direction one]
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    57
  If a language $\Lang$ is regular (i.e. $\Lang = L(M)$ for some DFA $M$), then 
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    58
  the partition induced by $\approx_\Lang$ is finite.
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    59
\end{Lem}
68
dceb84acdee8 Check in Myhill.thy before trying another way to explain DFA
zhang
parents: 64
diff changeset
    60
69
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    61
The other direction is:
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    62
\begin{Lem}[\mht , Direction two]\label{auto_mh_d2}
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    63
  If  the partition induced by $\approx_\Lang$ is finite, then
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    64
  $\Lang$ is regular (i.e. $\Lang = L(M)$ for some DFA $M$).
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    65
\end{Lem}
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    66
To prove this lemma, a DFA $M_\Lang$ is constructed out of $\approx_\Lang$ with:
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    67
\begin{subequations}
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    68
\begin{eqnarray}
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    69
  Q_{M_\Lang} ~ & \equiv & ~ \{ \cls{x}{\approx_\Lang}~|~ x \in \Sigma^* \}\\
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    70
  \Sigma_{M_\Lang} ~ & \equiv & ~ \Sigma_M \\
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    71
  \delta_{M_\Lang} ~ & \equiv & ~ \left (\lambda (\cls{x}{\approx_\Lang}, a).  \cls{xa}{\approx_\Lang} \right) \\
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    72
  s_{M_\Lang} ~ & \equiv & ~ \cls{[]}{\approx_\Lang} \\
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    73
  F_{M_\Lang} ~ & \equiv & ~ \{ \cls{x}{\approx_\Lang}~|~ x \in \Lang \}
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    74
\end{eqnarray}
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    75
\end{subequations}
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    76
From the assumption of lemma \ref{auto_mh_d2}, we have that $\{ \cls{x}{\approx_\Lang}~|~ x \in \Sigma^* \}$
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    77
is finite
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    78
It can be proved that $\Lang = L(M_\Lang)$.
68
dceb84acdee8 Check in Myhill.thy before trying another way to explain DFA
zhang
parents: 64
diff changeset
    79
69
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    80
*}
68
dceb84acdee8 Check in Myhill.thy before trying another way to explain DFA
zhang
parents: 64
diff changeset
    81
69
ecf6c61a4541 A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
zhang
parents: 68
diff changeset
    82
text {*
64
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
    83
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
    84
\begin{figure}[h!]
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
    85
\centering
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
    86
\begin{minipage}{0.5\textwidth}
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
    87
\scalebox{0.8}{
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
    88
\begin{tikzpicture}[ultra thick,>=latex]
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
    89
  \node[draw,circle,initial] (start) {$X_0$};
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
    90
  \node[draw,circle,accepting] at ($(start) + (3.5cm,2cm)$) (ac1) {$X_1$};
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
    91
  \node[draw,circle,accepting] at ($(start) + (3.5cm,-2cm)$) (ac2) {$X_2$};
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
    92
  \node[draw,circle] at ($(start) + (6.5cm,0cm)$) (ab) {$X_3$};
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
    93
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
    94
  \path[->] (start) edge [bend left] node [midway, above] {$a$} (ac1);
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
    95
  \path[->] (start) edge [bend right] node [midway, below] {$b$} (ac2);
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
    96
  \path[->] (ac1) edge [loop above] node [midway, above] {$b$} (ac1);
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
    97
  \path[->] (ac2) edge [loop below] node [midway, below] {$a$} (ac2);
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
    98
  \path[->] (ac1) edge [bend right] node [midway, left] {$c$} (ac2);
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
    99
  \path[->] (ac2) edge [bend right] node [midway, right] {$c$} (ac1);
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
   100
  \path[->] (ac1) edge node [midway, sloped, above] {$d$} (start);
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
   101
  \path[->] (ac2) edge node [midway, sloped, above] {$d$} (start);
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
   102
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
   103
  \path [draw, rounded corners,->,dashed] (start) -- ($(start) + (0cm, 3.7cm)$)
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
   104
         -- ($(ab) + (0cm, 3.7cm)$)  node[midway,above,sloped]{$\Sigma - \{a, b\}$} -- (ab);
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
   105
  \path[->,dashed] (ac1) edge node [midway, above, sloped] {$\Sigma - \{b,c,d\}$} (ab);
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
   106
  \path[->,dashed] (ac2) edge node [midway, below, sloped] {$\Sigma - \{a,c,d\}$} (ab);
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
   107
  \path[->,dashed] (ab) edge [loop right] node [midway, right] {$\Sigma$} (ab);
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
   108
\end{tikzpicture}}
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
   109
\end{minipage}
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
   110
\caption{The relationship between automata and finite partition}\label{fig_auto_part_rel}
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
   111
\end{figure}
b69d4e04e64a Added something to Myhill.thy, trying to explain the relationship between finite automata and language partition.
zhang
parents: 62
diff changeset
   112
55
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   113
  *}
d71424eb5d0c Myhill.thy IsabelleMakefile modified
zhang
parents: 49
diff changeset
   114
31
b6815473ee2e 1. Myhill.thy is proved to correct some typos. 2. Some sections are added to the first direction. 3. The small theory of list difference is now separated into Prefix_subtract.thy and the ROOT.ML
zhang
parents: 30
diff changeset
   115
end