Paper/Paper.thy
author urbanc
Tue, 08 Feb 2011 18:04:54 +0000
changeset 82 14b12b5de6d3
parent 79 bba9c80735f9
child 83 f438f4dbaada
permissions -rw-r--r--
added coments about functions
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
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
    17
  str_eq ("_ \<approx>\<^bsub>_\<^esub> _") and
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
    18
  Seq (infixr "\<cdot>" 100) and
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
    19
  Star ("_\<^bsup>\<star>\<^esup>") and
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
    20
  pow ("_\<^bsup>_\<^esup>" [100, 100] 100) and
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
    21
  Suc ("_+1" [100] 100) and
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
    22
  quotient ("_ \<^raw:\ensuremath{\!\sslash\!}> _" [90, 90] 90) and
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    23
  REL ("\<approx>") and
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
    24
  UPLUS ("_ \<^raw:\ensuremath{\uplus}> _" [90, 90] 90) and
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
    25
  L ("\<^raw:\ensuremath{\cal{L}}>'(_')" [0] 101) and
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
    26
  Lam ("\<lambda>'(_')" [100] 100) and 
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
    27
  Trn ("_, _" [100, 100] 100) and 
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
    28
  EClass ("\<lbrakk>_\<rbrakk>\<^bsub>_\<^esub>" [100, 100] 100) and
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
    29
  transition ("_ \<^raw:\ensuremath{\stackrel{\text{>_\<^raw:}}{\Longmapsto}}> _" [100, 100, 100] 100)
24
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
    30
(*>*)
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
    31
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
    32
24
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
    33
section {* Introduction *}
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
    34
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
    35
