Journal/Paper.thy
author urbanc
Tue, 16 Aug 2011 10:21:14 +0000
changeset 198 b300f2c5d51d
parent 197 cf1c17431dab
child 199 11c3c302fa2e
permissions -rw-r--r--
final(?) version of 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
182
560712a29a36 a version of the proof which dispenses with the notion of string-subtraction
urbanc
parents: 181
diff changeset
     3
imports "../Closures" "../Attic/Prefix_subtract" 
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
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
    15
abbreviation 
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 160
diff changeset
    16
  "Append_rexp2 r_itm r \<equiv> Append_rexp r r_itm"
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
    17
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
    18
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    19
abbreviation
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    20
  "pow" (infixl "\<up>" 100)
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    21
where
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    22
  "A \<up> n \<equiv> A ^^ n"
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    23
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    24
syntax (latex output)
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    25
  "_Collect" :: "pttrn => bool => 'a set"              ("(1{_ | _})")
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    26
  "_CollectIn" :: "pttrn => 'a set => bool => 'a set"   ("(1{_ \<in> _ | _})")
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    27
translations
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    28
  "_Collect p P"      <= "{p. P}"
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    29
  "_Collect p P"      <= "{p|xs. P}"
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    30
  "_CollectIn p A P"  <= "{p : A. P}"
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    31
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
    32
abbreviation "ZERO \<equiv> Zero"
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
    33
abbreviation "ONE \<equiv> One"
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
    34
abbreviation "ATOM \<equiv> Atom"
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
    35
abbreviation "PLUS \<equiv> Plus"
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
    36
abbreviation "TIMES \<equiv> Times"
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
    37
abbreviation "TIMESS \<equiv> Timess"
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    38
abbreviation "STAR \<equiv> Star"
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    39
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    40
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
    41
notation (latex output)
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
    42
  str_eq ("\<approx>\<^bsub>_\<^esub>") and
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
    43
  str_eq_applied ("_ \<approx>\<^bsub>_\<^esub> _") and
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    44
  conc (infixr "\<cdot>" 100) and
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
    45
  star ("_\<^bsup>\<star>\<^esup>" [101] 200) and
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
    46
  pow ("_\<^bsup>_\<^esup>" [100, 100] 100) and
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
    47
  Suc ("_+1" [100] 100) and
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
    48
  quotient ("_ \<^raw:\ensuremath{\!\sslash\!}> _" [90, 90] 90) and
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    49
  REL ("\<approx>") and
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
    50
  UPLUS ("_ \<^raw:\ensuremath{\uplus}> _" [90, 90] 90) and
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
    51
  lang ("\<^raw:\ensuremath{\cal{L}}>" 101) and
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    52
  lang ("\<^raw:\ensuremath{\cal{L}}>'(_')" [0] 101) and
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
    53
  lang_trm ("\<^raw:\ensuremath{\cal{L}}>'(_')" [0] 101) and
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
    54
  Lam ("\<lambda>'(_')" [100] 100) and 
89
42af13d194c9 a bit more on the paper
urbanc
parents: 88
diff changeset
    55
  Trn ("'(_, _')" [100, 100] 100) and 
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
    56
  EClass ("\<lbrakk>_\<rbrakk>\<^bsub>_\<^esub>" [100, 100] 100) and
88
1436fc451bb9 added something about Setalt and folds
urbanc
parents: 86
diff changeset
    57
  transition ("_ \<^raw:\ensuremath{\stackrel{\text{>_\<^raw:}}{\Longmapsto}}> _" [100, 100, 100] 100) and
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
    58
  Setalt ("\<^raw:\ensuremath{\bigplus}>_" [1000] 999) and
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 160
diff changeset
    59
  Append_rexp2 ("_ \<^raw:\ensuremath{\triangleleft}> _" [100, 100] 100) and
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 160
diff changeset
    60
  Append_rexp_rhs ("_ \<^raw:\ensuremath{\triangleleft}> _" [100, 100] 50) and
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    61
  
119
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
    62
  uminus ("\<^raw:\ensuremath{\overline{>_\<^raw:}}>" [100] 100) and
183
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
    63
  tag_Plus ("+tag _ _" [100, 100] 100) and
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
    64
  tag_Plus ("+tag _ _ _" [100, 100, 100] 100) and
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
    65
  tag_Times ("\<times>tag _ _" [100, 100] 100) and
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
    66
  tag_Times ("\<times>tag _ _ _" [100, 100, 100] 100) and
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
    67
  tag_Star ("\<star>tag _" [100] 100) and
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
    68
  tag_Star ("\<star>tag _ _" [100, 100] 100) and
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
    69
  tag_eq ("\<^raw:$\threesim$>\<^bsub>_\<^esub>" [100] 100) and
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
    70
  Delta ("\<Delta>'(_')") and
180
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
    71
  nullable ("\<delta>'(_')") and
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
    72
  Cons ("_ :: _" [100, 100] 100) and
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
    73
  Rev ("Rev _" [1000] 100) and
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
    74
  Der ("Der _ _" [1000, 1000] 100) and
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
    75
  ONE ("ONE" 999) and
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
    76
  ZERO ("ZERO" 999) and
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
    77
  pders_lang ("pdersl") and
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
    78
  UNIV1 ("UNIV\<^isup>+") and
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
    79
  Ders_lang ("Dersl")
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
    80
119
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
    81
lemma meta_eq_app:
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
    82
  shows "f \<equiv> \<lambda>x. g x \<Longrightarrow> f x \<equiv> g x"
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
    83
  by auto
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
    84
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
    85
lemma str_eq_def':
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
    86
  shows "x \<approx>A y \<equiv> (\<forall>z. x @ z \<in> A \<longleftrightarrow> y @ z \<in> A)"
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
    87
unfolding str_eq_def by simp
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
    88
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    89
lemma conc_def':
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    90
  "A \<cdot> B = {s\<^isub>1 @ s\<^isub>2 | s\<^isub>1 s\<^isub>2. s\<^isub>1 \<in> A \<and> s\<^isub>2 \<in> B}"
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    91
unfolding conc_def by simp
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    92
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    93
lemma conc_Union_left: 
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    94
  shows "B \<cdot> (\<Union>n. A \<up> n) = (\<Union>n. B \<cdot> (A \<up> n))"
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    95
unfolding conc_def by auto
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    96
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    97
lemma test:
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    98
  assumes X_in_eqs: "(X, rhs) \<in> Init (UNIV // \<approx>A)"
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    99
  shows "X = \<Union> (lang_trm `  rhs)"
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   100
using assms l_eq_r_in_eqs by (simp)
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   101
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   102
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   103
(* THEOREMS *)
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   104
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   105
notation (Rule output)
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   106
  "==>"  ("\<^raw:\mbox{}\inferrule{\mbox{>_\<^raw:}}>\<^raw:{\mbox{>_\<^raw:}}>")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   107
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   108
syntax (Rule output)
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   109
  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   110
  ("\<^raw:\mbox{}\inferrule{>_\<^raw:}>\<^raw:{\mbox{>_\<^raw:}}>")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   111
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   112
  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" 
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   113
  ("\<^raw:\mbox{>_\<^raw:}\\>/ _")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   114
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   115
  "_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   116
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   117
notation (Axiom output)
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   118
  "Trueprop"  ("\<^raw:\mbox{}\inferrule{\mbox{}}{\mbox{>_\<^raw:}}>")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   119
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   120
notation (IfThen output)
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   121
  "==>"  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   122
syntax (IfThen output)
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   123
  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   124
  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   125
  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}> /\<^raw:{\normalsize \,>and\<^raw:\,}>/ _")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   126
  "_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   127
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   128
notation (IfThenNoBox output)
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   129
  "==>"  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   130
