Paper/Paper.thy
author urbanc
Thu, 03 Feb 2011 09:54:19 +0000
changeset 61 070f543e2560
parent 60 fb08f41ca33d
child 66 828ea293b61f
permissions -rw-r--r--
more to the intro
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
     1
(*<*)
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
     2
theory Paper
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
     3
imports "../Myhill" "LaTeXsugar"
24
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
     4
begin
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
     5
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
     6
declare [[show_question_marks = false]]
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
     7
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
     8
consts
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
     9
 REL :: "(string \<times> string) \<Rightarrow> bool"
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
    10
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
    11
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
    12
notation (latex output)
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
    13
  str_eq_rel ("\<approx>\<^bsub>_\<^esub>") and
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
    14
  Seq (infixr "\<cdot>" 100) and
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
    15
  Star ("_\<^bsup>\<star>\<^esup>") and
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
    16
  pow ("_\<^bsup>_\<^esup>" [100, 100] 100) and
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
    17
  Suc ("_+1" [100] 100) and
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
    18
  quotient ("_ \<^raw:\ensuremath{\!\sslash\!}> _" [90, 90] 90) and
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
    19
  REL ("\<approx>")
52
4a517c6ac07d tuning of the syntax; needs the stmaryrd latex package
urbanc
parents: 51
diff changeset
    20
24
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
    21
(*>*)
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
    22
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
    23
section {* Introduction *}
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
    24
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
    25