text {*
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
    36
  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
    37
  Science, with many beautiful theorems and many useful algorithms. There is a
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    38
  wide range of textbooks on this subject, many of which are aimed at students
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    39
  and contain very detailed ``pencil-and-paper'' proofs
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    40
  (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
    41
  formalising these theorems and by verifying formally the algorithms.
59
fc35eb54fdc9 more on the intro
urbanc
parents: 58
diff changeset
    42
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    43
  There is however a problem: the typical approach to regular languages is to
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    44
  introduce finite automata and then define everything in terms of them.  For
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    45
  example, a regular language is normally defined as one whose strings are
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    46
  recognised by a finite deterministic automaton. This approach has many
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
    47
  benefits. Among them is the fact that it is easy to convince oneself that
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    48
  regular languages are closed under complementation: one just has to exchange
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    49
  the accepting and non-accepting states in the corresponding automaton to
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    50
  obtain an automaton for the complement language.  The problem, however, lies with
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
    51
  formalising such reasoning in a HOL-based theorem prover, in our case
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
    52
  Isabelle/HOL. Automata are build up from states and transitions that 
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
    53
  need to be represented as graphs, matrices or functions, none
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
    54
  of which can be defined as inductive datatype. 
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    55
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
    56
  In case of graphs and matrices, this means we have to build our own
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
    57
  reasoning infrastructure for them, as neither Isabelle/HOL nor HOL4 nor
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
    58
  HOLlight support them with libraries. Even worse, reasoning about graphs and
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
    59
  matrices can be a real hassle in HOL-based theorem provers.  Consider for
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
    60
  example the operation of sequencing two automata, say $A_1$ and $A_2$, by
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
    61
  connecting the accepting states of $A_1$ to the initial state of $A_2$:  
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    62
  
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
    63
  \begin{center}
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    64
  \begin{tabular}{ccc}
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    65
  \begin{tikzpicture}[scale=0.8]
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    66
  %\draw[step=2mm] (-1,-1) grid (1,1);
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    67
  
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    68
  \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
    69
  \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
    70
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    71
  \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
    72
  \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
    73
  
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    74
  \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
    75
  \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
    76
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    77
  \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
    78
  \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
    79
  \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
    80
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    81
  \draw (-0.6,0.0) node {\footnotesize$A_1$};
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    82
  \draw ( 0.6,0.0) node {\footnotesize$A_2$};
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    83
  \end{tikzpicture}
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
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    87
  \raisebox{1.1mm}{\bf\Large$\;\;\;\Rightarrow\,\;\;$}
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
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    91
  \begin{tikzpicture}[scale=0.8]
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    92
  %\draw[step=2mm] (-1,-1) grid (1,1);
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    93
  
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    94
  \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
    95
  \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
    96
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    97
  \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
    98
  \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
    99
  
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   100
  \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
   101
  \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
   102
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   103
  \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
   104
  \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
   105
  \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
   106
  
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   107
  \draw (C) to [very thick, bend left=45] (B);
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   108
  \draw (D) to [very thick, bend right=45] (B);
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   109
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   110
  \draw (-0.6,0.0) node {\footnotesize$A_1$};
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   111
  \draw ( 0.6,0.0) node {\footnotesize$A_2$};
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   112
  \end{tikzpicture}
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   113
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   114
  \end{tabular}
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
   115
  \end{center}
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
   116
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
   117
  \noindent
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   118
  On ``paper'' we can define the corresponding graph in terms of the disjoint 
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   119
  union of the state nodes. Unfortunately in HOL, the definition for disjoint 
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   120
  union, namely 
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   121
  %
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   122
  \begin{equation}\label{disjointunion}
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   123
  @{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
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   124
  \end{equation}
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
   125
61
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
   126
  \noindent
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   127
  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
   128
  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
   129
  and hence will not be able to state properties about \emph{all}
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   130
  automata, since there is no type quantification available in HOL. An
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   131
  alternative, which provides us with a single type for automata, is to give every 
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   132
  state node an identity, for example a natural
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   133
  number, and then be careful to rename these identities apart whenever
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   134
  connecting two automata. This results in clunky proofs
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   135
  establishing that properties are invariant under renaming. Similarly,
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   136
  connecting two automata represented as matrices results in very adhoc
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   137
  constructions, which are not pleasant to reason about.
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   138
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   139
  Functions are much better supported in Isabelle/HOL, but they still lead to similar
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   140
  problems as with graphs.  Composing two non-deterministic automata in parallel
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   141
  poses still the problem of how to implement disjoint unions. Nipkow \cite{Nipkow98} 
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   142
  dismisses the option using identities, because it leads to messy proofs. He
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   143
  opts for a variant of \eqref{disjointunion}, but writes
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   144
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   145
  \begin{quote}
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   146
  \it ``If the reader finds the above treatment in terms of bit lists revoltingly
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   147
  concrete, I cannot disagree.''
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   148
  \end{quote}
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   149
  
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   150
  \noindent
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   151
  Moreover, it is not so clear how to conveniently impose a finiteness condition 
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   152
  upon functions in order to represent \emph{finite} automata. The best is
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   153
  probably to resort to more advanced reasoning frameworks, such as \emph{locales}.
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   154
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   155
  Because of these problems to do with representing automata, there seems
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   156
  to be no substantial formalisation of automata theory and regular languages 
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   157
  carried out in a HOL-based theorem prover. Nipkow establishes in 
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   158
  \cite{Nipkow98} the link between regular expressions and automata in
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   159
  the context of lexing. The only larger formalisations of automata theory 
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   160
  are carried out in Nuprl \cite{Constable00} and in Coq (for example 
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   161
  \cite{Filliatre97}).
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   162
  
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   163
  In this paper, we will not attempt to formalise automata theory in
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   164
  Isabelle/HOL, but take a completely different approach to regular
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   165
  languages. Instead of defining a regular language as one where there exists
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   166
  an automaton that recognises all strings of the language, we define a
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   167
  regular language as:
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   168
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   169
  \begin{definition}
77
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   170
  A language @{text A} is \emph{regular}, provided there is a regular expression that matches all
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   171
  strings of @{text "A"}.
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   172
  \end{definition}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   173
  
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   174
  \noindent
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   175
  The reason is that regular expressions, unlike graphs and matrices, can
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   176
  be easily defined as inductive datatype. Consequently a corresponding reasoning 
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   177
  infrastructure comes for free. This has recently been exploited in HOL4 with a formalisation
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   178
  of regular expression matching based on derivatives \cite{OwensSlind08}.  The purpose of this paper is to
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   179
  show that a central result about regular languages---the Myhill-Nerode theorem---can 
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   180
  be recreated by only using regular expressions. This theorem gives necessary
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   181
  and sufficient conditions for when a language is regular. As a corollary of this
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   182
  theorem we can easily establish the usual closure properties, including 
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   183
  complementation, for regular languages.\smallskip
61
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
   184
  
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
   185
  \noindent
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   186
  {\bf Contributions:} To our knowledge, our proof of the Myhill-Nerode theorem is the
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   187
  first that is based on regular expressions, only. We prove the part of this theorem 
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   188
  stating that a regular expression has only finitely many partitions using certain 
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   189
  tagging-functions. Again to our best knowledge, these tagging functions have
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   190
  not been used before to establish the Myhill-Nerode theorem.
24
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   191
*}
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   192
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   193
section {* Preliminaries *}
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   194
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   195
text {*
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   196
  Strings in Isabelle/HOL are lists of characters with the \emph{empty string}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   197
  being represented by the empty list, written @{term "[]"}. \emph{Languages}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   198
  are sets of strings. The language containing all strings is written in
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   199
  Isabelle/HOL as @{term "UNIV::string set"}. The concatenation of two languages 
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   200
  is written @{term "A ;; B"} and a language raised to the power $n$ is written 
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   201
  @{term "A \<up> n"}. Their definitions are
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   202
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   203
  \begin{center}
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   204
  @{thm Seq_def[THEN eq_reflection, where A1="A" and B1="B"]}
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   205
  \hspace{7mm}
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   206
  @{thm pow.simps(1)[THEN eq_reflection, where A1="A"]}
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   207
  \hspace{7mm}
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   208
  @{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
   209
  \end{center}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   210
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   211
  \noindent
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   212
  where @{text "@"} is the usual list-append operation. The Kleene-star of a language @{text A}
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   213
  is defined as the union over all powers, namely @{thm Star_def}. In the paper
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   214
  we will often make use of the following properties.
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   215
  
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   216
  \begin{proposition}\label{langprops}\mbox{}\\
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   217
  \begin{tabular}{@ {}ll@ {\hspace{10mm}}ll}
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   218
  (i)   & @{thm star_cases}      & (ii)  & @{thm[mode=IfThen] pow_length}\\
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   219
  (iii) & @{thm seq_Union_left}  & 
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   220
  \end{tabular}
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   221
  \end{proposition}
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   222
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   223
  \noindent
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   224
  We omit the proofs of these properties, but invite the reader to consult
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   225
  our formalisation.\footnote{Available at ???}
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   226
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   227
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   228
  The notation for the quotient of a language @{text A} according to an 
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   229
  equivalence relation @{term REL} is @{term "A // REL"}. We will write 
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   230
  @{text "\<lbrakk>x\<rbrakk>\<^isub>\<approx>"} for the equivalence class defined 
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   231
  as @{text "{y | y \<approx> x}"}.
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   232
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   233
51
6cfb92de4654 some tuning of the paper
urbanc
parents: 50
diff changeset
   234
  Central to our proof will be the solution of equational systems
77
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   235
  involving sets of languages. For this we will use Arden's lemma \cite{Brzozowski64}
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   236
  which solves equations of the form @{term "X = A ;; X \<union> B"} provided
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   237
  @{term "[] \<notin> A"}. However we will need the following ``reverse'' 
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   238
  version of Arden's lemma.
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   239
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   240
  \begin{lemma}[Reverse Arden's Lemma]\label{arden}\mbox{}\\
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   241
  If @{thm (prem 1) ardens_revised} then
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   242
  @{thm (lhs) ardens_revised} has the unique solution
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   243
  @{thm (rhs) ardens_revised}.
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   244
  \end{lemma}
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   245
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   246
  \begin{proof}
51
6cfb92de4654 some tuning of the paper
urbanc
parents: 50
diff changeset
   247
  For the right-to-left direction we assume @{thm (rhs) ardens_revised} and show
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   248
  that @{thm (lhs) ardens_revised} holds. From Prop.~\ref{langprops}@{text "(i)"} 
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   249
  we have @{term "A\<star> = {[]} \<union> A ;; A\<star>"},
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   250
  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
   251
  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
   252
  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
   253
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   254
  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
   255
  on @{text n}, we can establish the property
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   256
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   257
  \begin{center}
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   258
  @{text "(*)"}\hspace{5mm} @{thm (concl) ardens_helper}
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   259
  \end{center}
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   260
  
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   261
  \noindent
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   262
  Using this property we can show that @{term "B ;; (A \<up> n) \<subseteq> X"} holds for
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   263
  all @{text n}. From this we can infer @{term "B ;; A\<star> \<subseteq> X"} using the definition
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   264
  of @{text "\<star>"}.
51
6cfb92de4654 some tuning of the paper
urbanc
parents: 50
diff changeset
   265
  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
   266
  with length @{text k} is element in @{text X}. Since @{thm (prem 1) ardens_revised}
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   267
  we know by Prop.~\ref{langprops}@{text "(ii)"} that 
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   268
  @{term "s \<notin> X ;; (A \<up> Suc k)"} since its length is only @{text k}
51
6cfb92de4654 some tuning of the paper
urbanc
parents: 50
diff changeset
   269
  (the strings in @{term "X ;; (A \<up> Suc k)"} are all longer). 
53
da85feadb8e3 small typo
urbanc
parents: 52
diff changeset
   270
  From @{text "(*)"} it follows then that
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   271
  @{term s} must be element in @{term "(\<Union>m\<in>{0..k}. B ;; (A \<up> m))"}. This in turn
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   272
  implies that @{term s} is in @{term "(\<Union>n. B ;; (A \<up> n))"}. Using Prop.~\ref{langprops}@{text "(iii)"} 
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   273
  this is equal to @{term "B ;; A\<star>"}, as we needed to show.\qed
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   274
  \end{proof}
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   275
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   276
  \noindent
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   277
  Regular expressions are defined as the following inductive datatype
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   278
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   279
  \begin{center}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   280
  @{text r} @{text "::="}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   281
  @{term NULL}\hspace{1.5mm}@{text"|"}\hspace{1.5mm} 
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   282
  @{term EMPTY}\hspace{1.5mm}@{text"|"}\hspace{1.5mm} 
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   283
  @{term "CHAR c"}\hspace{1.5mm}@{text"|"}\hspace{1.5mm} 
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   284
  @{term "SEQ r r"}\hspace{1.5mm}@{text"|"}\hspace{1.5mm} 
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   285
  @{term "ALT r r"}\hspace{1.5mm}@{text"|"}\hspace{1.5mm} 
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   286
  @{term "STAR r"}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   287
  \end{center}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   288
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   289
  \noindent
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   290
  and the language matched by a regular expression is defined as:
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   291
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   292
  \begin{center}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   293
  \begin{tabular}{c@ {\hspace{10mm}}c}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   294
  \begin{tabular}{rcl}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   295
  @{thm (lhs) L_rexp.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) L_rexp.simps(1)}\\
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   296
  @{thm (lhs) L_rexp.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) L_rexp.simps(2)}\\
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   297
  @{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
   298
  \end{tabular}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   299
  &
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   300
  \begin{tabular}{rcl}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   301
  @{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
   302
       @{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
   303
  @{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
   304
       @{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
   305
  @{thm (lhs) L_rexp.simps(6)[where r="r"]} & @{text "\<equiv>"} &
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   306
      @{thm (rhs) L_rexp.simps(6)[where r="r"]}\\
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   307
  \end{tabular}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   308
  \end{tabular}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   309
  \end{center}
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   310
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   311
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   312
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   313
*}
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   314
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   315
section {* Finite Partitions Imply Regularity of a Language *}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   316
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   317
text {*
77
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   318
  The key definition in the Myhill-Nerode theorem is the
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   319
  \emph{Myhill-Nerode relation}, which states that w.r.t.~a language two 
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   320
  strings are related, provided there is no distinguishing extension in this
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   321
  language. This can be defined as:
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   322
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   323
  \begin{definition}[Myhill-Nerode Relation]\mbox{}\\
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   324
  @{thm str_eq_def[simplified str_eq_rel_def Pair_Collect]}
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   325
  \end{definition}
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   326
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   327
  \noindent
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   328
  It is easy to see that @{term "\<approx>A"} is an equivalence relation, which
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   329
  partitions the set of all strings, @{text "UNIV"}, into a set of disjoint
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   330
  equivalence classes. One direction of the Myhill-Nerode theorem establishes 
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   331
  that if there are finitely many equivalence classes, then the language is 
77
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   332
  regular. In our setting we therefore have to show:
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   333
  
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   334
  \begin{theorem}\label{myhillnerodeone}
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   335
  @{thm[mode=IfThen] hard_direction}
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   336
  \end{theorem}
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   337
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   338
  \noindent
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   339
  To prove this theorem, we define the set @{term "finals A"} as those equivalence
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   340
  classes that contain strings of @{text A}, namely
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   341
  %
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   342
  \begin{equation} 
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   343
  @{thm finals_def}
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   344
  \end{equation}
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   345
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   346
  \noindent
77
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   347
  It is straightforward to show that @{thm lang_is_union_of_finals} and 
79
bba9c80735f9 started to define things more directly
urbanc
parents: 77
diff changeset
   348
  @{thm finals_in_partitions} hold. 
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   349
  Therefore if we know that there exists a regular expression for every
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   350
  equivalence class in @{term "finals A"} (which by assumption must be
77
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   351
  a finite set), then we can combine these regular expressions with @{const ALT}
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   352
  and obtain a regular expression that matches every string in @{text A}.
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   353
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   354
77
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   355
  We prove Thm.~\ref{myhillnerodeone} by giving a method that can calculate a
79
bba9c80735f9 started to define things more directly
urbanc
parents: 77
diff changeset
   356
  regular expression for \emph{every} equivalence class, not just the ones 
77
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   357
  in @{term "finals A"}. We
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   358
  first define a notion of \emph{transition} between equivalence classes
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   359
  %
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   360
  \begin{equation} 
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   361
  @{thm transition_def}
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   362
  \end{equation}
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   363
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   364
  \noindent
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   365
  which means that if we concatenate all strings matching the regular expression @{text r} 
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   366
  to the end of all strings in the equivalence class @{text Y}, we obtain a subset of 
77
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   367
  @{text X}. Note that we do not define an automaton here, we merely relate two sets
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   368
  (w.r.t.~a regular expression). 
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   369
  
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   370
  Next we build an equational system that
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   371
  contains an equation for each equivalence class. Suppose we have 
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   372
  the equivalence classes @{text "X\<^isub>1,\<dots>,X\<^isub>n"}, there must be one and only one that
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   373
  contains the empty string @{text "[]"} (since equivalence classes are disjoint).
77
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   374
  Let us assume @{text "[] \<in> X\<^isub>1"}. We build the following equational system
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   375
  
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   376
  \begin{center}
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   377
  \begin{tabular}{rcl}
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   378
  @{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)"} \\
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   379
  @{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)"} \\
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   380
  & $\vdots$ \\
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   381
  @{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)"}\\
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   382
  \end{tabular}
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   383
  \end{center}
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   384
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   385
  \noindent
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   386
  where the pairs @{text "(Y\<^isub>i\<^isub>j, CHAR c\<^isub>i\<^isub>j)"} stand for all transitions 
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   387
  @{term "Y\<^isub>i\<^isub>j \<Turnstile>(CHAR c\<^isub>i\<^isub>j)\<Rightarrow> X\<^isub>i"}.  The term @{text "\<lambda>(EMPTY)"} acts as a marker for the equivalence
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   388
  class containing @{text "[]"}. (Note that we mark, roughly speaking, the
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   389
  single ``initial'' state in the equational system, which is different from
77
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   390
  the method by Brzozowski \cite{Brzozowski64}, since for his purposes he needs to mark 
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   391
  the ``terminal'' states.) Overloading the function @{text L} for the two kinds of terms in the 
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   392
  equational system as follows
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   393
  
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   394
  \begin{center}
77
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   395
  @{thm L_rhs_e.simps(2)[where X="Y" and r="r", THEN eq_reflection]}\hspace{10mm}
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   396
  @{thm L_rhs_e.simps(1)[where r="r", THEN eq_reflection]}
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   397
  \end{center}
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   398
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   399
  \noindent
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   400
  we can prove for @{text "X\<^isub>2\<^isub>.\<^isub>.\<^isub>n"} that the following equations
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   401
  %
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   402
  \begin{equation}\label{inv1}
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   403
  @{text "X\<^isub>i = L(Y\<^isub>i\<^isub>1, CHAR c\<^isub>i\<^isub>1) \<union> \<dots> \<union> L(Y\<^isub>i\<^isub>q, CHAR c\<^isub>i\<^isub>q)"}.
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   404
  \end{equation}
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   405
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   406
  \noindent
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   407
  hold. Similarly for @{text "X\<^isub>1"} we can show the following equation
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   408
  %
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   409
  \begin{equation}\label{inv2}
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   410
  @{text "X\<^isub>1 = L(Y\<^isub>i\<^isub>1, CHAR c\<^isub>i\<^isub>1) \<union> \<dots> \<union> L(Y\<^isub>i\<^isub>p, CHAR c\<^isub>i\<^isub>p) \<union> L(\<lambda>(EMPTY))"}.
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   411
  \end{equation}
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   412
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   413
  \noindent
77
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   414
  The reason for adding the @{text \<lambda>}-marker to our equational system is 
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   415
  to obtain this equation, which only holds in this form since none of 
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   416
  the other terms contain the empty string. 
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   417
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   418
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   419
  Our proof of Thm.~\ref{myhillnerodeone}
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   420
  will be by transforming the equational system into a \emph{solved form}
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   421
  maintaining the invariants \eqref{inv1} and \eqref{inv2}. From the solved
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   422
  form we will be able to read off the regular expressions using our 
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   423
  variant of Arden's Lemma (Lem.~\ref{arden}).
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   424
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   425
*}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   426
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   427
section {* Regular Expressions Generate Finitely Many Partitions *}
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   428
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   429
text {*
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   430
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   431
  \begin{theorem}
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   432
  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
   433
  \end{theorem}  
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   434
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   435
  \begin{proof}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   436
  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
   437
  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
   438
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   439
  \begin{center}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   440
  \begin{tabular}{l}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   441
  @{thm quot_null_eq}\\
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   442
  @{thm quot_empty_subset}\\
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   443
  @{thm quot_char_subset}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   444
  \end{tabular}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   445
  \end{center}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   446
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   447
  \end{proof}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   448
*}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   449
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   450
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   451
section {* Conclusion and Related Work *}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   452
24
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   453
(*<*)
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   454
end
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   455
(*>*)