syntax (IfThenNoBox output)
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   131
  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   132
  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   133
  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("_ /\<^raw:{\normalsize \,>and\<^raw:\,}>/ _")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   134
  "_asm" :: "prop \<Rightarrow> asms" ("_")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   135
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   136
24
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   137
(*>*)
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   138
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   139
24
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   140
section {* Introduction *}
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   141
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   142
text {*
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   143
  \noindent
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   144
  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
   145
  Science, with many beautiful theorems and many useful algorithms. There is a
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   146
  wide range of textbooks on this subject, many of which are aimed at students
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   147
  and contain very detailed `pencil-and-paper' proofs (e.g.~\cite{Kozen97,
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   148
  HopcroftUllman69}). It seems natural to exercise theorem provers by
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   149
  formalising the theorems and by verifying formally the algorithms.  
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   150
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   151
  A popular choice for a theorem prover would be one based on Higher-Order
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   152
  Logic (HOL), for example HOL4, HOLlight or Isabelle/HOL. For the development
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   153
  presented in this paper we will use the latter. HOL is a predicate calculus
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   154
  that allows quantification over predicate variables. Its type system is
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   155
  based on Church's Simple Theory of Types \cite{Church40}.  Although many
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   156
  mathematical concepts can be conveniently expressed in HOL, there are some
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   157
  limitations that hurt badly, if one attempts a simple-minded formalisation
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
   158
  of regular languages in it. 
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
   159
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
   160
  The typical approach to regular languages is to
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   161
  introduce finite automata and then define everything in terms of them
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   162
  \cite{ HopcroftUllman69,Kozen97}.  For example, a regular language is 
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   163
  normally defined as:
59
fc35eb54fdc9 more on the intro
urbanc
parents: 58
diff changeset
   164
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   165
  \begin{dfntn}\label{baddef}
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   166
  A language @{text A} is \emph{regular}, provided there is a 
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   167
  finite deterministic automaton that recognises all strings of @{text "A"}.
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   168
  \end{dfntn}
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   169
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   170
  \noindent  
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   171
  This approach has many benefits. Among them is the fact that it is easy to
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   172
  convince oneself that regular languages are closed under complementation:
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   173
  one just has to exchange the accepting and non-accepting states in the
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   174
  corresponding automaton to obtain an automaton for the complement language.
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   175
  The problem, however, lies with formalising such reasoning in a HOL-based
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   176
  theorem prover. Automata are built up from states and transitions that need
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   177
  to be represented as graphs, matrices or functions, none of which can be
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   178
  defined as an inductive datatype.
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   179
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   180
  In case of graphs and matrices, this means we have to build our own
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   181
  reasoning infrastructure for them, as neither Isabelle/HOL nor HOL4 nor
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   182
  HOLlight support them with libraries. Even worse, reasoning about graphs and
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   183
  matrices can be a real hassle in HOL-based theorem provers, because
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   184
  we have to be able to combine automata.  Consider for
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   185
  example the operation of sequencing two automata, say $A_1$ and $A_2$, by
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   186
  connecting the accepting states of $A_1$ to the initial state of $A_2$:
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   187
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
   188
  \begin{center}
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   189
  \begin{tabular}{ccc}
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   190
  \begin{tikzpicture}[scale=1]
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   191
  %\draw[step=2mm] (-1,-1) grid (1,1);
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   192
  
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   193
  \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
   194
  \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
   195
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   196
  \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
   197
  \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
   198
  
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   199
  \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
   200
  \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
   201
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   202
  \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
   203
  \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
   204
  \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
   205
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   206
  \draw (-0.6,0.0) node {\small$A_1$};
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   207
  \draw ( 0.6,0.0) node {\small$A_2$};
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   208
  \end{tikzpicture}
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   209
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   210
  & 
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   211
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
   212
  \raisebox{2.1mm}{\bf\Large$\;\;\;\Rightarrow\,\;\;$}
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   213
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   214
  &
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   215
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   216
  \begin{tikzpicture}[scale=1]
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   217
  %\draw[step=2mm] (-1,-1) grid (1,1);
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   218
  
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   219
  \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
   220
  \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
   221
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   222
  \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
   223
  \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
   224
  
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   225
  \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
   226
  \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
   227
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   228
  \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
   229
  \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
   230
  \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
   231
  
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   232
  \draw (C) to [very thick, bend left=45] (B);
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   233
  \draw (D) to [very thick, bend right=45] (B);
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   234
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   235
  \draw (-0.6,0.0) node {\small$A_1$};
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   236
  \draw ( 0.6,0.0) node {\small$A_2$};
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   237
  \end{tikzpicture}
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   238
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   239
  \end{tabular}
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
   240
  \end{center}
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
   241
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
   242
  \noindent
178
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   243
  On `paper' we can define the corresponding 
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   244
  graph in terms of the disjoint 
88
1436fc451bb9 added something about Setalt and folds
urbanc
parents: 86
diff changeset
   245
  union of the state nodes. Unfortunately in HOL, the standard definition for disjoint 
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   246
  union, namely 
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   247
  %
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   248
  \begin{equation}\label{disjointunion}
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   249
  @{text "A\<^isub>1 \<uplus> A\<^isub>2 \<equiv> {(1, x) | x \<in> A\<^isub>1} \<union> {(2, y) | y \<in> A\<^isub>2}"}
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   250
  \end{equation}
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
   251
61
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
   252
  \noindent
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
   253
  changes the type---the disjoint union is not a set, but a set of
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
   254
  pairs. Using this definition for disjoint union means we do not have a
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   255
  single type for the states of automata. As a result we will not be able to define a
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
   256
  regular language as one for which there exists an automaton that recognises
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
   257
  all its strings (Definition~\ref{baddef}). This is because we cannot make a
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   258
  definition in HOL that is only polymorphic in the state type and there is no type
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
   259
  quantification available in HOL (unlike in Coq, for example).\footnote{Slind
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
   260
  already pointed out this problem in an email to the HOL4 mailing list on
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
   261
  21st April 2005.}
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
   262
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   263
  An alternative, which provides us with a single type for states of automata,
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   264
  is to give every state node an identity, for example a natural number, and
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   265
  then be careful to rename these identities apart whenever connecting two
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   266
  automata. This results in clunky proofs establishing that properties are
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   267
  invariant under renaming. Similarly, connecting two automata represented as
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   268
  matrices results in very adhoc constructions, which are not pleasant to
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   269
  reason about.
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   270
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   271
  Functions are much better supported in Isabelle/HOL, but they still lead to similar
88
1436fc451bb9 added something about Setalt and folds
urbanc
parents: 86
diff changeset
   272
  problems as with graphs.  Composing, for example, two non-deterministic automata in parallel
93
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   273
  requires also the formalisation of disjoint unions. Nipkow \cite{Nipkow98} 
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   274
  dismisses for this the option of using identities, because it leads according to 
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   275
  him to ``messy proofs''. Since he does not need to define what regular
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   276
  languages are, Nipkow opts for a variant of \eqref{disjointunion} using bit lists, but writes 
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   277
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   278
  \begin{quote}
93
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   279
  \it%
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   280
  \begin{tabular}{@ {}l@ {}p{0.88\textwidth}@ {}}
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   281
  `` & All lemmas appear obvious given a picture of the composition of automata\ldots
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   282
       Yet their proofs require a painful amount of detail.''
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   283
  \end{tabular}
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   284
  \end{quote}
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   285
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   286
  \noindent
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   287
  and
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   288
  
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   289
  \begin{quote}
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   290
  \it%
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   291
  \begin{tabular}{@ {}l@ {}p{0.88\textwidth}@ {}}
93
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   292
  `` & If the reader finds the above treatment in terms of bit lists revoltingly
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   293
       concrete, I cannot disagree. A more abstract approach is clearly desirable.''
93
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   294
  \end{tabular}
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   295
  \end{quote}
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   296
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   297
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   298
  \noindent
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   299
  Moreover, it is not so clear how to conveniently impose a finiteness
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   300
  condition upon functions in order to represent \emph{finite} automata. The
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   301
  best is probably to resort to more advanced reasoning frameworks, such as
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   302
  \emph{locales} or \emph{type classes}, which are \emph{not} available in all
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   303
  HOL-based theorem provers.
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   304
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   305
  Because of these problems to do with representing automata, there seems to
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   306
  be no substantial formalisation of automata theory and regular languages
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   307
  carried out in HOL-based theorem provers. Nipkow \cite{Nipkow98} establishes
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   308
  the link between regular expressions and automata in the context of
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   309
  lexing. Berghofer and Reiter \cite{BerghoferReiter09} formalise automata
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   310
  working over bit strings in the context of Presburger arithmetic.  The only
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   311
  larger formalisations of automata theory are carried out in Nuprl
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   312
  \cite{Constable00} and in Coq \cite{Filliatre97}.
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 160
diff changeset
   313
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   314
  Also one might consider automata theory and regular languages as a well-worn
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   315
  stock subject where everything is crystal clear. However, paper proofs about
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   316
  automata often involve subtle side-conditions which are easily overlooked,
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   317
  but which make formal reasoning rather painful. For example Kozen's proof of
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   318
  the Myhill-Nerode theorem requires that automata do not have inaccessible
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   319
  states \cite{Kozen97}. Another subtle side-condition is completeness of
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   320
  automata, that is automata need to have total transition functions and at
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   321
  most one `sink' state from which there is no connection to a final state
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   322
  (Brzozowski mentions this side-condition in the context of state complexity
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   323
  of automata \cite{Brzozowski10}). Such side-conditions mean that if we
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   324
  define a regular language as one for which there exists \emph{a} finite
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
   325
  automaton that recognises all its strings (see Definition~\ref{baddef}), then we
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   326
  need a lemma which ensures that another equivalent one can be found
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   327
  satisfying the side-condition. Unfortunately, such `little' and `obvious'
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   328
  lemmas make a formalisation of automata theory a hair-pulling experience.
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
   329
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   330
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   331
  In this paper, we will not attempt to formalise automata theory in
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
   332
  Isabelle/HOL nor will we attempt to formalise automata proofs from the
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   333
  literature, but take a different approach to regular languages than is
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   334
  usually taken. Instead of defining a regular language as one where there
178
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   335
  exists an automaton that recognises all its strings, we define a
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   336
  regular language as:
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   337
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
   338
  \begin{dfntn}\label{regular}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
   339
  A language @{text A} is \emph{regular}, provided there is a regular expression 
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
   340
  that matches all strings of @{text "A"}.
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   341
  \end{dfntn}
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   342
  
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   343
  \noindent
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   344
  The reason is that regular expressions, unlike graphs, matrices and
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   345
  functions, can be easily defined as an inductive datatype. A reasoning
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   346
  infrastructure (like induction and recursion) comes then for free in
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   347
  HOL. Moreover, no side-conditions will be needed for regular expressions,
178
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   348
  like we need for automata. This convenience of regular expressions has
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   349
  recently been exploited in HOL4 with a formalisation of regular expression
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   350
  matching based on derivatives \cite{OwensSlind08} and with an equivalence
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   351
  checker for regular expressions in Isabelle/HOL \cite{KraussNipkow11}.  The
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   352
  main purpose of this paper is to show that a central result about regular
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   353
  languages---the Myhill-Nerode theorem---can be recreated by only using
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   354
  regular expressions. This theorem gives necessary and sufficient conditions
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   355
  for when a language is regular. As a corollary of this theorem we can easily
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   356
  establish the usual closure properties, including complementation, for
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   357
  regular languages.\medskip
61
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
   358
  
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
   359
  \noindent 
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   360
  {\bf Contributions:} There is an extensive literature on regular languages.
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   361
  To our best knowledge, our proof of the Myhill-Nerode theorem is the first
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   362
  that is based on regular expressions, only. The part of this theorem stating
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   363
  that finitely many partitions imply regularity of the language is proved by
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   364
  an argument about solving equational systems.  This argument appears to be
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   365
  folklore. For the other part, we give two proofs: one direct proof using
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   366
  certain tagging-functions, and another indirect proof using Antimirov's
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   367
  partial derivatives \cite{Antimirov95}. Again to our best knowledge, the
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   368
  tagging-functions have not been used before to establish the Myhill-Nerode
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   369
  theorem. Derivatives of regular expressions have been used recently quite
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   370
  widely in the literature; partial derivatives, in contrast, attract much
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   371
  less attention. However, partial derivatives are more suitable in the
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   372
  context of the Myhill-Nerode theorem, since it is easier to establish
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   373
  formally their finiteness result. We are not aware of any proof that uses
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   374
  either of them for proving the Myhill-Nerode theorem.
24
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   375
*}
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   376
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   377
section {* Preliminaries *}
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   378
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   379
text {*
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   380
  \noindent
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   381
  Strings in Isabelle/HOL are lists of characters with the \emph{empty string}
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   382
  being represented by the empty list, written @{term "[]"}.  We assume there
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   383
  are only finitely many different characters.  \emph{Languages} are sets of
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   384
  strings. The language containing all strings is written in Isabelle/HOL as
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   385
  @{term "UNIV::string set"}. The concatenation of two languages is written
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   386
  @{term "A \<cdot> B"} and a language raised to the power @{text n} is written
93
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   387
  @{term "A \<up> n"}. They are defined as usual
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   388
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   389
  \begin{center}
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   390
  \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   391
  @{thm (lhs) conc_def'[THEN eq_reflection, where A1="A" and B1="B"]}
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   392
  & @{text "\<equiv>"} & @{thm (rhs) conc_def'[THEN eq_reflection, where A1="A" and B1="B"]}\\
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   393
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   394
  @{thm (lhs) lang_pow.simps(1)[THEN eq_reflection, where A1="A"]}
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   395
  & @{text "\<equiv>"} & @{thm (rhs) lang_pow.simps(1)[THEN eq_reflection, where A1="A"]}\\
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   396
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   397
  @{thm (lhs) lang_pow.simps(2)[THEN eq_reflection, where A1="A" and n1="n"]}
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   398
  & @{text "\<equiv>"} & @{thm (rhs) lang_pow.simps(2)[THEN eq_reflection, where A1="A" and n1="n"]}
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   399
  \end{tabular}
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   400
  \end{center}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   401
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   402
  \noindent
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
   403
  where @{text "@"} is the list-append operation. The Kleene-star of a language @{text A}
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   404
  is defined as the union over all powers, namely @{thm star_def}. In the paper
88
1436fc451bb9 added something about Setalt and folds
urbanc
parents: 86
diff changeset
   405
  we will make use of the following properties of these constructions.
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   406
  
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   407
  \begin{prpstn}\label{langprops}\mbox{}\\
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   408
  \begin{tabular}{@ {}lp{10cm}}
180
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
   409
  (i)   & @{thm star_unfold_left}     \\ 
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
   410
  (ii)  & @{thm[mode=IfThen] pow_length}\\
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   411
  (iii) & @{thm conc_Union_left} \\ 
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   412
  (iv)  & If @{thm (prem 1) star_decom} and @{thm (prem 2) star_decom} then
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   413
          there exists an @{text "x\<^isub>p"} and @{text "x\<^isub>s"} with @{text "x = x\<^isub>p @ x\<^isub>s"} 
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   414
          and @{term "x\<^isub>p \<noteq> []"} such that @{term "x\<^isub>p \<in> A"} and @{term "x\<^isub>s \<in> A\<star>"}.
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   415
  \end{tabular}
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   416
  \end{prpstn}
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   417
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   418
  \noindent
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   419
  In @{text "(ii)"} we use the notation @{term "length s"} for the length of a
156
fd39492b187c a few more changes
urbanc
parents: 154
diff changeset
   420
  string; this property states that if \mbox{@{term "[] \<notin> A"}} then the lengths of
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   421
  the strings in @{term "A \<up> (Suc n)"} must be longer than @{text n}.  
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   422
  Property @{text "(iv)"} states that a non-empty string in @{term "A\<star>"} can
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   423
  always be split up into a non-empty prefix belonging to @{text "A"} and the
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   424
  rest being in @{term "A\<star>"}. We omit
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   425
  the proofs for these properties, but invite the reader to consult our
123
23c0e6f2929d polished everywhere...two cases still missing
urbanc
parents: 122
diff changeset
   426
  formalisation.\footnote{Available at \url{http://www4.in.tum.de/~urbanc/regexp.html}}
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   427
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   428
  The notation in Isabelle/HOL for the quotient of a language @{text A}
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   429
  according to an equivalence relation @{term REL} is @{term "A // REL"}. We
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   430
  will write @{text "\<lbrakk>x\<rbrakk>\<^isub>\<approx>"} for the equivalence class defined as
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   431
  \mbox{@{text "{y | y \<approx> x}"}}, and have @{text "x \<approx> y"} if and only if @{text
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   432
  "\<lbrakk>x\<rbrakk>\<^isub>\<approx> = \<lbrakk>y\<rbrakk>\<^isub>\<approx>"}.
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   433
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   434
51
6cfb92de4654 some tuning of the paper
urbanc
parents: 50
diff changeset
   435
  Central to our proof will be the solution of equational systems
176
6969de1eb96b latest version of the journal paper
urbanc
parents: 175
diff changeset
   436
  involving equivalence classes of languages. For this we will use Arden's Lemma 
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   437
  (see for example \cite[Page 100]{Sakarovitch09}),
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   438
  which solves equations of the form @{term "X = A \<cdot> X \<union> B"} provided
115
c5f138b5fc88 added comment from Larry
urbanc
parents: 114
diff changeset
   439
  @{term "[] \<notin> A"}. However we will need the following `reverse' 
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   440
  version of Arden's Lemma (`reverse' in the sense of changing the order of @{term "A \<cdot> X"} to
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   441
  \mbox{@{term "X \<cdot> A"}}).
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   442
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   443
  \begin{lmm}[Reverse Arden's Lemma]\label{arden}\mbox{}\\
86
6457e668dee5 tuned comments and names in Myhill_1
urbanc
parents: 83
diff changeset
   444
  If @{thm (prem 1) arden} then
115
c5f138b5fc88 added comment from Larry
urbanc
parents: 114
diff changeset
   445
  @{thm (lhs) arden} if and only if
86
6457e668dee5 tuned comments and names in Myhill_1
urbanc
parents: 83
diff changeset
   446
  @{thm (rhs) arden}.
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   447
  \end{lmm}
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   448
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   449
  \begin{proof}
86
6457e668dee5 tuned comments and names in Myhill_1
urbanc
parents: 83
diff changeset
   450
  For the right-to-left direction we assume @{thm (rhs) arden} and show
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   451
  that @{thm (lhs) arden} holds. From Property~\ref{langprops}@{text "(i)"} 
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   452
  we have @{term "A\<star> = A \<cdot> A\<star> \<union> {[]}"},
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   453
  which is equal to @{term "A\<star> = A\<star> \<cdot> A \<union> {[]}"}. Adding @{text B} to both 
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   454
  sides gives @{term "B \<cdot> A\<star> = B \<cdot> (A\<star> \<cdot> A \<union> {[]})"}, whose right-hand side
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   455
  is equal to @{term "(B \<cdot> A\<star>) \<cdot> A \<union> B"}. This completes this direction. 
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   456
86
6457e668dee5 tuned comments and names in Myhill_1
urbanc
parents: 83
diff changeset
   457
  For the other direction we assume @{thm (lhs) arden}. By a simple induction
51
6cfb92de4654 some tuning of the paper
urbanc
parents: 50
diff changeset
   458
  on @{text n}, we can establish the property
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   459
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   460
  \begin{center}
86
6457e668dee5 tuned comments and names in Myhill_1
urbanc
parents: 83
diff changeset
   461
  @{text "(*)"}\hspace{5mm} @{thm (concl) arden_helper}
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   462
  \end{center}
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   463
  
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   464
  \noindent
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   465
  Using this property we can show that @{term "B \<cdot> (A \<up> n) \<subseteq> X"} holds for
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   466
  all @{text n}. From this we can infer @{term "B \<cdot> A\<star> \<subseteq> X"} using the definition
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   467
  of @{text "\<star>"}.
51
6cfb92de4654 some tuning of the paper
urbanc
parents: 50
diff changeset
   468
  For the inclusion in the other direction we assume a string @{text s}
134
08afbed1c8c7 chunhan's comments
urbanc
parents: 133
diff changeset
   469
  with length @{text k} is an element in @{text X}. Since @{thm (prem 1) arden}
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   470
  we know by Property~\ref{langprops}@{text "(ii)"} that 
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   471
  @{term "s \<notin> X \<cdot> (A \<up> Suc k)"} since its length is only @{text k}
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   472
  (the strings in @{term "X \<cdot> (A \<up> Suc k)"} are all longer). 
53
da85feadb8e3 small typo
urbanc
parents: 52
diff changeset
   473
  From @{text "(*)"} it follows then that
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   474
  @{term s} must be an element in @{term "(\<Union>m\<in>{0..k}. B \<cdot> (A \<up> m))"}. This in turn
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   475
  implies that @{term s} is in @{term "(\<Union>n. B \<cdot> (A \<up> n))"}. Using Property~\ref{langprops}@{text "(iii)"} 
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
   476
  this is equal to @{term "B \<cdot> A\<star>"}, as we needed to show.
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   477
  \end{proof}
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   478
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   479
  \noindent
88
1436fc451bb9 added something about Setalt and folds
urbanc
parents: 86
diff changeset
   480
  Regular expressions are defined as the inductive datatype
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   481
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   482
  \begin{center}
176
6969de1eb96b latest version of the journal paper
urbanc
parents: 175
diff changeset
   483
  \begin{tabular}{rcl}
6969de1eb96b latest version of the journal paper
urbanc
parents: 175
diff changeset
   484
  @{text r} & @{text "::="} & @{term ZERO}\\
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   485
   & @{text"|"} & @{term One}\\ 
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   486
   & @{text"|"} & @{term "Atom c"}\\
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   487
   & @{text"|"} & @{term "Times r r"}\\
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   488
   & @{text"|"} & @{term "Plus r r"}\\
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   489
   & @{text"|"} & @{term "Star r"}
176
6969de1eb96b latest version of the journal paper
urbanc
parents: 175
diff changeset
   490
  \end{tabular}
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   491
  \end{center}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   492
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   493
  \noindent
88
1436fc451bb9 added something about Setalt and folds
urbanc
parents: 86
diff changeset
   494
  and the language matched by a regular expression is defined as
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   495
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   496
  \begin{center}
176
6969de1eb96b latest version of the journal paper
urbanc
parents: 175
diff changeset
   497
  \begin{tabular}{r@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   498
  @{thm (lhs) lang.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) lang.simps(1)}\\
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   499
  @{thm (lhs) lang.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) lang.simps(2)}\\
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   500
  @{thm (lhs) lang.simps(3)[where a="c"]} & @{text "\<equiv>"} & @{thm (rhs) lang.simps(3)[where a="c"]}\\
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   501
  @{thm (lhs) lang.simps(4)[where ?r="r\<^isub>1" and ?s="r\<^isub>2"]} & @{text "\<equiv>"} &
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   502
       @{thm (rhs) lang.simps(4)[where ?r="r\<^isub>1" and ?s="r\<^isub>2"]}\\
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   503
  @{thm (lhs) lang.simps(5)[where ?r="r\<^isub>1" and ?s="r\<^isub>2"]} & @{text "\<equiv>"} &
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   504
       @{thm (rhs) lang.simps(5)[where ?r="r\<^isub>1" and ?s="r\<^isub>2"]}\\
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   505
  @{thm (lhs) lang.simps(6)[where r="r"]} & @{text "\<equiv>"} &
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   506
      @{thm (rhs) lang.simps(6)[where r="r"]}\\
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   507
  \end{tabular}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   508
  \end{center}
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   509
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   510
  Given a finite set of regular expressions @{text rs}, we will make use of the operation of generating 
132
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
   511
  a regular expression that matches the union of all languages of @{text rs}. We only need to know the 
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
   512
  existence
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
   513
  of such a regular expression and therefore we use Isabelle/HOL's @{const "fold_graph"} and Hilbert's
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
   514
  @{text "\<epsilon>"} to define @{term "\<Uplus>rs"}. This operation, roughly speaking, folds @{const PLUS} over the 
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
   515
  set @{text rs} with @{const ZERO} for the empty set. We can prove that for a finite set @{text rs}
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   516
  %
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   517
  \begin{equation}\label{uplus}
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   518
  \mbox{@{thm (lhs) folds_alt_simp} @{text "= \<Union> (\<calL> ` rs)"}}
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   519
  \end{equation}
88
1436fc451bb9 added something about Setalt and folds
urbanc
parents: 86
diff changeset
   520
1436fc451bb9 added something about Setalt and folds
urbanc
parents: 86
diff changeset
   521
  \noindent
90
97b783438316 added an example
urbanc
parents: 89
diff changeset
   522
  holds, whereby @{text "\<calL> ` rs"} stands for the 
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   523
  image of the set @{text rs} under function @{text "\<calL>"} defined as
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   524
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   525
  \begin{center}
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   526
  @{term "lang ` rs \<equiv> {lang r | r. r \<in> rs}"}
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   527
  \end{center}
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   528
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   529
  \noindent
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   530
  In what follows we shall use this convenient short-hand notation for images of sets 
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   531
  also with other functions.
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   532
*}
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   533
132
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
   534
133
3ab755a96cef minor change
urbanc
parents: 132
diff changeset
   535
section {* The Myhill-Nerode Theorem, First Part *}
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   536
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   537
text {*
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   538
  \noindent
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
   539
  \footnote{Folklore: Henzinger (arden-DFA-regexp.pdf); Hofmann}
77
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   540
  The key definition in the Myhill-Nerode theorem is the
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   541
  \emph{Myhill-Nerode relation}, which states that w.r.t.~a language two 
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   542
  strings are related, provided there is no distinguishing extension in this
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 149
diff changeset
   543
  language. This can be defined as a tertiary relation.
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   544
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
   545
  \begin{dfntn}[Myhill-Nerode Relation]\label{myhillneroderel} 
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
   546
  Given a language @{text A}, two strings @{text x} and
123
23c0e6f2929d polished everywhere...two cases still missing
urbanc
parents: 122
diff changeset
   547
  @{text y} are Myhill-Nerode related provided
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
   548
  \begin{center}
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   549
  @{thm str_eq_def'}
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
   550
  \end{center}
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   551
  \end{dfntn}
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   552
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   553
  \noindent
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   554
  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
   555
  partitions the set of all strings, @{text "UNIV"}, into a set of disjoint
108
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   556
  equivalence classes. To illustrate this quotient construction, let us give a simple 
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   557
  example: consider the regular language containing just
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
   558
  the string @{text "[c]"}. The relation @{term "\<approx>({[c]})"} partitions @{text UNIV}
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   559
  into three equivalence classes @{text "X\<^isub>1"}, @{text "X\<^isub>2"} and  @{text "X\<^isub>3"}
90
97b783438316 added an example
urbanc
parents: 89
diff changeset
   560
  as follows
97b783438316 added an example
urbanc
parents: 89
diff changeset
   561
  
97b783438316 added an example
urbanc
parents: 89
diff changeset
   562
  \begin{center}
176
6969de1eb96b latest version of the journal paper
urbanc
parents: 175
diff changeset
   563
  \begin{tabular}{l}
6969de1eb96b latest version of the journal paper
urbanc
parents: 175
diff changeset
   564
  @{text "X\<^isub>1 = {[]}"}\\
6969de1eb96b latest version of the journal paper
urbanc
parents: 175
diff changeset
   565
  @{text "X\<^isub>2 = {[c]}"}\\
90
97b783438316 added an example
urbanc
parents: 89
diff changeset
   566
  @{text "X\<^isub>3 = UNIV - {[], [c]}"}
176
6969de1eb96b latest version of the journal paper
urbanc
parents: 175
diff changeset
   567
  \end{tabular}
90
97b783438316 added an example
urbanc
parents: 89
diff changeset
   568
  \end{center}
97b783438316 added an example
urbanc
parents: 89
diff changeset
   569
97b783438316 added an example
urbanc
parents: 89
diff changeset
   570
  One direction of the Myhill-Nerode theorem establishes 
93
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   571
  that if there are finitely many equivalence classes, like in the example above, then 
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   572
  the language is regular. In our setting we therefore have to show:
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   573
  
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   574
  \begin{thrm}\label{myhillnerodeone}
96
3b9deda4f459 simplified a bit the proof
urbanc
parents: 95
diff changeset
   575
  @{thm[mode=IfThen] Myhill_Nerode1}
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   576
  \end{thrm}
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   577
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   578
  \noindent
90
97b783438316 added an example
urbanc
parents: 89
diff changeset
   579
  To prove this theorem, we first define the set @{term "finals A"} as those equivalence
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   580
  classes from @{term "UNIV // \<approx>A"} that contain strings of @{text A}, namely
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   581
  %
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   582
  \begin{equation} 
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   583
  @{thm finals_def}
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   584
  \end{equation}
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   585
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   586
  \noindent
132
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
   587
  In our running example, @{text "X\<^isub>2"} is the only 
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
   588
  equivalence class in @{term "finals {[c]}"}.
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
   589
  It is straightforward to show that in general 
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
   590
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   591
  \begin{equation}\label{finalprops}
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
   592
  @{thm lang_is_union_of_finals}\hspace{15mm} 
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
   593
  @{thm finals_in_partitions} 
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   594
  \end{equation}
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
   595
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
   596
  \noindent
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
   597
  hold. 
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   598
  Therefore if we know that there exists a regular expression for every
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   599
  equivalence class in \mbox{@{term "finals A"}} (which by assumption must be
93
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   600
  a finite set), then we can use @{text "\<bigplus>"} to obtain a regular expression 
98
36f9d19be0e6 included comments by Xingyuan
urbanc
parents: 96
diff changeset
   601
  that matches every string in @{text A}.
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   602
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   603
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   604
  Our proof of Theorem~\ref{myhillnerodeone} relies on a method that can calculate a
79
bba9c80735f9 started to define things more directly
urbanc
parents: 77
diff changeset
   605
  regular expression for \emph{every} equivalence class, not just the ones 
77
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   606
  in @{term "finals A"}. We
93
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   607
  first define the notion of \emph{one-character-transition} between 
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   608
  two equivalence classes
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   609
  %
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   610
  \begin{equation} 
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   611
  @{thm transition_def}
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   612
  \end{equation}
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   613
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   614
  \noindent
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
   615
  which means that if we concatenate the character @{text c} to the end of all 
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
   616
  strings in the equivalence class @{text Y}, we obtain a subset of 
77
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   617
  @{text X}. Note that we do not define an automaton here, we merely relate two sets
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   618
  (with the help of a character). In our concrete example we have 
178
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   619
  @{term "X\<^isub>1 \<Turnstile>c\<Rightarrow> X\<^isub>2"}, @{term "X\<^isub>1 \<Turnstile>d\<^isub>i\<Rightarrow> X\<^isub>3"} with @{text "d\<^isub>i"} being any 
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   620
  other character than @{text c}, and @{term "X\<^isub>3 \<Turnstile>c\<^isub>j\<Rightarrow> X\<^isub>3"} for any 
194
5347d7556487 some typos
urbanc
parents: 193
diff changeset
   621
  character @{text "c\<^isub>j"}.
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   622
  
156
fd39492b187c a few more changes
urbanc
parents: 154
diff changeset
   623
  Next we construct an \emph{initial equational system} that
fd39492b187c a few more changes
urbanc
parents: 154
diff changeset
   624
  contains an equation for each equivalence class. We first give
fd39492b187c a few more changes
urbanc
parents: 154
diff changeset
   625
  an informal description of this construction. Suppose we have 
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   626
  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
   627
  contains the empty string @{text "[]"} (since equivalence classes are disjoint).
77
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   628
  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
   629
  
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   630
  \begin{center}
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   631
  \begin{tabular}{rcl}
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
   632
  @{text "X\<^isub>1"} & @{text "="} & @{text "(Y\<^isub>1\<^isub>1, ATOM c\<^isub>1\<^isub>1) + \<dots> + (Y\<^isub>1\<^isub>p, ATOM c\<^isub>1\<^isub>p) + \<lambda>(ONE)"} \\
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
   633
  @{text "X\<^isub>2"} & @{text "="} & @{text "(Y\<^isub>2\<^isub>1, ATOM c\<^isub>2\<^isub>1) + \<dots> + (Y\<^isub>2\<^isub>o, ATOM c\<^isub>2\<^isub>o)"} \\
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   634
  & $\vdots$ \\
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
   635
  @{text "X\<^isub>n"} & @{text "="} & @{text "(Y\<^isub>n\<^isub>1, ATOM c\<^isub>n\<^isub>1) + \<dots> + (Y\<^isub>n\<^isub>q, ATOM c\<^isub>n\<^isub>q)"}\\
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   636
  \end{tabular}
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   637
  \end{center}
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   638
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   639
  \noindent
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
   640
  where the terms @{text "(Y\<^isub>i\<^isub>j, ATOM c\<^isub>i\<^isub>j)"}
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   641
  stand for all transitions @{term "Y\<^isub>i\<^isub>j \<Turnstile>c\<^isub>i\<^isub>j\<Rightarrow>
159
990c12ab1562 edits; sqeezed to 16 pages
urbanc
parents: 157
diff changeset
   642
  X\<^isub>i"}. 
990c12ab1562 edits; sqeezed to 16 pages
urbanc
parents: 157
diff changeset
   643
  %The intuition behind the equational system is that every 
990c12ab1562 edits; sqeezed to 16 pages
urbanc
parents: 157
diff changeset
   644
  %equation @{text "X\<^isub>i = rhs\<^isub>i"} in this system
990c12ab1562 edits; sqeezed to 16 pages
urbanc
parents: 157
diff changeset
   645
  %corresponds roughly to a state of an automaton whose name is @{text X\<^isub>i} and its predecessor states 
990c12ab1562 edits; sqeezed to 16 pages
urbanc
parents: 157
diff changeset
   646
  %are the @{text "Y\<^isub>i\<^isub>j"}; the @{text "c\<^isub>i\<^isub>j"} are the labels of the transitions from these 
990c12ab1562 edits; sqeezed to 16 pages
urbanc
parents: 157
diff changeset
   647
  %predecessor states to @{text X\<^isub>i}. 
990c12ab1562 edits; sqeezed to 16 pages
urbanc
parents: 157
diff changeset
   648
  There can only be
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
   649
  finitely many terms of the form @{text "(Y\<^isub>i\<^isub>j, ATOM c\<^isub>i\<^isub>j)"} in a right-hand side 
156
fd39492b187c a few more changes
urbanc
parents: 154
diff changeset
   650
  since by assumption there are only finitely many
159
990c12ab1562 edits; sqeezed to 16 pages
urbanc
parents: 157
diff changeset
   651
  equivalence classes and only finitely many characters.
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
   652
  The term @{text "\<lambda>(ONE)"} in the first equation acts as a marker for the initial state, that
159
990c12ab1562 edits; sqeezed to 16 pages
urbanc
parents: 157
diff changeset
   653
  is the equivalence class
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   654
  containing @{text "[]"}.\footnote{Note that we mark, roughly speaking, the
115
c5f138b5fc88 added comment from Larry
urbanc
parents: 114
diff changeset
   655
  single `initial' state in the equational system, which is different from
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   656
  the method by Brzozowski \cite{Brzozowski64}, where he marks the
115
c5f138b5fc88 added comment from Larry
urbanc
parents: 114
diff changeset
   657
  `terminal' states. We are forced to set up the equational system in our
c5f138b5fc88 added comment from Larry
urbanc
parents: 114
diff changeset
   658
  way, because the Myhill-Nerode relation determines the `direction' of the
123
23c0e6f2929d polished everywhere...two cases still missing
urbanc
parents: 122
diff changeset
   659
  transitions---the successor `state' of an equivalence class @{text Y} can
23c0e6f2929d polished everywhere...two cases still missing
urbanc
parents: 122
diff changeset
   660
  be reached by adding a character to the end of @{text Y}. This is also the
156
fd39492b187c a few more changes
urbanc
parents: 154
diff changeset
   661
  reason why we have to use our reverse version of Arden's Lemma.}
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   662
  In our running example we have the initial equational system
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   663
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   664
  \begin{equation}\label{exmpcs}
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   665
  \mbox{\begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   666
  @{term "X\<^isub>1"} & @{text "="} & @{text "\<lambda>(ONE)"}\\
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   667
  @{term "X\<^isub>2"} & @{text "="} & @{text "(X\<^isub>1, ATOM c)"}\\
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   668
  @{term "X\<^isub>3"} & @{text "="} & @{text "(X\<^isub>1, ATOM d\<^isub>1) + \<dots> + (X\<^isub>1, ATOM d\<^isub>n)"}\\
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
   669
               & & \mbox{}\hspace{10mm}@{text "+ (X\<^isub>3, ATOM c\<^isub>1) + \<dots> + (X\<^isub>3, ATOM c\<^isub>m)"}
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   670
  \end{tabular}}
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   671
  \end{equation}
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   672
  
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   673
  \noindent
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   674
  where @{text "d\<^isub>1\<dots>d\<^isub>n"} is the sequence of all characters
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   675
  but not containing @{text c}, and @{text "c\<^isub>1\<dots>c\<^isub>m"} is the sequence of all
178
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   676
  characters.  
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   677
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   678
  Overloading the function @{text \<calL>} for the two kinds of terms in the
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
   679
  equational system, we have
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   680
  
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   681
  \begin{center}
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
   682
  @{text "\<calL>(Y, r) \<equiv>"} %
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   683
  @{thm (rhs) lang_trm.simps(2)[where X="Y" and r="r", THEN eq_reflection]}\hspace{10mm}
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   684
  @{thm lang_trm.simps(1)[where r="r", THEN eq_reflection]}
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   685
  \end{center}
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   686
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   687
  \noindent
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   688
  and we can prove for @{text "X\<^isub>2\<^isub>.\<^isub>.\<^isub>n"} that the following equations
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   689
  %
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   690
  \begin{equation}\label{inv1}
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
   691
  @{text "X\<^isub>i = \<calL>(Y\<^isub>i\<^isub>1, ATOM c\<^isub>i\<^isub>1) \<union> \<dots> \<union> \<calL>(Y\<^isub>i\<^isub>q, ATOM c\<^isub>i\<^isub>q)"}.
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   692
  \end{equation}
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   693
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   694
  \noindent
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   695
  hold. Similarly for @{text "X\<^isub>1"} we can show the following equation
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   696
  %
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   697
  \begin{equation}\label{inv2}
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
   698
  @{text "X\<^isub>1 = \<calL>(Y\<^isub>1\<^isub>1, ATOM c\<^isub>1\<^isub>1) \<union> \<dots> \<union> \<calL>(Y\<^isub>1\<^isub>p, ATOM c\<^isub>1\<^isub>p) \<union> \<calL>(\<lambda>(ONE))"}.
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   699
  \end{equation}
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   700
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   701
  \noindent
160
ea2e5acbfe4a added comments from Chunhan
urbanc
parents: 159
diff changeset
   702
  holds. The reason for adding the @{text \<lambda>}-marker to our initial equational system is 
103
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
   703
  to obtain this equation: it only holds with the marker, since none of 
108
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   704
  the other terms contain the empty string. The point of the initial equational system is
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   705
  that solving it means we will be able to extract a regular expression for every equivalence class. 
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   706
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   707
  Our representation for the equations in Isabelle/HOL are pairs,
108
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   708
  where the first component is an equivalence class (a set of strings)
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   709
  and the second component
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   710
  is a set of terms. Given a set of equivalence
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   711
  classes @{text CS}, our initial equational system @{term "Init CS"} is thus 
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   712
  formally defined as
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   713
  %
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   714
  \begin{equation}\label{initcs}
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   715
  \mbox{\begin{tabular}{rcl}     
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   716
  @{thm (lhs) Init_rhs_def} & @{text "\<equiv>"} & 
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   717
  @{text "if"}~@{term "[] \<in> X"}\\
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
   718
  & & @{text "then"}~@{term "{Trn Y (ATOM c) | Y c. Y \<in> CS \<and> Y \<Turnstile>c\<Rightarrow> X} \<union> {Lam ONE}"}\\
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
   719
  & & @{text "else"}~@{term "{Trn Y (ATOM c)| Y c. Y \<in> CS \<and> Y \<Turnstile>c\<Rightarrow> X}"}\\
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   720
  @{thm (lhs) Init_def}     & @{text "\<equiv>"} & @{thm (rhs) Init_def}
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   721
  \end{tabular}}
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   722
  \end{equation}
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   723
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   724
  
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   725
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   726
  \noindent
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   727
  Because we use sets of terms 
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   728
  for representing the right-hand sides of equations, we can 
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   729
  prove \eqref{inv1} and \eqref{inv2} more concisely as
93
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   730
  %
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   731
  \begin{lmm}\label{inv}
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   732
  If @{thm (prem 1) test} then @{text "X = \<Union> \<calL> ` rhs"}.
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   733
  \end{lmm}
77
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   734
93
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   735
  \noindent
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   736
  Our proof of Theorem~\ref{myhillnerodeone} will proceed by transforming the
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   737
  initial equational system into one in \emph{solved form} maintaining the invariant
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   738
  in Lemma~\ref{inv}. From the solved form we will be able to read
89
42af13d194c9 a bit more on the paper
urbanc
parents: 88
diff changeset
   739
  off the regular expressions. 
42af13d194c9 a bit more on the paper
urbanc
parents: 88
diff changeset
   740
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   741
  In order to transform an equational system into solved form, we have two 
89
42af13d194c9 a bit more on the paper
urbanc
parents: 88
diff changeset
   742
  operations: one that takes an equation of the form @{text "X = rhs"} and removes
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   743
  any recursive occurrences of @{text X} in the @{text rhs} using our variant of Arden's 
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
   744
  Lemma. The other operation takes an equation @{text "X = rhs"}
89
42af13d194c9 a bit more on the paper
urbanc
parents: 88
diff changeset
   745
  and substitutes @{text X} throughout the rest of the equational system
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   746
  adjusting the remaining regular expressions appropriately. To define this adjustment 
108
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   747
  we define the \emph{append-operation} taking a term and a regular expression as argument
89
42af13d194c9 a bit more on the paper
urbanc
parents: 88
diff changeset
   748
42af13d194c9 a bit more on the paper
urbanc
parents: 88
diff changeset
   749
  \begin{center}
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   750
  \begin{tabular}{r@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   751
  @{thm (lhs) Append_rexp.simps(2)[where X="Y" and r="r\<^isub>1" and rexp="r\<^isub>2", THEN eq_reflection]}
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   752
  & @{text "\<equiv>"} & 
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   753
  @{thm (rhs) Append_rexp.simps(2)[where X="Y" and r="r\<^isub>1" and rexp="r\<^isub>2", THEN eq_reflection]}\\
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   754
  @{thm (lhs) Append_rexp.simps(1)[where r="r\<^isub>1" and rexp="r\<^isub>2", THEN eq_reflection]}
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   755
  & @{text "\<equiv>"} & 
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   756
  @{thm (rhs) Append_rexp.simps(1)[where r="r\<^isub>1" and rexp="r\<^isub>2", THEN eq_reflection]}
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   757
  \end{tabular}
89
42af13d194c9 a bit more on the paper
urbanc
parents: 88
diff changeset
   758
  \end{center}
42af13d194c9 a bit more on the paper
urbanc
parents: 88
diff changeset
   759
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
   760
  \noindent
108
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   761
  We lift this operation to entire right-hand sides of equations, written as
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 160
diff changeset
   762
  @{thm (lhs) Append_rexp_rhs_def[where rexp="r"]}. With this we can define
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   763
  the \emph{arden-operation} for an equation of the form @{text "X = rhs"} as:
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   764
  %
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   765
  \begin{equation}\label{arden_def}
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   766
  \mbox{\begin{tabular}{rc@ {\hspace{2mm}}r@ {\hspace{1mm}}l}
94
5b12cd0a3b3c latest on the paper
urbanc
parents: 93
diff changeset
   767
  @{thm (lhs) Arden_def} & @{text "\<equiv>"}~~\mbox{} & \multicolumn{2}{@ {\hspace{-2mm}}l}{@{text "let"}}\\ 
5b12cd0a3b3c latest on the paper
urbanc
parents: 93
diff changeset
   768
   & & @{text "rhs' ="} & @{term "rhs - {Trn X r | r. Trn X r \<in> rhs}"} \\
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   769
   & & @{text "r' ="}   & @{term "Star (\<Uplus> {r. Trn X r \<in> rhs})"}\\
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   770
   & &  \multicolumn{2}{@ {\hspace{-2mm}}l}{@{text "in"}~~@{term "Append_rexp_rhs rhs' r'"}}\\ 
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   771
  \end{tabular}}
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   772
  \end{equation}
93
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   773
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   774
  \noindent
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   775
  In this definition, we first delete all terms of the form @{text "(X, r)"} from @{text rhs};
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   776
  then we calculate the combined regular expressions for all @{text r} coming 
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   777
  from the deleted @{text "(X, r)"}, and take the @{const Star} of it;
178
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   778
  finally we append this regular expression to @{text rhs'}. If we apply this
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   779
  operation to the right-hand side of @{text "X\<^isub>3"} in \eqref{exmpcs}, we obtain
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   780
  the equation:
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   781
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   782
  \begin{center}
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   783
  \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   784
  @{term "X\<^isub>3"} & @{text "="} & 
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
   785
    @{text "(X\<^isub>1, TIMES (ATOM d\<^isub>1) (STAR \<^raw:\ensuremath{\bigplus}>{ATOM c\<^isub>1,\<dots>, ATOM c\<^isub>m})) + \<dots> "}\\
178
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   786
  & & \mbox{}\hspace{13mm}
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
   787
     @{text "\<dots> + (X\<^isub>1, TIMES (ATOM d\<^isub>n) (STAR \<^raw:\ensuremath{\bigplus}>{ATOM c\<^isub>1,\<dots>, ATOM c\<^isub>m}))"}
178
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   788
  \end{tabular}
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   789
  \end{center}
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   790
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   791
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   792
  \noindent
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   793
  That means we eliminated the dependency of @{text "X\<^isub>3"} on the
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   794
  right-hand side.  Note we used the abbreviation 
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
   795
  @{text "\<^raw:\ensuremath{\bigplus}>{ATOM c\<^isub>1,\<dots>, ATOM c\<^isub>m}"} 
178
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   796
  to stand for a regular expression that matches with every character. In 
183
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
   797
  our algorithm we are only interested in the existence of such a regular expression
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
   798
  and do not specify it any further. 
178
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   799
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   800
  It can be easily seen that the @{text "Arden"}-operation mimics Arden's
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   801
  Lemma on the level of equations. To ensure the non-emptiness condition of
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   802
  Arden's Lemma we say that a right-hand side is @{text ardenable} provided
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   803
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   804
  \begin{center}
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   805
  @{thm ardenable_def}
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   806
  \end{center}
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   807
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   808
  \noindent
156
fd39492b187c a few more changes
urbanc
parents: 154
diff changeset
   809
  This allows us to prove a version of Arden's Lemma on the level of equations.
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   810
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   811
  \begin{lmm}\label{ardenable}
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
   812
  Given an equation @{text "X = rhs"}.
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   813
  If @{text "X = \<Union>\<calL> ` rhs"},
179
edacc141060f small improvements
urbanc
parents: 178
diff changeset
   814
  @{thm (prem 2) Arden_preserves_soundness}, and
edacc141060f small improvements
urbanc
parents: 178
diff changeset
   815
  @{thm (prem 3) Arden_preserves_soundness}, then
135
604518f0127f final polished
urbanc
parents: 134
diff changeset
   816
  @{text "X = \<Union>\<calL> ` (Arden X rhs)"}.
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   817
  \end{lmm}
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   818
  
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   819
  \noindent
156
fd39492b187c a few more changes
urbanc
parents: 154
diff changeset
   820
  Our @{text ardenable} condition is slightly stronger than needed for applying Arden's Lemma,
194
5347d7556487 some typos
urbanc
parents: 193
diff changeset
   821
  but we can still ensure that it holds throughout our algorithm of transforming equations
156
fd39492b187c a few more changes
urbanc
parents: 154
diff changeset
   822
  into solved form. The \emph{substitution-operation} takes an equation
95
9540c2f2ea77 more things
urbanc
parents: 94
diff changeset
   823
  of the form @{text "X = xrhs"} and substitutes it into the right-hand side @{text rhs}.
94
5b12cd0a3b3c latest on the paper
urbanc
parents: 93
diff changeset
   824
5b12cd0a3b3c latest on the paper
urbanc
parents: 93
diff changeset
   825
  \begin{center}
95
9540c2f2ea77 more things
urbanc
parents: 94
diff changeset
   826
  \begin{tabular}{rc@ {\hspace{2mm}}r@ {\hspace{1mm}}l}
9540c2f2ea77 more things
urbanc
parents: 94
diff changeset
   827
  @{thm (lhs) Subst_def} & @{text "\<equiv>"}~~\mbox{} & \multicolumn{2}{@ {\hspace{-2mm}}l}{@{text "let"}}\\ 
9540c2f2ea77 more things
urbanc
parents: 94
diff changeset
   828
   & & @{text "rhs' ="} & @{term "rhs - {Trn X r | r. Trn X r \<in> rhs}"} \\
9540c2f2ea77 more things
urbanc
parents: 94
diff changeset
   829
   & & @{text "r' ="}   & @{term "\<Uplus> {r. Trn X r \<in> rhs}"}\\
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   830
   & &  \multicolumn{2}{@ {\hspace{-2mm}}l}{@{text "in"}~~@{term "rhs' \<union> Append_rexp_rhs xrhs r'"}}\\ 
95
9540c2f2ea77 more things
urbanc
parents: 94
diff changeset
   831
  \end{tabular}
94
5b12cd0a3b3c latest on the paper
urbanc
parents: 93
diff changeset
   832
  \end{center}
95
9540c2f2ea77 more things
urbanc
parents: 94
diff changeset
   833
9540c2f2ea77 more things
urbanc
parents: 94
diff changeset
   834
  \noindent
134
08afbed1c8c7 chunhan's comments
urbanc
parents: 133
diff changeset
   835
  We again delete first all occurrences of @{text "(X, r)"} in @{text rhs}; we then calculate
95
9540c2f2ea77 more things
urbanc
parents: 94
diff changeset
   836
  the regular expression corresponding to the deleted terms; finally we append this
9540c2f2ea77 more things
urbanc
parents: 94
diff changeset
   837
  regular expression to @{text "xrhs"} and union it up with @{text rhs'}. When we use
9540c2f2ea77 more things
urbanc
parents: 94
diff changeset
   838
  the substitution operation we will arrange it so that @{text "xrhs"} does not contain
178
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   839
  any occurrence of @{text X}. For example substituting the first equation in
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   840
  \eqref{exmpcs} into the right-hand side of the second, thus eliminating the equivalence 
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   841
  class @{text "X\<^isub>1"}, gives us the equation
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   842
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   843
  \begin{equation}\label{exmpresult}
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   844
  \mbox{\begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   845
  @{term "X\<^isub>2"} & @{text "="} & @{text "\<lambda>(TIMES ONE (ATOM c))"}\\
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   846
  \end{tabular}}
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   847
  \end{equation}
96
3b9deda4f459 simplified a bit the proof
urbanc
parents: 95
diff changeset
   848
134
08afbed1c8c7 chunhan's comments
urbanc
parents: 133
diff changeset
   849
  With these two operations in place, we can define the operation that removes one equation
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   850
  from an equational systems @{text ES}. The operation @{const Subst_all}
96
3b9deda4f459 simplified a bit the proof
urbanc
parents: 95
diff changeset
   851
  substitutes an equation @{text "X = xrhs"} throughout an equational system @{text ES}; 
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   852
  @{const Remove} then completely removes such an equation from @{text ES} by substituting 
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   853
  it to the rest of the equational system, but first eliminating all recursive occurrences
96
3b9deda4f459 simplified a bit the proof
urbanc
parents: 95
diff changeset
   854
  of @{text X} by applying @{const Arden} to @{text "xrhs"}.
3b9deda4f459 simplified a bit the proof
urbanc
parents: 95
diff changeset
   855
3b9deda4f459 simplified a bit the proof
urbanc
parents: 95
diff changeset
   856
  \begin{center}
3b9deda4f459 simplified a bit the proof
urbanc
parents: 95
diff changeset
   857
  \begin{tabular}{rcl}
3b9deda4f459 simplified a bit the proof
urbanc
parents: 95
diff changeset
   858
  @{thm (lhs) Subst_all_def} & @{text "\<equiv>"} & @{thm (rhs) Subst_all_def}\\
3b9deda4f459 simplified a bit the proof
urbanc
parents: 95
diff changeset
   859
  @{thm (lhs) Remove_def}    & @{text "\<equiv>"} & @{thm (rhs) Remove_def}
3b9deda4f459 simplified a bit the proof
urbanc
parents: 95
diff changeset
   860
  \end{tabular}
3b9deda4f459 simplified a bit the proof
urbanc
parents: 95
diff changeset
   861
  \end{center}
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   862
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   863
  \noindent
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   864
  Finally, we can define how an equational system should be solved. For this 
107
6f4f9b7b9891 updated paper
urbanc
parents: 106
diff changeset
   865
  we will need to iterate the process of eliminating equations until only one equation
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 149
diff changeset
   866
  will be left in the system. However, we do not just want to have any equation
107
6f4f9b7b9891 updated paper
urbanc
parents: 106
diff changeset
   867
  as being the last one, but the one involving the equivalence class for 
6f4f9b7b9891 updated paper
urbanc
parents: 106
diff changeset
   868
  which we want to calculate the regular 
108
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   869
  expression. Let us suppose this equivalence class is @{text X}. 
107
6f4f9b7b9891 updated paper
urbanc
parents: 106
diff changeset
   870
  Since @{text X} is the one to be solved, in every iteration step we have to pick an
108
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   871
  equation to be eliminated that is different from @{text X}. In this way 
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   872
  @{text X} is kept to the final step. The choice is implemented using Hilbert's choice 
107
6f4f9b7b9891 updated paper
urbanc
parents: 106
diff changeset
   873
  operator, written @{text SOME} in the definition below.
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   874
  
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   875
  \begin{center}
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   876
  \begin{tabular}{rc@ {\hspace{4mm}}r@ {\hspace{1mm}}l}
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   877
  @{thm (lhs) Iter_def} & @{text "\<equiv>"}~~\mbox{} & \multicolumn{2}{@ {\hspace{-4mm}}l}{@{text "let"}}\\ 
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   878
   & & @{text "(Y, yrhs) ="} & @{term "SOME (Y, yrhs). (Y, yrhs) \<in> ES \<and> X \<noteq> Y"} \\
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   879
   & &  \multicolumn{2}{@ {\hspace{-4mm}}l}{@{text "in"}~~@{term "Remove ES Y yrhs"}}\\ 
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   880
  \end{tabular}
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   881
  \end{center}
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   882
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   883
  \noindent
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   884
  The last definition we need applies @{term Iter} over and over until a condition 
159
990c12ab1562 edits; sqeezed to 16 pages
urbanc
parents: 157
diff changeset
   885
  @{text Cond} is \emph{not} satisfied anymore. This condition states that there
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   886
  are more than one equation left in the equational system @{text ES}. To solve
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   887
  an equational system we use Isabelle/HOL's @{text while}-operator as follows:
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   888
  
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   889
  \begin{center}
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   890
  @{thm Solve_def}
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   891
  \end{center}
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   892
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   893
  \noindent
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   894
  We are not concerned here with the definition of this operator (see
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   895
  Berghofer and Nipkow \cite{BerghoferNipkow00} for example), but note that we
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   896
  eliminate in each @{const Iter}-step a single equation, and therefore have a
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   897
  well-founded termination order by taking the cardinality of the equational
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   898
  system @{text ES}. This enables us to prove properties about our definition
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   899
  of @{const Solve} when we `call' it with the equivalence class @{text X} and
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   900
  the initial equational system @{term "Init (UNIV // \<approx>A)"} from
108
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   901
  \eqref{initcs} using the principle:
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   902
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   903
  \begin{equation}\label{whileprinciple}
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   904
  \mbox{\begin{tabular}{l}
103
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
   905
  @{term "invariant (Init (UNIV // \<approx>A))"} \\
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
   906
  @{term "\<forall>ES. invariant ES \<and> Cond ES \<longrightarrow> invariant (Iter X ES)"}\\
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
   907
  @{term "\<forall>ES. invariant ES \<and> Cond ES \<longrightarrow> card (Iter X ES) < card ES"}\\
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
   908
  @{term "\<forall>ES. invariant ES \<and> \<not> Cond ES \<longrightarrow> P ES"}\\
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
   909
  \hline
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
   910
  \multicolumn{1}{c}{@{term "P (Solve X (Init (UNIV // \<approx>A)))"}}
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   911
  \end{tabular}}
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   912
  \end{equation}
103
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
   913
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
   914
  \noindent
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   915
  This principle states that given an invariant (which we will specify below) 
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   916
  we can prove a property
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   917
  @{text "P"} involving @{const Solve}. For this we have to discharge the following
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   918
  proof obligations: first the
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
   919
  initial equational system satisfies the invariant; second the iteration
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 149
diff changeset
   920
  step @{text "Iter"} preserves the invariant as long as the condition @{term Cond} holds;
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
   921
  third @{text "Iter"} decreases the termination order, and fourth that
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   922
  once the condition does not hold anymore then the property @{text P} must hold.
103
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
   923
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   924
  The property @{term P} in our proof will state that @{term "Solve X (Init (UNIV // \<approx>A))"}
108
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   925
  returns with a single equation @{text "X = xrhs"} for some @{text "xrhs"}, and
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   926
  that this equational system still satisfies the invariant. In order to get
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   927
  the proof through, the invariant is composed of the following six properties:
103
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
   928
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
   929
  \begin{center}
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   930
  \begin{tabular}{@ {}rcl@ {\hspace{-13mm}}l @ {}}
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   931
  @{text "invariant ES"} & @{text "\<equiv>"} &
103
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
   932
      @{term "finite ES"} & @{text "(finiteness)"}\\
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
   933
  & @{text "\<and>"} & @{thm (rhs) finite_rhs_def} & @{text "(finiteness rhs)"}\\
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   934
  & @{text "\<and>"} & @{text "\<forall>(X, rhs)\<in>ES. X = \<Union>\<calL> ` rhs"} & @{text "(soundness)"}\\
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 160
diff changeset
   935
  & @{text "\<and>"} & @{thm (rhs) distinctness_def}\\
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   936
  &             &  & @{text "(distinctness)"}\\
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   937
  & @{text "\<and>"} & @{thm (rhs) ardenable_all_def} & @{text "(ardenable)"}\\   
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 160
diff changeset
   938
  & @{text "\<and>"} & @{thm (rhs) validity_def} & @{text "(validity)"}\\
103
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
   939
  \end{tabular}
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
   940
  \end{center}
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
   941
 
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   942
  \noindent
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   943
  The first two ensure that the equational system is always finite (number of equations
160
ea2e5acbfe4a added comments from Chunhan
urbanc
parents: 159
diff changeset
   944
  and number of terms in each equation); the third makes sure the `meaning' of the 
108
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   945
  equations is preserved under our transformations. The other properties are a bit more
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   946
  technical, but are needed to get our proof through. Distinctness states that every
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 149
diff changeset
   947
  equation in the system is distinct. @{text Ardenable} ensures that we can always
156
fd39492b187c a few more changes
urbanc
parents: 154
diff changeset
   948
  apply the @{text Arden} operation. 
108
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   949
  The last property states that every @{text rhs} can only contain equivalence classes
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   950
  for which there is an equation. Therefore @{text lhss} is just the set containing 
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   951
  the first components of an equational system,
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   952
  while @{text "rhss"} collects all equivalence classes @{text X} in the terms of the 
123
23c0e6f2929d polished everywhere...two cases still missing
urbanc
parents: 122
diff changeset
   953
  form @{term "Trn X r"}. That means formally @{thm (lhs) lhss_def}~@{text "\<equiv> {X | (X, rhs) \<in> ES}"} 
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   954
  and @{thm (lhs) rhss_def}~@{text "\<equiv> {X | (X, r) \<in> rhs}"}.
108
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   955
  
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   956
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   957
  It is straightforward to prove that the initial equational system satisfies the
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
   958
  invariant.
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
   959
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   960
  \begin{lmm}\label{invzero}
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   961
  @{thm[mode=IfThen] Init_ES_satisfies_invariant}
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   962
  \end{lmm}
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   963
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
   964
  \begin{proof}
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
   965
  Finiteness is given by the assumption and the way how we set up the 
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   966
  initial equational system. Soundness is proved in Lemma~\ref{inv}. Distinctness
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 149
diff changeset
   967
  follows from the fact that the equivalence classes are disjoint. The @{text ardenable}
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
   968
  property also follows from the setup of the initial equational system, as does 
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
   969
  validity.
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
   970
  \end{proof}
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
   971
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
   972
  \noindent
ec774952190c polished everything
urbanc
parents: 112
diff changeset
   973
  Next we show that @{text Iter} preserves the invariant.
ec774952190c polished everything
urbanc
parents: 112
diff changeset
   974
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   975
  \begin{lmm}\label{iterone}
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   976
  @{thm[mode=IfThen] iteration_step_invariant[where xrhs="rhs"]}
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   977
  \end{lmm}
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   978
107
6f4f9b7b9891 updated paper
urbanc
parents: 106
diff changeset
   979
  \begin{proof} 
156
fd39492b187c a few more changes
urbanc
parents: 154
diff changeset
   980
  The argument boils down to choosing an equation @{text "Y = yrhs"} to be eliminated
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   981
  and to show that @{term "Subst_all (ES - {(Y, yrhs)}) Y (Arden Y yrhs)"} 
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   982
  preserves the invariant.
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   983
  We prove this as follows:
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   984
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   985
  \begin{center}
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   986
  \begin{tabular}{@ {}l@ {}}
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   987
  @{text "\<forall> ES."}\\ \mbox{}\hspace{5mm}@{thm (prem 1) Subst_all_satisfies_invariant} implies
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   988
  @{thm (concl) Subst_all_satisfies_invariant}
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   989
  \end{tabular}
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   990
  \end{center}
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   991
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   992
  \noindent
156
fd39492b187c a few more changes
urbanc
parents: 154
diff changeset
   993
  Finiteness is straightforward, as the @{const Subst} and @{const Arden} operations 
116
342983676c8f included comments by Chunhan
urbanc
parents: 115
diff changeset
   994
  keep the equational system finite. These operations also preserve soundness 
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   995
  and distinctness (we proved soundness for @{const Arden} in Lemma~\ref{ardenable}).
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 149
diff changeset
   996
  The property @{text ardenable} is clearly preserved because the append-operation
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   997
  cannot make a regular expression to match the empty string. Validity is
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   998
  given because @{const Arden} removes an equivalence class from @{text yrhs}
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   999
  and then @{const Subst_all} removes @{text Y} from the equational system.
132
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
  1000
  Having proved the implication above, we can instantiate @{text "ES"} with @{text "ES - {(Y, yrhs)}"}
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1001
  which matches with our proof-obligation of @{const "Subst_all"}. Since
132
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
  1002
  \mbox{@{term "ES = ES - {(Y, yrhs)} \<union> {(Y, yrhs)}"}}, we can use the assumption 
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1003
  to complete the proof.
107
6f4f9b7b9891 updated paper
urbanc
parents: 106
diff changeset
  1004
  \end{proof}
6f4f9b7b9891 updated paper
urbanc
parents: 106
diff changeset
  1005
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1006
  \noindent
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1007
  We also need the fact that @{text Iter} decreases the termination measure.
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1008
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
  1009
  \begin{lmm}\label{itertwo}
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1010
  @{thm[mode=IfThen] iteration_step_measure[simplified (no_asm), where xrhs="rhs"]}
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
  1011
  \end{lmm}
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1012
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1013
  \begin{proof}
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1014
  By assumption we know that @{text "ES"} is finite and has more than one element.
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1015
  Therefore there must be an element @{term "(Y, yrhs) \<in> ES"} with 
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1016
  @{term "(Y, yrhs) \<noteq> (X, rhs)"}. Using the distinctness property we can infer
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1017
  that @{term "Y \<noteq> X"}. We further know that @{text "Remove ES Y yrhs"}
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1018
  removes the equation @{text "Y = yrhs"} from the system, and therefore 
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1019
  the cardinality of @{const Iter} strictly decreases.
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1020
  \end{proof}
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1021
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1022
  \noindent
134
08afbed1c8c7 chunhan's comments
urbanc
parents: 133
diff changeset
  1023
  This brings us to our property we want to establish for @{text Solve}.
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1024
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1025
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
  1026
  \begin{lmm}
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1027
  If @{thm (prem 1) Solve} and @{thm (prem 2) Solve} then there exists
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1028
  a @{text rhs} such that  @{term "Solve X (Init (UNIV // \<approx>A)) = {(X, rhs)}"}
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1029
  and @{term "invariant {(X, rhs)}"}.
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
  1030
  \end{lmm}
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1031
107
6f4f9b7b9891 updated paper
urbanc
parents: 106
diff changeset
  1032
  \begin{proof} 
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1033
  In order to prove this lemma using \eqref{whileprinciple}, we have to use a slightly
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1034
  stronger invariant since Lemma~\ref{iterone} and \ref{itertwo} have the precondition 
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1035
  that @{term "(X, rhs) \<in> ES"} for some @{text rhs}. This precondition is needed
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1036
  in order to choose in the @{const Iter}-step an equation that is not \mbox{@{term "X = rhs"}}.
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1037
  Therefore our invariant cannot be just @{term "invariant ES"}, but must be 
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1038
  @{term "invariant ES \<and> (\<exists>rhs. (X, rhs) \<in> ES)"}. By assumption 
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1039
  @{thm (prem 2) Solve} and Lemma~\ref{invzero}, the more general invariant holds for
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1040
  the initial equational system. This is premise 1 of~\eqref{whileprinciple}.
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1041
  Premise 2 is given by Lemma~\ref{iterone} and the fact that @{const Iter} might
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1042
  modify the @{text rhs} in the equation @{term "X = rhs"}, but does not remove it.
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1043
  Premise 3 of~\eqref{whileprinciple} is by Lemma~\ref{itertwo}. Now in premise 4
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1044
  we like to show that there exists a @{text rhs} such that @{term "ES = {(X, rhs)}"}
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1045
  and that @{text "invariant {(X, rhs)}"} holds, provided the condition @{text "Cond"}
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1046
  does not holds. By the stronger invariant we know there exists such a @{text "rhs"}
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1047
  with @{term "(X, rhs) \<in> ES"}. Because @{text Cond} is not true, we know the cardinality
123
23c0e6f2929d polished everywhere...two cases still missing
urbanc
parents: 122
diff changeset
  1048
  of @{text ES} is @{text 1}. This means @{text "ES"} must actually be the set @{text "{(X, rhs)}"},
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1049
  for which the invariant holds. This allows us to conclude that 
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1050
  @{term "Solve X (Init (UNIV // \<approx>A)) = {(X, rhs)}"} and @{term "invariant {(X, rhs)}"} hold,
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1051
  as needed.
107
6f4f9b7b9891 updated paper
urbanc
parents: 106
diff changeset
  1052
  \end{proof}
6f4f9b7b9891 updated paper
urbanc
parents: 106
diff changeset
  1053
106
91dc591de63f updated paper
urbanc
parents: 105
diff changeset
  1054
  \noindent
91dc591de63f updated paper
urbanc
parents: 105
diff changeset
  1055
  With this lemma in place we can show that for every equivalence class in @{term "UNIV // \<approx>A"}
91dc591de63f updated paper
urbanc
parents: 105
diff changeset
  1056
  there exists a regular expression.
91dc591de63f updated paper
urbanc
parents: 105
diff changeset
  1057
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
  1058
  \begin{lmm}\label{every_eqcl_has_reg}
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1059
  @{thm[mode=IfThen] every_eqcl_has_reg}
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
  1060
  \end{lmm}
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1061
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1062
  \begin{proof}
138
2dfe13bc1443 three typos
urbanc
parents: 137
diff changeset
  1063
  By the preceding lemma, we know that there exists a @{text "rhs"} such
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1064
  that @{term "Solve X (Init (UNIV // \<approx>A))"} returns the equation @{text "X = rhs"},
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1065
  and that the invariant holds for this equation. That means we 
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1066
  know @{text "X = \<Union>\<calL> ` rhs"}. We further know that
109
79b37ef9505f minor updated
urbanc
parents: 108
diff changeset
  1067
  this is equal to \mbox{@{text "\<Union>\<calL> ` (Arden X rhs)"}} using the properties of the 
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1068
  invariant and Lemma~\ref{ardenable}. Using the validity property for the equation @{text "X = rhs"},
156
fd39492b187c a few more changes
urbanc
parents: 154
diff changeset
  1069
  we can infer that @{term "rhss rhs \<subseteq> {X}"} and because the @{text Arden} operation
106
91dc591de63f updated paper
urbanc
parents: 105
diff changeset
  1070
  removes that @{text X} from @{text rhs}, that @{term "rhss (Arden X rhs) = {}"}.
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1071
  This means the right-hand side @{term "Arden X rhs"} can only consist of terms of the form @{term "Lam r"}.
176
6969de1eb96b latest version of the journal paper
urbanc
parents: 175
diff changeset
  1072
  So we can collect those (finitely many) regular expressions @{text rs} and have @{term "X = lang (\<Uplus>rs)"}.
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1073
  With this we can conclude the proof.
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1074
  \end{proof}
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1075
106
91dc591de63f updated paper
urbanc
parents: 105
diff changeset
  1076
  \noindent
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1077
  Lemma~\ref{every_eqcl_has_reg} allows us to finally give a proof for the first direction
106
91dc591de63f updated paper
urbanc
parents: 105
diff changeset
  1078
  of the Myhill-Nerode theorem.
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1079
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1080
  \begin{proof}[Proof of Theorem~\ref{myhillnerodeone}]
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1081
  By Lemma~\ref{every_eqcl_has_reg} we know that there exists a regular expression for
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1082
  every equivalence class in @{term "UNIV // \<approx>A"}. Since @{text "finals A"} is
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1083
  a subset of  @{term "UNIV // \<approx>A"}, we also know that for every equivalence class
123
23c0e6f2929d polished everywhere...two cases still missing
urbanc
parents: 122
diff changeset
  1084
  in @{term "finals A"} there exists a regular expression. Moreover by assumption 
106
91dc591de63f updated paper
urbanc
parents: 105
diff changeset
  1085
  we know that @{term "finals A"} must be finite, and therefore there must be a finite
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1086
  set of regular expressions @{text "rs"} such that
176
6969de1eb96b latest version of the journal paper
urbanc
parents: 175
diff changeset
  1087
  @{term "\<Union>(finals A) = lang (\<Uplus>rs)"}.
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1088
  Since the left-hand side is equal to @{text A}, we can use @{term "\<Uplus>rs"} 
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1089
  as the regular expression that is needed in the theorem.
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1090
  \end{proof}
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
  1091
*}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
  1092
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
  1093
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
  1094
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
  1095
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
  1096
section {* Myhill-Nerode, Second Part *}
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
  1097
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
  1098
text {*
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
  1099
  \noindent
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1100
  In this section we will give a proof for establishing the second 
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1101
  part of the Myhill-Nerode theorem. It can be formulated in our setting as follows:
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
  1102
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1103
  \begin{thrm}\label{myhillnerodetwo}
135
604518f0127f final polished
urbanc
parents: 134
diff changeset
  1104
  Given @{text "r"} is a regular expression, then @{thm Myhill_Nerode2}.
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
  1105
  \end{thrm}  
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
  1106
116
342983676c8f included comments by Chunhan
urbanc
parents: 115
diff changeset
  1107
  \noindent
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1108
  The proof will be by induction on the structure of @{text r}. It turns out
116
342983676c8f included comments by Chunhan
urbanc
parents: 115
diff changeset
  1109
  the base cases are straightforward.
342983676c8f included comments by Chunhan
urbanc
parents: 115
diff changeset
  1110
342983676c8f included comments by Chunhan
urbanc
parents: 115
diff changeset
  1111
342983676c8f included comments by Chunhan
urbanc
parents: 115
diff changeset
  1112
  \begin{proof}[Base Cases]
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
  1113
  The cases for @{const ZERO}, @{const ONE} and @{const ATOM} are routine, because 
149
e122cb146ecc added the most current versions of the theories.
urbanc
parents: 145
diff changeset
  1114
  we can easily establish that
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
  1115
114
c5eb5f3065ae updated bib
urbanc
parents: 113
diff changeset
  1116
  \begin{center}
c5eb5f3065ae updated bib
urbanc
parents: 113
diff changeset
  1117
  \begin{tabular}{l}
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
  1118
  @{thm quot_zero_eq}\\
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
  1119
  @{thm quot_one_subset}\\
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
  1120
  @{thm quot_atom_subset}
114
c5eb5f3065ae updated bib
urbanc
parents: 113
diff changeset
  1121
  \end{tabular}
c5eb5f3065ae updated bib
urbanc
parents: 113
diff changeset
  1122
  \end{center}
c5eb5f3065ae updated bib
urbanc
parents: 113
diff changeset
  1123
116
342983676c8f included comments by Chunhan
urbanc
parents: 115
diff changeset
  1124
  \noindent
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1125
  hold, which shows that @{term "UNIV // \<approx>(lang r)"} must be finite.
114
c5eb5f3065ae updated bib
urbanc
parents: 113
diff changeset
  1126
  \end{proof}
109
79b37ef9505f minor updated
urbanc
parents: 108
diff changeset
  1127
116
342983676c8f included comments by Chunhan
urbanc
parents: 115
diff changeset
  1128
  \noindent
183
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
  1129
  Much more interesting, however, are the inductive cases. They seem hard to be solved 
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1130
  directly. The reader is invited to try. 
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1131
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1132
  In order to see how our proof proceeds consider the following suggestive picture 
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1133
  taken from Constable et al \cite{Constable00}:
180
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1134
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1135
  \begin{equation}\label{pics}
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1136
  \mbox{\begin{tabular}{c@ {\hspace{10mm}}c@ {\hspace{10mm}}c}
180
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1137
  \begin{tikzpicture}[scale=1]
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1138
  %Circle
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1139
  \draw[thick] (0,0) circle (1.1);    
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1140
  \end{tikzpicture}
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1141
  &
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1142
  \begin{tikzpicture}[scale=1]
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1143
  %Circle
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1144
  \draw[thick] (0,0) circle (1.1);    
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1145
  %Main rays
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1146
  \foreach \a in {0, 90,...,359}
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1147
    \draw[very thick] (0, 0) -- (\a:1.1);
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1148
  \foreach \a / \l in {45/1, 135/2, 225/3, 315/4}
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1149
      \draw (\a: 0.65) node {$a_\l$};
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1150
  \end{tikzpicture}
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1151
  &
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1152
  \begin{tikzpicture}[scale=1]
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1153
  %Circle
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1154
  \draw[thick] (0,0) circle (1.1);    
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1155
  %Main rays
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1156
  \foreach \a in {0, 45,...,359}
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1157
     \draw[very thick] (0, 0) -- (\a:1.1);
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1158
  \foreach \a / \l in {22.5/1.1, 67.5/1.2, 112.5/2.1, 157.5/2.2, 202.4/3.1, 247.5/3.2, 292.5/4.1, 337.5/4.2}
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1159
      \draw (\a: 0.77) node {$a_{\l}$};
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1160
  \end{tikzpicture}\\
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1161
  @{term UNIV} & @{term "UNIV // (\<approx>(lang r))"} & @{term "UNIV // R"}
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1162
  \end{tabular}}
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1163
  \end{equation}
179
edacc141060f small improvements
urbanc
parents: 178
diff changeset
  1164
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1165
  \noindent
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1166
  The relation @{term "\<approx>(lang r)"} partitions the set of all strings, @{term UNIV}, into some
183
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
  1167
  equivalence classes. To show that there are only finitely many of them, it
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
  1168
  suffices to show in each induction step that another relation, say @{text
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1169
  R}, has finitely many equivalence classes and refines @{term "\<approx>(lang r)"}. 
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1170
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1171
  \begin{dfntn}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1172
  A relation @{text "R\<^isub>1"} is said to \emph{refine} @{text "R\<^isub>2"}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1173
  provided @{text "R\<^isub>1 \<subseteq> R\<^isub>2"}. 
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1174
  \end{dfntn}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1175
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1176
  \noindent
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1177
  For constructing @{text R}, will rely on some \emph{tagging-functions}
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1178
  defined over strings. Given the inductive hypothesis, it will be easy to
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1179
  prove that the \emph{range} of these tagging-functions is finite. The range
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1180
  of a function @{text f} is defined as
183
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
  1181
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1182
  \begin{center}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1183
  @{text "range f \<equiv> f ` UNIV"}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1184
  \end{center}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1185
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1186
  \noindent
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1187
  that means we take the image of @{text f} w.r.t.~all elements in the
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1188
  domain. With this we will be able to infer that the tagging-functions, seen
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1189
  as relations, give rise to finitely many equivalence classes. 
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1190
  Finally we will show that the tagging-relations are more refined than
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1191
  @{term "\<approx>(lang r)"}, which implies that @{term "UNIV // \<approx>(lang r)"} must
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1192
  also be finite.  We formally define the notion of a \emph{tagging-relation}
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1193
  as follows.
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1194
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1195
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
  1196
  \begin{dfntn}[Tagging-Relation] Given a tagging-function @{text tag}, then two strings @{text x}
119
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
  1197
  and @{text y} are \emph{tag-related} provided
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1198
  \begin{center}
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1199
  @{text "x \<^raw:$\threesim$>\<^bsub>tag\<^esub> y \<equiv> tag x = tag y"}\;.
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1200
  \end{center}
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
  1201
  \end{dfntn}
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1202
145
099e20f25b25 corrected small typo
urbanc
parents: 143
diff changeset
  1203
123
23c0e6f2929d polished everywhere...two cases still missing
urbanc
parents: 122
diff changeset
  1204
  In order to establish finiteness of a set @{text A}, we shall use the following powerful
118
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1205
  principle from Isabelle/HOL's library.
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1206
  %
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1207
  \begin{equation}\label{finiteimageD}
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1208
  @{thm[mode=IfThen] finite_imageD}
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1209
  \end{equation}
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1210
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1211
  \noindent
123
23c0e6f2929d polished everywhere...two cases still missing
urbanc
parents: 122
diff changeset
  1212
  It states that if an image of a set under an injective function @{text f} (injective over this set) 
131
6b4c20714b4f chunhan's comments
urbanc
parents: 130
diff changeset
  1213
  is finite, then the set @{text A} itself must be finite. We can use it to establish the following 
118
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1214
  two lemmas.
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1215
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
  1216
  \begin{lmm}\label{finone}
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1217
  @{thm[mode=IfThen] finite_eq_tag_rel}
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
  1218
  \end{lmm}
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1219
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1220
  \begin{proof}
119
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
  1221
  We set in \eqref{finiteimageD}, @{text f} to be @{text "X \<mapsto> tag ` X"}. We have
123
23c0e6f2929d polished everywhere...two cases still missing
urbanc
parents: 122
diff changeset
  1222
  @{text "range f"} to be a subset of @{term "Pow (range tag)"}, which we know must be
119
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
  1223
  finite by assumption. Now @{term "f (UNIV // =tag=)"} is a subset of @{text "range f"},
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
  1224
  and so also finite. Injectivity amounts to showing that @{text "X = Y"} under the
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
  1225
  assumptions that @{text "X, Y \<in> "}~@{term "UNIV // =tag="} and @{text "f X = f Y"}.
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1226
  From the assumptions we obtain \mbox{@{text "x \<in> X"}} and @{text "y \<in> Y"} with 
123
23c0e6f2929d polished everywhere...two cases still missing
urbanc
parents: 122
diff changeset
  1227
  @{text "tag x = tag y"}. Since @{text x} and @{text y} are tag-related, this in 
23c0e6f2929d polished everywhere...two cases still missing
urbanc
parents: 122
diff changeset
  1228
  turn means that the equivalence classes @{text X}
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1229
  and @{text Y} must be equal. Therefore \eqref{finiteimageD} allows us to conclude
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1230
  with @{thm (concl) finite_eq_tag_rel}.
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1231
  \end{proof}
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1232
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
  1233
  \begin{lmm}\label{fintwo} 
123
23c0e6f2929d polished everywhere...two cases still missing
urbanc
parents: 122
diff changeset
  1234
  Given two equivalence relations @{text "R\<^isub>1"} and @{text "R\<^isub>2"}, whereby
118
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1235
  @{text "R\<^isub>1"} refines @{text "R\<^isub>2"}.
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1236
  If @{thm (prem 1) refined_partition_finite[where ?R1.0="R\<^isub>1" and ?R2.0="R\<^isub>2"]}
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1237
  then @{thm (concl) refined_partition_finite[where ?R1.0="R\<^isub>1" and ?R2.0="R\<^isub>2"]}.
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
  1238
  \end{lmm}
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1239
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1240
  \begin{proof}
123
23c0e6f2929d polished everywhere...two cases still missing
urbanc
parents: 122
diff changeset
  1241
  We prove this lemma again using \eqref{finiteimageD}. This time we set @{text f} to
118
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1242
  be @{text "X \<mapsto>"}~@{term "{R\<^isub>1 `` {x} | x. x \<in> X}"}. It is easy to see that 
135
604518f0127f final polished
urbanc
parents: 134
diff changeset
  1243
  @{term "finite (f ` (UNIV // R\<^isub>2))"} because it is a subset of @{term "Pow (UNIV // R\<^isub>1)"},
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1244
  which must be finite by assumption. What remains to be shown is that @{text f} is injective
118
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1245
  on @{term "UNIV // R\<^isub>2"}. This is equivalent to showing that two equivalence 
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1246
  classes, say @{text "X"} and @{text Y}, in @{term "UNIV // R\<^isub>2"} are equal, provided
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1247
  @{text "f X = f Y"}. For @{text "X = Y"} to be equal, we have to find two elements
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1248
  @{text "x \<in> X"} and @{text "y \<in> Y"} such that they are @{text R\<^isub>2} related.
135
604518f0127f final polished
urbanc
parents: 134
diff changeset
  1249
  We know there exists a @{text "x \<in> X"} with \mbox{@{term "X = R\<^isub>2 `` {x}"}}. 
604518f0127f final polished
urbanc
parents: 134
diff changeset
  1250
  From the latter fact we can infer that @{term "R\<^isub>1 ``{x} \<in> f X"}
123
23c0e6f2929d polished everywhere...two cases still missing
urbanc
parents: 122
diff changeset
  1251
  and further @{term "R\<^isub>1 ``{x} \<in> f Y"}. This means we can obtain a @{text y}
23c0e6f2929d polished everywhere...two cases still missing
urbanc
parents: 122
diff changeset
  1252
  such that @{term "R\<^isub>1 `` {x} = R\<^isub>1 `` {y}"} holds. Consequently @{text x} and @{text y}
118
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1253
  are @{text "R\<^isub>1"}-related. Since by assumption @{text "R\<^isub>1"} refines @{text "R\<^isub>2"},
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1254
  they must also be @{text "R\<^isub>2"}-related, as we need to show.
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1255
  \end{proof}
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1256
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1257
  \noindent
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1258
  Chaining Lemma~\ref{finone} and \ref{fintwo} together, means in order to show
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1259
  that @{term "UNIV // \<approx>(lang r)"} is finite, we have to construct a tagging-function whose
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1260
  range can be shown to be finite and whose tagging-relation refines @{term "\<approx>(lang r)"}.
183
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
  1261
  Let us attempt the @{const PLUS}-case first. We take as tagging-function 
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
  1262
 
119
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
  1263
  \begin{center}
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1264
  @{thm tag_Plus_def[where A="A" and B="B", THEN meta_eq_app]}
119
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
  1265
  \end{center}
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1266
119
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
  1267
  \noindent
183
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
  1268
  where @{text "A"} and @{text "B"} are some arbitrary languages. The reason for this choice 
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1269
  is that we need to establish that @{term "=(tag_Plus A B)="} refines @{term "\<approx>(A \<union> B)"}. 
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1270
  This amounts to showing @{term "x \<approx>A y"} or @{term "x \<approx>B y"} under the assumption
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1271
  @{term "x"}~@{term "=(tag_Plus A B)="}~@{term y}. As we shall see, this definition will 
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1272
  provide us with just the right assumptions in order to get the proof through.
183
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
  1273
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
  1274
 \begin{proof}[@{const "PLUS"}-Case]
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
  1275
  We can show in general, if @{term "finite (UNIV // \<approx>A)"} and @{term "finite
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
  1276
  (UNIV // \<approx>B)"} then @{term "finite ((UNIV // \<approx>A) \<times> (UNIV // \<approx>B))"}
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
  1277
  holds. The range of @{term "tag_Plus A B"} is a subset of this product
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
  1278
  set---so finite. For the refinement proof-obligation, we know that @{term
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
  1279
  "(\<approx>A `` {x}, \<approx>B `` {x}) = (\<approx>A `` {y}, \<approx>B `` {y})"} holds by assumption. Then
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1280
  clearly either @{term "x \<approx>A y"} or @{term "x \<approx>B y"}, as we needed to
183
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
  1281
  show. Finally we can discharge this case by setting @{text A} to @{term
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
  1282
  "lang r\<^isub>1"} and @{text B} to @{term "lang r\<^isub>2"}.
119
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
  1283
  \end{proof}
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
  1284
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1285
  \noindent
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1286
  The @{const TIMES}-case is slightly more complicated. We first prove the
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1287
  following lemma, which will aid the proof about refinement.
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1288
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1289
  \begin{lmm}\label{refinement}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1290
  The relation @{text "\<^raw:$\threesim$>\<^bsub>tag\<^esub>"} refines @{term "\<approx>A"}, provided for
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1291
  all strings @{text x}, @{text y} and @{text z} we have that \mbox{@{text "x \<^raw:$\threesim$>\<^bsub>tag\<^esub> y"}}
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1292
  and @{term "x @ z \<in> A"} imply @{text "y @ z \<in> A"}.
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1293
  \end{lmm}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1294
109
79b37ef9505f minor updated
urbanc
parents: 108
diff changeset
  1295
121
1cf12a107b03 added directory with the small files and numbers of lines
urbanc
parents: 120
diff changeset
  1296
  \noindent
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1297
  We therefore can analyse how the strings @{text "x @ z"} are in the language
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1298
  @{text A} and then construct an appropriate tagging-function to infer that
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1299
  @{term "y @ z"} are also in @{text A}.  For this we will use the notion of
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1300
  the set of all possible \emph{partitions} of a string:
182
560712a29a36 a version of the proof which dispenses with the notion of string-subtraction
urbanc
parents: 181
diff changeset
  1301
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1302
  \begin{equation}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1303
  @{thm Partitions_def}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1304
  \end{equation}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1305
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1306
  \noindent
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1307
  If we know that @{text "(x\<^isub>p, x\<^isub>s) \<in> Partitions x"}, we will
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1308
  refer to @{text "x\<^isub>p"} as the \emph{prefix} of the string @{text x},
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1309
  and respectively to @{text "x\<^isub>s"} as the \emph{suffix}.
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1310
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1311
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1312
  Now assuming  @{term "x @ z \<in> A \<cdot> B"}, there are only two possible ways of how to `split' 
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
  1313
  this string to be in @{term "A \<cdot> B"}:
132
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
  1314
  %
125
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1315
  \begin{center}
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1316
  \begin{tabular}{c}
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1317
  \scalebox{1}{
125
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1318
  \begin{tikzpicture}
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1319
    \node[draw,minimum height=3.8ex] (x) { $\hspace{4.8em}@{text x}\hspace{4.8em}$ };
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1320
    \node[draw,minimum height=3.8ex, right=-0.03em of x] (za) { $\hspace{0.6em}@{text "z\<^isub>p"}\hspace{0.6em}$ };
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1321
    \node[draw,minimum height=3.8ex, right=-0.03em of za] (zza) { $\hspace{2.6em}@{text "z\<^isub>s"}\hspace{2.6em}$  };
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1322
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1323
    \draw[decoration={brace,transform={yscale=3}},decorate]
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1324
           (x.north west) -- ($(za.north west)+(0em,0em)$)
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1325
               node[midway, above=0.5em]{@{text x}};
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1326
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1327
    \draw[decoration={brace,transform={yscale=3}},decorate]
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1328
           ($(za.north west)+(0em,0ex)$) -- ($(zza.north east)+(0em,0ex)$)
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1329
               node[midway, above=0.5em]{@{text z}};
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1330
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1331
    \draw[decoration={brace,transform={yscale=3}},decorate]
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1332
           ($(x.north west)+(0em,3ex)$) -- ($(zza.north east)+(0em,3ex)$)
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1333
               node[midway, above=0.8em]{@{term "x @ z \<in> A \<cdot> B"}};
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1334
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1335
    \draw[decoration={brace,transform={yscale=3}},decorate]
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1336
           ($(za.south east)+(0em,0ex)$) -- ($(x.south west)+(0em,0ex)$)
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1337
               node[midway, below=0.5em]{@{text "x @ z\<^isub>p \<in> A"}};
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1338
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1339
    \draw[decoration={brace,transform={yscale=3}},decorate]
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1340
           ($(zza.south east)+(0em,0ex)$) -- ($(za.south east)+(0em,0ex)$)
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1341
               node[midway, below=0.5em]{@{text "z\<^isub>s \<in> B"}};
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1342
  \end{tikzpicture}}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1343
  \\[2mm]
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1344
  \scalebox{1}{
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1345
  \begin{tikzpicture}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1346
    \node[draw,minimum height=3.8ex] (xa) { $\hspace{3em}@{text "x\<^isub>p"}\hspace{3em}$ };
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1347
    \node[draw,minimum height=3.8ex, right=-0.03em of xa] (xxa) { $\hspace{0.2em}@{text "x\<^isub>s"}\hspace{0.2em}$ };
159
990c12ab1562 edits; sqeezed to 16 pages
urbanc
parents: 157
diff changeset
  1348
    \node[draw,minimum height=3.8ex, right=-0.03em of xxa] (z) { $\hspace{5em}@{text z}\hspace{5em}$ };
125
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1349
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1350
    \draw[decoration={brace,transform={yscale=3}},decorate]
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1351
           (xa.north west) -- ($(xxa.north east)+(0em,0em)$)
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1352
               node[midway, above=0.5em]{@{text x}};
125
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1353
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1354
    \draw[decoration={brace,transform={yscale=3}},decorate]
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1355
           (z.north west) -- ($(z.north east)+(0em,0em)$)
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1356
               node[midway, above=0.5em]{@{text z}};
125
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1357
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1358
    \draw[decoration={brace,transform={yscale=3}},decorate]
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1359
           ($(xa.north west)+(0em,3ex)$) -- ($(z.north east)+(0em,3ex)$)
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
  1360
               node[midway, above=0.8em]{@{term "x @ z \<in> A \<cdot> B"}};
125
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1361
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1362
    \draw[decoration={brace,transform={yscale=3}},decorate]
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1363
           ($(z.south east)+(0em,0ex)$) -- ($(xxa.south west)+(0em,0ex)$)
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1364
               node[midway, below=0.5em]{@{term "x\<^isub>s @ z \<in> B"}};
125
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1365
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1366
    \draw[decoration={brace,transform={yscale=3}},decorate]
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1367
           ($(xa.south east)+(0em,0ex)$) -- ($(xa.south west)+(0em,0ex)$)
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1368
               node[midway, below=0.5em]{@{term "x\<^isub>p \<in> A"}};
125
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1369
  \end{tikzpicture}}
159
990c12ab1562 edits; sqeezed to 16 pages
urbanc
parents: 157
diff changeset
  1370
  \end{tabular}
125
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1371
  \end{center}
132
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
  1372
  %
125
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1373
  \noindent
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1374
  Either @{text x} and a prefix of @{text "z"} is in @{text A} and the rest in @{text B} 
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1375
  (first picture) or there is a prefix of @{text x} in @{text A} and the rest is in @{text B} 
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1376
  (second picture). In both cases we have to show that @{term "y @ z \<in> A \<cdot> B"}. The first case
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1377
  we will only go through if we know that  @{term "x \<approx>A y"} holds @{text "(*)"}. Because then 
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1378
  we can infer from @{term "x @ z\<^isub>p \<in> A"} that @{term "y @ z\<^isub>p \<in> A"} holds for all @{text "z\<^isub>p"}.
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1379
  In the second case we only know that @{text "x\<^isub>p"} and @{text "x\<^isub>s"} is one possible partition
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1380
  of the string @{text x}. We have to know that both @{text "x\<^isub>p"} and the
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1381
  corresponding partition @{text "y\<^isub>p"} are in @{text "A"}, and that @{text "x\<^isub>s"} is `@{text B}-related' 
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1382
  to @{text "y\<^isub>s"} @{text "(**)"}. From the latter fact we can infer that @{text "y\<^isub>s @ z \<in> B"}.
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1383
  This will solve the second case.
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1384
  Taking the two requirements, @{text "(*)"} and @{text "(**)"}, together we define the
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1385
  tagging-function in the @{const Times}-case as:
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1386
121
1cf12a107b03 added directory with the small files and numbers of lines
urbanc
parents: 120
diff changeset
  1387
  \begin{center}
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1388
  @{thm (lhs) tag_Times_def[where ?A="A" and ?B="B"]}~@{text "\<equiv>"}~
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1389
  @{text "(\<lbrakk>x\<rbrakk>\<^bsub>\<approx>A\<^esub>, {\<lbrakk>x\<^isub>s\<rbrakk>\<^bsub>\<approx>B\<^esub> | x\<^isub>p \<in> A \<and> (x\<^isub>p, x\<^isub>s) \<in> Partitions x})"}
121
1cf12a107b03 added directory with the small files and numbers of lines
urbanc
parents: 120
diff changeset
  1390
  \end{center}
125
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1391
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1392
  \noindent
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1393
  Note that we have to make the assumption for all suffixes @{text "x\<^isub>s"}, since we do 
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1394
  not know anything about how the string @{term x} is partitioned.
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1395
  With this definition in place, let us prove the @{const "Times"}-case.
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1396
125
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1397
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
  1398
  \begin{proof}[@{const TIMES}-Case]
127
8440863a9900 seq case finished
urbanc
parents: 126
diff changeset
  1399
  If @{term "finite (UNIV // \<approx>A)"} and @{term "finite (UNIV // \<approx>B)"}
8440863a9900 seq case finished
urbanc
parents: 126
diff changeset
  1400
  then @{term "finite ((UNIV // \<approx>A) \<times> (Pow (UNIV // \<approx>B)))"} holds. The range of
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1401
  @{term "tag_Times A B"} is a subset of this product set, and therefore finite.
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1402
  For the refinement of @{term "\<approx>(A \<cdot> B)"} and @{text "\<^raw:$\threesim$>\<^bsub>\<times>tag A B\<^esub>"}, 
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1403
  we have by Lemma \ref{refinement} 
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1404
127
8440863a9900 seq case finished
urbanc
parents: 126
diff changeset
  1405
  \begin{center}
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1406
   @{term "tag_Times A B x = tag_Times A B y"}
127
8440863a9900 seq case finished
urbanc
parents: 126
diff changeset
  1407
  \end{center}
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1408
127
8440863a9900 seq case finished
urbanc
parents: 126
diff changeset
  1409
  \noindent
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1410
  and @{term "x @ z \<in> A \<cdot> B"}, and have to establish @{term "y @ z \<in> A \<cdot>
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1411
  B"}. As shown in the pictures above, there are two cases to be
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1412
  considered. First, there exists a @{text "z\<^isub>p"} and @{text
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1413
  "z\<^isub>s"} such that @{term "x @ z\<^isub>p \<in> A"} and @{text "z\<^isub>s
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1414
  \<in> B"}.  By the assumption about @{term "tag_Times A B"} we have @{term "\<approx>A
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1415
  `` {x} = \<approx>A `` {y}"} and thus @{term "x \<approx>A y"}. Hence by the Myhill-Nerode
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1416
  relation @{term "y @ z\<^isub>p \<in> A"} holds. Using @{text "z\<^isub>s \<in> B"},
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1417
  we can conclude in this case with @{term "y @ z \<in> A \<cdot> B"} (recall @{text
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1418
  "z\<^isub>p @ z\<^isub>s = z"}).
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1419
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1420
  Second there exists a partition @{text "x\<^isub>p"} and @{text "x\<^isub>s"} with 
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1421
  @{text "x\<^isub>p \<in> A"} and @{text "x\<^isub>s @ z \<in> B"}. We therefore have
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1422
  
127
8440863a9900 seq case finished
urbanc
parents: 126
diff changeset
  1423
  \begin{center}
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1424
  @{text "\<lbrakk>x\<^isub>s\<rbrakk>\<^bsub>\<approx>B\<^esub> \<in> {\<lbrakk>x\<^isub>s\<rbrakk>\<^bsub>\<approx>B\<^esub> | x\<^isub>p \<in> A \<and> (x\<^isub>p, x\<^isub>s) \<in> Partitions x}"}
127
8440863a9900 seq case finished
urbanc
parents: 126
diff changeset
  1425
  \end{center}
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1426
  
127
8440863a9900 seq case finished
urbanc
parents: 126
diff changeset
  1427
  \noindent
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1428
  and by the assumption about @{term "tag_Times A B"} also 
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1429
  
127
8440863a9900 seq case finished
urbanc
parents: 126
diff changeset
  1430
  \begin{center}
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1431
  @{text "\<lbrakk>x\<^isub>s\<rbrakk>\<^bsub>\<approx>B\<^esub> \<in> {\<lbrakk>y\<^isub>s\<rbrakk>\<^bsub>\<approx>B\<^esub> | y\<^isub>p \<in> A \<and> (y\<^isub>p, y\<^isub>s) \<in> Partitions y}"}
127
8440863a9900 seq case finished
urbanc
parents: 126
diff changeset
  1432
  \end{center}
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1433
  
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1434
  \noindent
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1435
  This means there must be a partition @{text "y\<^isub>p"} and @{text "y\<^isub>s"}
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1436
  such that @{term "y\<^isub>p \<in> A"} and @{term "\<approx>B `` {x\<^isub>s} = \<approx>B ``
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1437
  {y\<^isub>s}"}. Unfolding the Myhill-Nerode relation and together with the
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1438
  facts that @{text "x\<^isub>p \<in> A"} and \mbox{@{text "x\<^isub>s @ z \<in> B"}}, we
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1439
  obtain @{term "y\<^isub>p \<in> A"} and @{text "y\<^isub>s @ z \<in> B"}, as needed in
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1440
  this case.  We again can complete the @{const TIMES}-case by setting @{text
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1441
  A} to @{term "lang r\<^isub>1"} and @{text B} to @{term "lang r\<^isub>2"}.
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1442
  \end{proof}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1443
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1444
  \noindent
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1445
  The case for @{const Star} is similar to @{const TIMES}, but poses a few
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1446
  extra challenges.  To deal with them, we define first the notion of a \emph{string
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1447
  prefix} and a \emph{strict string prefix}:
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1448
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1449
  \begin{center}
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1450
  \begin{tabular}{l}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1451
  @{text "x \<le> y \<equiv> \<exists>z. y = x @ z"}\\
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1452
  @{text "x < y \<equiv> x \<le> y \<and> x \<noteq> y"}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1453
  \end{tabular}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1454
  \end{center}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1455
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1456
  When analysing the case of @{text "x @ z"} being an element in @{term "A\<star>"}
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1457
  and @{text x} is not the empty string, we have the following picture:
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1458
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1459
  \begin{center}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1460
  \scalebox{1}{
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1461
  \begin{tikzpicture}
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1462
    \node[draw,minimum height=3.8ex] (xa) { $\hspace{4em}@{text "x\<^bsub>pmax\<^esub>"}\hspace{4em}$ };
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1463
    \node[draw,minimum height=3.8ex, right=-0.03em of xa] (xxa) { $\hspace{0.5em}@{text "x\<^bsub>s\<^esub>"}\hspace{0.5em}$ };
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1464
    \node[draw,minimum height=3.8ex, right=-0.03em of xxa] (za) { $\hspace{2em}@{text "z\<^isub>a"}\hspace{2em}$ };
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1465
    \node[draw,minimum height=3.8ex, right=-0.03em of za] (zb) { $\hspace{7em}@{text "z\<^isub>b"}\hspace{7em}$ };
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1466
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1467
    \draw[decoration={brace,transform={yscale=3}},decorate]
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1468
           (xa.north west) -- ($(xxa.north east)+(0em,0em)$)
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1469
               node[midway, above=0.5em]{@{text x}};
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1470
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1471
    \draw[decoration={brace,transform={yscale=3}},decorate]
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1472
           (za.north west) -- ($(zb.north east)+(0em,0em)$)
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1473
               node[midway, above=0.5em]{@{text z}};
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1474
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1475
    \draw[decoration={brace,transform={yscale=3}},decorate]
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1476
           ($(xa.north west)+(0em,3ex)$) -- ($(zb.north east)+(0em,3ex)$)
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1477
               node[midway, above=0.8em]{@{term "x @ z \<in> A\<star>"}};
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1478
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1479
    \draw[decoration={brace,transform={yscale=3}},decorate]
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1480
           ($(za.south east)+(0em,0ex)$) -- ($(xxa.south west)+(0em,0ex)$)
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1481
               node[midway, below=0.5em]{@{term "x\<^isub>s @ z\<^isub>a \<in> A"}};
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1482
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1483
    \draw[decoration={brace,transform={yscale=3}},decorate]
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1484
           ($(xa.south east)+(0em,0ex)$) -- ($(xa.south west)+(0em,0ex)$)
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1485
               node[midway, below=0.5em]{@{text "x\<^bsub>pmax\<^esub> \<in> A\<^isup>\<star>"}};
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1486
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1487
    \draw[decoration={brace,transform={yscale=3}},decorate]
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1488
           ($(zb.south east)+(0em,0ex)$) -- ($(zb.south west)+(0em,0ex)$)
136
13b0f3dac9a2 final final polishing
urbanc
parents: 135
diff changeset
  1489
               node[midway, below=0.5em]{@{term "z\<^isub>b \<in> A\<star>"}};
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1490
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1491
    \draw[decoration={brace,transform={yscale=3}},decorate]
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1492
           ($(zb.south east)+(0em,-4ex)$) -- ($(xxa.south west)+(0em,-4ex)$)
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1493
               node[midway, below=0.5em]{@{term "x\<^isub>s @ z \<in> A\<star>"}};
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1494
  \end{tikzpicture}}
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1495
  \end{center}
132
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
  1496
  %
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1497
  \noindent
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1498
  We can find a strict prefix @{text "x\<^isub>p"} of @{text x} such that @{term "x\<^isub>p \<in> A\<star>"},
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1499
  @{text "x\<^isub>p < x"} and the rest @{term "x\<^isub>s @ z \<in> A\<star>"}. For example the empty string 
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1500
  @{text "[]"} would do (recall @{term "x \<noteq> []"}).
135
604518f0127f final polished
urbanc
parents: 134
diff changeset
  1501
  There are potentially many such prefixes, but there can only be finitely many of them (the 
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1502
  string @{text x} is finite). Let us therefore choose the longest one and call it 
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1503
  @{text "x\<^bsub>pmax\<^esub>"}. Now for the rest of the string @{text "x\<^isub>s @ z"} we
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1504
  know it is in @{term "A\<star>"} and cannot be the empty string. By Property~\ref{langprops}@{text "(iv)"}, 
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1505
  we can separate
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1506
  this string into two parts, say @{text "a"} and @{text "b"}, such that @{text "a \<noteq> []"}, @{text "a \<in> A"}
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1507
  and @{term "b \<in> A\<star>"}. Now @{text a} must be strictly longer than @{text "x\<^isub>s"},
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1508
  otherwise @{text "x\<^bsub>pmax\<^esub>"} is not the longest prefix. That means @{text a}
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1509
  `overlaps' with @{text z}, splitting it into two components @{text "z\<^isub>a"} and
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1510
   @{text "z\<^isub>b"}. For this we know that @{text "x\<^isub>s @ z\<^isub>a \<in> A"} and
135
604518f0127f final polished
urbanc
parents: 134
diff changeset
  1511
  @{term "z\<^isub>b \<in> A\<star>"}. To cut a story short, we have divided @{term "x @ z \<in> A\<star>"}
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1512
  such that we have a string @{text a} with @{text "a \<in> A"} that lies just on the
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1513
  `border' of @{text x} and @{text z}. This string is @{text "x\<^isub>s @ z\<^isub>a"}.
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1514
135
604518f0127f final polished
urbanc
parents: 134
diff changeset
  1515
  In order to show that @{term "x @ z \<in> A\<star>"} implies @{term "y @ z \<in> A\<star>"}, we use
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1516
  the following tagging-function:
132
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
  1517
  %
121
1cf12a107b03 added directory with the small files and numbers of lines
urbanc
parents: 120
diff changeset
  1518
  \begin{center}
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1519
  @{thm (lhs) tag_Star_def[where ?A="A", THEN meta_eq_app]}~@{text "\<equiv>"}~
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1520
  @{text "{\<lbrakk>x\<^isub>s\<rbrakk>\<^bsub>\<approx>A\<^esub> | x\<^isub>p < x \<and> x\<^isub>p \<in> A\<^isup>\<star> \<and> (x\<^isub>s, x\<^isub>p) \<in> Partitions x}"}
121
1cf12a107b03 added directory with the small files and numbers of lines
urbanc
parents: 120
diff changeset
  1521
  \end{center}
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1522
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
  1523
  \begin{proof}[@{const Star}-Case]
130
3e4ad22193e7 pre-final version
urbanc
parents: 129
diff changeset
  1524
  If @{term "finite (UNIV // \<approx>A)"} 
3e4ad22193e7 pre-final version
urbanc
parents: 129
diff changeset
  1525
  then @{term "finite (Pow (UNIV // \<approx>A))"} holds. The range of
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1526
  @{term "tag_Star A"} is a subset of this set, and therefore finite.
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1527
  Again we have to show under the assumption @{term "x"}~@{term "=(tag_Star A)="}~@{term y}
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1528
  that @{term "x @ z \<in> A\<star>"} implies @{term "y @ z \<in> A\<star>"}.
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1529
130
3e4ad22193e7 pre-final version
urbanc
parents: 129
diff changeset
  1530
  We first need to consider the case that @{text x} is the empty string.
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1531
  From the assumption about strict prefixes in @{text "\<^raw:$\threesim$>\<^bsub>\<star>tag A\<^esub>"}, we 
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1532
  can infer @{text y} is the empty string and
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1533
  then clearly have @{term "y @ z \<in> A\<star>"}. In case @{text x} is not the empty
134
08afbed1c8c7 chunhan's comments
urbanc
parents: 133
diff changeset
  1534
  string, we can divide the string @{text "x @ z"} as shown in the picture 
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1535
  above. By the tagging-function and the facts @{text "x\<^bsub>pmax\<^esub> \<in> A\<^isup>\<star>"} and @{text "x\<^bsub>pmax\<^esub> < x"}, 
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1536
  we have
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1537
130
3e4ad22193e7 pre-final version
urbanc
parents: 129
diff changeset
  1538
  \begin{center}
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1539
  @{text "\<lbrakk>x\<^isub>s\<rbrakk>\<^bsub>\<approx>A\<^esub> \<in> {\<lbrakk>x\<^isub>s\<rbrakk>\<^bsub>\<approx>A\<^esub> | x\<^bsub>pmax\<^esub> < x \<and> x\<^bsub>pmax\<^esub> \<in> A\<^isup>\<star> \<and> (x\<^bsub>pmax\<^esub>, x\<^isub>s) \<in> Partitions x}"}
130
3e4ad22193e7 pre-final version
urbanc
parents: 129
diff changeset
  1540
  \end{center}
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1541
  
130
3e4ad22193e7 pre-final version
urbanc
parents: 129
diff changeset
  1542
  \noindent
3e4ad22193e7 pre-final version
urbanc
parents: 129
diff changeset
  1543
  which by assumption is equal to
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1544
  
130
3e4ad22193e7 pre-final version
urbanc
parents: 129
diff changeset
  1545
  \begin{center}
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1546
  @{text "\<lbrakk>x\<^isub>s\<rbrakk>\<^bsub>\<approx>A\<^esub> \<in> {\<lbrakk>y\<^isub>s\<rbrakk>\<^bsub>\<approx>A\<^esub> | y\<^bsub>p\<^esub> < y \<and> y\<^bsub>p\<^esub> \<in> A\<^isup>\<star> \<and> (y\<^bsub>p\<^esub>, y\<^isub>s) \<in> Partitions y}"}
130
3e4ad22193e7 pre-final version
urbanc
parents: 129
diff changeset
  1547
  \end{center}
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1548
  
130
3e4ad22193e7 pre-final version
urbanc
parents: 129
diff changeset
  1549
  \noindent 
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1550
  From this we know there exist a partition @{text "y\<^isub>p"} and @{text
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1551
  "y\<^isub>s"} with @{term "y\<^isub>p \<in> A\<star>"} and also @{term "x\<^isub>s \<approx>A
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1552
  y\<^isub>s"}. Unfolding the Myhill-Nerode relation we know @{term
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1553
  "y\<^isub>s @ z\<^isub>a \<in> A"}. We also know that @{term "z\<^isub>b \<in> A\<star>"}.
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1554
  Therefore @{term "y\<^isub>p @ (y\<^isub>s @ z\<^isub>a) @ z\<^isub>b \<in>
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1555
  A\<star>"}, which means @{term "y @ z \<in> A\<star>"}. The last step is to set
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1556
  @{text "A"} to @{term "lang r"} and thus complete the proof.
121
1cf12a107b03 added directory with the small files and numbers of lines
urbanc
parents: 120
diff changeset
  1557
  \end{proof}
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
  1558
*}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
  1559
194
5347d7556487 some typos
urbanc
parents: 193
diff changeset
  1560
section {* Second Part proved using Partial Derivatives\label{derivatives} *}
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 160
diff changeset
  1561
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 160
diff changeset
  1562
text {*
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
  1563
  \noindent
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
  1564
  As we have seen in the previous section, in order to establish
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
  1565
  the second direction of the Myhill-Nerode theorem, we need to find 
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1566
  a more refined relation than @{term "\<approx>(lang r)"} for which we can
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1567
  show that there are only finitely many equivalence classes. So far we 
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1568
  showed this directly by induction on @{text "r"} using tagging-functions. 
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1569
  However, there is also  an indirect method to come up with such a refined 
197
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1570
  relation by using derivatives of regular expressions~\cite{Brzozowski64}. 
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1571
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1572
  Assume the following two definitions for the \emph{left-quotient} of a language,
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1573
  which we write as @{term "Der c A"} and @{term "Ders s A"} where @{text c}
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1574
  is a character and @{text s} a string, respectively:
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1575
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1576
  \begin{center}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1577
  \begin{tabular}{r@ {\hspace{1mm}}c@ {\hspace{2mm}}l}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1578
  @{thm (lhs) Der_def}  & @{text "\<equiv>"} & @{thm (rhs) Der_def}\\
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1579
  @{thm (lhs) Ders_def} & @{text "\<equiv>"} & @{thm (rhs) Ders_def}\\
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1580
  \end{tabular}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1581
  \end{center}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1582
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1583
  \noindent
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1584
  In order to aid readability, we shall make use of the following abbreviation
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1585
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1586
  \begin{center}
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1587
  @{abbrev "Derss s As"}
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1588
  \end{center}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1589
  
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1590
  \noindent
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1591
  where we apply the left-quotient to a set of languages and then combine the results.
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1592
  Clearly we have the following equivalence between the Myhill-Nerode relation
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1593
  (Definition~\ref{myhillneroderel}) and left-quotients
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1594
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1595
  \begin{equation}\label{mhders}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1596
  @{term "x \<approx>A y"} \hspace{4mm}\text{if and only if}\hspace{4mm} @{term "Ders x A = Ders y A"}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1597
  \end{equation}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1598
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1599
  \noindent
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1600
  It is also straightforward to establish the following properties of left-quotients
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1601
  
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1602
  \begin{equation}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1603
  \mbox{\begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{2mm}}l}
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1604
  @{thm (lhs) Der_simps(1)} & $=$ & @{thm (rhs) Der_simps(1)}\\
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1605
  @{thm (lhs) Der_simps(2)} & $=$ & @{thm (rhs) Der_simps(2)}\\
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1606
  @{thm (lhs) Der_simps(3)} & $=$ & @{thm (rhs) Der_simps(3)}\\
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1607
  @{thm (lhs) Der_simps(4)} & $=$ & @{thm (rhs) Der_simps(4)}\\
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1608
  @{thm (lhs) Der_conc}  & $=$ & @{thm (rhs) Der_conc}\\
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1609
  @{thm (lhs) Der_star}  & $=$ & @{thm (rhs) Der_star}\\
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1610
  @{thm (lhs) Ders_simps(1)} & $=$ & @{thm (rhs) Ders_simps(1)}\\
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1611
  @{thm (lhs) Ders_simps(2)} & $=$ & @{thm (rhs) Ders_simps(2)}\\
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1612
  %@{thm (lhs) Ders_simps(3)[where ?s1.0="s\<^isub>1" and ?s2.0="s\<^isub>2"]}  & $=$ 
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1613
  %   & @{thm (rhs) Ders_simps(3)[where ?s1.0="s\<^isub>1" and ?s2.0="s\<^isub>2"]}\\
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1614
  \end{tabular}}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1615
  \end{equation}
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1616
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1617
  \noindent
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1618
  where @{text "\<Delta>"} in the fifth line is a function that tests whether the
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1619
  empty string is in the language and returns @{term "{[]}"} or @{term "{}"},
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1620
  accordingly.  In the last equation we use the list-cons operator written
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1621
  \mbox{@{text "_ :: _"}}.  The only interesting case is the @{text "A\<star>"}-case
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1622
  where we use Property~\ref{langprops}@{text "(i)"} in order to infer that
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1623
  @{term "Der c (A\<star>) = Der c (A \<cdot> A\<star>)"}. We can then complete the proof by
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1624
  using the fifth equation and noting that @{term "Delta A \<cdot> Der c (A\<star>) \<subseteq> (Der
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1625
  c A) \<cdot> A\<star>"}. 
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1626
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1627
  Brzozowski observed that the left-quotients for languages of
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1628
  regular expressions can be calculated directly using the notion of
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1629
  \emph{derivatives of a regular expression} \cite{Brzozowski64}. We define
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1630
  this notion in Isabelle/HOL as follows:
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1631
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1632
  \begin{center}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1633
  \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1.5mm}}l@ {}}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1634
  @{thm (lhs) der.simps(1)}  & @{text "\<equiv>"} & @{thm (rhs) der.simps(1)}\\
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1635
  @{thm (lhs) der.simps(2)}  & @{text "\<equiv>"} & @{thm (rhs) der.simps(2)}\\
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1636
  @{thm (lhs) der.simps(3)[where c'="d"]}  & @{text "\<equiv>"} & @{thm (rhs) der.simps(3)[where c'="d"]}\\
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1637
  @{thm (lhs) der.simps(4)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]}  
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1638
     & @{text "\<equiv>"} & @{thm (rhs) der.simps(4)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]}\\
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1639
  @{thm (lhs) der.simps(5)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]}  
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
  1640
     & @{text "\<equiv>"} & @{text "if"}~@{term "nullable r\<^isub>1"}~@{text "then"}~%
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
  1641
       @{term "Plus (Times (der c r\<^isub>1) r\<^isub>2) (der c r\<^isub>2)"}\\
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
  1642
     &             & \phantom{@{text "if"}~@{term "nullable r\<^isub>1"}~}@{text "else"}~%  
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
  1643
                    @{term "Times (der c r\<^isub>1) r\<^isub>2"}\\ 
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1644
  @{thm (lhs) der.simps(6)}  & @{text "\<equiv>"} & @{thm (rhs) der.simps(6)}\smallskip\\
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1645
  @{thm (lhs) ders.simps(1)}  & @{text "\<equiv>"} & @{thm (rhs) ders.simps(1)}\\
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1646
  @{thm (lhs) ders.simps(2)}  & @{text "\<equiv>"} & @{thm (rhs) ders.simps(2)}\\
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1647
  \end{tabular}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1648
  \end{center}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1649
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1650
  \noindent
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1651
  The last two clauses extend derivatives from characters to strings. The
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1652
  boolean function @{term "nullable r"} needed in the @{const Times}-case tests
197
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1653
  whether a regular expression can recognise the empty string. It can be defined as 
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1654
  follows.
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1655
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1656
  \begin{center}
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
  1657
  \begin{tabular}{c@ {\hspace{10mm}}c}
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1658
  \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1.5mm}}l@ {}}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1659
  @{thm (lhs) nullable.simps(1)}  & @{text "\<equiv>"} & @{thm (rhs) nullable.simps(1)}\\
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1660
  @{thm (lhs) nullable.simps(2)}  & @{text "\<equiv>"} & @{thm (rhs) nullable.simps(2)}\\
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1661
  @{thm (lhs) nullable.simps(3)}  & @{text "\<equiv>"} & @{thm (rhs) nullable.simps(3)}\\
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1662
  \end{tabular} &
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1663
  \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1.5mm}}l@ {}}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1664
  @{thm (lhs) nullable.simps(4)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]}  
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1665
     & @{text "\<equiv>"} & @{thm (rhs) nullable.simps(4)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]}\\
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1666
  @{thm (lhs) nullable.simps(5)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]}  
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1667
     & @{text "\<equiv>"} & @{thm (rhs) nullable.simps(5)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]}\\
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1668
  @{thm (lhs) nullable.simps(6)}  & @{text "\<equiv>"} & @{thm (rhs) nullable.simps(6)}\\
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1669
  \end{tabular}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1670
  \end{tabular}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1671
  \end{center}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1672
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1673
  \noindent
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1674
  By induction on the regular expression @{text r}, respectively on the string @{text s}, 