text {*
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
    26
  Regular languages are an important and well-understood subject in Computer
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    27
  Science, with many beautiful theorems and many useful algorithms. There is a
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    28
  wide range of textbooks on this subject, many of which are aimed at
61
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    29
  students and contain very detailed ``pencil-and-paper'' proofs
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    30
  (e.g.~\cite{Kozen97}). It seems natural to exercise theorem provers by
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    31
  formalising these theorems and by verifying formally the algorithms.
59
fc35eb54fdc9 more on the intro
urbanc
parents: 58
diff changeset
    32
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    33
  There is however a problem with this: the typical approach to regular
61
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    34
  languages is to introduce finite automata and then define everything in terms of
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    35
  them.  For example, a regular language is normally defined as one where
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    36
  there is a finite deterministic automaton that recognises all the strings of
61
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    37
  the language. This approach has many benefits. One is that it is easy to convince
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    38
  oneself from the fact that regular languages are closed under
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    39
  complementation: one just has to exchange the accepting and non-accepting
61
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    40
  states in the corresponding automaton to obtain an automaton for the complement language.
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    41
  The problem lies with formalising such reasoning in a theorem
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    42
  prover, in our case Isabelle/HOL. Automata need to be represented as graphs 
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    43
  or matrices, neither of which can be defined as inductive datatype.\footnote{In 
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    44
  some works functions are used to represent transitions, but they are also not 
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    45
  inductive datatypes.} This means we have to build our own reasoning infrastructure
61
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    46
  for them, as neither Isabelle nor HOL4 nor HOLlight support them with libraries.
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    47
  
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    48
  Even worse, reasoning about graphs in typed languages can be a real hassle. 
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    49
  Consider for example the operation of combining 
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    50
  two automata into a new automaton by connecting their
61
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    51
  initial states to a new initial state (similarly with the accepting states):
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    52
  
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    53
  \begin{center}
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    54
  picture
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    55
  \end{center}
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    56
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    57
  \noindent
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    58
  How should we implement this operation? On paper we can just
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    59
  form the disjoint union of the state nodes and add two more nodes---one for the
61
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    60
  new initial state, the other for the new accepting state. In a theorem
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    61
  prover based on set-theory, this operaton can be more or less
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    62
  straightforwardly implemented. But in a HOL-based theorem prover the
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    63
  standard definition of disjoint unions as pairs
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    64
61
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    65
  \begin{center}
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    66
  definition
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    67
  \end{center}
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    68
61
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    69
  \noindent
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    70
  changes the type (from sets of nodes to sets of pairs). This means we
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    71
  cannot formulate in this represeantation any property about \emph{all} 
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    72
  automata---since there is no type quantification available in HOL-based
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    73
  theorem provers. A working alternative is to give every state node an 
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    74
  identity, for example a natural number, and then be careful to rename
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    75
  these indentities appropriately when connecting two automata together. 
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    76
  This results in very clunky side-proofs establishing that properties
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    77
  are invariant under renaming. We are only aware of the formalisation 
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    78
  of automata theory in Nuprl that carries this approach trough and is
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    79
  quite substantial. 
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
    80
  
61
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    81
  We will take a completely different approach to formalising theorems
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    82
  about regular languages. Instead of defining a regular language as one 
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    83
  where there exists an automaton that recognises all of its strings, we 
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    84
  define a regular language as
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
    85
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
    86
  \begin{definition}[A Regular Language]
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
    87
  A language @{text A} is regular, if there is a regular expression that matches all
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
    88
  strings of @{text "A"}.
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
    89
  \end{definition}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
    90
  
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
    91
  \noindent
61
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    92
  The reason is that regular expressinons, unlike graphs and metrices, can
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    93
  be eaily defined as inductive datatype and this means a reasoning infrastructre
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    94
  comes for them in Isabelle for free. The purpose of this paper is to
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    95
  show that a central and highly non-trivisl result about regular languages, 
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    96
  namely the Myhill-Nerode theorem,  can be recreated only using regular 
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    97
  expressions. In our approach we do not need to formalise graps or
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    98
  metrices.
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    99
  
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
   100
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
   101
  \noindent
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
   102
  {\bf Contributions:} A proof of the Myhill-Nerode Theorem based on regular expressions. The 
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   103
  finiteness part of this theorem is proved using tagging-functions (which to our knowledge
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   104
  are novel in this context).
24
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   105
  
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   106
*}
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   107
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   108
section {* Preliminaries *}
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   109
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   110
text {*
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   111
  Strings in Isabelle/HOL are lists of characters and the
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   112
  \emph{empty string} is the empty list, written @{term "[]"}. \emph{Languages} are sets of 
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   113
  strings. The language containing all strings is written in Isabelle/HOL as @{term "UNIV::string set"}.
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   114
  The notation for the quotient of a language @{text A} according to a relation @{term REL} is
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   115
  @{term "A // REL"}. The concatenation of two languages is written @{term "A ;; B"}; a language 
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   116
  raised  tow the power $n$ is written @{term "A \<up> n"}. Both concepts are defined as
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   117
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   118
  \begin{center}
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   119
  @{thm Seq_def[THEN eq_reflection, where A1="A" and B1="B"]}
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   120
  \hspace{7mm}
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   121
  @{thm pow.simps(1)[THEN eq_reflection, where A1="A"]}
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   122
  \hspace{7mm}
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   123
  @{thm pow.simps(2)[THEN eq_reflection, where A1="A" and n1="n"]}
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   124
  \end{center}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   125
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   126
  \noindent
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   127
  where @{text "@"} is the usual list-append operation. The Kleene-star of a language @{text A}
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   128
  is defined as the union over all powers, namely @{thm Star_def}.
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   129
  
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   130
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   131
  Regular expressions are defined as the following datatype
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   132
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   133
  \begin{center}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   134
  @{text r} @{text "::="}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   135
  @{term NULL}\hspace{1.5mm}@{text"|"}\hspace{1.5mm} 
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   136
  @{term EMPTY}\hspace{1.5mm}@{text"|"}\hspace{1.5mm} 
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   137
  @{term "CHAR c"}\hspace{1.5mm}@{text"|"}\hspace{1.5mm} 
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   138
  @{term "SEQ r r"}\hspace{1.5mm}@{text"|"}\hspace{1.5mm} 
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   139
  @{term "ALT r r"}\hspace{1.5mm}@{text"|"}\hspace{1.5mm} 
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   140
  @{term "STAR r"}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   141
  \end{center}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   142
51
6cfb92de4654 some tuning of the paper
urbanc
parents: 50
diff changeset
   143
  Central to our proof will be the solution of equational systems
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   144
  involving regular expressions. For this we will use the following ``reverse'' 
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   145
  version of Arden's lemma.
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   146
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   147
  \begin{lemma}[Reverse Arden's Lemma]\mbox{}\\
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   148
  If @{thm (prem 1) ardens_revised} then
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   149
  @{thm (lhs) ardens_revised} has the unique solution
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   150
  @{thm (rhs) ardens_revised}.
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   151
  \end{lemma}
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   152
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   153
  \begin{proof}
51
6cfb92de4654 some tuning of the paper
urbanc
parents: 50
diff changeset
   154
  For the right-to-left direction we assume @{thm (rhs) ardens_revised} and show
6cfb92de4654 some tuning of the paper
urbanc
parents: 50
diff changeset
   155
  that @{thm (lhs) ardens_revised} holds. From Lemma ??? we have @{term "A\<star> = {[]} \<union> A ;; A\<star>"},
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   156
  which is equal to @{term "A\<star> = {[]} \<union> A\<star> ;; A"}. Adding @{text B} to both 
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   157
  sides gives @{term "B ;; A\<star> = B ;; ({[]} \<union> A\<star> ;; A)"}, whose right-hand side
51
6cfb92de4654 some tuning of the paper
urbanc
parents: 50
diff changeset
   158
  is equal to @{term "(B ;; A\<star>) ;; A \<union> B"}. This completes this direction. 
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   159
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   160
  For the other direction we assume @{thm (lhs) ardens_revised}. By a simple induction
51
6cfb92de4654 some tuning of the paper
urbanc
parents: 50
diff changeset
   161
  on @{text n}, we can establish the property
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   162
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   163
  \begin{center}
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   164
  @{text "(*)"}\hspace{5mm} @{thm (concl) ardens_helper}
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   165
  \end{center}
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   166
  
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   167
  \noindent
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   168
  Using this property we can show that @{term "B ;; (A \<up> n) \<subseteq> X"} holds for
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   169
  all @{text n}. From this we can infer @{term "B ;; A\<star> \<subseteq> X"} using Lemma ???.
51
6cfb92de4654 some tuning of the paper
urbanc
parents: 50
diff changeset
   170
  For the inclusion in the other direction we assume a string @{text s}
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   171
  with length @{text k} is element in @{text X}. Since @{thm (prem 1) ardens_revised}
51
6cfb92de4654 some tuning of the paper
urbanc
parents: 50
diff changeset
   172
  we know that @{term "s \<notin> X ;; (A \<up> Suc k)"} since its length is only @{text k}
6cfb92de4654 some tuning of the paper
urbanc
parents: 50
diff changeset
   173
  (the strings in @{term "X ;; (A \<up> Suc k)"} are all longer). 
53
da85feadb8e3 small typo
urbanc
parents: 52
diff changeset
   174
  From @{text "(*)"} it follows then that
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   175
  @{term s} must be element in @{term "(\<Union>m\<in>{0..k}. B ;; (A \<up> m))"}. This in turn
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   176
  implies that @{term s} is in @{term "(\<Union>n. B ;; (A \<up> n))"}. Using Lemma ??? this
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   177
  is equal to @{term "B ;; A\<star>"}, as we needed to show.\qed
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   178
  \end{proof}
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   179
*}
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   180
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   181
section {* Finite Partitions Imply Regularity of a Language *}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   182
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   183
text {*
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   184
  \begin{theorem}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   185
  Given a language @{text A}.
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   186
  @{thm[mode=IfThen] hard_direction[where Lang="A"]}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   187
  \end{theorem}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   188
*}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   189
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   190
section {* Regular Expressions Generate Finitely Many Partitions *}
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   191
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   192
text {*
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   193
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   194
  \begin{theorem}
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   195
  Given @{text "r"} is a regular expressions, then @{thm rexp_imp_finite}.
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   196
  \end{theorem}  
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   197
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   198
  \begin{proof}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   199
  By induction on the structure of @{text r}. The cases for @{const NULL}, @{const EMPTY}
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   200
  and @{const CHAR} are straightforward, because we can easily establish
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   201
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   202
  \begin{center}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   203
  \begin{tabular}{l}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   204
  @{thm quot_null_eq}\\
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   205
  @{thm quot_empty_subset}\\
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   206
  @{thm quot_char_subset}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   207
  \end{tabular}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   208
  \end{center}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   209
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   210
  \end{proof}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   211
*}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   212
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   213
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   214
section {* Conclusion and Related Work *}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   215
24
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   216
(*<*)
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   217
end
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   218
(*>*)