24
|
1 |
(*<*)
|
|
2 |
theory Paper
|
94
|
3 |
imports "../Myhill" "LaTeXsugar"
|
24
|
4 |
begin
|
39
|
5 |
|
|
6 |
declare [[show_question_marks = false]]
|
|
7 |
|
54
|
8 |
consts
|
|
9 |
REL :: "(string \<times> string) \<Rightarrow> bool"
|
66
|
10 |
UPLUS :: "'a set \<Rightarrow> 'a set \<Rightarrow> (nat \<times> 'a) set"
|
54
|
11 |
|
70
|
12 |
abbreviation
|
|
13 |
"EClass x R \<equiv> R `` {x}"
|
54
|
14 |
|
92
|
15 |
abbreviation
|
|
16 |
"append_rexp2 r_itm r \<equiv> append_rexp r r_itm"
|
|
17 |
|
|
18 |
|
39
|
19 |
notation (latex output)
|
50
|
20 |
str_eq_rel ("\<approx>\<^bsub>_\<^esub>") and
|
75
|
21 |
str_eq ("_ \<approx>\<^bsub>_\<^esub> _") and
|
50
|
22 |
Seq (infixr "\<cdot>" 100) and
|
|
23 |
Star ("_\<^bsup>\<star>\<^esup>") and
|
|
24 |
pow ("_\<^bsup>_\<^esup>" [100, 100] 100) and
|
58
|
25 |
Suc ("_+1" [100] 100) and
|
54
|
26 |
quotient ("_ \<^raw:\ensuremath{\!\sslash\!}> _" [90, 90] 90) and
|
66
|
27 |
REL ("\<approx>") and
|
67
|
28 |
UPLUS ("_ \<^raw:\ensuremath{\uplus}> _" [90, 90] 90) and
|
82
|
29 |
L ("\<^raw:\ensuremath{\cal{L}}>'(_')" [0] 101) and
|
75
|
30 |
Lam ("\<lambda>'(_')" [100] 100) and
|
89
|
31 |
Trn ("'(_, _')" [100, 100] 100) and
|
71
|
32 |
EClass ("\<lbrakk>_\<rbrakk>\<^bsub>_\<^esub>" [100, 100] 100) and
|
88
|
33 |
transition ("_ \<^raw:\ensuremath{\stackrel{\text{>_\<^raw:}}{\Longmapsto}}> _" [100, 100, 100] 100) and
|
92
|
34 |
Setalt ("\<^raw:\ensuremath{\bigplus}>_" [1000] 999) and
|
|
35 |
append_rexp2 ("_ \<^raw:\ensuremath{\triangleleft}> _" [100, 100] 100) and
|
95
|
36 |
append_rhs_rexp ("_ \<^raw:\ensuremath{\triangleleft}> _" [100, 100] 50)
|
92
|
37 |
|
24
|
38 |
(*>*)
|
|
39 |
|
70
|
40 |
|
24
|
41 |
section {* Introduction *}
|
|
42 |
|
|
43 |
text {*
|
58
|
44 |
Regular languages are an important and well-understood subject in Computer
|
60
|
45 |
Science, with many beautiful theorems and many useful algorithms. There is a
|
66
|
46 |
wide range of textbooks on this subject, many of which are aimed at students
|
|
47 |
and contain very detailed ``pencil-and-paper'' proofs
|
60
|
48 |
(e.g.~\cite{Kozen97}). It seems natural to exercise theorem provers by
|
101
|
49 |
formalising the theorems and by verifying formally the algorithms.
|
59
|
50 |
|
66
|
51 |
There is however a problem: the typical approach to regular languages is to
|
|
52 |
introduce finite automata and then define everything in terms of them. For
|
|
53 |
example, a regular language is normally defined as one whose strings are
|
|
54 |
recognised by a finite deterministic automaton. This approach has many
|
71
|
55 |
benefits. Among them is the fact that it is easy to convince oneself that
|
66
|
56 |
regular languages are closed under complementation: one just has to exchange
|
|
57 |
the accepting and non-accepting states in the corresponding automaton to
|
|
58 |
obtain an automaton for the complement language. The problem, however, lies with
|
67
|
59 |
formalising such reasoning in a HOL-based theorem prover, in our case
|
70
|
60 |
Isabelle/HOL. Automata are build up from states and transitions that
|
82
|
61 |
need to be represented as graphs, matrices or functions, none
|
|
62 |
of which can be defined as inductive datatype.
|
66
|
63 |
|
82
|
64 |
In case of graphs and matrices, this means we have to build our own
|
|
65 |
reasoning infrastructure for them, as neither Isabelle/HOL nor HOL4 nor
|
|
66 |
HOLlight support them with libraries. Even worse, reasoning about graphs and
|
|
67 |
matrices can be a real hassle in HOL-based theorem provers. Consider for
|
|
68 |
example the operation of sequencing two automata, say $A_1$ and $A_2$, by
|
|
69 |
connecting the accepting states of $A_1$ to the initial state of $A_2$:
|
60
|
70 |
|
|
71 |
\begin{center}
|
66
|
72 |
\begin{tabular}{ccc}
|
|
73 |
\begin{tikzpicture}[scale=0.8]
|
|
74 |
%\draw[step=2mm] (-1,-1) grid (1,1);
|
|
75 |
|
|
76 |
\draw[rounded corners=1mm, very thick] (-1.0,-0.3) rectangle (-0.2,0.3);
|
|
77 |
\draw[rounded corners=1mm, very thick] ( 0.2,-0.3) rectangle ( 1.0,0.3);
|
|
78 |
|
|
79 |
\node (A) at (-1.0,0.0) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
|
|
80 |
\node (B) at ( 0.2,0.0) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
|
|
81 |
|
|
82 |
\node (C) at (-0.2, 0.13) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
|
|
83 |
\node (D) at (-0.2,-0.13) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
|
|
84 |
|
|
85 |
\node (E) at (1.0, 0.2) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
|
|
86 |
\node (F) at (1.0,-0.0) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
|
|
87 |
\node (G) at (1.0,-0.2) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
|
|
88 |
|
|
89 |
\draw (-0.6,0.0) node {\footnotesize$A_1$};
|
|
90 |
\draw ( 0.6,0.0) node {\footnotesize$A_2$};
|
|
91 |
\end{tikzpicture}
|
|
92 |
|
|
93 |
&
|
|
94 |
|
|
95 |
\raisebox{1.1mm}{\bf\Large$\;\;\;\Rightarrow\,\;\;$}
|
|
96 |
|
|
97 |
&
|
|
98 |
|
|
99 |
\begin{tikzpicture}[scale=0.8]
|
|
100 |
%\draw[step=2mm] (-1,-1) grid (1,1);
|
|
101 |
|
|
102 |
\draw[rounded corners=1mm, very thick] (-1.0,-0.3) rectangle (-0.2,0.3);
|
|
103 |
\draw[rounded corners=1mm, very thick] ( 0.2,-0.3) rectangle ( 1.0,0.3);
|
|
104 |
|
|
105 |
\node (A) at (-1.0,0.0) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
|
|
106 |
\node (B) at ( 0.2,0.0) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
|
|
107 |
|
|
108 |
\node (C) at (-0.2, 0.13) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
|
|
109 |
\node (D) at (-0.2,-0.13) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
|
|
110 |
|
|
111 |
\node (E) at (1.0, 0.2) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
|
|
112 |
\node (F) at (1.0,-0.0) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
|
|
113 |
\node (G) at (1.0,-0.2) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
|
|
114 |
|
|
115 |
\draw (C) to [very thick, bend left=45] (B);
|
|
116 |
\draw (D) to [very thick, bend right=45] (B);
|
|
117 |
|
|
118 |
\draw (-0.6,0.0) node {\footnotesize$A_1$};
|
|
119 |
\draw ( 0.6,0.0) node {\footnotesize$A_2$};
|
|
120 |
\end{tikzpicture}
|
|
121 |
|
|
122 |
\end{tabular}
|
60
|
123 |
\end{center}
|
|
124 |
|
|
125 |
\noindent
|
67
|
126 |
On ``paper'' we can define the corresponding graph in terms of the disjoint
|
88
|
127 |
union of the state nodes. Unfortunately in HOL, the standard definition for disjoint
|
66
|
128 |
union, namely
|
82
|
129 |
%
|
|
130 |
\begin{equation}\label{disjointunion}
|
66
|
131 |
@{term "UPLUS A\<^isub>1 A\<^isub>2 \<equiv> {(1, x) | x. x \<in> A\<^isub>1} \<union> {(2, y) | y. y \<in> A\<^isub>2}"}
|
82
|
132 |
\end{equation}
|
60
|
133 |
|
61
|
134 |
\noindent
|
66
|
135 |
changes the type---the disjoint union is not a set, but a set of pairs.
|
|
136 |
Using this definition for disjoint unions means we do not have a single type for automata
|
92
|
137 |
and hence will not be able to state certain properties about \emph{all}
|
67
|
138 |
automata, since there is no type quantification available in HOL. An
|
|
139 |
alternative, which provides us with a single type for automata, is to give every
|
|
140 |
state node an identity, for example a natural
|
70
|
141 |
number, and then be careful to rename these identities apart whenever
|
67
|
142 |
connecting two automata. This results in clunky proofs
|
66
|
143 |
establishing that properties are invariant under renaming. Similarly,
|
67
|
144 |
connecting two automata represented as matrices results in very adhoc
|
66
|
145 |
constructions, which are not pleasant to reason about.
|
|
146 |
|
82
|
147 |
Functions are much better supported in Isabelle/HOL, but they still lead to similar
|
88
|
148 |
problems as with graphs. Composing, for example, two non-deterministic automata in parallel
|
93
|
149 |
requires also the formalisation of disjoint unions. Nipkow \cite{Nipkow98}
|
101
|
150 |
dismisses for this the option of using identities, because it leads according to
|
|
151 |
him to ``messy proofs''. He
|
103
|
152 |
opts for a variant of \eqref{disjointunion} using bit lists, but writes
|
82
|
153 |
|
|
154 |
\begin{quote}
|
93
|
155 |
\it%
|
|
156 |
\begin{tabular}{@ {}l@ {}p{0.88\textwidth}@ {}}
|
101
|
157 |
`` & All lemmas appear obvious given a picture of the composition of automata\ldots
|
|
158 |
Yet their proofs require a painful amount of detail.''
|
|
159 |
\end{tabular}
|
|
160 |
\end{quote}
|
|
161 |
|
|
162 |
\noindent
|
|
163 |
and
|
|
164 |
|
|
165 |
\begin{quote}
|
|
166 |
\it%
|
|
167 |
\begin{tabular}{@ {}l@ {}p{0.88\textwidth}@ {}}
|
93
|
168 |
`` & If the reader finds the above treatment in terms of bit lists revoltingly
|
101
|
169 |
concrete, I cannot disagree. A more abstract approach is clearly desirable.''
|
93
|
170 |
\end{tabular}
|
82
|
171 |
\end{quote}
|
101
|
172 |
|
|
173 |
|
82
|
174 |
\noindent
|
|
175 |
Moreover, it is not so clear how to conveniently impose a finiteness condition
|
|
176 |
upon functions in order to represent \emph{finite} automata. The best is
|
92
|
177 |
probably to resort to more advanced reasoning frameworks, such as \emph{locales}
|
|
178 |
or \emph{type classes},
|
93
|
179 |
which are not avaiable in \emph{all} HOL-based theorem provers.
|
82
|
180 |
|
66
|
181 |
Because of these problems to do with representing automata, there seems
|
|
182 |
to be no substantial formalisation of automata theory and regular languages
|
93
|
183 |
carried out in HOL-based theorem provers. Nipkow establishes in
|
82
|
184 |
\cite{Nipkow98} the link between regular expressions and automata in
|
100
|
185 |
the context of lexing. The only larger formalisations of automata theory
|
82
|
186 |
are carried out in Nuprl \cite{Constable00} and in Coq (for example
|
|
187 |
\cite{Filliatre97}).
|
58
|
188 |
|
82
|
189 |
In this paper, we will not attempt to formalise automata theory in
|
|
190 |
Isabelle/HOL, but take a completely different approach to regular
|
|
191 |
languages. Instead of defining a regular language as one where there exists
|
|
192 |
an automaton that recognises all strings of the language, we define a
|
|
193 |
regular language as:
|
54
|
194 |
|
82
|
195 |
\begin{definition}
|
77
|
196 |
A language @{text A} is \emph{regular}, provided there is a regular expression that matches all
|
54
|
197 |
strings of @{text "A"}.
|
|
198 |
\end{definition}
|
|
199 |
|
|
200 |
\noindent
|
88
|
201 |
The reason is that regular expressions, unlike graphs, matrices and functons, can
|
71
|
202 |
be easily defined as inductive datatype. Consequently a corresponding reasoning
|
|
203 |
infrastructure comes for free. This has recently been exploited in HOL4 with a formalisation
|
101
|
204 |
of regular expression matching based on derivatives \cite{OwensSlind08} and
|
|
205 |
with an equivalence checker for regular expressions in Isabelle/HOL \cite{KraussNipkow11}.
|
|
206 |
The purpose of this paper is to
|
71
|
207 |
show that a central result about regular languages---the Myhill-Nerode theorem---can
|
|
208 |
be recreated by only using regular expressions. This theorem gives necessary
|
|
209 |
and sufficient conditions for when a language is regular. As a corollary of this
|
67
|
210 |
theorem we can easily establish the usual closure properties, including
|
|
211 |
complementation, for regular languages.\smallskip
|
61
|
212 |
|
|
213 |
\noindent
|
88
|
214 |
{\bf Contributions:}
|
|
215 |
There is an extensive literature on regular languages.
|
|
216 |
To our knowledge, our proof of the Myhill-Nerode theorem is the
|
67
|
217 |
first that is based on regular expressions, only. We prove the part of this theorem
|
|
218 |
stating that a regular expression has only finitely many partitions using certain
|
|
219 |
tagging-functions. Again to our best knowledge, these tagging functions have
|
|
220 |
not been used before to establish the Myhill-Nerode theorem.
|
24
|
221 |
*}
|
|
222 |
|
50
|
223 |
section {* Preliminaries *}
|
|
224 |
|
|
225 |
text {*
|
67
|
226 |
Strings in Isabelle/HOL are lists of characters with the \emph{empty string}
|
92
|
227 |
being represented by the empty list, written @{term "[]"}. \emph{Languages}
|
67
|
228 |
are sets of strings. The language containing all strings is written in
|
71
|
229 |
Isabelle/HOL as @{term "UNIV::string set"}. The concatenation of two languages
|
90
|
230 |
is written @{term "A ;; B"} and a language raised to the power @{text n} is written
|
93
|
231 |
@{term "A \<up> n"}. They are defined as usual
|
54
|
232 |
|
|
233 |
\begin{center}
|
58
|
234 |
@{thm Seq_def[THEN eq_reflection, where A1="A" and B1="B"]}
|
|
235 |
\hspace{7mm}
|
|
236 |
@{thm pow.simps(1)[THEN eq_reflection, where A1="A"]}
|
|
237 |
\hspace{7mm}
|
|
238 |
@{thm pow.simps(2)[THEN eq_reflection, where A1="A" and n1="n"]}
|
54
|
239 |
\end{center}
|
|
240 |
|
|
241 |
\noindent
|
58
|
242 |
where @{text "@"} is the usual list-append operation. The Kleene-star of a language @{text A}
|
71
|
243 |
is defined as the union over all powers, namely @{thm Star_def}. In the paper
|
88
|
244 |
we will make use of the following properties of these constructions.
|
58
|
245 |
|
71
|
246 |
\begin{proposition}\label{langprops}\mbox{}\\
|
92
|
247 |
\begin{tabular}{@ {}ll}
|
|
248 |
(i) & @{thm star_cases} \\
|
|
249 |
(ii) & @{thm[mode=IfThen] pow_length}\\
|
|
250 |
(iii) & @{thm seq_Union_left} \\
|
71
|
251 |
\end{tabular}
|
|
252 |
\end{proposition}
|
|
253 |
|
|
254 |
\noindent
|
100
|
255 |
In @{text "(ii)"} we use the notation @{term "length s"} for the length of a
|
|
256 |
string. This property states that if @{term "[] \<notin> A"} then the lengths of
|
|
257 |
the strings in @{term "A \<up> (Suc n)"} must be longer than @{text n}. We omit
|
|
258 |
the proofs for these properties, but invite the reader to consult our
|
|
259 |
formalisation.\footnote{Available at ???}
|
71
|
260 |
|
90
|
261 |
The notation in Isabelle/HOL for the quotient of a language @{text A} according to an
|
|
262 |
equivalence relation @{term REL} is @{term "A // REL"}. We will write
|
71
|
263 |
@{text "\<lbrakk>x\<rbrakk>\<^isub>\<approx>"} for the equivalence class defined
|
|
264 |
as @{text "{y | y \<approx> x}"}.
|
|
265 |
|
|
266 |
|
51
|
267 |
Central to our proof will be the solution of equational systems
|
101
|
268 |
involving equivalence classes of languages. For this we will use Arden's lemma \cite{Brzozowski64}
|
93
|
269 |
which solves equations of the form @{term "X = A ;; X \<union> B"} provided
|
71
|
270 |
@{term "[] \<notin> A"}. However we will need the following ``reverse''
|
50
|
271 |
version of Arden's lemma.
|
|
272 |
|
75
|
273 |
\begin{lemma}[Reverse Arden's Lemma]\label{arden}\mbox{}\\
|
86
|
274 |
If @{thm (prem 1) arden} then
|
|
275 |
@{thm (lhs) arden} has the unique solution
|
|
276 |
@{thm (rhs) arden}.
|
50
|
277 |
\end{lemma}
|
|
278 |
|
|
279 |
\begin{proof}
|
86
|
280 |
For the right-to-left direction we assume @{thm (rhs) arden} and show
|
|
281 |
that @{thm (lhs) arden} holds. From Prop.~\ref{langprops}@{text "(i)"}
|
71
|
282 |
we have @{term "A\<star> = {[]} \<union> A ;; A\<star>"},
|
50
|
283 |
which is equal to @{term "A\<star> = {[]} \<union> A\<star> ;; A"}. Adding @{text B} to both
|
|
284 |
sides gives @{term "B ;; A\<star> = B ;; ({[]} \<union> A\<star> ;; A)"}, whose right-hand side
|
51
|
285 |
is equal to @{term "(B ;; A\<star>) ;; A \<union> B"}. This completes this direction.
|
50
|
286 |
|
86
|
287 |
For the other direction we assume @{thm (lhs) arden}. By a simple induction
|
51
|
288 |
on @{text n}, we can establish the property
|
50
|
289 |
|
|
290 |
\begin{center}
|
86
|
291 |
@{text "(*)"}\hspace{5mm} @{thm (concl) arden_helper}
|
50
|
292 |
\end{center}
|
|
293 |
|
|
294 |
\noindent
|
|
295 |
Using this property we can show that @{term "B ;; (A \<up> n) \<subseteq> X"} holds for
|
71
|
296 |
all @{text n}. From this we can infer @{term "B ;; A\<star> \<subseteq> X"} using the definition
|
|
297 |
of @{text "\<star>"}.
|
51
|
298 |
For the inclusion in the other direction we assume a string @{text s}
|
86
|
299 |
with length @{text k} is element in @{text X}. Since @{thm (prem 1) arden}
|
75
|
300 |
we know by Prop.~\ref{langprops}@{text "(ii)"} that
|
71
|
301 |
@{term "s \<notin> X ;; (A \<up> Suc k)"} since its length is only @{text k}
|
51
|
302 |
(the strings in @{term "X ;; (A \<up> Suc k)"} are all longer).
|
53
|
303 |
From @{text "(*)"} it follows then that
|
50
|
304 |
@{term s} must be element in @{term "(\<Union>m\<in>{0..k}. B ;; (A \<up> m))"}. This in turn
|
75
|
305 |
implies that @{term s} is in @{term "(\<Union>n. B ;; (A \<up> n))"}. Using Prop.~\ref{langprops}@{text "(iii)"}
|
71
|
306 |
this is equal to @{term "B ;; A\<star>"}, as we needed to show.\qed
|
50
|
307 |
\end{proof}
|
67
|
308 |
|
|
309 |
\noindent
|
88
|
310 |
Regular expressions are defined as the inductive datatype
|
67
|
311 |
|
|
312 |
\begin{center}
|
|
313 |
@{text r} @{text "::="}
|
|
314 |
@{term NULL}\hspace{1.5mm}@{text"|"}\hspace{1.5mm}
|
|
315 |
@{term EMPTY}\hspace{1.5mm}@{text"|"}\hspace{1.5mm}
|
|
316 |
@{term "CHAR c"}\hspace{1.5mm}@{text"|"}\hspace{1.5mm}
|
|
317 |
@{term "SEQ r r"}\hspace{1.5mm}@{text"|"}\hspace{1.5mm}
|
|
318 |
@{term "ALT r r"}\hspace{1.5mm}@{text"|"}\hspace{1.5mm}
|
|
319 |
@{term "STAR r"}
|
|
320 |
\end{center}
|
|
321 |
|
|
322 |
\noindent
|
88
|
323 |
and the language matched by a regular expression is defined as
|
67
|
324 |
|
|
325 |
\begin{center}
|
|
326 |
\begin{tabular}{c@ {\hspace{10mm}}c}
|
|
327 |
\begin{tabular}{rcl}
|
|
328 |
@{thm (lhs) L_rexp.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) L_rexp.simps(1)}\\
|
|
329 |
@{thm (lhs) L_rexp.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) L_rexp.simps(2)}\\
|
|
330 |
@{thm (lhs) L_rexp.simps(3)[where c="c"]} & @{text "\<equiv>"} & @{thm (rhs) L_rexp.simps(3)[where c="c"]}\\
|
|
331 |
\end{tabular}
|
|
332 |
&
|
|
333 |
\begin{tabular}{rcl}
|
|
334 |
@{thm (lhs) L_rexp.simps(4)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]} & @{text "\<equiv>"} &
|
|
335 |
@{thm (rhs) L_rexp.simps(4)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]}\\
|
|
336 |
@{thm (lhs) L_rexp.simps(5)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]} & @{text "\<equiv>"} &
|
|
337 |
@{thm (rhs) L_rexp.simps(5)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]}\\
|
|
338 |
@{thm (lhs) L_rexp.simps(6)[where r="r"]} & @{text "\<equiv>"} &
|
|
339 |
@{thm (rhs) L_rexp.simps(6)[where r="r"]}\\
|
|
340 |
\end{tabular}
|
|
341 |
\end{tabular}
|
|
342 |
\end{center}
|
70
|
343 |
|
100
|
344 |
Given a finite set of regular expressions @{text rs}, we will make use of the operation of generating
|
92
|
345 |
a regular expression that matches all languages of @{text rs}. We only need to know the existence
|
|
346 |
of such a regular expression and therefore we use Isabelle/HOL's @{const "fold_graph"} and Hilbert's
|
93
|
347 |
@{text "\<epsilon>"} to define @{term "\<Uplus>rs"}. This operation, roughly speaking, folds @{const ALT} over the
|
100
|
348 |
set @{text rs} with @{const NULL} for the empty set. We can prove that for a finite set @{text rs}
|
82
|
349 |
|
88
|
350 |
\begin{center}
|
93
|
351 |
@{thm (lhs) folds_alt_simp} @{text "= \<Union> (\<calL> ` rs)"}
|
88
|
352 |
\end{center}
|
|
353 |
|
|
354 |
\noindent
|
90
|
355 |
holds, whereby @{text "\<calL> ` rs"} stands for the
|
|
356 |
image of the set @{text rs} under function @{text "\<calL>"}.
|
50
|
357 |
*}
|
39
|
358 |
|
100
|
359 |
section {* The Myhill-Nerode Theorem, First Part *}
|
54
|
360 |
|
|
361 |
text {*
|
77
|
362 |
The key definition in the Myhill-Nerode theorem is the
|
75
|
363 |
\emph{Myhill-Nerode relation}, which states that w.r.t.~a language two
|
|
364 |
strings are related, provided there is no distinguishing extension in this
|
|
365 |
language. This can be defined as:
|
|
366 |
|
70
|
367 |
\begin{definition}[Myhill-Nerode Relation]\mbox{}\\
|
75
|
368 |
@{thm str_eq_def[simplified str_eq_rel_def Pair_Collect]}
|
70
|
369 |
\end{definition}
|
|
370 |
|
71
|
371 |
\noindent
|
75
|
372 |
It is easy to see that @{term "\<approx>A"} is an equivalence relation, which
|
|
373 |
partitions the set of all strings, @{text "UNIV"}, into a set of disjoint
|
101
|
374 |
equivalence classes. To illustrate this quotient construction, let us give an
|
|
375 |
example: consider the regular language containing just
|
92
|
376 |
the string @{text "[c]"}. The relation @{term "\<approx>({[c]})"} partitions @{text UNIV}
|
101
|
377 |
into three equivalence classes @{text "X\<^isub>1"}, @{text "X\<^isub>2"} and @{text "X\<^isub>3"}
|
90
|
378 |
as follows
|
|
379 |
|
|
380 |
\begin{center}
|
|
381 |
@{text "X\<^isub>1 = {[]}"}\hspace{5mm}
|
|
382 |
@{text "X\<^isub>2 = {[c]}"}\hspace{5mm}
|
|
383 |
@{text "X\<^isub>3 = UNIV - {[], [c]}"}
|
|
384 |
\end{center}
|
|
385 |
|
|
386 |
One direction of the Myhill-Nerode theorem establishes
|
93
|
387 |
that if there are finitely many equivalence classes, like in the example above, then
|
|
388 |
the language is regular. In our setting we therefore have to show:
|
75
|
389 |
|
|
390 |
\begin{theorem}\label{myhillnerodeone}
|
96
|
391 |
@{thm[mode=IfThen] Myhill_Nerode1}
|
75
|
392 |
\end{theorem}
|
71
|
393 |
|
75
|
394 |
\noindent
|
90
|
395 |
To prove this theorem, we first define the set @{term "finals A"} as those equivalence
|
100
|
396 |
classes from @{term "UNIV // \<approx>A"} that contain strings of @{text A}, namely
|
75
|
397 |
%
|
71
|
398 |
\begin{equation}
|
70
|
399 |
@{thm finals_def}
|
71
|
400 |
\end{equation}
|
|
401 |
|
|
402 |
\noindent
|
92
|
403 |
In our running example, @{text "X\<^isub>2"} is the only equivalence class in @{term "finals {[c]}"}.
|
90
|
404 |
It is straightforward to show that in general @{thm lang_is_union_of_finals} and
|
79
|
405 |
@{thm finals_in_partitions} hold.
|
75
|
406 |
Therefore if we know that there exists a regular expression for every
|
100
|
407 |
equivalence class in \mbox{@{term "finals A"}} (which by assumption must be
|
93
|
408 |
a finite set), then we can use @{text "\<bigplus>"} to obtain a regular expression
|
98
|
409 |
that matches every string in @{text A}.
|
70
|
410 |
|
75
|
411 |
|
90
|
412 |
Our proof of Thm.~\ref{myhillnerodeone} relies on a method that can calculate a
|
79
|
413 |
regular expression for \emph{every} equivalence class, not just the ones
|
77
|
414 |
in @{term "finals A"}. We
|
93
|
415 |
first define the notion of \emph{one-character-transition} between
|
|
416 |
two equivalence classes
|
75
|
417 |
%
|
71
|
418 |
\begin{equation}
|
|
419 |
@{thm transition_def}
|
|
420 |
\end{equation}
|
70
|
421 |
|
71
|
422 |
\noindent
|
92
|
423 |
which means that if we concatenate the character @{text c} to the end of all
|
|
424 |
strings in the equivalence class @{text Y}, we obtain a subset of
|
77
|
425 |
@{text X}. Note that we do not define an automaton here, we merely relate two sets
|
98
|
426 |
(with respect to a character). In our concrete example we have
|
92
|
427 |
@{term "X\<^isub>1 \<Turnstile>c\<Rightarrow> X\<^isub>2"}, @{term "X\<^isub>1 \<Turnstile>d\<Rightarrow> X\<^isub>3"} with @{text d} being any
|
93
|
428 |
other character than @{text c}, and @{term "X\<^isub>3 \<Turnstile>d\<Rightarrow> X\<^isub>3"} for any @{text d}.
|
75
|
429 |
|
100
|
430 |
Next we build an \emph{initial} equational system that
|
75
|
431 |
contains an equation for each equivalence class. Suppose we have
|
|
432 |
the equivalence classes @{text "X\<^isub>1,\<dots>,X\<^isub>n"}, there must be one and only one that
|
|
433 |
contains the empty string @{text "[]"} (since equivalence classes are disjoint).
|
77
|
434 |
Let us assume @{text "[] \<in> X\<^isub>1"}. We build the following equational system
|
75
|
435 |
|
|
436 |
\begin{center}
|
|
437 |
\begin{tabular}{rcl}
|
|
438 |
@{text "X\<^isub>1"} & @{text "="} & @{text "(Y\<^isub>1\<^isub>1, CHAR c\<^isub>1\<^isub>1) + \<dots> + (Y\<^isub>1\<^isub>p, CHAR c\<^isub>1\<^isub>p) + \<lambda>(EMPTY)"} \\
|
|
439 |
@{text "X\<^isub>2"} & @{text "="} & @{text "(Y\<^isub>2\<^isub>1, CHAR c\<^isub>2\<^isub>1) + \<dots> + (Y\<^isub>2\<^isub>o, CHAR c\<^isub>2\<^isub>o)"} \\
|
|
440 |
& $\vdots$ \\
|
|
441 |
@{text "X\<^isub>n"} & @{text "="} & @{text "(Y\<^isub>n\<^isub>1, CHAR c\<^isub>n\<^isub>1) + \<dots> + (Y\<^isub>n\<^isub>q, CHAR c\<^isub>n\<^isub>q)"}\\
|
|
442 |
\end{tabular}
|
|
443 |
\end{center}
|
70
|
444 |
|
75
|
445 |
\noindent
|
100
|
446 |
where the terms @{text "(Y\<^isub>i\<^isub>j, CHAR c\<^isub>i\<^isub>j)"}
|
|
447 |
stand for all transitions @{term "Y\<^isub>i\<^isub>j \<Turnstile>c\<^isub>i\<^isub>j\<Rightarrow>
|
|
448 |
X\<^isub>i"}. There can only be
|
|
449 |
finitely many such terms in a right-hand side since there are only finitely many
|
|
450 |
equivalence classes and only finitely many characters. The term @{text
|
|
451 |
"\<lambda>(EMPTY)"} in the first equation acts as a marker for the equivalence class
|
|
452 |
containing @{text "[]"}.\footnote{Note that we mark, roughly speaking, the
|
75
|
453 |
single ``initial'' state in the equational system, which is different from
|
100
|
454 |
the method by Brzozowski \cite{Brzozowski64}, where he marks the
|
|
455 |
``terminal'' states. We are forced to set up the equational system in our
|
|
456 |
way, because the Myhill-Nerode relation determines the ``direction'' of the
|
|
457 |
transitions. The successor ``state'' of an equivalence class @{text Y} can
|
|
458 |
be reached by adding characters to the end of @{text Y}. This is also the
|
|
459 |
reason why we have to use our reverse version of Arden's lemma.}
|
|
460 |
Overloading the function @{text \<calL>} for the two kinds of terms in the
|
92
|
461 |
equational system, we have
|
75
|
462 |
|
|
463 |
\begin{center}
|
92
|
464 |
@{text "\<calL>(Y, r) \<equiv>"} %
|
|
465 |
@{thm (rhs) L_rhs_item.simps(2)[where X="Y" and r="r", THEN eq_reflection]}\hspace{10mm}
|
86
|
466 |
@{thm L_rhs_item.simps(1)[where r="r", THEN eq_reflection]}
|
75
|
467 |
\end{center}
|
|
468 |
|
|
469 |
\noindent
|
100
|
470 |
and we can prove for @{text "X\<^isub>2\<^isub>.\<^isub>.\<^isub>n"} that the following equations
|
75
|
471 |
%
|
|
472 |
\begin{equation}\label{inv1}
|
83
|
473 |
@{text "X\<^isub>i = \<calL>(Y\<^isub>i\<^isub>1, CHAR c\<^isub>i\<^isub>1) \<union> \<dots> \<union> \<calL>(Y\<^isub>i\<^isub>q, CHAR c\<^isub>i\<^isub>q)"}.
|
75
|
474 |
\end{equation}
|
|
475 |
|
|
476 |
\noindent
|
|
477 |
hold. Similarly for @{text "X\<^isub>1"} we can show the following equation
|
|
478 |
%
|
|
479 |
\begin{equation}\label{inv2}
|
83
|
480 |
@{text "X\<^isub>1 = \<calL>(Y\<^isub>i\<^isub>1, CHAR c\<^isub>i\<^isub>1) \<union> \<dots> \<union> \<calL>(Y\<^isub>i\<^isub>p, CHAR c\<^isub>i\<^isub>p) \<union> \<calL>(\<lambda>(EMPTY))"}.
|
75
|
481 |
\end{equation}
|
|
482 |
|
|
483 |
\noindent
|
101
|
484 |
The reason for adding the @{text \<lambda>}-marker to our initial equational system is
|
103
|
485 |
to obtain this equation: it only holds with the marker, since none of
|
100
|
486 |
the other terms contain the empty string.
|
|
487 |
|
101
|
488 |
Our representation for the equations in Isabelle/HOL are pairs,
|
100
|
489 |
where the first component is an equivalence class and the second component
|
101
|
490 |
is a set of terms. Given a set of equivalence
|
100
|
491 |
classes @{text CS}, our initial equational system @{term "Init CS"} is thus
|
101
|
492 |
formally defined as
|
104
|
493 |
%
|
|
494 |
\begin{equation}\label{initcs}
|
|
495 |
\mbox{\begin{tabular}{rcl}
|
100
|
496 |
@{thm (lhs) Init_rhs_def} & @{text "\<equiv>"} &
|
|
497 |
@{text "if"}~@{term "[] \<in> X"}\\
|
|
498 |
& & @{text "then"}~@{term "{Trn Y (CHAR c) | Y c. Y \<in> CS \<and> Y \<Turnstile>c\<Rightarrow> X} \<union> {Lam EMPTY}"}\\
|
|
499 |
& & @{text "else"}~@{term "{Trn Y (CHAR c)| Y c. Y \<in> CS \<and> Y \<Turnstile>c\<Rightarrow> X}"}\\
|
|
500 |
@{thm (lhs) Init_def} & @{text "\<equiv>"} & @{thm (rhs) Init_def}
|
104
|
501 |
\end{tabular}}
|
|
502 |
\end{equation}
|
100
|
503 |
|
|
504 |
|
|
505 |
|
|
506 |
\noindent
|
|
507 |
Because we use sets of terms
|
101
|
508 |
for representing the right-hand sides of equations, we can
|
100
|
509 |
prove \eqref{inv1} and \eqref{inv2} more concisely as
|
93
|
510 |
%
|
100
|
511 |
\begin{lemma}\label{inv}
|
|
512 |
If @{thm (prem 1) test} then @{text "X = \<Union> \<calL> ` rhs"}.
|
|
513 |
\end{lemma}
|
77
|
514 |
|
93
|
515 |
\noindent
|
92
|
516 |
Our proof of Thm.~\ref{myhillnerodeone} will proceed by transforming the
|
100
|
517 |
initial equational system into one in \emph{solved form} maintaining the invariant
|
|
518 |
in Lemma \ref{inv}. From the solved form we will be able to read
|
89
|
519 |
off the regular expressions.
|
|
520 |
|
100
|
521 |
In order to transform an equational system into solved form, we have two
|
89
|
522 |
operations: one that takes an equation of the form @{text "X = rhs"} and removes
|
93
|
523 |
the recursive occurences of @{text X} in the @{text rhs} using our variant of Arden's
|
92
|
524 |
Lemma. The other operation takes an equation @{text "X = rhs"}
|
89
|
525 |
and substitutes @{text X} throughout the rest of the equational system
|
92
|
526 |
adjusting the remaining regular expressions approriately. To define this adjustment
|
|
527 |
we define the \emph{append-operation}
|
89
|
528 |
|
|
529 |
\begin{center}
|
92
|
530 |
@{thm append_rexp.simps(2)[where X="Y" and r="r\<^isub>1" and rexp="r\<^isub>2", THEN eq_reflection]}\hspace{10mm}
|
|
531 |
@{thm append_rexp.simps(1)[where r="r\<^isub>1" and rexp="r\<^isub>2", THEN eq_reflection]}
|
89
|
532 |
\end{center}
|
|
533 |
|
92
|
534 |
\noindent
|
|
535 |
which we also lift to entire right-hand sides of equations, written as
|
93
|
536 |
@{thm (lhs) append_rhs_rexp_def[where rexp="r"]}. With this we can define
|
101
|
537 |
the \emph{arden-operation} for an equation of the form @{text "X = rhs"} as:
|
94
|
538 |
|
92
|
539 |
\begin{center}
|
94
|
540 |
\begin{tabular}{rc@ {\hspace{2mm}}r@ {\hspace{1mm}}l}
|
|
541 |
@{thm (lhs) Arden_def} & @{text "\<equiv>"}~~\mbox{} & \multicolumn{2}{@ {\hspace{-2mm}}l}{@{text "let"}}\\
|
|
542 |
& & @{text "rhs' ="} & @{term "rhs - {Trn X r | r. Trn X r \<in> rhs}"} \\
|
|
543 |
& & @{text "r' ="} & @{term "STAR (\<Uplus> {r. Trn X r \<in> rhs})"}\\
|
|
544 |
& & \multicolumn{2}{@ {\hspace{-2mm}}l}{@{text "in"}~~@{term "append_rhs_rexp rhs' r'"}}\\
|
|
545 |
\end{tabular}
|
92
|
546 |
\end{center}
|
93
|
547 |
|
|
548 |
\noindent
|
101
|
549 |
In this definition, we first delete all terms of the form @{text "(X, r)"} from @{text rhs};
|
94
|
550 |
then we calculate the combinded regular expressions for all @{text r} coming
|
|
551 |
from the deleted @{text "(X, r)"}, and take the @{const STAR} of it;
|
|
552 |
finally we append this regular expression to @{text rhs'}. It can be easily seen
|
95
|
553 |
that this operation mimics Arden's lemma on the level of equations.
|
|
554 |
The \emph{substituion-operation} takes an equation
|
|
555 |
of the form @{text "X = xrhs"} and substitutes it into the right-hand side @{text rhs}.
|
94
|
556 |
|
|
557 |
\begin{center}
|
95
|
558 |
\begin{tabular}{rc@ {\hspace{2mm}}r@ {\hspace{1mm}}l}
|
|
559 |
@{thm (lhs) Subst_def} & @{text "\<equiv>"}~~\mbox{} & \multicolumn{2}{@ {\hspace{-2mm}}l}{@{text "let"}}\\
|
|
560 |
& & @{text "rhs' ="} & @{term "rhs - {Trn X r | r. Trn X r \<in> rhs}"} \\
|
|
561 |
& & @{text "r' ="} & @{term "\<Uplus> {r. Trn X r \<in> rhs}"}\\
|
|
562 |
& & \multicolumn{2}{@ {\hspace{-2mm}}l}{@{text "in"}~~@{term "rhs' \<union> append_rhs_rexp xrhs r'"}}\\
|
|
563 |
\end{tabular}
|
94
|
564 |
\end{center}
|
95
|
565 |
|
|
566 |
\noindent
|
|
567 |
We again delete first all occurence of @{text "(X, r)"} in @{text rhs}; we then calculate
|
|
568 |
the regular expression corresponding to the deleted terms; finally we append this
|
|
569 |
regular expression to @{text "xrhs"} and union it up with @{text rhs'}. When we use
|
|
570 |
the substitution operation we will arrange it so that @{text "xrhs"} does not contain
|
|
571 |
any occurence of @{text X}.
|
96
|
572 |
|
100
|
573 |
With these two operation in place, we can define the operation that removes one equation
|
|
574 |
from an equational systems @{text ES}. The operation @{const Subst_all}
|
96
|
575 |
substitutes an equation @{text "X = xrhs"} throughout an equational system @{text ES};
|
100
|
576 |
@{const Remove} then completely removes such an equation from @{text ES} by substituting
|
|
577 |
it to the rest of the equational system, but first eliminating all recursive occurences
|
96
|
578 |
of @{text X} by applying @{const Arden} to @{text "xrhs"}.
|
|
579 |
|
|
580 |
\begin{center}
|
|
581 |
\begin{tabular}{rcl}
|
|
582 |
@{thm (lhs) Subst_all_def} & @{text "\<equiv>"} & @{thm (rhs) Subst_all_def}\\
|
|
583 |
@{thm (lhs) Remove_def} & @{text "\<equiv>"} & @{thm (rhs) Remove_def}
|
|
584 |
\end{tabular}
|
|
585 |
\end{center}
|
100
|
586 |
|
|
587 |
\noindent
|
|
588 |
Finially, we can define how an equational system should be solved. For this
|
101
|
589 |
we will need to iterate the elimination of an equation until only one equation
|
100
|
590 |
will be left in the system. However, we not just want to have any equation
|
|
591 |
as being the last one, but the one for which we want to calculate the regular
|
|
592 |
expression. Therefore we define the iteration step so that it chooses an
|
|
593 |
equation with an equivalence class that is not @{text X}. This allows us to
|
101
|
594 |
control, which equation will be the last. We use Hilbert's choice operator,
|
103
|
595 |
written @{text SOME}, to choose an equation to be eliminated in @{text ES}.
|
100
|
596 |
|
|
597 |
\begin{center}
|
|
598 |
\begin{tabular}{rc@ {\hspace{4mm}}r@ {\hspace{1mm}}l}
|
|
599 |
@{thm (lhs) Iter_def} & @{text "\<equiv>"}~~\mbox{} & \multicolumn{2}{@ {\hspace{-4mm}}l}{@{text "let"}}\\
|
|
600 |
& & @{text "(Y, yrhs) ="} & @{term "SOME (Y, yrhs). (Y, yrhs) \<in> ES \<and> X \<noteq> Y"} \\
|
|
601 |
& & \multicolumn{2}{@ {\hspace{-4mm}}l}{@{text "in"}~~@{term "Remove ES Y yrhs"}}\\
|
|
602 |
\end{tabular}
|
|
603 |
\end{center}
|
|
604 |
|
|
605 |
\noindent
|
103
|
606 |
The last definition we need applies @{term Iter} over and over again until a condition
|
101
|
607 |
@{text COND} is \emph{not} satisfied anymore. The condition states that there
|
103
|
608 |
are more than one equation left in the equational system @{text ES}. For this
|
|
609 |
we use Isabelle/HOL's @{text while}-operator as follows:
|
101
|
610 |
|
100
|
611 |
\begin{center}
|
|
612 |
@{thm Solve_def}
|
|
613 |
\end{center}
|
|
614 |
|
101
|
615 |
\noindent
|
103
|
616 |
We are not concerned here with the definition of this operator
|
|
617 |
(see \cite{BerghoferNipkow00}), but note that we eliminate
|
|
618 |
in each @{const Iter}-step a single equation, and therefore
|
|
619 |
have a well-founded termination order by taking the cardinality
|
|
620 |
of the equational system @{text ES}. This enables us to prove
|
104
|
621 |
properties about our definition of @{const Solve} when we ``call'' it with
|
|
622 |
the equivalence class @{text X} and the initial equational system
|
|
623 |
@{term "Init (UNIV // \<approx>A)"} from
|
|
624 |
\eqref{initcs}:
|
103
|
625 |
|
100
|
626 |
|
|
627 |
\begin{center}
|
103
|
628 |
\begin{tabular}{l}
|
|
629 |
@{term "invariant (Init (UNIV // \<approx>A))"} \\
|
|
630 |
@{term "\<forall>ES. invariant ES \<and> Cond ES \<longrightarrow> invariant (Iter X ES)"}\\
|
|
631 |
@{term "\<forall>ES. invariant ES \<and> Cond ES \<longrightarrow> card (Iter X ES) < card ES"}\\
|
|
632 |
@{term "\<forall>ES. invariant ES \<and> \<not> Cond ES \<longrightarrow> P ES"}\\
|
|
633 |
\hline
|
|
634 |
\multicolumn{1}{c}{@{term "P (Solve X (Init (UNIV // \<approx>A)))"}}
|
|
635 |
\end{tabular}
|
100
|
636 |
\end{center}
|
103
|
637 |
|
|
638 |
\noindent
|
104
|
639 |
This principle states that given an invariant (which we will specify below)
|
|
640 |
we can prove a property
|
|
641 |
@{text "P"} involving @{const Solve}. For this we have to discharge the following
|
|
642 |
proof obligations: first the
|
103
|
643 |
initial equational system satisfies the invariant; second that the iteration
|
104
|
644 |
step @{text "Iter"} preserves the the invariant as long as the condition @{term Cond} holds;
|
103
|
645 |
third that @{text "Iter"} decreases the termination order, and fourth that
|
104
|
646 |
once the condition does not hold anymore then the property @{text P} must hold.
|
103
|
647 |
|
104
|
648 |
The property @{term P} in our proof will state that @{term "Solve X (Init (UNIV // \<approx>A))"}
|
103
|
649 |
returns with a single equation @{text "X = xrhs"}, for some @{text "xrhs"} and
|
104
|
650 |
that this equational system still satisfies the invariant. In order to get
|
|
651 |
the proof through, the invariant is composed of the following six properties:
|
103
|
652 |
|
|
653 |
\begin{center}
|
104
|
654 |
\begin{tabular}{@ {}rcl@ {\hspace{-13mm}}l @ {}}
|
|
655 |
@{text "invariant ES"} & @{text "\<equiv>"} &
|
103
|
656 |
@{term "finite ES"} & @{text "(finiteness)"}\\
|
|
657 |
& @{text "\<and>"} & @{thm (rhs) finite_rhs_def} & @{text "(finiteness rhs)"}\\
|
104
|
658 |
& @{text "\<and>"} & @{text "\<forall>(X, rhs)\<in>ES. X = \<Union>\<calL> ` rhs"} & @{text "(soundness)"}\\
|
|
659 |
& @{text "\<and>"} & @{thm (rhs) distinct_equas_def}\\
|
|
660 |
& & & @{text "(distinctness)"}\\
|
103
|
661 |
& @{text "\<and>"} & @{thm (rhs) ardenable_def} & @{text "(ardenable)"}\\
|
104
|
662 |
& @{text "\<and>"} & @{thm (rhs) valid_eqs_def} & @{text "(validity)"}\\
|
103
|
663 |
\end{tabular}
|
|
664 |
\end{center}
|
|
665 |
|
104
|
666 |
\noindent
|
|
667 |
The first two ensure that the equational system is always finite (number of equations
|
|
668 |
and number of terms in each equation); \ldots
|
|
669 |
|
|
670 |
\begin{lemma}
|
|
671 |
@{thm[mode=IfThen] Init_ES_satisfies_invariant}
|
|
672 |
\end{lemma}
|
|
673 |
|
|
674 |
\begin{lemma}
|
|
675 |
@{thm[mode=IfThen] iteration_step_invariant[where xrhs="rhs"]}
|
|
676 |
\end{lemma}
|
|
677 |
|
|
678 |
\begin{lemma}
|
|
679 |
@{thm[mode=IfThen] iteration_step_measure[simplified (no_asm), where xrhs="rhs"]}
|
|
680 |
\end{lemma}
|
|
681 |
|
|
682 |
\begin{lemma}
|
|
683 |
If @{thm (prem 1) Solve} and @{thm (prem 2) Solve} then there exists
|
|
684 |
a @{text rhs} such that @{term "Solve X (Init (UNIV // \<approx>A)) = {(X, rhs)}"}
|
|
685 |
and @{term "invariant {(X, rhs)}"}.
|
|
686 |
\end{lemma}
|
|
687 |
|
|
688 |
\begin{theorem}
|
|
689 |
@{thm[mode=IfThen] Myhill_Nerode1}
|
|
690 |
\end{theorem}
|
54
|
691 |
*}
|
|
692 |
|
100
|
693 |
|
|
694 |
|
|
695 |
|
|
696 |
section {* Myhill-Nerode, Second Part *}
|
39
|
697 |
|
|
698 |
text {*
|
|
699 |
|
54
|
700 |
\begin{theorem}
|
39
|
701 |
Given @{text "r"} is a regular expressions, then @{thm rexp_imp_finite}.
|
54
|
702 |
\end{theorem}
|
39
|
703 |
|
|
704 |
\begin{proof}
|
|
705 |
By induction on the structure of @{text r}. The cases for @{const NULL}, @{const EMPTY}
|
50
|
706 |
and @{const CHAR} are straightforward, because we can easily establish
|
39
|
707 |
|
|
708 |
\begin{center}
|
|
709 |
\begin{tabular}{l}
|
|
710 |
@{thm quot_null_eq}\\
|
|
711 |
@{thm quot_empty_subset}\\
|
|
712 |
@{thm quot_char_subset}
|
|
713 |
\end{tabular}
|
|
714 |
\end{center}
|
|
715 |
|
|
716 |
\end{proof}
|
|
717 |
*}
|
|
718 |
|
|
719 |
|
54
|
720 |
section {* Conclusion and Related Work *}
|
|
721 |
|
92
|
722 |
text {*
|
|
723 |
In this paper we took the view that a regular language as one where there exists
|
|
724 |
a regular expression that matches all its strings. For us it was important to find
|
|
725 |
out how far we can push this point of view. Having formalised the Myhill-Nerode
|
|
726 |
theorem means pushed very far. Having the Myhill-Nerode theorem means we can
|
|
727 |
formalise much of the textbook results in this subject.
|
|
728 |
|
|
729 |
|
|
730 |
*}
|
|
731 |
|
|
732 |
|
24
|
733 |
(*<*)
|
|
734 |
end
|
|
735 |
(*>*) |