197
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1675
  one can easily show that left-quotients and derivatives of regular expressions 
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1676
  relate as follows (see for example~\cite{Sakarovitch09}):
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1677
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1678
  \begin{equation}\label{Dersders}
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1679
  \mbox{\begin{tabular}{c}
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1680
  @{thm Der_der}\\
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1681
  @{thm Ders_ders}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1682
  \end{tabular}}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1683
  \end{equation}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1684
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1685
  \noindent
197
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1686
  The importance of this fact in the context of the Myhill-Nerode theorem is that 
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1687
  we can use \eqref{mhders} and \eqref{Dersders} in order to 
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1688
  establish that @{term "x \<approx>(lang r) y"} is equivalent to
197
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1689
  \mbox{@{term "lang (ders x r) = lang (ders y r)"}}. Hence
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1690
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1691
  \begin{equation}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1692
  @{term "x \<approx>(lang r) y"}\hspace{4mm}\mbox{provided}\hspace{4mm}@{term "ders x r = ders y r"}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1693
  \end{equation}
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1694
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1695
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1696
  \noindent
197
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1697
  which means the right-hand side (seen as relation) refines the Myhill-Nerode
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1698
  relation.  Consequently, we can use @{text
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1699
  "\<^raw:$\threesim$>\<^bsub>(\<lambda>x. ders x r)\<^esub>"} as a
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1700
  tagging-relation. However, in order to be useful for the second part of the
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1701
  Myhill-Nerode theorem, we have to be able to establish that for the
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1702
  corresponding language there are only finitely many derivatives---thus
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1703
  ensuring that there are only finitely many equivalence
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1704
  classes. Unfortunately, this is not true in general. Sakarovitch gives an
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1705
  example where a regular expression has infinitely many derivatives
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1706
  w.r.t.~the language \mbox{@{term "({a} \<cdot> {b})\<star> \<union> ({a} \<cdot> {b})\<star> \<cdot> {a}"}}
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1707
  (see \cite[Page~141]{Sakarovitch09}).
