Paper/Paper.thy
author urbanc
Sun, 06 Feb 2011 11:21:12 +0000
changeset 70 8ab3a06577cf
parent 67 7478be786f87
child 71 426070e68b21
permissions -rw-r--r--
slightly more on the paper
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"
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    10
 UPLUS :: "'a set \<Rightarrow> 'a set \<Rightarrow> (nat \<times> 'a) set"
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
    11
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
    12
abbreviation
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
    13
  "EClass x R \<equiv> R `` {x}"
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
    14
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
    15
notation (latex output)
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
    16
  str_eq_rel ("\<approx>\<^bsub>_\<^esub>") and
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
    17
  Seq (infixr "\<cdot>" 100) and
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
    18
  Star ("_\<^bsup>\<star>\<^esup>") and
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
    19
  pow ("_\<^bsup>_\<^esup>" [100, 100] 100) and
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
    20
  Suc ("_+1" [100] 100) and
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
    21
  quotient ("_ \<^raw:\ensuremath{\!\sslash\!}> _" [90, 90] 90) and
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    22
  REL ("\<approx>") and
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
    23
  UPLUS ("_ \<^raw:\ensuremath{\uplus}> _" [90, 90] 90) and
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
    24
  L ("L '(_')" [0] 101) and
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
    25
  EClass ("\<lbrakk>_\<rbrakk>\<^bsub>_\<^esub>" [100, 100] 100) 
24
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
    26
(*>*)
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
    27
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
    28
24
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
    29
section {* Introduction *}
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
    30
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
    31
