A formal presentation of automaton and Myhill-Nerrode theorem added to serve as a basis for further discussion
theory Myhill imports Myhill_2beginsection {* Direction @{text "regular language \<Rightarrow>finite partition"} *}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. Now, we get one directionof \mht:\begin{Lem}[\mht , Direction one] 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 two]\label{auto_mh_d2} If the partition induced by $\approx_\Lang$ is finite, then $\Lang$ is regular (i.e. $\Lang = L(M)$ for some DFA $M$).\end{Lem}To prove this lemma, a DFA $M_\Lang$ is constructed out of $\approx_\Lang$ with:\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}From the assumption of lemma \ref{auto_mh_d2}, we have that $\{ \cls{x}{\approx_\Lang}~|~ x \in \Sigma^* \}$is finiteIt can be proved that $\Lang = L(M_\Lang)$.*}text {*\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{The relationship between automata and finite partition}\label{fig_auto_part_rel}\end{figure} *}end