197
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1708
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1709
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1710
  What Brzozowski \cite{Brzozowski64} established is that for every language there
197
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1711
  \emph{are} only finitely `dissimilar' derivatives of a regular
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1712
  expression. Two regular expressions are said to be \emph{similar} provided
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1713
  they can be identified using the using the @{text "ACI"}-identities:
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1714
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1715
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1716
  \begin{equation}\label{ACI}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1717
  \mbox{\begin{tabular}{cl}
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1718
  (@{text A}) & @{term "Plus (Plus r\<^isub>1 r\<^isub>2) r\<^isub>3"} $\equiv$ @{term "Plus r\<^isub>1 (Plus r\<^isub>2 r\<^isub>3)"}\\
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1719
  (@{text C}) & @{term "Plus r\<^isub>1 r\<^isub>2"} $\equiv$ @{term "Plus r\<^isub>2 r\<^isub>1"}\\
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1720
  (@{text I}) & @{term "Plus r r"} $\equiv$ @{term "r"}\\
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1721
  \end{tabular}}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1722
  \end{equation}
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1723
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1724
  \noindent
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1725
  Carrying this idea through, we must not consider the set of all derivatives,
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1726
  but the ones modulo @{text "ACI"}.  In principle, this can be done formally, 
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1727
  but it is very painful in a theorem prover (since there is no
194
5347d7556487 some typos
urbanc
parents: 193
diff changeset
  1728
  direct characterisation of the set of dissimilar derivatives).
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1729
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1730
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1731
  Fortunately, there is a much simpler approach using \emph{partial
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1732
  derivatives}. They were introduced by Antimirov \cite{Antimirov95} and can be defined
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1733
  in Isabelle/HOL as follows:
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
  1734
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
  1735
  \begin{center}
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
  1736
  \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1.5mm}}l@ {}}
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
  1737
  @{thm (lhs) pder.simps(1)}  & @{text "\<equiv>"} & @{thm (rhs) pder.simps(1)}\\
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
  1738
  @{thm (lhs) pder.simps(2)}  & @{text "\<equiv>"} & @{thm (rhs) pder.simps(2)}\\
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
  1739
  @{thm (lhs) pder.simps(3)[where c'="d"]}  & @{text "\<equiv>"} & @{thm (rhs) pder.simps(3)[where c'="d"]}\\
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
  1740
  @{thm (lhs) pder.simps(4)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]}  
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
  1741
     & @{text "\<equiv>"} & @{thm (rhs) pder.simps(4)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]}\\
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
  1742
  @{thm (lhs) pder.simps(5)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]}  
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
  1743
     & @{text "\<equiv>"} & @{text "if"}~@{term "nullable r\<^isub>1"}~@{text "then"}~%
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1744
       @{term "(Timess (pder c r\<^isub>1) r\<^isub>2) \<union> (pder c r\<^isub>2)"}\\
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
  1745
     & & \phantom{@{text "if"}~@{term "nullable r\<^isub>1"}~}@{text "else"}~%  
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1746
                    @{term "Timess (pder c r\<^isub>1) r\<^isub>2"}\\ 
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
  1747
  @{thm (lhs) pder.simps(6)}  & @{text "\<equiv>"} & @{thm (rhs) pder.simps(6)}\smallskip\\
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
  1748
  @{thm (lhs) pders.simps(1)}  & @{text "\<equiv>"} & @{thm (rhs) pders.simps(1)}\\
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1749
  @{thm (lhs) pders.simps(2)}  & @{text "\<equiv>"} & @{text "\<Union> (pders s) ` (pder c r)"}\\
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
  1750
  \end{tabular}
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
  1751
  \end{center}
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
  1752
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1753
  \noindent
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1754
  Again the last two clauses extend partial derivatives from characters to strings. 
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1755
  Unlike `simple' derivatives, the functions for partial derivatives return sets of regular
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1756
  expressions. In the @{const Times} and @{const Star} cases we therefore use the
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1757
  auxiliary definition
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1758
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1759
  \begin{center}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1760
  @{text "TIMESS rs r \<equiv> {TIMES r' r | r' \<in> rs}"}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1761
  \end{center}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1762
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1763
  \noindent
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1764
  in order to `sequence' a regular expression with a set of regular
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1765
  expressions. Note that in the last clause we first build the set of partial
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1766
  derivatives w.r.t~the character @{text c}, then build the image of this set under the
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1767
  function @{term "pders s"} and finally `union up' all resulting sets. It will be
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1768
  convenient to introduce for this the following abbreviation
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1769
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1770
  \begin{center}
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1771
  @{abbrev "pderss s rs"}
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1772
  \end{center}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1773
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1774
  \noindent
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1775
  which simplifies the last clause of @{const "pders"} to
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1776
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1777
  \begin{center}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1778
  \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1.5mm}}l@ {}}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1779
  @{thm (lhs) pders.simps(2)}  & @{text "\<equiv>"} & @{thm (rhs) pders.simps(2)}\\
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1780
  \end{tabular}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1781
  \end{center}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1782
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1783
  Partial derivatives can be seen as having the @{text "ACI"}-identities already built in: 
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1784
  taking the partial derivatives of the
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1785
  regular expressions in \eqref{ACI} gives us in each case
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1786
  equal sets.  Antimirov \cite{Antimirov95} showed a similar result to
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1787
  \eqref{Dersders} for partial derivatives, namely
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1788
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1789
  \begin{equation}\label{Derspders}
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1790
  \mbox{\begin{tabular}{lc}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1791
  @{text "(i)"}  & @{thm Der_pder}\\
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1792
  @{text "(ii)"} & @{thm Ders_pders}
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1793
  \end{tabular}}
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1794
  \end{equation} 
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1795
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1796
  \begin{proof}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1797
  The first fact is by a simple induction on @{text r}. For the second we slightly
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1798
  modify Antimirov's proof by performing an induction on @{text s} where we
194
5347d7556487 some typos
urbanc
parents: 193
diff changeset
  1799
  generalise over all @{text r}. That means in the @{text "cons"}-case the 
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1800
  induction hypothesis is
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1801
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1802
  \begin{center}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1803
  @{text "(IH)"}\hspace{3mm}@{term "\<forall>r. Ders s (lang r) = \<Union> lang ` (pders s r)"}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1804
  \end{center}
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1805
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1806
  \noindent
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1807
  With this we can establish
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1808
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1809
  \begin{center}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1810
  \begin{tabular}{r@ {\hspace{1.5mm}}c@ {\hspace{1.5mm}}ll}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1811
  @{term "Ders (c # s) (lang r)"} 
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1812
    & @{text "="} & @{term "Ders s (Der c (lang r))"} & by def.\\
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1813
    & @{text "="} & @{term "Ders s (\<Union> lang ` (pder c r))"} & by @{text "("}\ref{Derspders}@{text ".i)"}\\
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1814
    & @{text "="} & @{term "\<Union> (Ders s) ` (lang ` (pder c r))"} & by def.~of @{text "Ders"}\\
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1815
    & @{text "="} & @{term "\<Union> lang ` (\<Union> pders s ` (pder c r))"} & by IH\\
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1816
    & @{text "="} & @{term "\<Union> lang ` (pders (c # s) r)"} & by def.\\
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1817
  \end{tabular}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1818
  \end{center}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1819
  
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1820
  \noindent
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1821
  Note that in order to apply the induction hypothesis in the fourth equation, we
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1822
  need the generalisation over all regular expressions @{text r}. The case for
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1823
  the empty string is routine and omitted.
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1824
  \end{proof}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1825
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1826
  \noindent
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1827
  Taking \eqref{Dersders} and \eqref{Derspders} together gives the relationship 
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1828
  between languages of derivatives and partial derivatives
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1829
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1830
  \begin{equation}
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1831
  \mbox{\begin{tabular}{lc}
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1832
  @{text "(i)"}  & @{thm der_pder[symmetric]}\\
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1833
  @{text "(ii)"} & @{thm ders_pders[symmetric]}
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1834
  \end{tabular}}
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1835
  \end{equation} 
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1836
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1837
  \noindent
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1838
  These two properties confirm the observation made earlier 
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1839
  that by using sets, partial derivatives have the @{text "ACI"}-identities
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1840
  of derivatives already built in. 
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1841
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1842
  Antimirov also proved that for every language and regular expression 
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1843
  there are only finitely many partial derivatives, whereby the partial
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1844
  derivatives of @{text r} w.r.t.~a language @{text A} is defined as
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1845
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1846
  \begin{equation}\label{Pdersdef}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1847
  @{thm pders_lang_def}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1848
  \end{equation}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1849
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1850
  \begin{thrm}[Antimirov \cite{Antimirov95}]\label{antimirov}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1851
  For every language @{text A} and every regular expression @{text r}, 
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1852
  \mbox{@{thm finite_pders_lang}}.
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1853
  \end{thrm}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1854
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1855
  \noindent
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1856
  Antimirov's argument first shows this theorem for the language @{term UNIV1}, 
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1857
  which is the set of all non-empty strings. For this he proves:
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1858
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1859
  \begin{equation}\label{pdersunivone}
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1860
  \mbox{\begin{tabular}{l}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1861
  @{thm pders_lang_Zero}\\
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1862
  @{thm pders_lang_One}\\
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1863
  @{thm pders_lang_Atom}\\
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1864
  @{thm pders_lang_Plus[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]}\\
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1865
  @{thm pders_lang_Times[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]}\\
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1866
  @{thm pders_lang_Star}\\
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1867
  \end{tabular}}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1868
  \end{equation}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1869
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1870
  \noindent
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1871
  from which one can deduce by induction on @{text r} that
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1872
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1873
  \begin{center}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1874
  @{thm finite_pders_lang_UNIV1}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1875
  \end{center}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1876
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1877
  \noindent
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1878
  holds. Now Antimirov's theorem follows because
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1879
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1880
  \begin{center}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1881
  @{thm pders_lang_UNIV}\\
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1882
  \end{center}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1883
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1884
  \noindent
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1885
  and for all languages @{text "A"}, @{thm pders_lang_subset[where B="UNIV",
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1886
  simplified]} holds.  Since we follow Antimirov's proof quite closely in our
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1887
  formalisation (only the last two cases of \eqref{pdersunivone} involve some
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1888
  non-routine induction argument), we omit the details.
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1889
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1890
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1891
  Let us now return to our proof about the second direction in the Myhill-Nerode
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1892
  theorem. The point of the above calculations is to use 
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1893
  @{text "\<^raw:$\threesim$>\<^bsub>(\<lambda>x. ders x r)\<^esub>"} as tagging-relation. 
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1894
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1895
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  1896
  \begin{proof}[Proof of Theorem~\ref{myhillnerodetwo} (second version)]
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1897
  Using \eqref{mhders}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1898
  and \eqref{Derspders} we can easily infer that
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1899
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1900
  \begin{center}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1901
  @{term "x \<approx>(lang r) y"}\hspace{4mm}\mbox{provided}\hspace{4mm}@{term "pders x r = pders y r"}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1902
  \end{center}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1903
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1904
  \noindent
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1905
  which means the tagging-relation @{text "\<^raw:$\threesim$>\<^bsub>(\<lambda>x. ders x r)\<^esub>"} refines @{term "\<approx>(lang r)"}.
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1906
  So we know by Lemma~\ref{fintwo}, \mbox{@{term "finite (UNIV // (\<approx>(lang r)))"}} 
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1907
  holds if @{term "finite (UNIV // (=(\<lambda>x. pders x r)=))"}. In order to establish 
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1908
  the latter, we can use Lemma~\ref{finone} and show that the range of the 
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1909
  tagging-function \mbox{@{term "\<lambda>x. pders x r"}} is finite. For this recall Definition
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1910
  \ref{Pdersdef}, which gives us that 
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1911
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1912
  \begin{center}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1913
  @{thm pders_lang_def[where A="UNIV", simplified]}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1914
  \end{center}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1915
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1916
  \noindent
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1917
  Now the range of @{term "\<lambda>x. pders x r"} is a subset of @{term "Pow (pders_lang UNIV r)"},
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1918
  which we know is finite by Theorem~\ref{antimirov}. This means there 
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1919
  are only finitely many equivalence classes of @{text "\<^raw:$\threesim$>\<^bsub>(\<lambda>x. ders x r)\<^esub>"},
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1920
  which refines @{term "\<approx>(lang r)"}, and consequently we can again conclude the 
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1921
  second part of the Myhill-Nerode theorem.
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1922
  \end{proof}
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 160
diff changeset
  1923