text {*
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
    32
  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
    33
  Science, with many beautiful theorems and many useful algorithms. There is a
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    34
  wide range of textbooks on this subject, many of which are aimed at students
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    35
  and contain very detailed ``pencil-and-paper'' proofs
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    36
  (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
    37
  formalising these theorems and by verifying formally the algorithms.
59
fc35eb54fdc9 more on the intro
urbanc
parents: 58
diff changeset
    38
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    39
  There is however a problem: the typical approach to regular languages is to
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    40
  introduce finite automata and then define everything in terms of them.  For
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    41
  example, a regular language is normally defined as one whose strings are
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    42
  recognised by a finite deterministic automaton. This approach has many
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    43
  benefits. Among them is that it is easy to convince oneself from the fact that
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    44
  regular languages are closed under complementation: one just has to exchange
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    45
  the accepting and non-accepting states in the corresponding automaton to
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    46
  obtain an automaton for the complement language.  The problem, however, lies with
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
    47
  formalising such reasoning in a HOL-based theorem prover, in our case
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
    48
  Isabelle/HOL. Automata are build up from states and transitions that 
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
    49
  need to be represented as graphs or matrices, neither
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    50
  of which can be defined as inductive datatype.\footnote{In some works
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    51
  functions are used to represent state transitions, but also they are not
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    52
  inductive datatypes.} This means we have to build our own reasoning
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    53
  infrastructure for them, as neither Isabelle/HOL nor HOL4 nor HOLlight support
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    54
  them with libraries.
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    55
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    56
  Even worse, reasoning about graphs and matrices can be a real hassle in HOL-based
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    57
  theorem provers.  Consider for example the operation of sequencing 
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    58
  two automata, say $A_1$ and $A_2$, by connecting the
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
    59
  accepting states of $A_1$ to the initial state of $A_2$:
61
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
    60
  
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    61
  
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    62
  \begin{center}
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    63
  \begin{tabular}{ccc}
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    64
  \begin{tikzpicture}[scale=0.8]
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    65
  %\draw[step=2mm] (-1,-1) grid (1,1);
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    66
  
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    67
  \draw[rounded corners=1mm, very thick] (-1.0,-0.3) rectangle (-0.2,0.3);
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    68
  \draw[rounded corners=1mm, very thick] ( 0.2,-0.3) rectangle ( 1.0,0.3);
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    69
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    70
  \node (A) at (-1.0,0.0) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    71
  \node (B) at ( 0.2,0.0) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    72
  
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    73
  \node (C) at (-0.2, 0.13) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    74
  \node (D) at (-0.2,-0.13) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    75
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    76
  \node (E) at (1.0, 0.2) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    77
  \node (F) at (1.0,-0.0) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    78
  \node (G) at (1.0,-0.2) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    79
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    80
  \draw (-0.6,0.0) node {\footnotesize$A_1$};
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    81
  \draw ( 0.6,0.0) node {\footnotesize$A_2$};
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    82
  \end{tikzpicture}
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    83
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    84
  & 
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    85
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    86
  \raisebox{1.1mm}{\bf\Large$\;\;\;\Rightarrow\,\;\;$}
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    87
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    88
  &
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    89
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    90
  \begin{tikzpicture}[scale=0.8]
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    91
  %\draw[step=2mm] (-1,-1) grid (1,1);
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    92
  
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    93
  \draw[rounded corners=1mm, very thick] (-1.0,-0.3) rectangle (-0.2,0.3);
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    94
  \draw[rounded corners=1mm, very thick] ( 0.2,-0.3) rectangle ( 1.0,0.3);
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    95
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    96
  \node (A) at (-1.0,0.0) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    97
  \node (B) at ( 0.2,0.0) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    98
  
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    99
  \node (C) at (-0.2, 0.13) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   100
  \node (D) at (-0.2,-0.13) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   101
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   102
  \node (E) at (1.0, 0.2) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   103
  \node (F) at (1.0,-0.0) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   104
  \node (G) at (1.0,-0.2) [circle, very thick, draw, fill=white, inner sep=0.4mm] {};
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   105
  
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   106
  \draw (C) to [very thick, bend left=45] (B);
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   107
  \draw (D) to [very thick, bend right=45] (B);
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   108
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   109
  \draw (-0.6,0.0) node {\footnotesize$A_1$};
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   110
  \draw ( 0.6,0.0) node {\footnotesize$A_2$};
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   111
  \end{tikzpicture}
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   112
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   113
  \end{tabular}
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
   114
  \end{center}
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
   115
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
   116
  \noindent
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   117
  On ``paper'' we can define the corresponding graph in terms of the disjoint 
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   118
  union of the state nodes. Unfortunately in HOL, the definition for disjoint 
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   119
  union, namely 
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
   120
61
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
   121
  \begin{center}
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   122
  @{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}"}
61
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
   123
  \end{center}
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
   124
61
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
   125
  \noindent
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   126
  changes the type---the disjoint union is not a set, but a set of pairs. 
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   127
  Using this definition for disjoint unions means we do not have a single type for automata
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   128
  and hence will not be able to state properties about \emph{all}
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   129
  automata, since there is no type quantification available in HOL. An
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   130
  alternative, which provides us with a single type for automata, is to give every 
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   131
  state node an identity, for example a natural
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   132
  number, and then be careful to rename these identities apart whenever
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   133
  connecting two automata. This results in clunky proofs
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   134
  establishing that properties are invariant under renaming. Similarly,
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   135
  connecting two automata represented as matrices results in very adhoc
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   136
  constructions, which are not pleasant to reason about.
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   137
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   138
  Because of these problems to do with representing automata, there seems
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   139
  to be no substantial formalisation of automata theory and regular languages 
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   140
  carried out in a HOL-based theorem prover. We are only aware of the 
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   141
  large formalisation of automata theory in Nuprl \cite{Constable00} and 
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   142
  some smaller formalisations in Coq, for example \cite{Filliatre97}.
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   143
  
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   144
  In this paper, we will not attempt to formalise automata theory, but take a completely 
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   145
  different approach to regular languages. Instead of defining a regular language as one 
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   146
  where there exists an automaton that recognises all strings of the language, we define 
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   147
  a regular language as:
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   148
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   149
  \begin{definition}[A Regular Language]
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   150
  A language @{text A} is regular, provided there is a regular expression that matches all
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   151
  strings of @{text "A"}.
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   152
  \end{definition}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   153
  
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   154
  \noindent
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   155
  The reason is that regular expressions, unlike graphs and matrices, can
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   156
  be easily defined as inductive datatype. Therefore a corresponding reasoning 
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   157
  infrastructure comes for free. This has recently been used in HOL4 for formalising regular 
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   158
  expression matching based on derivatives \cite{OwensSlind08}.  The purpose of this paper is to
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   159
  show that a central result about regular languages, the Myhill-Nerode theorem,  
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   160
  can be recreated by only using regular expressions. This theorem gives a necessary
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   161
  and sufficient condition for when a language is regular. As a corollary of this
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   162
  theorem we can easily establish the usual closure properties, including 
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   163
  complementation, for regular languages.\smallskip
61
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
   164
  
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
   165
  \noindent
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   166
  {\bf Contributions:} To our knowledge, our proof of the Myhill-Nerode theorem is the
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   167
  first that is based on regular expressions, only. We prove the part of this theorem 
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   168
  stating that a regular expression has only finitely many partitions using certain 
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   169
  tagging-functions. Again to our best knowledge, these tagging functions have
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   170
  not been used before to establish the Myhill-Nerode theorem.
24
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   171
*}
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   172
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   173
section {* Preliminaries *}
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   174
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   175
text {*
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   176
  Strings in Isabelle/HOL are lists of characters with the \emph{empty string}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   177
  being represented by the empty list, written @{term "[]"}. \emph{Languages}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   178
  are sets of strings. The language containing all strings is written in
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   179
  Isabelle/HOL as @{term "UNIV::string set"}.  The notation for the quotient
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   180
  of a language @{text A} according to a relation @{term REL} is @{term "A //
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   181
  REL"}. The concatenation of two languages is written @{term "A ;; B"}; a
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   182
  language raised to the power $n$ is written @{term "A \<up> n"}. Both concepts
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   183
  are defined as
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   184
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   185
  \begin{center}
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   186
  @{thm Seq_def[THEN eq_reflection, where A1="A" and B1="B"]}
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   187
  \hspace{7mm}
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   188
  @{thm pow.simps(1)[THEN eq_reflection, where A1="A"]}
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   189
  \hspace{7mm}
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   190
  @{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
   191
  \end{center}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   192
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   193
  \noindent
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   194
  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
   195
  is defined as the union over all powers, namely @{thm Star_def}.
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   196
  
51
6cfb92de4654 some tuning of the paper
urbanc
parents: 50
diff changeset
   197
  Central to our proof will be the solution of equational systems
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   198
  involving regular expressions. For this we will use the following ``reverse'' 
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   199
  version of Arden's lemma.
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   200
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   201
  \begin{lemma}[Reverse Arden's Lemma]\mbox{}\\
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   202
  If @{thm (prem 1) ardens_revised} then
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   203
  @{thm (lhs) ardens_revised} has the unique solution
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   204
  @{thm (rhs) ardens_revised}.
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   205
  \end{lemma}
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   206
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   207
  \begin{proof}
51
6cfb92de4654 some tuning of the paper
urbanc
parents: 50
diff changeset
   208
  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
   209
  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
   210
  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
   211
  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
   212
  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
   213
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   214
  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
   215
  on @{text n}, we can establish the property
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   216
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   217
  \begin{center}
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   218
  @{text "(*)"}\hspace{5mm} @{thm (concl) ardens_helper}
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   219
  \end{center}
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   220
  
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   221
  \noindent
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   222
  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
   223
  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
   224
  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
   225
  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
   226
  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
   227
  (the strings in @{term "X ;; (A \<up> Suc k)"} are all longer). 
53
da85feadb8e3 small typo
urbanc
parents: 52
diff changeset
   228
  From @{text "(*)"} it follows then that
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   229
  @{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
   230
  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
   231
  is equal to @{term "B ;; A\<star>"}, as we needed to show.\qed
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   232
  \end{proof}
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   233
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   234
  \noindent
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   235
  Regular expressions are defined as the following inductive datatype
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   236
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   237
  \begin{center}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   238
  @{text r} @{text "::="}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   239
  @{term NULL}\hspace{1.5mm}@{text"|"}\hspace{1.5mm} 
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   240
  @{term EMPTY}\hspace{1.5mm}@{text"|"}\hspace{1.5mm} 
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   241
  @{term "CHAR c"}\hspace{1.5mm}@{text"|"}\hspace{1.5mm} 
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   242
  @{term "SEQ r r"}\hspace{1.5mm}@{text"|"}\hspace{1.5mm} 
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   243
  @{term "ALT r r"}\hspace{1.5mm}@{text"|"}\hspace{1.5mm} 
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   244
  @{term "STAR r"}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   245
  \end{center}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   246
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   247
  \noindent
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   248
  The language matched by a regular expression is defined as usual:
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   249
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   250
  \begin{center}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   251
  \begin{tabular}{c@ {\hspace{10mm}}c}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   252
  \begin{tabular}{rcl}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   253
  @{thm (lhs) L_rexp.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) L_rexp.simps(1)}\\
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   254
  @{thm (lhs) L_rexp.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) L_rexp.simps(2)}\\
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   255
  @{thm (lhs) L_rexp.simps(3)[where c="c"]} & @{text "\<equiv>"} & @{thm (rhs) L_rexp.simps(3)[where c="c"]}\\
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   256
  \end{tabular}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   257
  &
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   258
  \begin{tabular}{rcl}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   259
  @{thm (lhs) L_rexp.simps(4)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]} & @{text "\<equiv>"} &
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   260
       @{thm (rhs) L_rexp.simps(4)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]}\\
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   261
  @{thm (lhs) L_rexp.simps(5)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]} & @{text "\<equiv>"} &
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   262
       @{thm (rhs) L_rexp.simps(5)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]}\\
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   263
  @{thm (lhs) L_rexp.simps(6)[where r="r"]} & @{text "\<equiv>"} &
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   264
      @{thm (rhs) L_rexp.simps(6)[where r="r"]}\\
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   265
  \end{tabular}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   266
  \end{tabular}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   267
  \end{center}
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   268
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   269
*}
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   270
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   271
section {* Finite Partitions Imply Regularity of a Language *}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   272
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   273
text {*
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   274
  \begin{definition}[Myhill-Nerode Relation]\mbox{}\\
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   275
  @{thm str_eq_rel_def[simplified]}
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   276
  \end{definition}
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   277
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   278
  \begin{definition} @{text "finals A"} are the equivalence classes that contain
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   279
  strings from @{text A}\\
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   280
  @{thm finals_def}
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   281
  \end{definition}
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   282
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   283
  @{thm lang_is_union_of_finals}
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   284
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   285
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   286
  \begin{theorem}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   287
  Given a language @{text A}.
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   288
  @{thm[mode=IfThen] hard_direction}
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   289
  \end{theorem}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   290
*}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   291
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   292
section {* Regular Expressions Generate Finitely Many Partitions *}
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   293
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   294
text {*
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   295
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   296
  \begin{theorem}
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   297
  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
   298
  \end{theorem}  
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   299
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   300
  \begin{proof}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   301
  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
   302
  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
   303
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   304
  \begin{center}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   305
  \begin{tabular}{l}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   306
  @{thm quot_null_eq}\\
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   307
  @{thm quot_empty_subset}\\
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   308
  @{thm quot_char_subset}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   309
  \end{tabular}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   310
  \end{center}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   311
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   312
  \end{proof}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   313
*}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   314
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   315
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   316
section {* Conclusion and Related Work *}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   317
24
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   318
(*<*)
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   319
end
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   320
(*>*)