*}
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 160
diff changeset
  1924
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  1925
section {* Closure Properties of Regular Languages *}
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
  1926
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1927
text {*
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1928
  \noindent
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  1929
  The beauty of regular languages is that they are closed under many set
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  1930
  operations. Closure under union, concatenation and Kleene-star are trivial
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  1931
  to establish given our definition of regularity (recall Definition~\ref{regular}).
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  1932
  More interesting is the closure under complement, because it seems difficult
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  1933
  to construct a regular expression for the complement language by direct
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  1934
  means. However the existence of such a regular expression can now be easily
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1935
  proved using both parts of the Myhill-Nerode theorem, since
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  1936
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1937
  \begin{center}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1938
  @{term "s\<^isub>1 \<approx>A s\<^isub>2"} if and only if @{term "s\<^isub>1 \<approx>(-A) s\<^isub>2"}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1939
  \end{center}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1940
  
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1941
  \noindent
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1942
  holds for any strings @{text "s\<^isub>1"} and @{text
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1943
  "s\<^isub>2"}. Therefore @{text A} and the complement language @{term "-A"}
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  1944
  give rise to the same partitions. So if one is finite, the other is too, and
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1945
  vice versa. Proving the existence of such a regular expression via
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  1946
  automata using the standard method would be quite involved. It includes the
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  1947
  steps: regular expression @{text "\<Rightarrow>"} non-deterministic automaton @{text
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  1948
  "\<Rightarrow>"} deterministic automaton @{text "\<Rightarrow>"} complement automaton @{text "\<Rightarrow>"}
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  1949
  regular expression. Clearly not something you want to formalise in a theorem
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  1950
  prover in which it is cumbersome to reason about automata.
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1951
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1952
  Once closure under complement is established, closure under intersection
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1953
  and set difference is also easy, because
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1954
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1955
  \begin{center}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1956
  \begin{tabular}{c}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1957
  @{term "A \<inter> B = - (- A \<union> - B)"}\\
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1958
  @{term "A - B = - (- A \<union> B)"}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1959
  \end{tabular}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1960
  \end{center}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1961
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1962
  \noindent
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  1963
  Closure of regular languages under reversal, that is
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1964
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1965
  \begin{center}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1966
  @{text "A\<^bsup>-1\<^esup> \<equiv> {s\<^bsup>-1\<^esup> | s \<in> A}"}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1967
  \end{center}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1968
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1969
  \noindent 
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  1970
  can be shown with the help of the following operation defined recursively over 
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  1971
  regular expressions
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1972
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1973
  \begin{center}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1974
  \begin{tabular}{r@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1975
  @{thm (lhs) Rev.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) Rev.simps(1)}\\
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1976
  @{thm (lhs) Rev.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) Rev.simps(2)}\\
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1977
  @{thm (lhs) Rev.simps(3)} & @{text "\<equiv>"} & @{thm (rhs) Rev.simps(3)}\\
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1978
  @{thm (lhs) Rev.simps(4)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]} & @{text "\<equiv>"} & 
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1979
      @{thm (rhs) Rev.simps(4)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]}\\
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1980
  @{thm (lhs) Rev.simps(5)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]} & @{text "\<equiv>"} & 
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1981
      @{thm (rhs) Rev.simps(5)[where ?r1.0="r\<^isub>1" and ?r2.0="r\<^isub>2"]}\\
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1982
  @{thm (lhs) Rev.simps(6)} & @{text "\<equiv>"} & @{thm (rhs) Rev.simps(6)}\\
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1983
  \end{tabular}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1984
  \end{center}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1985
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1986
  \noindent
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  1987
  For this operation we can show
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1988
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1989
  \begin{center}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1990
  @{text "(\<calL>(r))\<^bsup>-1\<^esup>"}~@{text "="}~@{thm (rhs) rev_lang}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1991
  \end{center}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1992
  
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1993
  \noindent
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  1994
  from which closure under reversal of regular languages follows.
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1995
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  1996
  A perhaps surprising fact is that regular languages are closed under any
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  1997
  left-quotient. Define
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1998
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1999
  \begin{center}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2000
  @{abbrev "Ders_lang B A"}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2001
  \end{center}
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2002
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2003
  \noindent
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2004
  and assume @{text B} is any language and @{text A} is regular, then @{term
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2005
  "Ders_lang B A"} is regular. To see this consider the following argument
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2006
  using partial derivatives: From @{text A} being regular we know there exists
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2007
  a regular expression @{text r} such that @{term "A = lang r"}. We also know
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2008
  that @{term "pders_lang B r"} is finite for every language @{text B} and 
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2009
  regular expression @{text r} (recall Theorem~\ref{antimirov}). By definition 
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2010
  and \eqref{Derspders} therefore
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2011
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2012
  
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2013
  \begin{equation}\label{eqq}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2014
  @{term "Ders_lang B (lang r) = (\<Union> lang ` (pders_lang B r))"}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2015
  \end{equation}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2016
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2017
  \noindent
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2018
  Since there are only finitely many regular expressions in @{term "pders_lang
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2019
  B r"}, we know by \eqref{uplus} that there exists a regular expression so that
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2020
  the right-hand side of \eqref{eqq} is equal to the language \mbox{@{term "lang (\<Uplus>(pders_lang B
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2021
  r))"}}. Thus the regular expression @{term "\<Uplus>(pders_lang B r)"} verifies that
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2022
  @{term "Ders_lang B A"} is regular.
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2023
*}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2024
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  2025
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
  2026
section {* Conclusion and Related Work *}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
  2027
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
  2028
text {*
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2029
  \noindent
112
62fdb4bf7239 more on the conclusion
urbanc
parents: 111
diff changeset
  2030
  In this paper we took the view that a regular language is one where there
115
c5f138b5fc88 added comment from Larry
urbanc
parents: 114
diff changeset
  2031
  exists a regular expression that matches all of its strings. Regular
145
099e20f25b25 corrected small typo
urbanc
parents: 143
diff changeset
  2032
  expressions can conveniently be defined as a datatype in HOL-based theorem
099e20f25b25 corrected small typo
urbanc
parents: 143
diff changeset
  2033
  provers. For us it was therefore interesting to find out how far we can push
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 149
diff changeset
  2034
  this point of view. We have established in Isabelle/HOL both directions 
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 149
diff changeset
  2035
  of the Myhill-Nerode theorem.
132
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
  2036
  %
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
  2037
  \begin{thrm}[The Myhill-Nerode Theorem]\mbox{}\\
132
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
  2038
  A language @{text A} is regular if and only if @{thm (rhs) Myhill_Nerode}.
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
  2039
  \end{thrm}
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2040
  
132
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
  2041
  \noindent
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2042
  Having formalised this theorem means we pushed our point of view quite
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2043
  far. Using this theorem we can obviously prove when a language is \emph{not}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2044
  regular---by establishing that it has infinitely many equivalence classes
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2045
  generated by the Myhill-Nerode relation (this is usually the purpose of the
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2046
  Pumping Lemma \cite{Kozen97}).  We can also use it to establish the standard
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2047
  textbook results about closure properties of regular languages. Interesting
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2048
  is the case of closure under complement, because it seems difficult to
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2049
  construct a regular expression for the complement language by direct
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2050
  means. However the existence of such a regular expression can be easily
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2051
  proved using the Myhill-Nerode theorem.  
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2052
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2053
  Our insistence on regular expressions for proving the Myhill-Nerode theorem
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2054
  arose from the limitations of HOL, on which the popular theorem provers HOL4,
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2055
  HOLlight and Isabelle/HOL are based. In order to guarantee consistency,
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2056
  formalisations can only extend HOL by definitions that introduce a new concept in
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2057
  terms of already existing notions. A convenient definition for automata
197
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  2058
  (based on graphs) uses a polymorphic type for the state nodes. This allows
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  2059
  us to use the standard operation of disjoint union in order to compose two
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2060
  automata. Unfortunately, we cannot use such a polymorphic definition
197
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  2061
  in HOL as part of the definition for regularity of a language (a
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  2062
  set of strings).  Consider the following attempt
112
62fdb4bf7239 more on the conclusion
urbanc
parents: 111
diff changeset
  2063
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2064
  \begin{center}
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2065
  @{text "is_regular A \<equiv> \<exists>M(\<alpha>). is_finite_automata (M) \<and> \<calL>(M) = A"}
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2066
  \end{center}
116
342983676c8f included comments by Chunhan
urbanc
parents: 115
diff changeset
  2067
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2068
  \noindent
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2069
  which means the definiens is polymorphic in the type of the automata @{text
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2070
  "M"}, but the definiendum @{text "is_regular"} is not. Such definitions are
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2071
  excluded in HOL, because they can lead easily to inconsistencies (see
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2072
  \cite{PittsHOL4} for a simple example). Also HOL does not contain
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2073
  type-quantifiers which would allow us to get rid of the polymorphism by
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2074
  quantifying over the type-variable @{text "\<alpha>"}. Therefore when defining
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2075
  regularity in terms of automata, the only natural way out in HOL is to use
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2076
  state nodes with an identity, for example a natural number. Unfortunatly,
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2077
  the consequence is that we have to be careful when combining two automata so
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2078
  that there is no clash between two such states. This makes formalisations
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2079
  quite fiddly and rather unpleasant. Regular expressions proved much more
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2080
  convenient for reasoning in HOL and we showed they can be used for
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2081
  establishing the central result in regular language theory: the Myhill-Nerode 
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2082
  theorem.
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2083
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2084
  While regular expressions are convenient, they have some limitations. One is
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2085
  that there seems to be no method of calculating a minimal regular expression
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2086
  (for example in terms of length) for a regular language, like there is for
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2087
  automata. On the other hand, efficient regular expression matching, without
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2088
  using automata, poses no problem \cite{OwensReppyTuron09}.  For an
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2089
  implementation of a simple regular expression matcher, whose correctness has
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2090
  been formally established, we refer the reader to Owens and Slind
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2091
  \cite{OwensSlind08}.
116
342983676c8f included comments by Chunhan
urbanc
parents: 115
diff changeset
  2092
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2093
  Our proof of the first direction is very much inspired by \emph{Brzozowski's
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2094
  algebraic method} used to convert a finite automaton to a regular expression
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2095
  \cite{Brzozowski64}. The close connection can be seen by considering the
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2096
  equivalence classes as the states of the minimal automaton for the regular
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2097
  language.  However there are some subtle differences. Since we identify
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2098
  equivalence classes with the states of the automaton, then the most natural
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2099
  choice is to characterise each state with the set of strings starting from
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2100
  the initial state leading up to that state. Usually, however, the states are
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2101
  characterised as the strings starting from that state leading to the
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2102
  terminal states.  The first choice has consequences about how the initial
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2103
  equational system is set up. We have the $\lambda$-term on our `initial
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2104
  state', while Brzozowski has it on the terminal states. This means we also
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2105
  need to reverse the direction of Arden's Lemma. We have not found anything
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2106
  in the literature about this way of proving the first direction of the
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2107
  Myhill-Nerode theorem, but it appears to be folklore.
112
62fdb4bf7239 more on the conclusion
urbanc
parents: 111
diff changeset
  2108
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2109
  We presented two proofs for the second direction of the Myhill-Nerode
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2110
  theorem. One direct proof using tagging-functions and another using partial
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2111
  derivatives. This part of our work is where our method using regular
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2112
  expressions shines, because we can completely side-step the standard
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2113
  argument \cite{Kozen97} where automata need to be composed. However, it is
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2114
  also the direction where we had to spend most of the `conceptual' time, as
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2115
  our first proof based on tagging-functions is new for establishing the
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2116
  Myhill-Nerode theorem. All standard proofs of this direction proceed by
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2117
  arguments over automata.
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2118
  
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2119
  The indirect proof for the second direction arose from our interest in
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2120
  Brzozowski's derivatives for regular expression matching. A corresponding
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2121
  regular expression matcher has been formalised by Owens and Slind in HOL4
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2122
  \cite{OwensSlind08}. In our opinion, their formalisation is considerably
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2123
  slicker than for example the approach to regular expression matching taken
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2124
  in \cite{Harper99} and \cite{Yi06}. While Brzozowski's derivatives lead to a
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2125
  simple regular expression matcher and he established that there are only
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2126
  finitely many dissimilar derivatives for every regular expression, this
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2127
  result is not as straightforward to formalise in a theorem prover. The
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2128
  reason is that the set of dissimilar derivatives is not defined inductively,
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2129
  but in terms of an ACI-equivalence relation. This difficulty prevented for
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2130
  example Krauss and Nipkow to prove termination of their equivalence checker
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2131
  for regular expressions \cite{KraussNipkow11}. Their checker is based on
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2132
  derivatives and for their argument the lack of a formal proof of termination
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2133
  is not crucial (it merely lets them ``sleep better'' \cite{KraussNipkow11}).
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2134
  We expect that their development simplifies by using partial derivatives,
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2135
  instead of derivatives, and that termination of the algorithm can be
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2136
  formally established. However, since partial derivatives use sets of regular
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2137
  expressions, one needs to carefully analyse whether the resulting algorithm
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2138
  is still executable. Given the existing infrastructure for executable sets
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2139
  in Isabelle/HOL, it should.
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2140
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2141
  Our formalisation of the Myhill-Nerode theorem consists of 780 lines of
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2142
  Isabelle/Isar code for the first direction and 460 for the second (the one
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2143
  based on tagging functions), plus around 300 lines of standard material
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2144
  about regular languages. The formalisation about derivatives and partial
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2145
  derivatives shown in Section~\ref{derivatives} consists of 390 lines of
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2146
  code.  The algorithm for solving equational systems, which we used in the
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2147
  first direction, is conceptually relatively simple. Still the use of sets
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2148
  over which the algorithm operates means it is not as easy to formalise as
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2149
  one might hope. However, it seems sets cannot be avoided since the `input'
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2150
  of the algorithm consists of equivalence classes and we cannot see how to
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2151
  reformulate the theory so that we can use lists. Lists would be much easier
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2152
  to reason about, since we can define function over them by recursion. For
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2153
  sets we have to use set-comprehensions.
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
  2154
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2155
  While our formalisation might be seen large, it should be seen
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2156
  in the context of the work done by Constable at al \cite{Constable00} who
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2157
  formalised the Myhill-Nerode theorem in Nuprl using automata. They write
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2158
  that their four-member team needed something on the magnitude of 18 months
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2159
  for their formalisation. The estimate for our formalisation is that we
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2160
  needed approximately 3 months and this included the time to find our proof
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2161
  arguments. Unlike Constable et al, who were able to follow the proofs from
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2162
  \cite{HopcroftUllman69}, we had to find our own arguments.  So for us the
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2163
  formalisation was not the bottleneck. It is hard to gauge the size of a
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2164
  formalisation in Nurpl, but from what is shown in the Nuprl Math Library
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2165
  about their development it seems substantially larger than ours. The code of
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2166
  ours can be found in the Mercurial Repository at
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2167
  \mbox{\url{http://www4.in.tum.de/~urbanc/regexp.html}}.\medskip
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 160
diff changeset
  2168
  
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 160
diff changeset
  2169
  \noindent
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
  2170
  {\bf Acknowledgements:}
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
  2171
  We are grateful for the comments we received from Larry
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 160
diff changeset
  2172
  Paulson.
111
d65d071798ff first ideas about conclusion
urbanc
parents: 110
diff changeset
  2173
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
  2174
*}
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
  2175
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
  2176
24
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
  2177
(*<*)
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
  2178
end
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
  2179
(*>*)