Journal/Paper.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Tue, 17 Sep 2013 11:21:58 +0100
changeset 388 0da31edd95b9
parent 387 288637d9dcde
child 392 87d3306acca8
permissions -rw-r--r--
new version
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
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
     3
imports "../Closures2" "../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
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
     6
lemma nullable_iff:
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
     7
  shows "nullable r \<longleftrightarrow> [] \<in> lang r"
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
     8
by (induct r) (auto simp add: conc_def split: if_splits)
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
     9
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
    10
lemma Deriv_deriv:
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
    11
  shows "Deriv c (lang r) = lang (deriv c r)"
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
    12
by (induct r) (simp_all add: nullable_iff)
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
    13
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
    14
lemma Derivs_derivs:
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
    15
  shows "Derivs s (lang r) = lang (derivs s r)"
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
    16
by (induct s arbitrary: r) (simp_all add: Deriv_deriv)
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
    17
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
    18
declare [[show_question_marks = false]]
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
    19
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
    20
consts
334
d47c2143ab8a partially updated conference paper; slightly tuned journal paper
urbanc
parents: 259
diff changeset
    21
 REL :: "(string \<times> string) set"
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    22
 UPLUS :: "'a set \<Rightarrow> 'a set \<Rightarrow> (nat \<times> 'a) set"
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
    23
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
    24
abbreviation
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
    25
  "EClass x R \<equiv> R `` {x}"
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
    26
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
    27
abbreviation 
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 160
diff changeset
    28
  "Append_rexp2 r_itm r \<equiv> Append_rexp r r_itm"
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
    29
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
    30
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    31
abbreviation
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    32
  "pow" (infixl "\<up>" 100)
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    33
where
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    34
  "A \<up> n \<equiv> A ^^ n"
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    35
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    36
syntax (latex output)
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    37
  "_Collect" :: "pttrn => bool => 'a set"              ("(1{_ | _})")
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    38
  "_CollectIn" :: "pttrn => 'a set => bool => 'a set"   ("(1{_ \<in> _ | _})")
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    39
translations
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    40
  "_Collect p P"      <= "{p. P}"
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    41
  "_Collect p P"      <= "{p|xs. P}"
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    42
  "_CollectIn p A P"  <= "{p : A. P}"
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    43
242
093e45c44d91 polished proposal
urbanc
parents: 240
diff changeset
    44
syntax (latex output)
093e45c44d91 polished proposal
urbanc
parents: 240
diff changeset
    45
  "_UNION_le"   :: "'a \<Rightarrow> 'a => 'b set => 'b set"       ("(3\<Union>(00\<^bsub>_ \<le> _\<^esub>)/ _)" [0, 0, 10] 10)
093e45c44d91 polished proposal
urbanc
parents: 240
diff changeset
    46
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
    47
abbreviation "ZERO \<equiv> Zero"
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
    48
abbreviation "ONE \<equiv> One"
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
    49
abbreviation "ATOM \<equiv> Atom"
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
    50
abbreviation "PLUS \<equiv> Plus"
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
    51
abbreviation "TIMES \<equiv> Times"
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
    52
abbreviation "TIMESS \<equiv> Timess"
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    53
abbreviation "STAR \<equiv> Star"
239
13de6a49294e polished SUBSEQ
urbanc
parents: 238
diff changeset
    54
abbreviation "ALLS \<equiv> Star Allreg"
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    55
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
    56
definition
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
    57
  Delta :: "'a lang \<Rightarrow> 'a lang"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
    58
where
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
    59
  "Delta A = (if [] \<in> A then {[]} else {})"
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    60
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
    61
notation (latex output)
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
    62
  str_eq ("\<approx>\<^bsub>_\<^esub>") and
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
    63
  str_eq_applied ("_ \<approx>\<^bsub>_\<^esub> _") and
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    64
  conc (infixr "\<cdot>" 100) and
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
    65
  star ("_\<^bsup>\<star>\<^esup>" [101] 200) and
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
    66
  pow ("_\<^bsup>_\<^esup>" [100, 100] 100) and
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
    67
  Suc ("_+1" [100] 100) and
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
    68
  quotient ("_ \<^raw:\ensuremath{\!\sslash\!}> _" [90, 90] 90) and
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
    69
  REL ("\<approx>") and
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
    70
  UPLUS ("_ \<^raw:\ensuremath{\uplus}> _" [90, 90] 90) and
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
    71
  lang ("\<^raw:\ensuremath{\cal{L}}>" 101) and
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    72
  lang ("\<^raw:\ensuremath{\cal{L}}>'(_')" [0] 101) and
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
    73
  lang_trm ("\<^raw:\ensuremath{\cal{L}}>'(_')" [0] 101) and
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
    74
  Lam ("\<lambda>'(_')" [100] 100) and 
89
42af13d194c9 a bit more on the paper
urbanc
parents: 88
diff changeset
    75
  Trn ("'(_, _')" [100, 100] 100) and 
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
    76
  EClass ("\<lbrakk>_\<rbrakk>\<^bsub>_\<^esub>" [100, 100] 100) and
88
1436fc451bb9 added something about Setalt and folds
urbanc
parents: 86
diff changeset
    77
  transition ("_ \<^raw:\ensuremath{\stackrel{\text{>_\<^raw:}}{\Longmapsto}}> _" [100, 100, 100] 100) and
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
    78
  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
    79
  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
    80
  Append_rexp_rhs ("_ \<^raw:\ensuremath{\triangleleft}> _" [100, 100] 50) and
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
    81
  
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
    82
  uminus ("\<^raw:\ensuremath{\overline{\isa{>_\<^raw:}}}>" [100] 100) and
183
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
    83
  tag_Plus ("+tag _ _" [100, 100] 100) and
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
    84
  tag_Plus ("+tag _ _ _" [100, 100, 100] 100) and
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
    85
  tag_Times ("\<times>tag _ _" [100, 100] 100) and
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
    86
  tag_Times ("\<times>tag _ _ _" [100, 100, 100] 100) and
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
    87
  tag_Star ("\<star>tag _" [100] 100) and
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
    88
  tag_Star ("\<star>tag _ _" [100, 100] 100) and
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
    89
  tag_eq ("\<^raw:$\threesim$>\<^bsub>_\<^esub>" [100] 100) and
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
    90
  Delta ("\<Delta>'(_')") and
180
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
    91
  nullable ("\<delta>'(_')") and
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
    92
  Cons ("_ :: _" [100, 100] 100) and
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
    93
  Rev ("Rev _" [1000] 100) and
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
    94
  Deriv ("Der _ _" [1000, 1000] 100) and
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
    95
  Derivs ("Ders") and
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
    96
  ONE ("ONE" 999) and
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
    97
  ZERO ("ZERO" 999) and
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
    98
  pderivs_lang ("pdersl") and
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
    99
  UNIV1 ("UNIV\<^sup>+") and
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   100
  Deriv_lang ("Dersl") and
217
05da74214979 a few bits on the journal paper
urbanc
parents: 203
diff changeset
   101
  Derivss ("Derss") and
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   102
  deriv ("der") and
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   103
  derivs ("ders") and
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   104
  pderiv ("pder") and
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   105
  pderivs ("pders") and
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
   106
  pderivs_set ("pderss") and
240
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
   107
  SUBSEQ ("Sub") and
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
   108
  SUPSEQ ("Sup") and
239
13de6a49294e polished SUBSEQ
urbanc
parents: 238
diff changeset
   109
  UP ("'(_')\<up>") and
13de6a49294e polished SUBSEQ
urbanc
parents: 238
diff changeset
   110
  ALLS ("ALL")
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   111
  
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   112
  
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   113
lemmas Deriv_simps = Deriv_empty Deriv_epsilon Deriv_char Deriv_union
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   114
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   115
definition
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   116
  Der :: "'a \<Rightarrow> 'a lang \<Rightarrow> 'a lang"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   117
where
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   118
  "Der c A \<equiv> {s. [c] @ s \<in> A}"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   119
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   120
definition
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   121
  Ders :: "'a list \<Rightarrow> 'a lang \<Rightarrow> 'a lang"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   122
where
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   123
  "Ders s A \<equiv> {s'. s @ s' \<in> A}"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   124
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   125
119
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
   126
lemma meta_eq_app:
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
   127
  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
   128
  by auto
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
   129
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   130
lemma str_eq_def':
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   131
  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
   132
unfolding str_eq_def by simp
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   133
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   134
lemma conc_def':
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   135
  "A \<cdot> B = {s\<^sub>1 @ s\<^sub>2 | s\<^sub>1 s\<^sub>2. s\<^sub>1 \<in> A \<and> s\<^sub>2 \<in> B}"
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   136
unfolding conc_def by simp
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   137
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   138
lemma conc_Union_left: 
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   139
  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
   140
unfolding conc_def by auto
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   141
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   142
lemma test:
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   143
  assumes X_in_eqs: "(X, rhs) \<in> Init (UNIV // \<approx>A)"
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   144
  shows "X = \<Union> (lang_trm `  rhs)"
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   145
using assms l_eq_r_in_eqs by (simp)
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   146
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
   147
abbreviation
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
   148
  notprec ("_ \<^raw:\mbox{$\not\preceq$}>_")
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
   149
where
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
   150
 "notprec x y \<equiv> \<not>(x \<preceq> y)"
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
   151
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
   152
abbreviation
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
   153
  minimal_syn ("min\<^bsub>_\<^esub> _")
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
   154
where
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
   155
  "minimal_syn A x \<equiv> minimal x A"   
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
   156
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   157
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   158
(* THEOREMS *)
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   159
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   160
notation (Rule output)
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   161
  "==>"  ("\<^raw:\mbox{}\inferrule{\mbox{>_\<^raw:}}>\<^raw:{\mbox{>_\<^raw:}}>")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   162
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   163
syntax (Rule output)
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   164
  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   165
  ("\<^raw:\mbox{}\inferrule{>_\<^raw:}>\<^raw:{\mbox{>_\<^raw:}}>")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   166
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   167
  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" 
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   168
  ("\<^raw:\mbox{>_\<^raw:}\\>/ _")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   169
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   170
  "_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   171
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   172
notation (Axiom output)
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   173
  "Trueprop"  ("\<^raw:\mbox{}\inferrule{\mbox{}}{\mbox{>_\<^raw:}}>")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   174
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   175
notation (IfThen output)
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   176
  "==>"  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   177
syntax (IfThen output)
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   178
  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   179
  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   180
  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}> /\<^raw:{\normalsize \,>and\<^raw:\,}>/ _")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   181
  "_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   182
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   183
notation (IfThenNoBox output)
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   184
  "==>"  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   185
syntax (IfThenNoBox output)
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   186
  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   187
  ("\<^raw:{\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\normalsize \,>then\<^raw:\,}>/ _.")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   188
  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("_ /\<^raw:{\normalsize \,>and\<^raw:\,}>/ _")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   189
  "_asm" :: "prop \<Rightarrow> asms" ("_")
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   190
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   191
lemma pow_length:
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   192
  assumes a: "[] \<notin> A"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   193
  and     b: "s \<in> A \<up> Suc n"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   194
  shows "n < length s"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   195
using b
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   196
proof (induct n arbitrary: s)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   197
  case 0
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   198
  have "s \<in> A \<up> Suc 0" by fact
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   199
  with a have "s \<noteq> []" by auto
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   200
  then show "0 < length s" by auto
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   201
next
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   202
  case (Suc n)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   203
  have ih: "\<And>s. s \<in> A \<up> Suc n \<Longrightarrow> n < length s" by fact
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   204
  have "s \<in> A \<up> Suc (Suc n)" by fact
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   205
  then obtain s1 s2 where eq: "s = s1 @ s2" and *: "s1 \<in> A" and **: "s2 \<in> A \<up> Suc n"
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   206
    by (auto simp add: Seq_def)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   207
  from ih ** have "n < length s2" by simp
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   208
  moreover have "0 < length s1" using * a by auto
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   209
  ultimately show "Suc n < length s" unfolding eq 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   210
    by (simp only: length_append)
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   211
qed
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   212
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
   213
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
   214
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
   215
24
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   216
(*>*)
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   217
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   218
24
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   219
section {* Introduction *}
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   220
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   221
text {*
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   222
  \noindent
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   223
  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
   224
  Science, with many beautiful theorems and many useful algorithms. There is a
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   225
  wide range of textbooks on this subject, many of which are aimed at students
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
   226
  and contain very detailed `pencil-and-paper' proofs 
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   227
  (e.g.~the textbooks by \citeN{HopcroftUllman69} and by \citeN{Kozen97}). 
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   228
  It seems natural to exercise theorem provers by
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   229
  formalising the theorems and by verifying formally the algorithms.  
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   230
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   231
  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
   232
  Logic (HOL), for example HOL4, HOLlight or Isabelle/HOL. For the development
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   233
  presented in this paper we will use the Isabelle/HOL system. HOL is a predicate calculus
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   234
  that allows quantification over predicate variables. Its type system is
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   235
  based on the Simple Theory of Types by \citeN{Church40}.  Although many
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   236
  mathematical concepts can be conveniently expressed in HOL, there are some
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   237
  limitations that hurt when attempting a simple-minded formalisation
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
   238
  of regular languages in it. 
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
   239
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   240
  The typical approach to
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   241
  regular languages, taken for example by \citeN{HopcroftUllman69}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   242
  and by \citeN{Kozen97},  is to introduce finite deterministic automata and then
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   243
  define most notions in terms of them.  For example, a regular language is
201
9fbf6d9f85ae added comments by Xingyuan
urbanc
parents: 200
diff changeset
   244
  normally defined as:
59
fc35eb54fdc9 more on the intro
urbanc
parents: 58
diff changeset
   245
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   246
  \begin{definition}\label{baddef}
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   247
  A language @{text A} is \emph{regular}, provided there is a 
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   248
  finite deterministic automaton that recognises all strings of @{text "A"}.
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   249
  \end{definition}
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   250
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   251
  \noindent  
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   252
  This approach has many benefits. Among them is the fact that it is easy to
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   253
  convince oneself that regular languages are closed under complementation:
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   254
  one just has to exchange the accepting and non-accepting states in the
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   255
  corresponding automaton to obtain an automaton for the complement language.
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
   256
  The problem, however, lies with formalising such reasoning in a 
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
   257
  theorem prover. Automata are built up from states and transitions that are 
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   258
  commonly represented as graphs, matrices or functions, none of which, unfortunately, 
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   259
  can be defined as an inductive datatype.
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   260
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   261
  In case of graphs and matrices, this means we have to build our own
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   262
  reasoning infrastructure for them, as neither Isabelle/HOL nor HOL4 nor
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   263
  HOLlight support them with libraries. Also, reasoning about graphs and
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   264
  matrices can be a hassle in theorem provers, because
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   265
  we have to be able to combine automata.  Consider for
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   266
  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
   267
  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
   268
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
   269
  \begin{center}
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   270
  \begin{tabular}{ccc}
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   271
  \begin{tikzpicture}[scale=1]
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   272
  %\draw[step=2mm] (-1,-1) grid (1,1);
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   273
  
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   274
  \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
   275
  \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
   276
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   277
  \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
   278
  \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
   279
  
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   280
  \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
   281
  \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
   282
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   283
  \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
   284
  \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
   285
  \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
   286
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   287
  \draw (-0.6,0.0) node {\small$A_1$};
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   288
  \draw ( 0.6,0.0) node {\small$A_2$};
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   289
  \end{tikzpicture}
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   290
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   291
  & 
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   292
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
   293
  \raisebox{2.1mm}{\bf\Large$\;\;\;\Rightarrow\,\;\;$}
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   294
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   295
  &
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   296
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   297
  \begin{tikzpicture}[scale=1]
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   298
  %\draw[step=2mm] (-1,-1) grid (1,1);
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   299
  
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   300
  \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
   301
  \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
   302
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   303
  \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
   304
  \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
   305
  
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   306
  \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
   307
  \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
   308
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   309
  \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
   310
  \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
   311
  \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
   312
  
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   313
  \draw (C) to [very thick, bend left=45] (B);
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   314
  \draw (D) to [very thick, bend right=45] (B);
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   315
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   316
  \draw (-0.6,0.0) node {\small$A_1$};
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   317
  \draw ( 0.6,0.0) node {\small$A_2$};
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   318
  \end{tikzpicture}
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   319
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   320
  \end{tabular}
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
   321
  \end{center}
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
   322
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
   323
  \noindent
178
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   324
  On `paper' we can define the corresponding 
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   325
  graph in terms of the disjoint 
88
1436fc451bb9 added something about Setalt and folds
urbanc
parents: 86
diff changeset
   326
  union of the state nodes. Unfortunately in HOL, the standard definition for disjoint 
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   327
  union, namely 
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   328
  %
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   329
  \begin{equation}\label{disjointunion}
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   330
  @{text "A\<^sub>1 \<uplus> A\<^sub>2 \<equiv> {(1, x) | x \<in> A\<^sub>1} \<union> {(2, y) | y \<in> A\<^sub>2}"}
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   331
  \end{equation}
60
fb08f41ca33d a bit more tuning on the introduction
urbanc
parents: 59
diff changeset
   332
61
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
   333
  \noindent
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
   334
  changes the type---the disjoint union is not a set, but a set of
386
92ca56c1a199 soem small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 385
diff changeset
   335
  pairs. Using this definition for disjoint union means we do not have
92ca56c1a199 soem small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 385
diff changeset
   336
  a single type for the states of automata. As a result we will not be
92ca56c1a199 soem small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 385
diff changeset
   337
  able to define a regular language as one for which there exists an
92ca56c1a199 soem small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 385
diff changeset
   338
  automaton that recognises all its strings
92ca56c1a199 soem small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 385
diff changeset
   339
  (Definition~\ref{baddef}). This is because we cannot make a
92ca56c1a199 soem small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 385
diff changeset
   340
  definition in HOL that is only polymorphic in the state type, but
92ca56c1a199 soem small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 385
diff changeset
   341
  not in the predicate for regularity; and there is no type
92ca56c1a199 soem small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 385
diff changeset
   342
  quantification available in HOL.\footnote{Slind already pointed out
92ca56c1a199 soem small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 385
diff changeset
   343
  this problem in an email to the HOL4 mailing list on 21st April
387
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   344
  2005.} Systems such as Coq permit quantification over types and thus can
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   345
  state such a definition.  This has been recently exploited in an
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   346
  elegant and short formalisation of the Myhill-Nerode theorem in Coq by
388
0da31edd95b9 new version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 387
diff changeset
   347
  \citeN{Smolka13}, but as said this is not available to us working in Isabelle/HOL.
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
   348
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   349
  An alternative, which provides us with a single type for states of automata,
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   350
  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
   351
  then be careful to rename these identities apart whenever connecting two
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   352
  automata. This results in clunky proofs establishing that properties are
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   353
  invariant under renaming. Similarly, connecting two automata represented as
372
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
   354
  matrices results in messy constructions, which are not pleasant to
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
   355
  formally reason about. \citeN[Page 67]{Braibant12}, for example, writes that there are no 
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
   356
  problems with reasoning about matrices, but that there is an
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
   357
  ``intrinsic difficulty of working with rectangular matrices'' in some 
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
   358
  parts of his formalisation of formal languages in Coq.
66
828ea293b61f more on the introduction
urbanc
parents: 61
diff changeset
   359
200
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
   360
  Functions are much better supported in Isabelle/HOL, but they still lead to
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
   361
  similar problems as with graphs.  Composing, for example, two
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
   362
  non-deterministic automata in parallel requires also the formalisation of
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   363
  disjoint unions. \citeN{Nipkow98} dismisses for this the option of
200
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
   364
  using identities, because it leads according to him to ``messy
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
   365
  proofs''. Since he does not need to define what regular languages are,
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
   366
  Nipkow opts for a variant of \eqref{disjointunion} using bit lists, but
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
   367
  writes\smallskip
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
   368
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   369
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   370
  \begin{quote}
93
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   371
  \it%
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   372
  \begin{tabular}{@ {}l@ {}p{0.88\textwidth}@ {}}
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   373
  `` & All lemmas appear obvious given a picture of the composition of automata\ldots
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   374
       Yet their proofs require a painful amount of detail.''
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   375
  \end{tabular}
200
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
   376
  \end{quote}\smallskip
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   377
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   378
  \noindent
200
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
   379
  and\smallskip
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   380
  
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   381
  \begin{quote}
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   382
  \it%
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   383
  \begin{tabular}{@ {}l@ {}p{0.88\textwidth}@ {}}
93
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   384
  `` & If the reader finds the above treatment in terms of bit lists revoltingly
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   385
       concrete, I cannot disagree. A more abstract approach is clearly desirable.''
93
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   386
  \end{tabular}
200
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
   387
  \end{quote}\smallskip
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   388
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   389
  %\noindent
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   390
  %Moreover, it is not so clear how to conveniently impose a finiteness
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   391
  %condition upon functions in order to represent \emph{finite} automata. The
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   392
  %best is probably to resort to more advanced reasoning frameworks, such as
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   393
  %\emph{locales} or \emph{type classes}, which are \emph{not} available in all
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   394
  %HOL-based theorem provers.
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   395
372
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
   396
  Because of these problems to do with representing automata, formalising 
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
   397
  automata theory is surprisingly not as easy as one might think, despite the 
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
   398
  sometimes very detailed, but informal, textbook proofs. \citeN{LammichTuerk12}
374
01d223421ba0 slight changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 372
diff changeset
   399
  formalised Hopcroft's algorithm using an automata library of 14 kloc in
372
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
   400
  Isabelle/HOL. There they use matrices for representing automata. 
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
   401
  Functions are used by \citeN{Nipkow98}, who establishes
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   402
  the link between regular expressions and automata in the context of
372
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
   403
  lexing. \citeN{BerghoferReiter09} use functions as well for formalising automata
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
   404
  working over bit strings in the context of Presburger arithmetic.  
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   405
  A larger formalisation of automata theory, including the Myhill-Nerode theorem, 
372
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
   406
  was carried out in Nuprl by \citeN{Constable00} which also uses functions.
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   407
  Other large formalisations of automata theory in Coq are by 
372
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
   408
  \citeN{Filliatre97} who essentially uses graphs and by \citeN{Almeidaetal10}
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
   409
  and \citeN{Braibant12}, who both use matrices. Many of these works,
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
   410
  like \citeN{Nipkow98} or \citeN{Braibant12}, mention intrinsic problems with 
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
   411
  their representation of
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
   412
  automata which made them `fight' their respective  theorem prover.
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 160
diff changeset
   413
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   414
  %Also, one might consider automata as just convenient `vehicles' for
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   415
  %establishing properties about regular languages.  However, paper proofs
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   416
  %about automata often involve subtle side-conditions which are easily
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   417
  %overlooked, but which make formal reasoning rather painful. For example
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   418
  %Kozen's proof of the Myhill-Nerode Theorem requires that automata do not
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   419
  %have inaccessible states \cite{Kozen97}. Another subtle side-condition is
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   420
  %completeness of automata, that is automata need to have total transition
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   421
  %functions and at most one `sink' state from which there is no connection to
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   422
  %a final state (Brzozowski mentions this side-condition in the context of
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   423
  %state complexity of automata \cite{Brzozowski10}, but it is also needed
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   424
  %in the usual construction of the complement automaton). Such side-conditions mean
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   425
  %that if we define a regular language as one for which there exists \emph{a}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   426
  %finite automaton that recognises all its strings (see
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   427
  %Definition~\ref{baddef}), then we need a lemma which ensures that another
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   428
  %equivalent one can be found satisfying the side-condition, and also need to
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   429
  %make sure our operations on automata preserve them. Unfortunately, such
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   430
  %`little' and `obvious' lemmas make formalisations of automata theory 
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   431
  %hair-pulling experiences.
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   432
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   433
  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
   434
  Isabelle/HOL nor will we attempt to formalise automata proofs from the
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   435
  literature, but take a different approach to regular languages than is
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   436
  usually taken. Instead of defining a regular language as one where there
178
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   437
  exists an automaton that recognises all its strings, we define a
82
14b12b5de6d3 added coments about functions
urbanc
parents: 79
diff changeset
   438
  regular language as:
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   439
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   440
  \begin{definition}\label{regular}
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
   441
  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
   442
  that matches all strings of @{text "A"}.
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   443
  \end{definition}
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   444
  
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   445
  \noindent
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   446
  We then want to `forget' automata completely and see how far we come
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
   447
  with formalising results from regular language theory by only using regular 
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
   448
  expressions.  The reason 
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   449
  is that regular expressions, unlike graphs, matrices and
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   450
  functions, can be defined as an inductive datatype and a reasoning
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   451
  infrastructure for them (like induction and recursion) comes for
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
   452
  free in HOL. 
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
   453
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
   454
  While our choice of using regular expressions is motivated by
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
   455
  shortcomings in theorem provers, the question whether formal
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
   456
  language theory can be done without automata crops up also in
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
   457
  non-theorem prover contexts. For example, Gasarch asked in the
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
   458
  Computational Complexity blog by \citeN{GasarchBlog} whether the
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
   459
  complementation of regular-expression languages can be proved
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
   460
  without using automata. He concluded
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   461
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   462
  \begin{quote}
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   463
  ``{\it \ldots it can't be done}''
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   464
  \end{quote} 
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   465
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   466
  \noindent
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   467
  and even pondered
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   468
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   469
  \begin{quote}
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   470
  ``{\it \ldots [b]ut is there a rigorous way to even state that?}'' 
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   471
  \end{quote} 
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   472
387
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   473
  \noindent
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   474
  We shall give an answer to these questions with our paper.
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   475
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   476
  %Moreover, no side-conditions will be needed for regular expressions,
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   477
  %like we need for automata. 
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
   478
 
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   479
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   480
  The convenience of regular expressions has
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   481
  recently been exploited in HOL4 with a formalisation of regular expression
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   482
  matching based on derivatives by \citeN{OwensSlind08}, and with an equivalence
372
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
   483
  checker for regular expressions in Isabelle/HOL by \citeN{KraussNipkow11}
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   484
  and in Matita by \citeN{Asperti12} and in Coq by \citeN{CoquandSiles12}.  The
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   485
  main purpose of this paper is to show that a central result about regular
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
   486
  languages---the Myhill-Nerode Theorem---can be recreated by only using
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   487
  regular expressions. This theorem gives necessary and sufficient conditions
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   488
  for when a language is regular. As a corollary of this theorem we can easily
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   489
  establish the usual closure properties, including complementation, for
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   490
  regular languages. We use the Continuation Lemma, which is also a corollary 
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   491
  of the Myhill-Nerode Theorem, for establishing 
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   492
  the non-regularity of the language @{text "a\<^sup>nb\<^sup>n"}.\medskip 
61
070f543e2560 more to the intro
urbanc
parents: 60
diff changeset
   493
  
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
   494
  \noindent 
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   495
  {\bf Contributions:} There is an extensive literature on regular languages.
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
   496
  To our best knowledge, our proof of the Myhill-Nerode Theorem is the first
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   497
  that is based on regular expressions, only. The part of this theorem stating
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   498
  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
   499
  an argument about solving equational systems.  This argument appears to be
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   500
  folklore. For the other part, we give two proofs: one direct proof using
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
   501
  certain tagging-functions, and another indirect proof using Antimirov's
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   502
  partial derivatives (\citeyear{Antimirov95}). Again to our best knowledge, the
245
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
   503
  tagging-functions have not been used before for establishing the Myhill-Nerode
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
   504
  Theorem. Derivatives of regular expressions have been used recently quite
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   505
  widely in the literature; partial derivatives, in contrast, attract much
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   506
  less attention. However, partial derivatives are more suitable in the
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   507
  context of the Myhill-Nerode Theorem, since it is easier to 
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   508
  formally establish their finiteness result. We are not aware of any proof that uses
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
   509
  either of them for proving the Myhill-Nerode Theorem.
24
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   510
*}
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
   511
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   512
section {* Preliminaries *}
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   513
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   514
text {*
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   515
  \noindent
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   516
  Strings in Isabelle/HOL are lists of characters with the \emph{empty string}
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   517
  being represented by the empty list, written @{term "[]"}.  We assume there
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   518
  are only finitely many different characters.  \emph{Languages} are sets of
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   519
  strings. The language containing all strings is written in Isabelle/HOL as
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   520
  @{term "UNIV::string set"}. The concatenation of two languages is written
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   521
  @{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
   522
  @{term "A \<up> n"}. They are defined as usual
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   523
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   524
  \begin{center}
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   525
  \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   526
  @{thm (lhs) conc_def'[THEN eq_reflection, where A1="A" and B1="B"]}
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   527
  & @{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
   528
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   529
  @{thm (lhs) lang_pow.simps(1)[THEN eq_reflection, where A1="A"]}
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   530
  & @{text "\<equiv>"} & @{thm (rhs) lang_pow.simps(1)[THEN eq_reflection, where A1="A"]}\\
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   531
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   532
  @{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
   533
  & @{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
   534
  \end{tabular}
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   535
  \end{center}
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
  \noindent
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
   538
  where @{text "@"} is the list-append operation. The Kleene-star of a language @{text A}
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   539
  is defined as the union over all powers, namely @{thm star_def[where A="A", THEN eq_reflection]}. 
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   540
  In the paper
88
1436fc451bb9 added something about Setalt and folds
urbanc
parents: 86
diff changeset
   541
  we will make use of the following properties of these constructions.
58
0d4d5bb321dc a little bit in the introduction
urbanc
parents: 54
diff changeset
   542
  
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   543
  \begin{proposition}\label{langprops}\mbox{}\\
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   544
  \begin{tabular}{@ {}lp{10cm}}
180
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
   545
  (i)   & @{thm star_unfold_left}     \\ 
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
   546
  (ii)  & @{thm[mode=IfThen] pow_length}\\
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   547
  (iii) & @{thm conc_Union_left} \\ 
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
   548
  (iv)  & If @{thm (prem 1) star_decom} and @{thm (prem 2) star_decom} then
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   549
          there exists an @{text "x\<^sub>p"} and @{text "x\<^sub>s"} with @{text "x = x\<^sub>p @ x\<^sub>s"} 
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   550
          and \mbox{@{term "x\<^sub>p \<noteq> []"}} such that @{term "x\<^sub>p \<in> A"} and @{term "x\<^sub>s \<in> A\<star>"}.
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   551
  \end{tabular}
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   552
  \end{proposition}
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   553
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   554
  \noindent
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   555
  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
   556
  string; this property states that if \mbox{@{term "[] \<notin> A"}} then the lengths of
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   557
  the strings in @{term "A \<up> (Suc n)"} must be longer than @{text n}.  
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   558
  Property @{text "(iv)"} states that a non-empty string in @{term "A\<star>"} can
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   559
  always be split up into a non-empty prefix belonging to @{text "A"} and the
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   560
  rest being in @{term "A\<star>"}. We omit
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   561
  the proofs for these properties, but invite the reader to consult our
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   562
  formalisation.\footnote{Available under \citeN{myhillnerodeafp11} in the Archive of Formal Proofs at\\ 
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   563
  \url{http://afp.sourceforge.net/entries/Myhill-Nerode.shtml}.}
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   564
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   565
  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
   566
  according to an equivalence relation @{term REL} is @{term "A // REL"}. We
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   567
  will write @{text "\<lbrakk>x\<rbrakk>\<^sub>\<approx>"} for the equivalence class defined as
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   568
  \mbox{@{text "{y | y \<approx> x}"}}, and have @{text "x \<approx> y"} if and only if @{text
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   569
  "\<lbrakk>x\<rbrakk>\<^sub>\<approx> = \<lbrakk>y\<rbrakk>\<^sub>\<approx>"}.
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   570
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   571
51
6cfb92de4654 some tuning of the paper
urbanc
parents: 50
diff changeset
   572
  Central to our proof will be the solution of equational systems
176
6969de1eb96b latest version of the journal paper
urbanc
parents: 175
diff changeset
   573
  involving equivalence classes of languages. For this we will use Arden's Lemma 
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
   574
  (see for example \citet[Page 100]{Sakarovitch09}),
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
   575
  which solves equations of the form @{term "X = A \<cdot> X \<union> B"} provided
201
9fbf6d9f85ae added comments by Xingyuan
urbanc
parents: 200
diff changeset
   576
  @{term "[] \<notin> A"}. However we will need the following `reversed' 
9fbf6d9f85ae added comments by Xingyuan
urbanc
parents: 200
diff changeset
   577
  version of Arden's Lemma (`reversed' in the sense of changing the order of @{term "A \<cdot> X"} to
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   578
  \mbox{@{term "X \<cdot> A"}}).
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   579
  %\footnote{The details of the proof for the reversed Arden's Lemma
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   580
  %are given in the Appendix.}
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   581
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   582
  \begin{lemma}[Reversed Arden's Lemma]\label{arden}\mbox{}\\
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   583
  If @{thm (prem 1) reversed_Arden} then
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   584
  @{thm (lhs) reversed_Arden} if and only if
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   585
  @{thm (rhs) reversed_Arden}.
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   586
  \end{lemma}
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   587
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   588
  \noindent
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   589
  The proof of this reversed version of Arden's lemma follows the proof of the
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   590
  original version.
88
1436fc451bb9 added something about Setalt and folds
urbanc
parents: 86
diff changeset
   591
  Regular expressions are defined as the inductive datatype
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   592
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   593
  \begin{center}
176
6969de1eb96b latest version of the journal paper
urbanc
parents: 175
diff changeset
   594
  \begin{tabular}{rcl}
6969de1eb96b latest version of the journal paper
urbanc
parents: 175
diff changeset
   595
  @{text r} & @{text "::="} & @{term ZERO}\\
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   596
   & @{text"|"} & @{term One}\\ 
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   597
   & @{text"|"} & @{term "Atom c"}\\
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   598
   & @{text"|"} & @{term "Times r r"}\\
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   599
   & @{text"|"} & @{term "Plus r r"}\\
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   600
   & @{text"|"} & @{term "Star r"}
176
6969de1eb96b latest version of the journal paper
urbanc
parents: 175
diff changeset
   601
  \end{tabular}
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   602
  \end{center}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   603
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   604
  \noindent
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   605
  and the language matched by a regular expression is defined by recursion as
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   606
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   607
  \begin{center}
176
6969de1eb96b latest version of the journal paper
urbanc
parents: 175
diff changeset
   608
  \begin{tabular}{r@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   609
  @{thm (lhs) lang.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) lang.simps(1)}\\
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   610
  @{thm (lhs) lang.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) lang.simps(2)}\\
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   611
  @{thm (lhs) lang.simps(3)[where a="c"]} & @{text "\<equiv>"} & @{thm (rhs) lang.simps(3)[where a="c"]}\\
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   612
  @{thm (lhs) lang.simps(4)[where ?r="r\<^sub>1" and ?s="r\<^sub>2"]} & @{text "\<equiv>"} &
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   613
       @{thm (rhs) lang.simps(4)[where ?r="r\<^sub>1" and ?s="r\<^sub>2"]}\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   614
  @{thm (lhs) lang.simps(5)[where ?r="r\<^sub>1" and ?s="r\<^sub>2"]} & @{text "\<equiv>"} &
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   615
       @{thm (rhs) lang.simps(5)[where ?r="r\<^sub>1" and ?s="r\<^sub>2"]}\\
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   616
  @{thm (lhs) lang.simps(6)[where r="r"]} & @{text "\<equiv>"} &
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   617
      @{thm (rhs) lang.simps(6)[where r="r"]}\\
67
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   618
  \end{tabular}
7478be786f87 more intro
urbanc
parents: 66
diff changeset
   619
  \end{center}
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   620
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   621
  Given a finite set of regular expressions @{text rs}, we will make use of the operation of generating 
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
   622
  a regular expression that matches the union of all languages of @{text rs}. 
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   623
  This definition is not trivial in a theorem prover, because @{text rs} (being a set) is unordered, 
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   624
  but the regular expression needs an order. Since 
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
   625
  we only need to know the 
132
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
   626
  existence
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
   627
  of such a regular expression, we can use Isabelle/HOL's @{const "fold_graph"} and Hilbert's
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   628
  choice operator, written @{text "SOME"} in Isabelle/HOL, for defining @{term "\<Uplus>rs"}. 
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   629
  This operation, roughly speaking, folds @{const PLUS} over the 
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
   630
  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
   631
  %
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   632
  \begin{equation}\label{uplus}
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
   633
  \mbox{@{thm (lhs) folds_plus_simp} @{text "= \<Union> (\<calL> ` rs)"}}
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   634
  \end{equation}
88
1436fc451bb9 added something about Setalt and folds
urbanc
parents: 86
diff changeset
   635
1436fc451bb9 added something about Setalt and folds
urbanc
parents: 86
diff changeset
   636
  \noindent
90
97b783438316 added an example
urbanc
parents: 89
diff changeset
   637
  holds, whereby @{text "\<calL> ` rs"} stands for the 
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   638
  image of the set @{text rs} under function @{text "\<calL>"} defined as
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   639
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   640
  \begin{center}
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   641
  @{term "lang ` rs \<equiv> {lang r | r. r \<in> rs}"}
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   642
  \end{center}
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   643
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   644
  \noindent
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   645
  In what follows we shall use this convenient short-hand notation for images of sets 
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
   646
  also with other functions.
50
32bff8310071 revised proof of Ardens lemma
urbanc
parents: 39
diff changeset
   647
*}
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
   648
133
3ab755a96cef minor change
urbanc
parents: 132
diff changeset
   649
section {* The Myhill-Nerode Theorem, First Part *}
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   650
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
   651
text {*
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   652
  \noindent
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
   653
  The key definition in the Myhill-Nerode Theorem is the
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   654
  \emph{Myhill-Nerode Relation}, which states that  two 
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   655
  strings are related w.r.t.~a language, provided there is no distinguishing extension in this
338
e7504bfdbd50 made the changes thes 2nd referee suggested and made it to compile again
urbanc
parents: 334
diff changeset
   656
  language. This can be defined as a ternary relation.
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   657
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   658
  \begin{definition}[Myhill-Nerode Relation]\label{myhillneroderel} 
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
   659
  Given a language @{text A}, two strings @{text x} and
123
23c0e6f2929d polished everywhere...two cases still missing
urbanc
parents: 122
diff changeset
   660
  @{text y} are Myhill-Nerode related provided
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
   661
  \begin{center}
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
   662
  @{thm str_eq_def'}
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
   663
  \end{center}
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   664
  \end{definition}
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   665
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   666
  \noindent
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   667
  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
   668
  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
   669
  equivalence classes. To illustrate this quotient construction, let us give a simple 
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
   670
  example: consider the regular language built up over the alphabet @{term "{a, b}"} and 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
   671
  containing just the string two strings @{text "[a]"} and @{text "[a, b]"}. The 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
   672
  relation @{term "\<approx>({[a], [a, b]})"} partitions @{text UNIV}
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   673
  into four equivalence classes @{text "X\<^sub>1"}, @{text "X\<^sub>2"}, @{text "X\<^sub>3"} and @{text "X\<^sub>4"}
90
97b783438316 added an example
urbanc
parents: 89
diff changeset
   674
  as follows
97b783438316 added an example
urbanc
parents: 89
diff changeset
   675
  
97b783438316 added an example
urbanc
parents: 89
diff changeset
   676
  \begin{center}
176
6969de1eb96b latest version of the journal paper
urbanc
parents: 175
diff changeset
   677
  \begin{tabular}{l}
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   678
  @{text "X\<^sub>1 = {[]}"}\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   679
  @{text "X\<^sub>2 = {[a]}"}\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   680
  @{text "X\<^sub>3 = {[a, b]}"}\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   681
  @{text "X\<^sub>4 = UNIV - {[], [a], [a, b]}"}
176
6969de1eb96b latest version of the journal paper
urbanc
parents: 175
diff changeset
   682
  \end{tabular}
90
97b783438316 added an example
urbanc
parents: 89
diff changeset
   683
  \end{center}
97b783438316 added an example
urbanc
parents: 89
diff changeset
   684
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
   685
  One direction of the Myhill-Nerode Theorem establishes 
93
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   686
  that if there are finitely many equivalence classes, like in the example above, then 
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   687
  the language is regular. In our setting we therefore have to show:
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   688
  
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   689
  \begin{theorem}\label{myhillnerodeone}
96
3b9deda4f459 simplified a bit the proof
urbanc
parents: 95
diff changeset
   690
  @{thm[mode=IfThen] Myhill_Nerode1}
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   691
  \end{theorem}
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   692
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   693
  \noindent
90
97b783438316 added an example
urbanc
parents: 89
diff changeset
   694
  To prove this theorem, we first define the set @{term "finals A"} as those equivalence
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   695
  classes from \mbox{@{term "UNIV // \<approx>A"}} that contain strings of @{text A}, namely
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   696
  %
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   697
  \begin{equation} 
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   698
  @{thm finals_def}
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   699
  \end{equation}
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   700
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   701
  \noindent
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   702
  In our running example, @{text "X\<^sub>2"} and @{text "X\<^sub>3"} are the only 
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
   703
  equivalence classes in @{term "finals {[a], [a, b]}"}.
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
   704
  It is straightforward to show that in general 
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   705
  %
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   706
  \begin{equation}\label{finalprops}
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
   707
  @{thm lang_is_union_of_finals}\hspace{15mm} 
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
   708
  @{thm finals_in_partitions} 
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   709
  \end{equation}
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
   710
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
   711
  \noindent
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
   712
  hold. 
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   713
  Therefore if we know that there exists a regular expression for every
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   714
  equivalence class in \mbox{@{term "finals A"}} (which by assumption must be
93
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   715
  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
   716
  that matches every string in @{text A}.
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   717
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   718
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   719
  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
   720
  regular expression for \emph{every} equivalence class, not just the ones 
77
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   721
  in @{term "finals A"}. We
93
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   722
  first define the notion of \emph{one-character-transition} between 
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   723
  two equivalence classes
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   724
  %
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   725
  \begin{equation} 
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   726
  @{thm transition_def}
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   727
  \end{equation}
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   728
71
426070e68b21 more on the paper
urbanc
parents: 70
diff changeset
   729
  \noindent
338
e7504bfdbd50 made the changes thes 2nd referee suggested and made it to compile again
urbanc
parents: 334
diff changeset
   730
  which means that if we append the character @{text c} to the end of all 
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
   731
  strings in the equivalence class @{text Y}, we obtain a subset of 
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   732
  @{text X}. 
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   733
  %Note that we do not define formally an automaton here, 
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   734
  %we merely relate two sets
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   735
  %(with the help of a character). 
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
   736
  In our concrete example we have
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
   737
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
   738
  \begin{center} 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
   739
  \begin{tabular}{l}
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   740
  @{term "X\<^sub>1 \<Turnstile>a\<Rightarrow> X\<^sub>2"},\; @{term "X\<^sub>1 \<Turnstile>b\<Rightarrow> X\<^sub>4"};\\ 
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   741
  @{term "X\<^sub>2 \<Turnstile>b\<Rightarrow> X\<^sub>3"},\; @{term "X\<^sub>2 \<Turnstile>a\<Rightarrow> X\<^sub>4"};\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   742
  @{term "X\<^sub>3 \<Turnstile>a\<Rightarrow> X\<^sub>4"},\; @{term "X\<^sub>3 \<Turnstile>b\<Rightarrow> X\<^sub>4"} and\\ 
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   743
  @{term "X\<^sub>4 \<Turnstile>a\<Rightarrow> X\<^sub>4"},\; @{term "X\<^sub>4 \<Turnstile>b\<Rightarrow> X\<^sub>4"}.
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
   744
  \end{tabular}
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
   745
  \end{center}
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   746
  
156
fd39492b187c a few more changes
urbanc
parents: 154
diff changeset
   747
  Next we construct an \emph{initial equational system} that
fd39492b187c a few more changes
urbanc
parents: 154
diff changeset
   748
  contains an equation for each equivalence class. We first give
fd39492b187c a few more changes
urbanc
parents: 154
diff changeset
   749
  an informal description of this construction. Suppose we have 
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   750
  the equivalence classes @{text "X\<^sub>1,\<dots>,X\<^sub>n"}, there must be one and only one that
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   751
  contains the empty string @{text "[]"} (since equivalence classes are disjoint).
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   752
  Let us assume @{text "[] \<in> X\<^sub>1"}. We build the following initial equational system
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   753
  
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   754
  \begin{center}
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   755
  \begin{tabular}{rcl}
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   756
  @{text "X\<^sub>1"} & @{text "="} & @{text "(Y\<^sub>1\<^sub>1, ATOM c\<^sub>1\<^sub>1) + \<dots> + (Y\<^sub>1\<^sub>p, ATOM c\<^sub>1\<^sub>p) + \<lambda>(ONE)"} \\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   757
  @{text "X\<^sub>2"} & @{text "="} & @{text "(Y\<^sub>2\<^sub>1, ATOM c\<^sub>2\<^sub>1) + \<dots> + (Y\<^sub>2\<^sub>o, ATOM c\<^sub>2\<^sub>o)"} \\
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   758
  & $\vdots$ \\
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   759
  @{text "X\<^sub>n"} & @{text "="} & @{text "(Y\<^sub>n\<^sub>1, ATOM c\<^sub>n\<^sub>1) + \<dots> + (Y\<^sub>n\<^sub>q, ATOM c\<^sub>n\<^sub>q)"}\\
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   760
  \end{tabular}
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   761
  \end{center}
70
8ab3a06577cf slightly more on the paper
urbanc
parents: 67
diff changeset
   762
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   763
  \noindent
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   764
  where the terms @{text "(Y\<^sub>i\<^sub>j, ATOM c\<^sub>i\<^sub>j)"} are pairs consisting of an equivalence class and
338
e7504bfdbd50 made the changes thes 2nd referee suggested and made it to compile again
urbanc
parents: 334
diff changeset
   765
  a regular expression. In the initial equational system, they
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   766
  stand for all transitions @{term "Y\<^sub>i\<^sub>j \<Turnstile>c\<^sub>i\<^sub>j\<Rightarrow>
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   767
  X\<^sub>i"}. 
159
990c12ab1562 edits; sqeezed to 16 pages
urbanc
parents: 157
diff changeset
   768
  %The intuition behind the equational system is that every 
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   769
  %equation @{text "X\<^sub>i = rhs\<^sub>i"} in this system
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   770
  %corresponds roughly to a state of an automaton whose name is @{text X\<^sub>i} and its predecessor states 
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   771
  %are the @{text "Y\<^sub>i\<^sub>j"}; the @{text "c\<^sub>i\<^sub>j"} are the labels of the transitions from these 
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   772
  %predecessor states to @{text X\<^sub>i}. 
159
990c12ab1562 edits; sqeezed to 16 pages
urbanc
parents: 157
diff changeset
   773
  There can only be
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   774
  finitely many terms of the form @{text "(Y\<^sub>i\<^sub>j, ATOM c\<^sub>i\<^sub>j)"} in a right-hand side 
156
fd39492b187c a few more changes
urbanc
parents: 154
diff changeset
   775
  since by assumption there are only finitely many
159
990c12ab1562 edits; sqeezed to 16 pages
urbanc
parents: 157
diff changeset
   776
  equivalence classes and only finitely many characters.
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
   777
  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
   778
  is the equivalence class
387
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   779
  containing the empty string @{text "[]"}.
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   780
  In our running example we have the initial equational system
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   781
  %
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   782
  \begin{equation}\label{exmpcs}
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   783
  \mbox{\begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   784
  @{term "X\<^sub>1"} & @{text "="} & @{text "\<lambda>(ONE)"}\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   785
  @{term "X\<^sub>2"} & @{text "="} & @{text "(X\<^sub>1, ATOM a)"}\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   786
  @{term "X\<^sub>3"} & @{text "="} & @{text "(X\<^sub>2, ATOM b)"}\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   787
  @{term "X\<^sub>4"} & @{text "="} & @{text "(X\<^sub>1, ATOM b) + (X\<^sub>2, ATOM a) + (X\<^sub>3, ATOM a)"}\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   788
          & & \mbox{}\hspace{0mm}@{text "+ (X\<^sub>3, ATOM b) + (X\<^sub>4, ATOM a) + (X\<^sub>4, ATOM b)"}\\
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   789
  \end{tabular}}
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   790
  \end{equation}
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
   791
387
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   792
  Note that we mark, roughly speaking, the
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   793
  single `initial' state in the equational system, which is different from
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   794
  the method by \citeN{Brzozowski64}, where he marks the
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   795
  `terminal' states. We are forced to set up the equational system in our
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   796
  way, because the Myhill-Nerode Relation determines the `direction' of the
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   797
  transitions---the successor `state' of an equivalence class @{text Y} can
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   798
  be reached by adding a character to the end of @{text Y}. If we had defined
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   799
  our equations the `Brzozowski-way', then our variables do not correspond to 
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   800
  the equivalence classes generated by the Myhill-Nerode relation. This need of reverse marking is also the
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   801
  reason why we have to use our reversed version of Arden's Lemma.
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   802
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   803
  Overloading the function @{text \<calL>} for the two kinds of terms in the
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
   804
  equational system, we have
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   805
  
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   806
  \begin{center}
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
   807
  @{text "\<calL>(Y, r) \<equiv>"} %
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
   808
  @{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
   809
  @{thm lang_trm.simps(1)[where r="r", THEN eq_reflection]}
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   810
  \end{center}
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   811
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   812
  \noindent
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   813
  and we can prove for @{text "X\<^sub>2\<^sub>.\<^sub>.\<^sub>n"} that the following equations
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   814
  %
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   815
  \begin{equation}\label{inv1}
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   816
  @{text "X\<^sub>i = \<calL>(Y\<^sub>i\<^sub>1, ATOM c\<^sub>i\<^sub>1) \<union> \<dots> \<union> \<calL>(Y\<^sub>i\<^sub>q, ATOM c\<^sub>i\<^sub>q)"}.
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   817
  \end{equation}
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   818
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   819
  \noindent
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   820
  hold. Similarly for @{text "X\<^sub>1"} we can show the following equation
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   821
  %
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   822
  \begin{equation}\label{inv2}
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   823
  @{text "X\<^sub>1 = \<calL>(Y\<^sub>1\<^sub>1, ATOM c\<^sub>1\<^sub>1) \<union> \<dots> \<union> \<calL>(Y\<^sub>1\<^sub>p, ATOM c\<^sub>1\<^sub>p) \<union> \<calL>(\<lambda>(ONE))"}.
75
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   824
  \end{equation}
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   825
d63baacbdb16 parts of the 3 section
urbanc
parents: 71
diff changeset
   826
  \noindent
387
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   827
  holds. Again note that the reason for adding the @{text \<lambda>}-marker to our initial equational system is 
103
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
   828
  to obtain this equation: it only holds with the marker, since none of 
387
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   829
  the other terms contain the empty string. 
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   830
  The point of the initial equational system is
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   831
  that solving it means we will be able to extract a regular expression for every equivalence class
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
   832
  generated by the Myhill-Nerode relation. 
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   833
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   834
  Our representation for the equations in Isabelle/HOL are pairs,
108
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   835
  where the first component is an equivalence class (a set of strings)
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   836
  and the second component
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   837
  is a set of terms. Given a set of equivalence
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   838
  classes @{text CS}, our initial equational system @{term "Init CS"} is thus 
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   839
  formally defined as
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   840
  %
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   841
  \begin{equation}\label{initcs}
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   842
  \mbox{\begin{tabular}{rcl}     
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   843
  @{thm (lhs) Init_rhs_def} & @{text "\<equiv>"} & 
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   844
  @{text "if"}~@{term "[] \<in> X"}\\
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
   845
  & & @{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
   846
  & & @{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
   847
  @{thm (lhs) Init_def}     & @{text "\<equiv>"} & @{thm (rhs) Init_def}
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   848
  \end{tabular}}
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
   849
  \end{equation}
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   850
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   851
  \noindent
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   852
  Because we use sets of terms 
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   853
  for representing the right-hand sides of equations, we can 
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   854
  prove \eqref{inv1} and \eqref{inv2} more concisely as
93
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   855
  %
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   856
  \begin{lemma}\label{inv}
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   857
  If @{thm (prem 1) test} then @{text "X = \<Union> \<calL> ` rhs"}.
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   858
  \end{lemma}
77
63bc9f9d96ba small additions
urbanc
parents: 75
diff changeset
   859
93
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   860
  \noindent
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   861
  Our proof of Theorem~\ref{myhillnerodeone} will proceed by transforming the
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   862
  initial equational system into one in \emph{solved form} maintaining the invariant
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
   863
  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
   864
  off the regular expressions. 
42af13d194c9 a bit more on the paper
urbanc
parents: 88
diff changeset
   865
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   866
  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
   867
  operations: one that takes an equation of the form @{text "X = rhs"} and removes
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   868
  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
   869
  Lemma. The other operation takes an equation @{text "X = rhs"}
89
42af13d194c9 a bit more on the paper
urbanc
parents: 88
diff changeset
   870
  and substitutes @{text X} throughout the rest of the equational system
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   871
  adjusting the remaining regular expressions appropriately. To define this adjustment 
108
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   872
  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
   873
42af13d194c9 a bit more on the paper
urbanc
parents: 88
diff changeset
   874
  \begin{center}
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   875
  \begin{tabular}{r@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   876
  @{thm (lhs) Append_rexp.simps(2)[where X="Y" and r="r\<^sub>1" and rexp="r\<^sub>2", THEN eq_reflection]}
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   877
  & @{text "\<equiv>"} & 
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   878
  @{thm (rhs) Append_rexp.simps(2)[where X="Y" and r="r\<^sub>1" and rexp="r\<^sub>2", THEN eq_reflection]}\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   879
  @{thm (lhs) Append_rexp.simps(1)[where r="r\<^sub>1" and rexp="r\<^sub>2", THEN eq_reflection]}
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   880
  & @{text "\<equiv>"} & 
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   881
  @{thm (rhs) Append_rexp.simps(1)[where r="r\<^sub>1" and rexp="r\<^sub>2", THEN eq_reflection]}
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   882
  \end{tabular}
89
42af13d194c9 a bit more on the paper
urbanc
parents: 88
diff changeset
   883
  \end{center}
42af13d194c9 a bit more on the paper
urbanc
parents: 88
diff changeset
   884
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
   885
  \noindent
108
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   886
  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
   887
  @{thm (lhs) Append_rexp_rhs_def[where rexp="r"]}. With this we can define
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   888
  the \emph{arden-operation} for an equation of the form @{text "X = rhs"} as:
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   889
  %
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   890
  \begin{equation}\label{arden_def}
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   891
  \mbox{\begin{tabular}{rc@ {\hspace{2mm}}r@ {\hspace{1mm}}l}
94
5b12cd0a3b3c latest on the paper
urbanc
parents: 93
diff changeset
   892
  @{thm (lhs) Arden_def} & @{text "\<equiv>"}~~\mbox{} & \multicolumn{2}{@ {\hspace{-2mm}}l}{@{text "let"}}\\ 
5b12cd0a3b3c latest on the paper
urbanc
parents: 93
diff changeset
   893
   & & @{text "rhs' ="} & @{term "rhs - {Trn X r | r. Trn X r \<in> rhs}"} \\
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   894
   & & @{text "r' ="}   & @{term "Star (\<Uplus> {r. Trn X r \<in> rhs})"}\\
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   895
   & &  \multicolumn{2}{@ {\hspace{-2mm}}l}{@{text "in"}~~@{term "Append_rexp_rhs rhs' r'"}}\\ 
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   896
  \end{tabular}}
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   897
  \end{equation}
93
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   898
2aa3756dcc9f more on the paper
urbanc
parents: 92
diff changeset
   899
  \noindent
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
   900
  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
   901
  then we calculate the combined regular expressions for all @{text r} coming 
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   902
  from the deleted @{text "(X, r)"}, and take the @{const Star} of it;
178
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   903
  finally we append this regular expression to @{text rhs'}. If we apply this
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   904
  operation to the right-hand side of @{text "X\<^sub>4"} in \eqref{exmpcs}, we obtain
178
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   905
  the equation:
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   906
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   907
  \begin{center}
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   908
  \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   909
  @{term "X\<^sub>4"} & @{text "="} & 
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   910
    @{text "(X\<^sub>1, TIMES (ATOM b) (STAR \<^raw:\ensuremath{\bigplus}>{ATOM a, ATOM b})) +"}\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   911
  & & @{text "(X\<^sub>2, TIMES (ATOM a) (STAR \<^raw:\ensuremath{\bigplus}>{ATOM a, ATOM b})) +"}\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   912
  & & @{text "(X\<^sub>3, TIMES (ATOM a) (STAR \<^raw:\ensuremath{\bigplus}>{ATOM a, ATOM b})) +"}\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   913
  & & @{text "(X\<^sub>3, TIMES (ATOM b) (STAR \<^raw:\ensuremath{\bigplus}>{ATOM a, ATOM b}))"}
178
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   914
  \end{tabular}
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   915
  \end{center}
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   916
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   917
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   918
  \noindent
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   919
  That means we eliminated the recursive occurrence of @{text "X\<^sub>4"} on the
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
   920
  right-hand side. 
178
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   921
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   922
  It can be easily seen that the @{text "Arden"}-operation mimics Arden's
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   923
  Lemma on the level of equations. To ensure the non-emptiness condition of
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   924
  Arden's Lemma we say that a right-hand side is @{text ardenable} provided
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   925
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   926
  \begin{center}
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   927
  @{thm ardenable_def}
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   928
  \end{center}
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   929
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   930
  \noindent
156
fd39492b187c a few more changes
urbanc
parents: 154
diff changeset
   931
  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
   932
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   933
  \begin{lemma}\label{ardenable}
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
   934
  Given an equation @{text "X = rhs"}.
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   935
  If @{text "X = \<Union>\<calL> ` rhs"},
179
edacc141060f small improvements
urbanc
parents: 178
diff changeset
   936
  @{thm (prem 2) Arden_preserves_soundness}, and
edacc141060f small improvements
urbanc
parents: 178
diff changeset
   937
  @{thm (prem 3) Arden_preserves_soundness}, then
135
604518f0127f final polished
urbanc
parents: 134
diff changeset
   938
  @{text "X = \<Union>\<calL> ` (Arden X rhs)"}.
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
   939
  \end{lemma}
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   940
  
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   941
  \noindent
156
fd39492b187c a few more changes
urbanc
parents: 154
diff changeset
   942
  Our @{text ardenable} condition is slightly stronger than needed for applying Arden's Lemma,
194
5347d7556487 some typos
urbanc
parents: 193
diff changeset
   943
  but we can still ensure that it holds throughout our algorithm of transforming equations
338
e7504bfdbd50 made the changes thes 2nd referee suggested and made it to compile again
urbanc
parents: 334
diff changeset
   944
  into solved form. 
e7504bfdbd50 made the changes thes 2nd referee suggested and made it to compile again
urbanc
parents: 334
diff changeset
   945
e7504bfdbd50 made the changes thes 2nd referee suggested and made it to compile again
urbanc
parents: 334
diff changeset
   946
  The \emph{substitution-operation} takes an equation
95
9540c2f2ea77 more things
urbanc
parents: 94
diff changeset
   947
  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
   948
5b12cd0a3b3c latest on the paper
urbanc
parents: 93
diff changeset
   949
  \begin{center}
95
9540c2f2ea77 more things
urbanc
parents: 94
diff changeset
   950
  \begin{tabular}{rc@ {\hspace{2mm}}r@ {\hspace{1mm}}l}
9540c2f2ea77 more things
urbanc
parents: 94
diff changeset
   951
  @{thm (lhs) Subst_def} & @{text "\<equiv>"}~~\mbox{} & \multicolumn{2}{@ {\hspace{-2mm}}l}{@{text "let"}}\\ 
9540c2f2ea77 more things
urbanc
parents: 94
diff changeset
   952
   & & @{text "rhs' ="} & @{term "rhs - {Trn X r | r. Trn X r \<in> rhs}"} \\
9540c2f2ea77 more things
urbanc
parents: 94
diff changeset
   953
   & & @{text "r' ="}   & @{term "\<Uplus> {r. Trn X r \<in> rhs}"}\\
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
   954
   & &  \multicolumn{2}{@ {\hspace{-2mm}}l}{@{text "in"}~~@{term "rhs' \<union> Append_rexp_rhs xrhs r'"}}\\ 
95
9540c2f2ea77 more things
urbanc
parents: 94
diff changeset
   955
  \end{tabular}
94
5b12cd0a3b3c latest on the paper
urbanc
parents: 93
diff changeset
   956
  \end{center}
95
9540c2f2ea77 more things
urbanc
parents: 94
diff changeset
   957
9540c2f2ea77 more things
urbanc
parents: 94
diff changeset
   958
  \noindent
134
08afbed1c8c7 chunhan's comments
urbanc
parents: 133
diff changeset
   959
  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
   960
  the regular expression corresponding to the deleted terms; finally we append this
9540c2f2ea77 more things
urbanc
parents: 94
diff changeset
   961
  regular expression to @{text "xrhs"} and union it up with @{text rhs'}. When we use
9540c2f2ea77 more things
urbanc
parents: 94
diff changeset
   962
  the substitution operation we will arrange it so that @{text "xrhs"} does not contain
178
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   963
  any occurrence of @{text X}. For example substituting the first equation in
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   964
  \eqref{exmpcs} into the right-hand side of the second, thus eliminating the equivalence 
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   965
  class @{text "X\<^sub>1"}, gives us the equation
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
   966
  %
178
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   967
  \begin{equation}\label{exmpresult}
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   968
  \mbox{\begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
   969
  @{term "X\<^sub>2"} & @{text "="} & @{text "\<lambda>(TIMES ONE (ATOM a))"}
178
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   970
  \end{tabular}}
c6ebfe052109 added more examles
urbanc
parents: 177
diff changeset
   971
  \end{equation}
96
3b9deda4f459 simplified a bit the proof
urbanc
parents: 95
diff changeset
   972
134
08afbed1c8c7 chunhan's comments
urbanc
parents: 133
diff changeset
   973
  With these two operations in place, we can define the operation that removes one equation
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   974
  from an equational systems @{text ES}. The operation @{const Subst_all}
96
3b9deda4f459 simplified a bit the proof
urbanc
parents: 95
diff changeset
   975
  substitutes an equation @{text "X = xrhs"} throughout an equational system @{text ES}; 
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   976
  @{const Remove} then completely removes such an equation from @{text ES} by substituting 
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   977
  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
   978
  of @{text X} by applying @{const Arden} to @{text "xrhs"}.
3b9deda4f459 simplified a bit the proof
urbanc
parents: 95
diff changeset
   979
3b9deda4f459 simplified a bit the proof
urbanc
parents: 95
diff changeset
   980
  \begin{center}
3b9deda4f459 simplified a bit the proof
urbanc
parents: 95
diff changeset
   981
  \begin{tabular}{rcl}
3b9deda4f459 simplified a bit the proof
urbanc
parents: 95
diff changeset
   982
  @{thm (lhs) Subst_all_def} & @{text "\<equiv>"} & @{thm (rhs) Subst_all_def}\\
3b9deda4f459 simplified a bit the proof
urbanc
parents: 95
diff changeset
   983
  @{thm (lhs) Remove_def}    & @{text "\<equiv>"} & @{thm (rhs) Remove_def}
3b9deda4f459 simplified a bit the proof
urbanc
parents: 95
diff changeset
   984
  \end{tabular}
3b9deda4f459 simplified a bit the proof
urbanc
parents: 95
diff changeset
   985
  \end{center}
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   986
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   987
  \noindent
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
   988
  Finally, we can define how an equational system should be solved. For this 
107
6f4f9b7b9891 updated paper
urbanc
parents: 106
diff changeset
   989
  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
   990
  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
   991
  as being the last one, but the one involving the equivalence class for 
6f4f9b7b9891 updated paper
urbanc
parents: 106
diff changeset
   992
  which we want to calculate the regular 
108
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   993
  expression. Let us suppose this equivalence class is @{text X}. 
107
6f4f9b7b9891 updated paper
urbanc
parents: 106
diff changeset
   994
  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
   995
  equation to be eliminated that is different from @{text X}. In this way 
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
   996
  @{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
   997
  operator, written @{text SOME} in the definition below.
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   998
  
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
   999
  \begin{center}
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
  1000
  \begin{tabular}{rc@ {\hspace{4mm}}r@ {\hspace{1mm}}l}
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
  1001
  @{thm (lhs) Iter_def} & @{text "\<equiv>"}~~\mbox{} & \multicolumn{2}{@ {\hspace{-4mm}}l}{@{text "let"}}\\ 
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
  1002
   & & @{text "(Y, yrhs) ="} & @{term "SOME (Y, yrhs). (Y, yrhs) \<in> ES \<and> X \<noteq> Y"} \\
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
  1003
   & &  \multicolumn{2}{@ {\hspace{-4mm}}l}{@{text "in"}~~@{term "Remove ES Y yrhs"}}\\ 
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
  1004
  \end{tabular}
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
  1005
  \end{center}
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
  1006
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
  1007
  \noindent
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1008
  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
  1009
  @{text Cond} is \emph{not} satisfied anymore. This condition states that there
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1010
  are more than one equation left in the equational system @{text ES}. To solve
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1011
  an equational system we use Isabelle/HOL's @{text while}-operator as follows:
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
  1012
  
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
  1013
  \begin{center}
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
  1014
  @{thm Solve_def}
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
  1015
  \end{center}
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
  1016
101
d3fe0597080a updated paper
urbanc
parents: 100
diff changeset
  1017
  \noindent
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1018
  We are not concerned here with the definition of this operator (see
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1019
  \cite{BerghoferNipkow00} for example), but note that we
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1020
  eliminate in each @{const Iter}-step a single equation, and therefore have a
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1021
  well-founded termination order by taking the cardinality of the equational
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1022
  system @{text ES}. This enables us to prove properties about our definition
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1023
  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
  1024
  the initial equational system @{term "Init (UNIV // \<approx>A)"} from
108
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
  1025
  \eqref{initcs} using the principle:
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1026
  %
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1027
  \begin{equation}\label{whileprinciple}
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1028
  \mbox{\begin{tabular}{l}
103
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
  1029
  @{term "invariant (Init (UNIV // \<approx>A))"} \\
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
  1030
  @{term "\<forall>ES. invariant ES \<and> Cond ES \<longrightarrow> invariant (Iter X ES)"}\\
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
  1031
  @{term "\<forall>ES. invariant ES \<and> Cond ES \<longrightarrow> card (Iter X ES) < card ES"}\\
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
  1032
  @{term "\<forall>ES. invariant ES \<and> \<not> Cond ES \<longrightarrow> P ES"}\\
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
  1033
  \hline
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
  1034
  \multicolumn{1}{c}{@{term "P (Solve X (Init (UNIV // \<approx>A)))"}}
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1035
  \end{tabular}}
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1036
  \end{equation}
103
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
  1037
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
  1038
  \noindent
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1039
  This principle states that given an invariant (which we will specify below) 
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1040
  we can prove a property
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1041
  @{text "P"} involving @{const Solve}. For this we have to discharge the following
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1042
  proof obligations: first the
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1043
  initial equational system satisfies the invariant; second the iteration
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 149
diff changeset
  1044
  step @{text "Iter"} preserves the invariant as long as the condition @{term Cond} holds;
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1045
  third @{text "Iter"} decreases the termination order, and fourth that
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1046
  once the condition does not hold anymore then the property @{text P} must hold.
103
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
  1047
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1048
  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
  1049
  returns with a single equation @{text "X = xrhs"} for some @{text "xrhs"}, and
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1050
  that this equational system still satisfies the invariant. In order to get
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1051
  the proof through, the invariant is composed of the following six properties:
103
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
  1052
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
  1053
  \begin{center}
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1054
  \begin{tabular}{@ {}rcl@ {\hspace{-13mm}}l @ {}}
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1055
  @{text "invariant ES"} & @{text "\<equiv>"} &
103
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
  1056
      @{term "finite ES"} & @{text "(finiteness)"}\\
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
  1057
  & @{text "\<and>"} & @{thm (rhs) finite_rhs_def} & @{text "(finiteness rhs)"}\\
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1058
  & @{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
  1059
  & @{text "\<and>"} & @{thm (rhs) distinctness_def}\\
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1060
  &             &  & @{text "(distinctness)"}\\
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1061
  & @{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
  1062
  & @{text "\<and>"} & @{thm (rhs) validity_def} & @{text "(validity)"}\\
103
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
  1063
  \end{tabular}
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
  1064
  \end{center}
f460d5f75cb5 updated
urbanc
parents: 101
diff changeset
  1065
 
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1066
  \noindent
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1067
  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
  1068
  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
  1069
  equations is preserved under our transformations. The other properties are a bit more
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
  1070
  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
  1071
  equation in the system is distinct. @{text Ardenable} ensures that we can always
156
fd39492b187c a few more changes
urbanc
parents: 154
diff changeset
  1072
  apply the @{text Arden} operation. 
108
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
  1073
  The last property states that every @{text rhs} can only contain equivalence classes
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
  1074
  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
  1075
  the first components of an equational system,
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
  1076
  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
  1077
  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
  1078
  and @{thm (lhs) rhss_def}~@{text "\<equiv> {X | (X, r) \<in> rhs}"}.
108
212bfa431fa5 filled details in one place
urbanc
parents: 107
diff changeset
  1079
  
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1080
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1081
  It is straightforward to prove that the initial equational system satisfies the
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1082
  invariant.
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1083
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1084
  \begin{lemma}\label{invzero}
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1085
  @{thm[mode=IfThen] Init_ES_satisfies_invariant}
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1086
  \end{lemma}
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1087
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1088
  \begin{proof}
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1089
  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
  1090
  initial equational system. Soundness is proved in Lemma~\ref{inv}. Distinctness
154
7c68b9ad4486 implemented most suggestions from the reviewers
urbanc
parents: 149
diff changeset
  1091
  follows from the fact that the equivalence classes are disjoint. The @{text ardenable}
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1092
  property also follows from the setup of the initial equational system, as does 
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  1093
  validity.\qed
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1094
  \end{proof}
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1095
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1096
  \noindent
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1097
  Next we show that @{text Iter} preserves the invariant.
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1098
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1099
  \begin{lemma}\label{iterone} If
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1100
  @{thm (prem 1) iteration_step_invariant[where xrhs="rhs"]},
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1101
  @{thm (prem 2) iteration_step_invariant[where xrhs="rhs"]} and
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1102
  @{thm (prem 3) iteration_step_invariant[where xrhs="rhs"]}, then
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1103
  @{thm (concl) iteration_step_invariant[where xrhs="rhs"]}.
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1104
  \end{lemma}
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1105
107
6f4f9b7b9891 updated paper
urbanc
parents: 106
diff changeset
  1106
  \begin{proof} 
156
fd39492b187c a few more changes
urbanc
parents: 154
diff changeset
  1107
  The argument boils down to choosing an equation @{text "Y = yrhs"} to be eliminated
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1108
  and to show that @{term "Subst_all (ES - {(Y, yrhs)}) Y (Arden Y yrhs)"} 
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1109
  preserves the invariant.
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1110
  We prove this as follows:
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1111
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1112
  \begin{center}
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
  1113
  \begin{tabular}{@ {}l@ {}}
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1114
  @{text "\<forall>ES."}~@{thm (prem 1) Subst_all_satisfies_invariant} implies
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1115
  @{thm (concl) Subst_all_satisfies_invariant}
177
50cc1a39c990 more one the paper
urbanc
parents: 176
diff changeset
  1116
  \end{tabular}
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1117
  \end{center}
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1118
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1119
  \noindent
156
fd39492b187c a few more changes
urbanc
parents: 154
diff changeset
  1120
  Finiteness is straightforward, as the @{const Subst} and @{const Arden} operations 
116
342983676c8f included comments by Chunhan
urbanc
parents: 115
diff changeset
  1121
  keep the equational system finite. These operations also preserve soundness 
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1122
  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
  1123
  The property @{text ardenable} is clearly preserved because the append-operation
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1124
  cannot make a regular expression to match the empty string. Validity is
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1125
  given because @{const Arden} removes an equivalence class from @{text yrhs}
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1126
  and then @{const Subst_all} removes @{text Y} from the equational system.
132
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
  1127
  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
  1128
  which matches with our proof-obligation of @{const "Subst_all"}. Since
132
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
  1129
  \mbox{@{term "ES = ES - {(Y, yrhs)} \<union> {(Y, yrhs)}"}}, we can use the assumption 
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  1130
  to complete the proof.\qed
107
6f4f9b7b9891 updated paper
urbanc
parents: 106
diff changeset
  1131
  \end{proof}
6f4f9b7b9891 updated paper
urbanc
parents: 106
diff changeset
  1132
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1133
  \noindent
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1134
  We also need the fact that @{text Iter} decreases the termination measure.
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1135
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1136
  \begin{lemma}\label{itertwo} If
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1137
  @{thm (prem 1) iteration_step_measure[simplified (no_asm), where xrhs="rhs"]},
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1138
  @{thm (prem 2) iteration_step_measure[simplified (no_asm), where xrhs="rhs"]} and
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1139
  @{thm (prem 3) iteration_step_measure[simplified (no_asm), where xrhs="rhs"]}, then\\
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1140
  \mbox{@{thm (concl) iteration_step_measure[simplified (no_asm), where xrhs="rhs"]}}.
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1141
  \end{lemma}
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1142
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1143
  \begin{proof}
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1144
  By assumption we know that @{text "ES"} is finite and has more than one element.
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1145
  Therefore there must be an element @{term "(Y, yrhs) \<in> ES"} with 
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1146
  @{term "(Y, yrhs) \<noteq> (X, rhs)"}. Using the distinctness property we can infer
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1147
  that @{term "Y \<noteq> X"}. We further know that @{text "Remove ES Y yrhs"}
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1148
  removes the equation @{text "Y = yrhs"} from the system, and therefore 
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  1149
  the cardinality of @{const Iter} strictly decreases.\qed
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1150
  \end{proof}
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1151
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1152
  \noindent
134
08afbed1c8c7 chunhan's comments
urbanc
parents: 133
diff changeset
  1153
  This brings us to our property we want to establish for @{text Solve}.
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1154
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1155
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1156
  \begin{lemma}
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1157
  If @{thm (prem 1) Solve} and @{thm (prem 2) Solve} then there exists
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1158
  a @{text rhs} such that  @{term "Solve X (Init (UNIV // \<approx>A)) = {(X, rhs)}"}
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1159
  and @{term "invariant {(X, rhs)}"}.
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1160
  \end{lemma}
104
5bd73aa805a7 updated paper
urbanc
parents: 103
diff changeset
  1161
107
6f4f9b7b9891 updated paper
urbanc
parents: 106
diff changeset
  1162
  \begin{proof} 
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1163
  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
  1164
  stronger invariant since Lemma~\ref{iterone} and \ref{itertwo} have the precondition 
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1165
  that @{term "(X, rhs) \<in> ES"} for some @{text rhs}. This precondition is needed
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1166
  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
  1167
  Therefore our invariant cannot be just @{term "invariant ES"}, but must be 
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1168
  @{term "invariant ES \<and> (\<exists>rhs. (X, rhs) \<in> ES)"}. By assumption 
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1169
  @{thm (prem 2) Solve} and Lemma~\ref{invzero}, the more general invariant holds for
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1170
  the initial equational system. This is premise 1 of~\eqref{whileprinciple}.
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1171
  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
  1172
  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
  1173
  Premise 3 of~\eqref{whileprinciple} is by Lemma~\ref{itertwo}. Now in premise 4
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1174
  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
  1175
  and that @{text "invariant {(X, rhs)}"} holds, provided the condition @{text "Cond"}
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1176
  does not hold. By the stronger invariant we know there exists such a @{text "rhs"}
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1177
  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
  1178
  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
  1179
  for which the invariant holds. This allows us to conclude that 
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1180
  @{term "Solve X (Init (UNIV // \<approx>A)) = {(X, rhs)}"} and @{term "invariant {(X, rhs)}"} hold,
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  1181
  as needed.\qed
107
6f4f9b7b9891 updated paper
urbanc
parents: 106
diff changeset
  1182
  \end{proof}
6f4f9b7b9891 updated paper
urbanc
parents: 106
diff changeset
  1183
106
91dc591de63f updated paper
urbanc
parents: 105
diff changeset
  1184
  \noindent
91dc591de63f updated paper
urbanc
parents: 105
diff changeset
  1185
  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
  1186
  there exists a regular expression.
91dc591de63f updated paper
urbanc
parents: 105
diff changeset
  1187
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1188
  \begin{lemma}\label{every_eqcl_has_reg}
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1189
  @{thm[mode=IfThen] every_eqcl_has_reg}
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1190
  \end{lemma}
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1191
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1192
  \begin{proof}
138
2dfe13bc1443 three typos
urbanc
parents: 137
diff changeset
  1193
  By the preceding lemma, we know that there exists a @{text "rhs"} such
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1194
  that @{term "Solve X (Init (UNIV // \<approx>A))"} returns the equation @{text "X = rhs"},
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1195
  and that the invariant holds for this equation. That means we 
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1196
  know @{text "X = \<Union>\<calL> ` rhs"}. We further know that
109
79b37ef9505f minor updated
urbanc
parents: 108
diff changeset
  1197
  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
  1198
  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
  1199
  we can infer that @{term "rhss rhs \<subseteq> {X}"} and because the @{text Arden} operation
106
91dc591de63f updated paper
urbanc
parents: 105
diff changeset
  1200
  removes that @{text X} from @{text rhs}, that @{term "rhss (Arden X rhs) = {}"}.
113
ec774952190c polished everything
urbanc
parents: 112
diff changeset
  1201
  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
  1202
  So we can collect those (finitely many) regular expressions @{text rs} and have @{term "X = lang (\<Uplus>rs)"}.
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  1203
  With this we can conclude the proof.\qed
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1204
  \end{proof}
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1205
106
91dc591de63f updated paper
urbanc
parents: 105
diff changeset
  1206
  \noindent
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1207
  Lemma~\ref{every_eqcl_has_reg} allows us to finally give a proof for the first direction
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
  1208
  of the Myhill-Nerode Theorem.
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1209
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  1210
  \begin{proof}[of Theorem~\ref{myhillnerodeone}]
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1211
  By Lemma~\ref{every_eqcl_has_reg} we know that there exists a regular expression for
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1212
  every equivalence class in @{term "UNIV // \<approx>A"}:
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1213
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1214
  \[
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1215
  \mbox{if}\;@{term "X \<in> (UNIV // \<approx>A)"}\;\mbox{then there exists a}\; 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1216
  @{text "r"}\;\mbox{such that}\;@{term "X = lang r"}
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1217
  \]\smallskip
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1218
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1219
  \noindent
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1220
  Since @{text "finals A"} is
110
e500cab16be4 completed first direction
urbanc
parents: 109
diff changeset
  1221
  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
  1222
  in @{term "finals A"} there exists a regular expression. Moreover by assumption 
106
91dc591de63f updated paper
urbanc
parents: 105
diff changeset
  1223
  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
  1224
  set of regular expressions @{text "rs"} such that
176
6969de1eb96b latest version of the journal paper
urbanc
parents: 175
diff changeset
  1225
  @{term "\<Union>(finals A) = lang (\<Uplus>rs)"}.
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1226
  Since the left-hand side is equal to @{text A}, we can use @{term "\<Uplus>rs"} 
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  1227
  as the regular expression that is needed in the theorem.\qed
105
ae6ad1363eb9 updated paper
urbanc
parents: 104
diff changeset
  1228
  \end{proof}
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  1229
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  1230
  \noindent
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1231
  Solving the equational system of our running example gives as solution for the 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1232
  two final equivalence classes:
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1233
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1234
  \begin{center}
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1235
  \begin{tabular}{r@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1236
  @{text "X\<^sub>2"} & @{text "="} & @{text "TIMES ONE (ATOM a)"}\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1237
  @{text "X\<^sub>3"} & @{text "="} & @{text "TIMES (TIMES ONE (ATOM a)) (ATOM b)"}
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1238
  \end{tabular}
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1239
  \end{center}
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1240
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1241
  \noindent
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1242
  Combining them with $\bigplus$ gives us a regular expression for the language @{text "{[a], [a, b]}"}.
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1243
245
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  1244
  Note that our algorithm for solving equational systems provides also a method for
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  1245
  calculating a regular expression for the complement of a regular language:
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  1246
  if we combine all regular
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1247
  expressions corresponding to equivalence classes not in @{term "finals A"} 
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1248
  (in the running example @{text "X\<^sub>1"} and @{text "X\<^sub>4"}),
245
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  1249
  then we obtain a regular expression for the complement language @{term "- A"}.
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  1250
  This is similar to the usual construction of a `complement automaton'.
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1251
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
  1252
*}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
  1253
100
2409827d8eb8 updated
urbanc
parents: 98
diff changeset
  1254
section {* Myhill-Nerode, Second Part *}
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
  1255
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
  1256
text {*
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
  1257
  \noindent
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1258
  In this section we will give a proof for establishing the second 
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
  1259
  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
  1260
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1261
  \begin{theorem}\label{myhillnerodetwo}
135
604518f0127f final polished
urbanc
parents: 134
diff changeset
  1262
  Given @{text "r"} is a regular expression, then @{thm Myhill_Nerode2}.
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1263
  \end{theorem}  
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
  1264
116
342983676c8f included comments by Chunhan
urbanc
parents: 115
diff changeset
  1265
  \noindent
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1266
  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
  1267
  the base cases are straightforward.
342983676c8f included comments by Chunhan
urbanc
parents: 115
diff changeset
  1268
342983676c8f included comments by Chunhan
urbanc
parents: 115
diff changeset
  1269
375
44c4450152e3 adapted to JAR
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 374
diff changeset
  1270
  \begin{proof}[Base Cases]
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
  1271
  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
  1272
  we can easily establish that
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
  1273
114
c5eb5f3065ae updated bib
urbanc
parents: 113
diff changeset
  1274
  \begin{center}
c5eb5f3065ae updated bib
urbanc
parents: 113
diff changeset
  1275
  \begin{tabular}{l}
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
  1276
  @{thm quot_zero_eq}\\
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
  1277
  @{thm quot_one_subset}\\
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
  1278
  @{thm quot_atom_subset}
114
c5eb5f3065ae updated bib
urbanc
parents: 113
diff changeset
  1279
  \end{tabular}
c5eb5f3065ae updated bib
urbanc
parents: 113
diff changeset
  1280
  \end{center}
c5eb5f3065ae updated bib
urbanc
parents: 113
diff changeset
  1281
116
342983676c8f included comments by Chunhan
urbanc
parents: 115
diff changeset
  1282
  \noindent
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  1283
  hold, which shows that @{term "UNIV // \<approx>(lang r)"} must be finite.\qed
114
c5eb5f3065ae updated bib
urbanc
parents: 113
diff changeset
  1284
  \end{proof}
109
79b37ef9505f minor updated
urbanc
parents: 108
diff changeset
  1285
116
342983676c8f included comments by Chunhan
urbanc
parents: 115
diff changeset
  1286
  \noindent
183
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
  1287
  Much more interesting, however, are the inductive cases. They seem hard to be solved 
372
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
  1288
  directly. The reader is invited to try.\footnote{The induction hypothesis is not strong enough 
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1289
  to make any progress with the @{text TIMES} and @{text STAR} cases.} 
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1290
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1291
  In order to see how our proof proceeds consider the following suggestive picture 
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1292
  given by \citeN{Constable00}:
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1293
  %
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1294
  \begin{equation}\label{pics}
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1295
  \mbox{\begin{tabular}{c@ {\hspace{10mm}}c@ {\hspace{10mm}}c}
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1296
  \begin{tikzpicture}[scale=0.95]
180
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1297
  %Circle
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1298
  \draw[thick] (0,0) circle (1.1);    
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1299
  \end{tikzpicture}
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1300
  &
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1301
  \begin{tikzpicture}[scale=0.95]
180
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1302
  %Circle
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1303
  \draw[thick] (0,0) circle (1.1);    
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1304
  %Main rays
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1305
  \foreach \a in {0, 90,...,359}
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1306
    \draw[very thick] (0, 0) -- (\a:1.1);
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1307
  \foreach \a / \l in {45/1, 135/2, 225/3, 315/4}
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1308
      \draw (\a: 0.65) node {$a_\l$};
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1309
  \end{tikzpicture}
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1310
  &
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1311
  \begin{tikzpicture}[scale=0.95]
180
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1312
  %Circle
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1313
  \draw[thick] (0,0) circle (1.1);    
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1314
  %Main rays
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1315
  \foreach \a in {0, 45,...,359}
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1316
     \draw[very thick] (0, 0) -- (\a:1.1);
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1317
  \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
  1318
      \draw (\a: 0.77) node {$a_{\l}$};
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1319
  \end{tikzpicture}\\
b755090d0f3d added a picture
urbanc
parents: 179
diff changeset
  1320
  @{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
  1321
  \end{tabular}}
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1322
  \end{equation}
179
edacc141060f small improvements
urbanc
parents: 178
diff changeset
  1323
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1324
  \noindent
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1325
  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
  1326
  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
  1327
  suffices to show in each induction step that another relation, say @{text
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1328
  R}, has finitely many equivalence classes and refines @{term "\<approx>(lang r)"}. 
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1329
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1330
  \begin{definition}
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1331
  A relation @{text "R\<^sub>1"} \emph{refines} @{text "R\<^sub>2"}
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1332
  provided @{text "R\<^sub>1 \<subseteq> R\<^sub>2"}. 
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1333
  \end{definition}
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1334
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1335
  \noindent
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
  1336
  For constructing @{text R}, we will rely on some \emph{tagging-functions}
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  1337
  defined over strings, see Fig.~\ref{tagfig}. These functions are parameterised by sets
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  1338
  of strings @{text A} and @{text B} standing for arbitrary languages and will
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  1339
  by instantiated with the induction hypotheses.
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1340
  Given the inductive hypotheses, it will be easy to
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1341
  prove that the \emph{range} of these tagging-functions is finite. The range
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1342
  of a function @{text f} is defined as
183
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
  1343
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1344
 
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1345
  \begin{center}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1346
  @{text "range f \<equiv> f ` UNIV"}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1347
  \end{center}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1348
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1349
  \noindent
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1350
  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
  1351
  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
  1352
  as relations, give rise to finitely many equivalence classes. 
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1353
  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
  1354
  @{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
  1355
  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
  1356
  as follows.
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1357
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1358
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1359
  \begin{definition}[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
  1360
  and @{text y} are \emph{tag-related} provided
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  1361
  \[
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1362
  @{text "x \<^raw:$\threesim$>\<^bsub>tag\<^esub> y \<equiv> tag x = tag y"}\;.
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  1363
  \]
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1364
  \end{definition}
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1365
145
099e20f25b25 corrected small typo
urbanc
parents: 143
diff changeset
  1366
123
23c0e6f2929d polished everywhere...two cases still missing
urbanc
parents: 122
diff changeset
  1367
  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
  1368
  principle from Isabelle/HOL's library.
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1369
  %
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1370
  \begin{equation}\label{finiteimageD}
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1371
  @{thm[mode=IfThen] finite_imageD}
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1372
  \end{equation}
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1373
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1374
  \noindent
123
23c0e6f2929d polished everywhere...two cases still missing
urbanc
parents: 122
diff changeset
  1375
  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
  1376
  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
  1377
  two lemmas.
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1378
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1379
  \begin{figure}[t]
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1380
  \normalsize
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1381
  \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1382
  @{thm (lhs) tag_Plus_def[where A="A" and B="B", THEN meta_eq_app]} &
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1383
  @{text "\<equiv>"} &
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1384
  @{thm (rhs) tag_Plus_def[where A="A" and B="B", THEN meta_eq_app]} \\
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1385
  @{thm (lhs) tag_Times_def[where ?A="A" and ?B="B"]}\;@{text "x"} & @{text "\<equiv>"} &
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1386
  @{text "(\<lbrakk>x\<rbrakk>\<^bsub>\<approx>A\<^esub>, {\<lbrakk>x\<^sub>s\<rbrakk>\<^bsub>\<approx>B\<^esub> | x\<^sub>p \<in> A \<and> (x\<^sub>p, x\<^sub>s) \<in> Partitions x})"}\\
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1387
  @{thm (lhs) tag_Star_def[where ?A="A", THEN meta_eq_app]} & @{text "\<equiv>"} &
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1388
  @{text "{\<lbrakk>x\<^sub>s\<rbrakk>\<^bsub>\<approx>A\<^esub> | x\<^sub>p < x \<and> x\<^sub>p \<in> A\<^sup>\<star> \<and> (x\<^sub>p, x\<^sub>s) \<in> Partitions x}"}
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1389
  \end{tabular}
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1390
  \caption{Three tagging functions used the cases for @{term "PLUS"}, @{term "TIMES"} and @{term "STAR"} 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1391
  regular expressions. The sets @{text A} and @{text B} are arbitrary languages instantiated
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1392
  in the proof with the induction hypotheses. \emph{Partitions} is a function
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1393
  that generates all possible partitions of a string.\label{tagfig}}
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1394
  \end{figure}
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1395
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1396
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1397
  \begin{lemma}\label{finone}
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1398
  @{thm[mode=IfThen] finite_eq_tag_rel}
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1399
  \end{lemma}
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1400
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1401
  \begin{proof}
119
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
  1402
  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
  1403
  @{text "range f"} to be a subset of @{term "Pow (range tag)"}, which we know must be
201
9fbf6d9f85ae added comments by Xingyuan
urbanc
parents: 200
diff changeset
  1404
  finite by assumption. Now @{term "f ` (UNIV // =tag=)"} is a subset of @{text "range f"},
119
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
  1405
  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
  1406
  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
  1407
  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
  1408
  @{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
  1409
  turn means that the equivalence classes @{text X}
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1410
  and @{text Y} must be equal. Therefore \eqref{finiteimageD} allows us to conclude
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  1411
  with @{thm (concl) finite_eq_tag_rel}.\qed
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1412
  \end{proof}
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1413
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1414
  \begin{lemma}\label{fintwo} 
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1415
  Given two equivalence relations @{text "R\<^sub>1"} and @{text "R\<^sub>2"}, whereby
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1416
  @{text "R\<^sub>1"} refines @{text "R\<^sub>2"}.
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1417
  If @{thm (prem 1) refined_partition_finite[where ?R1.0="R\<^sub>1" and ?R2.0="R\<^sub>2"]}
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1418
  then @{thm (concl) refined_partition_finite[where ?R1.0="R\<^sub>1" and ?R2.0="R\<^sub>2"]}.
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1419
  \end{lemma}
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1420
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1421
  \begin{proof}
123
23c0e6f2929d polished everywhere...two cases still missing
urbanc
parents: 122
diff changeset
  1422
  We prove this lemma again using \eqref{finiteimageD}. This time we set @{text f} to
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1423
  be @{text "X \<mapsto>"}~@{term "{R\<^sub>1 `` {x} | x. x \<in> X}"}. It is easy to see that 
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1424
  @{term "finite (f ` (UNIV // R\<^sub>2))"} because it is a subset of @{term "Pow (UNIV // R\<^sub>1)"},
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1425
  which must be finite by assumption. What remains to be shown is that @{text f} is injective
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1426
  on @{term "UNIV // R\<^sub>2"}. This is equivalent to showing that two equivalence 
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1427
  classes, say @{text "X"} and @{text Y}, in @{term "UNIV // R\<^sub>2"} are equal, provided
118
c3fa11ee776e first proof
urbanc
parents: 117
diff changeset
  1428
  @{text "f X = f Y"}. For @{text "X = Y"} to be equal, we have to find two elements
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1429
  @{text "x \<in> X"} and @{text "y \<in> Y"} such that they are @{text R\<^sub>2} related.
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1430
  We know there exists a @{text "x \<in> X"} with \mbox{@{term "X = R\<^sub>2 `` {x}"}}. 
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1431
  From the latter fact we can infer that @{term "R\<^sub>1 ``{x} \<in> f X"}
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1432
  and further @{term "R\<^sub>1 ``{x} \<in> f Y"}. This means we can obtain a @{text y}
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1433
  such that @{term "R\<^sub>1 `` {x} = R\<^sub>1 `` {y}"} holds. Consequently @{text x} and @{text y}
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1434
  are @{text "R\<^sub>1"}-related. Since by assumption @{text "R\<^sub>1"} refines @{text "R\<^sub>2"},
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1435
  they must also be @{text "R\<^sub>2"}-related, as we need to show.\qed
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1436
  \end{proof}
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1437
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1438
  \noindent
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1439
  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
  1440
  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
  1441
  range can be shown to be finite and whose tagging-relation refines @{term "\<approx>(lang r)"}.
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  1442
  Let us attempt the @{const PLUS}-case first. We take from Fig.~\ref{tagfig}
183
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
  1443
 
119
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
  1444
  \begin{center}
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1445
  @{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
  1446
  \end{center}
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  1447
119
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
  1448
  \noindent
183
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
  1449
  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
  1450
  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
  1451
  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
  1452
  @{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
  1453
  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
  1454
375
44c4450152e3 adapted to JAR
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 374
diff changeset
  1455
 \begin{proof}[@{const "PLUS"}-Case]
183
c4893e84c88e cleaned up the proofs in Myhill_2
urbanc
parents: 182
diff changeset
  1456
  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
  1457
  (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
  1458
  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
  1459
  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
  1460
  "(\<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
  1461
  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
  1462
  show. Finally we can discharge this case by setting @{text A} to @{term
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1463
  "lang r\<^sub>1"} and @{text B} to @{term "lang r\<^sub>2"}.\qed
119
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
  1464
  \end{proof}
ece3f197b92b first two proofs in 2 direction
urbanc
parents: 118
diff changeset
  1465
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1466
  \noindent
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1467
  The @{const TIMES}-case is slightly more complicated. We first prove the
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1468
  following lemma, which will aid the proof about refinement.
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1469
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1470
  \begin{lemma}\label{refinement}
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1471
  The relation @{text "\<^raw:$\threesim$>\<^bsub>tag\<^esub>"} refines @{term "\<approx>A"}, provided for
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1472
  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
  1473
  and @{term "x @ z \<in> A"} imply @{text "y @ z \<in> A"}.
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1474
  \end{lemma}
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1475
109
79b37ef9505f minor updated
urbanc
parents: 108
diff changeset
  1476
121
1cf12a107b03 added directory with the small files and numbers of lines
urbanc
parents: 120
diff changeset
  1477
  \noindent
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1478
  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
  1479
  @{text A} and then construct an appropriate tagging-function to infer that
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1480
  @{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
  1481
  the set of all possible \emph{partitions} of a string:
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1482
  %
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1483
  \begin{equation}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1484
  @{thm Partitions_def}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1485
  \end{equation}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1486
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1487
  \noindent
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1488
  If we know that @{text "(x\<^sub>p, x\<^sub>s) \<in> Partitions x"}, we will
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1489
  refer to @{text "x\<^sub>p"} as the \emph{prefix} of the string @{text x},
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1490
  and respectively to @{text "x\<^sub>s"} as the \emph{suffix}.
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1491
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1492
*}
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1493
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1494
text {*
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1495
  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
  1496
  this string to be in @{term "A \<cdot> B"}:
132
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
  1497
  %
125
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1498
  \begin{center}
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1499
  \begin{tabular}{c}
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1500
  \scalebox{1}{
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  1501
  \begin{tikzpicture}[scale=0.8,fill=gray!20]
200
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
  1502
    \node[draw,minimum height=3.8ex, fill] (x) { $\hspace{4.8em}@{text x}\hspace{4.8em}$ };
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1503
    \node[draw,minimum height=3.8ex, right=-0.03em of x, fill] (za) { $\hspace{0.6em}@{text "z\<^sub>p"}\hspace{0.6em}$ };
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1504
    \node[draw,minimum height=3.8ex, right=-0.03em of za, fill] (zza) { $\hspace{2.6em}@{text "z\<^sub>s"}\hspace{2.6em}$  };
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1505
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1506
    \draw[decoration={brace,transform={yscale=3}},decorate]
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1507
           (x.north west) -- ($(za.north west)+(0em,0em)$)
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1508
               node[midway, above=0.5em]{@{text x}};
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1509
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1510
    \draw[decoration={brace,transform={yscale=3}},decorate]
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1511
           ($(za.north west)+(0em,0ex)$) -- ($(zza.north east)+(0em,0ex)$)
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1512
               node[midway, above=0.5em]{@{text z}};
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1513
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1514
    \draw[decoration={brace,transform={yscale=3}},decorate]
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1515
           ($(x.north west)+(0em,3ex)$) -- ($(zza.north east)+(0em,3ex)$)
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1516
               node[midway, above=0.8em]{@{term "x @ z \<in> A \<cdot> B"}};
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1517
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1518
    \draw[decoration={brace,transform={yscale=3}},decorate]
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1519
           ($(za.south east)+(0em,0ex)$) -- ($(x.south west)+(0em,0ex)$)
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1520
               node[midway, below=0.5em]{@{text "x @ z\<^sub>p \<in> A"}};
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1521
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1522
    \draw[decoration={brace,transform={yscale=3}},decorate]
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1523
           ($(zza.south east)+(0em,0ex)$) -- ($(za.south east)+(0em,0ex)$)
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1524
               node[midway, below=0.5em]{@{text "z\<^sub>s \<in> B"}};
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1525
  \end{tikzpicture}}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1526
  \\[2mm]
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1527
  \scalebox{1}{
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  1528
  \begin{tikzpicture}[scale=0.8,fill=gray!20]
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1529
    \node[draw,minimum height=3.8ex, fill] (xa) { $\hspace{3em}@{text "x\<^sub>p"}\hspace{3em}$ };
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1530
    \node[draw,minimum height=3.8ex, right=-0.03em of xa, fill] (xxa) { $\hspace{0.2em}@{text "x\<^sub>s"}\hspace{0.2em}$ };
200
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
  1531
    \node[draw,minimum height=3.8ex, right=-0.03em of xxa, fill] (z) { $\hspace{5em}@{text z}\hspace{5em}$ };
125
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1532
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1533
    \draw[decoration={brace,transform={yscale=3}},decorate]
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1534
           (xa.north west) -- ($(xxa.north east)+(0em,0em)$)
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1535
               node[midway, above=0.5em]{@{text x}};
125
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1536
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1537
    \draw[decoration={brace,transform={yscale=3}},decorate]
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1538
           (z.north west) -- ($(z.north east)+(0em,0em)$)
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1539
               node[midway, above=0.5em]{@{text z}};
125
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1540
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1541
    \draw[decoration={brace,transform={yscale=3}},decorate]
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1542
           ($(xa.north west)+(0em,3ex)$) -- ($(z.north east)+(0em,3ex)$)
167
61d0a412a3ae added a journal version
urbanc
parents: 162
diff changeset
  1543
               node[midway, above=0.8em]{@{term "x @ z \<in> A \<cdot> B"}};
125
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1544
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1545
    \draw[decoration={brace,transform={yscale=3}},decorate]
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1546
           ($(z.south east)+(0em,0ex)$) -- ($(xxa.south west)+(0em,0ex)$)
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1547
               node[midway, below=0.5em]{@{term "x\<^sub>s @ z \<in> B"}};
125
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1548
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1549
    \draw[decoration={brace,transform={yscale=3}},decorate]
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1550
           ($(xa.south east)+(0em,0ex)$) -- ($(xa.south west)+(0em,0ex)$)
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1551
               node[midway, below=0.5em]{@{term "x\<^sub>p \<in> A"}};
125
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1552
  \end{tikzpicture}}
159
990c12ab1562 edits; sqeezed to 16 pages
urbanc
parents: 157
diff changeset
  1553
  \end{tabular}
125
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1554
  \end{center}
132
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
  1555
  %
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1556
125
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1557
  \noindent
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1558
  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
  1559
  (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
  1560
  (second picture). In both cases we have to show that @{term "y @ z \<in> A \<cdot> B"}. The first case
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1561
  we will only go through if we know that  @{term "x \<approx>A y"} holds @{text "("}@{text "*"}@{text ")"}.  
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1562
  Because then 
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1563
  we can infer from @{term "x @ z\<^sub>p \<in> A"} that @{term "y @ z\<^sub>p \<in> A"} holds for all @{text "z\<^sub>p"}.
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1564
  In the second case we only know that @{text "x\<^sub>p"} and @{text "x\<^sub>s"} is one possible partition
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1565
  of the string @{text x}. We have to know that both @{text "x\<^sub>p"} and the
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1566
  corresponding partition @{text "y\<^sub>p"} are in @{text "A"}, and that @{text "x\<^sub>s"} is `@{text B}-related' 
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1567
  to @{text "y\<^sub>s"} @{text "("}@{text "**"}@{text ")"}. From the latter fact we can infer that @{text "y\<^sub>s @ z \<in> B"}.
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1568
  This will solve the second case.
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  1569
  Taking the two requirements, @{text "("}@{text "*"}@{text ")"} and @{text "(**)"}, together we define the
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  1570
  tagging-function in the @{const Times}-case as (see Fig.~\ref{tagfig}):
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1571
121
1cf12a107b03 added directory with the small files and numbers of lines
urbanc
parents: 120
diff changeset
  1572
  \begin{center}
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1573
  @{thm (lhs) tag_Times_def[where ?A="A" and ?B="B"]}\;@{text "x"}~@{text "\<equiv>"}~
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1574
  @{text "(\<lbrakk>x\<rbrakk>\<^bsub>\<approx>A\<^esub>, {\<lbrakk>x\<^sub>s\<rbrakk>\<^bsub>\<approx>B\<^esub> | x\<^sub>p \<in> A \<and> (x\<^sub>p, x\<^sub>s) \<in> Partitions x})"}
121
1cf12a107b03 added directory with the small files and numbers of lines
urbanc
parents: 120
diff changeset
  1575
  \end{center}
125
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1576
62925473bf6b added pictures for seq-case
urbanc
parents: 124
diff changeset
  1577
  \noindent
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1578
  Note that we have to make the assumption for all suffixes @{text "x\<^sub>s"}, since we do 
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1579
  not know anything about how the string @{term x} is partitioned.
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1580
  With this definition in place, let us prove the @{const "Times"}-case.
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1581
375
44c4450152e3 adapted to JAR
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 374
diff changeset
  1582
  \begin{proof}[@{const TIMES}-Case]
127
8440863a9900 seq case finished
urbanc
parents: 126
diff changeset
  1583
  If @{term "finite (UNIV // \<approx>A)"} and @{term "finite (UNIV // \<approx>B)"}
8440863a9900 seq case finished
urbanc
parents: 126
diff changeset
  1584
  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
  1585
  @{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
  1586
  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
  1587
  we have by Lemma \ref{refinement} 
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1588
127
8440863a9900 seq case finished
urbanc
parents: 126
diff changeset
  1589
  \begin{center}
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1590
   @{term "tag_Times A B x = tag_Times A B y"}
127
8440863a9900 seq case finished
urbanc
parents: 126
diff changeset
  1591
  \end{center}
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1592
127
8440863a9900 seq case finished
urbanc
parents: 126
diff changeset
  1593
  \noindent
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1594
  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
  1595
  B"}. As shown in the pictures above, there are two cases to be
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1596
  considered. First, there exists a @{text "z\<^sub>p"} and @{text
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1597
  "z\<^sub>s"} such that @{term "x @ z\<^sub>p \<in> A"} and @{text "z\<^sub>s
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1598
  \<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
  1599
  `` {x} = \<approx>A `` {y}"} and thus @{term "x \<approx>A y"}. Hence by the Myhill-Nerode
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1600
  Relation @{term "y @ z\<^sub>p \<in> A"} holds. Using @{text "z\<^sub>s \<in> B"},
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1601
  we can conclude in this case with @{term "y @ z \<in> A \<cdot> B"} (recall @{text
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1602
  "z\<^sub>p @ z\<^sub>s = z"}).
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1603
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1604
  Second there exists a partition @{text "x\<^sub>p"} and @{text "x\<^sub>s"} with 
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1605
  @{text "x\<^sub>p \<in> A"} and @{text "x\<^sub>s @ z \<in> B"}. We therefore have
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1606
  
127
8440863a9900 seq case finished
urbanc
parents: 126
diff changeset
  1607
  \begin{center}
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1608
  @{text "\<lbrakk>x\<^sub>s\<rbrakk>\<^bsub>\<approx>B\<^esub> \<in> {\<lbrakk>x\<^sub>s\<rbrakk>\<^bsub>\<approx>B\<^esub> | x\<^sub>p \<in> A \<and> (x\<^sub>p, x\<^sub>s) \<in> Partitions x}"}
127
8440863a9900 seq case finished
urbanc
parents: 126
diff changeset
  1609
  \end{center}
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1610
  
127
8440863a9900 seq case finished
urbanc
parents: 126
diff changeset
  1611
  \noindent
181
97090fc7aa9f some experiments with the proofs in Myhill_2
urbanc
parents: 180
diff changeset
  1612
  and by the assumption about @{term "tag_Times A B"} also 
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1613
  
127
8440863a9900 seq case finished
urbanc
parents: 126
diff changeset
  1614
  \begin{center}
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1615
  @{text "\<lbrakk>x\<^sub>s\<rbrakk>\<^bsub>\<approx>B\<^esub> \<in> {\<lbrakk>y\<^sub>s\<rbrakk>\<^bsub>\<approx>B\<^esub> | y\<^sub>p \<in> A \<and> (y\<^sub>p, y\<^sub>s) \<in> Partitions y}"}
127
8440863a9900 seq case finished
urbanc
parents: 126
diff changeset
  1616
  \end{center}
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1617
  
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1618
  \noindent
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1619
  This means there must be a partition @{text "y\<^sub>p"} and @{text "y\<^sub>s"}
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1620
  such that @{term "y\<^sub>p \<in> A"} and @{term "\<approx>B `` {x\<^sub>s} = \<approx>B ``
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1621
  {y\<^sub>s}"}. Unfolding the Myhill-Nerode Relation and together with the
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1622
  facts that @{text "x\<^sub>p \<in> A"} and \mbox{@{text "x\<^sub>s @ z \<in> B"}}, we
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1623
  obtain @{term "y\<^sub>p \<in> A"} and @{text "y\<^sub>s @ z \<in> B"}, as needed in
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1624
  this case.  We again can complete the @{const TIMES}-case by setting @{text
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1625
  A} to @{term "lang r\<^sub>1"} and @{text B} to @{term "lang r\<^sub>2"}.\qed
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1626
  \end{proof}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1627
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1628
  \noindent
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1629
  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
  1630
  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
  1631
  prefix} and a \emph{strict string prefix}:
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1632
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1633
  \begin{center}
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1634
  \begin{tabular}{l}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1635
  @{text "x \<le> y \<equiv> \<exists>z. y = x @ z"}\\
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1636
  @{text "x < y \<equiv> x \<le> y \<and> x \<noteq> y"}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1637
  \end{tabular}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1638
  \end{center}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1639
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1640
  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
  1641
  and @{text x} is not the empty string, we have the following picture:
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1642
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1643
  \begin{center}
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1644
  \scalebox{1}{
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  1645
  \begin{tikzpicture}[scale=0.8,fill=gray!20]
200
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
  1646
    \node[draw,minimum height=3.8ex, fill] (xa) { $\hspace{4em}@{text "x\<^bsub>pmax\<^esub>"}\hspace{4em}$ };
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
  1647
    \node[draw,minimum height=3.8ex, right=-0.03em of xa, fill] (xxa) { $\hspace{0.5em}@{text "x\<^bsub>s\<^esub>"}\hspace{0.5em}$ };
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1648
    \node[draw,minimum height=3.8ex, right=-0.03em of xxa, fill] (za) { $\hspace{2em}@{text "z\<^sub>a"}\hspace{2em}$ };
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1649
    \node[draw,minimum height=3.8ex, right=-0.03em of za, fill] (zb) { $\hspace{7em}@{text "z\<^sub>b"}\hspace{7em}$ };
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1650
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1651
    \draw[decoration={brace,transform={yscale=3}},decorate]
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1652
           (xa.north west) -- ($(xxa.north east)+(0em,0em)$)
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1653
               node[midway, above=0.5em]{@{text x}};
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1654
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1655
    \draw[decoration={brace,transform={yscale=3}},decorate]
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1656
           (za.north west) -- ($(zb.north east)+(0em,0em)$)
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1657
               node[midway, above=0.5em]{@{text z}};
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1658
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1659
    \draw[decoration={brace,transform={yscale=3}},decorate]
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1660
           ($(xa.north west)+(0em,3ex)$) -- ($(zb.north east)+(0em,3ex)$)
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1661
               node[midway, above=0.8em]{@{term "x @ z \<in> A\<star>"}};
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1662
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1663
    \draw[decoration={brace,transform={yscale=3}},decorate]
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1664
           ($(za.south east)+(0em,0ex)$) -- ($(xxa.south west)+(0em,0ex)$)
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1665
               node[midway, below=0.5em]{@{term "x\<^sub>s @ z\<^sub>a \<in> A"}};
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1666
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1667
    \draw[decoration={brace,transform={yscale=3}},decorate]
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1668
           ($(xa.south east)+(0em,0ex)$) -- ($(xa.south west)+(0em,0ex)$)
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1669
               node[midway, below=0.5em]{@{text "x\<^bsub>pmax\<^esub> \<in> A\<^sup>\<star>"}};
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1670
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1671
    \draw[decoration={brace,transform={yscale=3}},decorate]
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1672
           ($(zb.south east)+(0em,0ex)$) -- ($(zb.south west)+(0em,0ex)$)
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1673
               node[midway, below=0.5em]{@{term "z\<^sub>b \<in> A\<star>"}};
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1674
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1675
    \draw[decoration={brace,transform={yscale=3}},decorate]
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1676
           ($(zb.south east)+(0em,-4ex)$) -- ($(xxa.south west)+(0em,-4ex)$)
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1677
               node[midway, below=0.5em]{@{term "x\<^sub>s @ z \<in> A\<star>"}};
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1678
  \end{tikzpicture}}
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1679
  \end{center}
132
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
  1680
  %
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1681
  \noindent
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1682
  We can find a strict prefix @{text "x\<^sub>p"} of @{text x} such that @{term "x\<^sub>p \<in> A\<star>"},
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1683
  @{text "x\<^sub>p < x"} and the rest @{term "x\<^sub>s @ z \<in> A\<star>"}. For example the empty string 
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1684
  @{text "[]"} would do (recall @{term "x \<noteq> []"}).
135
604518f0127f final polished
urbanc
parents: 134
diff changeset
  1685
  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
  1686
  string @{text x} is finite). Let us therefore choose the longest one and call it 
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1687
  @{text "x\<^bsub>pmax\<^esub>"}. Now for the rest of the string @{text "x\<^sub>s @ z"} we
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1688
  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
  1689
  we can separate
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1690
  this string into two parts, say @{text "a"} and @{text "b"}, such that @{text "a \<noteq> []"}, @{text "a \<in> A"}
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1691
  and @{term "b \<in> A\<star>"}. Now @{text a} must be strictly longer than @{text "x\<^sub>s"},
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1692
  otherwise @{text "x\<^bsub>pmax\<^esub>"} is not the longest prefix. That means @{text a}
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1693
  `overlaps' with @{text z}, splitting it into two components @{text "z\<^sub>a"} and
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1694
   @{text "z\<^sub>b"}. For this we know that @{text "x\<^sub>s @ z\<^sub>a \<in> A"} and
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1695
  @{term "z\<^sub>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
  1696
  such that we have a string @{text a} with @{text "a \<in> A"} that lies just on the
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1697
  `border' of @{text x} and @{text z}. This string is @{text "x\<^sub>s @ z\<^sub>a"}.
184
2455db3b06ac more on the paper
urbanc
parents: 183
diff changeset
  1698
135
604518f0127f final polished
urbanc
parents: 134
diff changeset
  1699
  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
  1700
  the following tagging-function:
132
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
  1701
  %
121
1cf12a107b03 added directory with the small files and numbers of lines
urbanc
parents: 120
diff changeset
  1702
  \begin{center}
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1703
  @{thm (lhs) tag_Star_def[where ?A="A", THEN meta_eq_app]}~@{text "\<equiv>"}~
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1704
  @{text "{\<lbrakk>x\<^sub>s\<rbrakk>\<^bsub>\<approx>A\<^esub> | x\<^sub>p < x \<and> x\<^sub>p \<in> A\<^sup>\<star> \<and> (x\<^sub>p, x\<^sub>s) \<in> Partitions x}"}
121
1cf12a107b03 added directory with the small files and numbers of lines
urbanc
parents: 120
diff changeset
  1705
  \end{center}
128
6d2693c78c37 finished picture
urbanc
parents: 127
diff changeset
  1706
375
44c4450152e3 adapted to JAR
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 374
diff changeset
  1707
  \begin{proof}[@{const Star}-Case]
130
3e4ad22193e7 pre-final version
urbanc
parents: 129
diff changeset
  1708
  If @{term "finite (UNIV // \<approx>A)"} 
3e4ad22193e7 pre-final version
urbanc
parents: 129
diff changeset
  1709
  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
  1710
  @{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
  1711
  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
  1712
  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
  1713
130
3e4ad22193e7 pre-final version
urbanc
parents: 129
diff changeset
  1714
  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
  1715
  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
  1716
  can infer @{text y} is the empty string and
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1717
  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
  1718
  string, we can divide the string @{text "x @ z"} as shown in the picture 
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1719
  above. By the tagging-function and the facts @{text "x\<^bsub>pmax\<^esub> \<in> A\<^sup>\<star>"} and @{text "x\<^bsub>pmax\<^esub> < x"}, 
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1720
  we have
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1721
130
3e4ad22193e7 pre-final version
urbanc
parents: 129
diff changeset
  1722
  \begin{center}
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1723
  @{text "\<lbrakk>x\<^sub>s\<rbrakk>\<^bsub>\<approx>A\<^esub> \<in> {\<lbrakk>x\<^sub>s\<rbrakk>\<^bsub>\<approx>A\<^esub> | x\<^bsub>pmax\<^esub> < x \<and> x\<^bsub>pmax\<^esub> \<in> A\<^sup>\<star> \<and> (x\<^bsub>pmax\<^esub>, x\<^sub>s) \<in> Partitions x}"}
130
3e4ad22193e7 pre-final version
urbanc
parents: 129
diff changeset
  1724
  \end{center}
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1725
  
130
3e4ad22193e7 pre-final version
urbanc
parents: 129
diff changeset
  1726
  \noindent
3e4ad22193e7 pre-final version
urbanc
parents: 129
diff changeset
  1727
  which by assumption is equal to
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1728
  
130
3e4ad22193e7 pre-final version
urbanc
parents: 129
diff changeset
  1729
  \begin{center}
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1730
  @{text "\<lbrakk>x\<^sub>s\<rbrakk>\<^bsub>\<approx>A\<^esub> \<in> {\<lbrakk>y\<^sub>s\<rbrakk>\<^bsub>\<approx>A\<^esub> | y\<^bsub>p\<^esub> < y \<and> y\<^bsub>p\<^esub> \<in> A\<^sup>\<star> \<and> (y\<^bsub>p\<^esub>, y\<^sub>s) \<in> Partitions y}"}
130
3e4ad22193e7 pre-final version
urbanc
parents: 129
diff changeset
  1731
  \end{center}
185
8749db46d5e6 completed the taging-function section
urbanc
parents: 184
diff changeset
  1732
  
130
3e4ad22193e7 pre-final version
urbanc
parents: 129
diff changeset
  1733
  \noindent 
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1734
  From this we know there exist a partition @{text "y\<^sub>p"} and @{text
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1735
  "y\<^sub>s"} with @{term "y\<^sub>p \<in> A\<star>"} and also @{term "x\<^sub>s \<approx>A
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1736
  y\<^sub>s"}. Unfolding the Myhill-Nerode Relation we know @{term
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1737
  "y\<^sub>s @ z\<^sub>a \<in> A"}. We also know that @{term "z\<^sub>b \<in> A\<star>"}.
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1738
  Therefore @{term "y\<^sub>p @ (y\<^sub>s @ z\<^sub>a) @ z\<^sub>b \<in>
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1739
  A\<star>"}, which means @{term "y @ z \<in> A\<star>"}. The last step is to set
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  1740
  @{text "A"} to @{term "lang r"} and thus complete the proof.\qed
121
1cf12a107b03 added directory with the small files and numbers of lines
urbanc
parents: 120
diff changeset
  1741
  \end{proof}
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1742
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1743
  \begin{rmk}
388
0da31edd95b9 new version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 387
diff changeset
  1744
  While our proof using tagging functions might seem like a `rabbit pulled 
0da31edd95b9 new version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 387
diff changeset
  1745
  out of a hat', the intuition behind can be somewhat rationalised by taking the 
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1746
  view that the equivalence classes @{term "UNIV // \<approx>(lang r)"} stand for the 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1747
  states of the minimal automaton for the language @{term "lang r"}. The theorem 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1748
  amounts to showing that this minimal automaton has finitely many states. 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1749
  However, by showing that our @{text "\<^raw:$\threesim$>\<^bsub>tag\<^esub>"} relation 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1750
  refines @{term "\<approx>A"} we do not actually have to show that the minimal automata
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1751
  has finitely many states, but only that we can show finiteness for an 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1752
  automaton with at least as many states and which can recognise the same 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1753
  language. By performing the induction over the regular expression @{text r},
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1754
  this means we have to construct inductively an automaton in
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1755
  the cases for @{term PLUS}, @{term TIMES} and @{term STAR}.
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1756
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1757
  In the @{text PLUS}-case, we can assume finitely many equivalence classes of the form
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1758
  @{text "\<lbrakk>_\<rbrakk>\<^bsub>\<approx>A\<^esub>"} and @{text "\<lbrakk>_\<rbrakk>\<^bsub>\<approx>B\<^esub>"}, each standing for an automaton recognising the 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1759
  languages @{text A} and @{text B} respectively. 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1760
  The @{text "+tag"} function constructs pairs of the form @{text "(\<lbrakk>_\<rbrakk>\<^bsub>\<approx>A\<^esub>, \<lbrakk>_\<rbrakk>\<^bsub>\<approx>B\<^esub>)"} 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1761
  which can be seen as the states of an automaton recognising the language 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1762
  \mbox{@{term "A \<union> B"}}. This is the usual product construction of automata: 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1763
  Given a string @{text x}, the first automata will be in the state  @{text "\<lbrakk>x\<rbrakk>\<^bsub>\<approx>A\<^esub>"}
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1764
  after recognising @{text x} (similarly @{text "\<lbrakk>x\<rbrakk>\<^bsub>\<approx>B\<^esub>"} for the other automaton). The product
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1765
  automaton will be in the state \mbox{@{text "(\<lbrakk>x\<rbrakk>\<^bsub>\<approx>A\<^esub>, \<lbrakk>x\<rbrakk>\<^bsub>\<approx>B\<^esub>)"}}, which is accepting
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1766
  if at least one component is an accepting state.
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1767
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1768
  The @{text "TIMES"}-case is a bit more complicated---essentially we 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1769
  need to sequentially compose the two automata with the states @{text "\<lbrakk>_\<rbrakk>\<^bsub>\<approx>A\<^esub>"}, 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1770
  respectively @{text "\<lbrakk>_\<rbrakk>\<^bsub>\<approx>B\<^esub>"}. We achieve this sequential composition by
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1771
  taking the first automaton @{text "\<lbrakk>_\<rbrakk>\<^bsub>\<approx>A\<^esub>"} and append on each of its accepting
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1772
  state the automaton @{text "\<lbrakk>_\<rbrakk>\<^bsub>\<approx>B\<^esub>"}. Unfortunately, this does not lead directly
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1773
  to a correct composition, since the automaton @{text "\<lbrakk>_\<rbrakk>\<^bsub>\<approx>A\<^esub>"} might have consumed
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1774
  some of the input needed for the automaton @{text "\<lbrakk>_\<rbrakk>\<^bsub>\<approx>B\<^esub>"} to succeed. The 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1775
  textbook construction for sequentially composing two automata circumvents this
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1776
  problem by connecting the terminating states of the first automaton via
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1777
  an epsilon-transition to the initial state of the second automaton (see for
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1778
  example \citeN{Kozen97}). In the absence of any epsilon-transition analogue in our work, 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1779
  we let the second automaton
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1780
  start in all possible states that may have been reached if the first
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1781
  automaton had not consumed the input meant for the second. We achieve this
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1782
  by having a set of states as the second component generated by our 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1783
  @{text "\<times>tag"} function (see second clause in Fig.~\ref{tagfig}).
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1784
  A state of this sequentially composed automaton is accepting, if the first 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1785
  component is accepting and at least one state in the set is also accepting.
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1786
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1787
  The idea behind the @{text "STAR"}-case is similar to the @{text "TIMES"}-case.
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1788
  We assume some automaton has consumed some strictly smaller part of the input in @{text "A\<^sup>\<star>"};
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1789
  we need to check that from the state we ended up in a terminal state in the 
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  1790
  automaton @{text "\<lbrakk>_\<rbrakk>\<^bsub>\<approx>A\<^esub>"} can be reached. Since we do not know from which state this will 
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1791
  succeed, we need to run the automaton from all possible states we could have 
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  1792
  ended up in. Therefore the @{text "\<star>tag"} function generates again a set of states.
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  1793
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  1794
  However, note that while the automata view sheds some light behind the idea of the 
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  1795
  tagging functions, our proof only works because we can perform a structural 
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  1796
  induction on the regular expression @{text r}.
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  1797
  \end{rmk}
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
  1798
*}
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
  1799
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  1800
section {* Second Part proved using Partial Derivatives *}
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 160
diff changeset
  1801
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 160
diff changeset
  1802
text {*
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  1803
  \label{derivatives} 
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
  1804
  \noindent
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
  1805
  As we have seen in the previous section, in order to establish
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
  1806
  the second direction of the Myhill-Nerode Theorem, it is sufficient to find 
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1807
  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
  1808
  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
  1809
  showed this directly by induction on @{text "r"} using tagging-functions. 
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1810
  However, there is also  an indirect method to come up with such a refined 
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1811
  relation by using derivatives of regular expressions introduced by \citeN{Brzozowski64}. 
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1812
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1813
  Assume the following two definitions for the \emph{left-quotient} of a language,
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1814
  which we write as @{term "Deriv c A"} and @{term "Derivs s A"} where @{text c}
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1815
  is a character and @{text s} a string, respectively:
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1816
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1817
  \begin{center}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1818
  \begin{tabular}{r@ {\hspace{1mm}}c@ {\hspace{2mm}}l}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1819
  @{thm (lhs) Der_def}  & @{text "\<equiv>"} & @{thm (rhs) Der_def}\\
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1820
  @{thm (lhs) Ders_def} & @{text "\<equiv>"} & @{thm (rhs) Ders_def}\\
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1821
  \end{tabular}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1822
  \end{center}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1823
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1824
  \noindent
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1825
  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
  1826
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1827
  \begin{center}
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1828
  @{abbrev "Derivss s As"}
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1829
  \end{center}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1830
  
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1831
  \noindent
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1832
  where we apply the left-quotient to a set of languages and then combine the results.
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
  1833
  Clearly we have the following equivalence between the Myhill-Nerode Relation
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1834
  (Definition~\ref{myhillneroderel}) and left-quotients
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1835
  %
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1836
  \begin{equation}\label{mhders}
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1837
  @{term "x \<approx>A y"} \hspace{4mm}\text{if and only if}\hspace{4mm} @{term "Derivs x A = Derivs y A"}
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1838
  \end{equation}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1839
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1840
  \noindent
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1841
  It is also straightforward to establish the following properties of left-quotients
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1842
  % 
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1843
  \begin{equation}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1844
  \mbox{\begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{2mm}}l}
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1845
  @{thm (lhs) Deriv_simps(1)} & $=$ & @{thm (rhs) Deriv_simps(1)}\\
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1846
  @{thm (lhs) Deriv_simps(2)} & $=$ & @{thm (rhs) Deriv_simps(2)}\\
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1847
  @{thm (lhs) Deriv_simps(3)} & $=$ & @{thm (rhs) Deriv_simps(3)}\\
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1848
  @{thm (lhs) Deriv_simps(4)} & $=$ & @{thm (rhs) Deriv_simps(4)}\\
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1849
  @{thm (lhs) Der_conc}  & $=$ & @{thm (rhs) Der_conc}\\
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1850
  @{thm (lhs) Deriv_star}  & $=$ & @{thm (rhs) Deriv_star}\\
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1851
  @{thm (lhs) Derivs_simps(1)} & $=$ & @{thm (rhs) Derivs_simps(1)}\\
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1852
  @{thm (lhs) Derivs_simps(2)} & $=$ & @{thm (rhs) Derivs_simps(2)}\\
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1853
  %@{thm (lhs) Derivs_simps(3)[where ?s1.0="s\<^sub>1" and ?s2.0="s\<^sub>2"]}  & $=$ 
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1854
  %   & @{thm (rhs) Derivs_simps(3)[where ?s1.0="s\<^sub>1" and ?s2.0="s\<^sub>2"]}\\
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1855
  \end{tabular}}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1856
  \end{equation}
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1857
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1858
  \noindent
245
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  1859
  Note that in the last equation we use the list-cons operator written
199
11c3c302fa2e a little tuning
urbanc
parents: 198
diff changeset
  1860
  \mbox{@{text "_ :: _"}}.  The only interesting case is the case of @{term "A\<star>"}
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1861
  where we use Property~\ref{langprops}@{text "(i)"} in order to infer that
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1862
  @{term "Deriv c (A\<star>) = Deriv c (A \<cdot> A\<star>)"}. We can then complete the proof by
245
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  1863
  using the fifth equation and noting that @{term "Deriv c (A\<star>) \<subseteq> (Deriv
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  1864
  c A) \<cdot> A\<star>"} provided @{text "[] \<in> A"}. 
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1865
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1866
  \citeN{Brzozowski64} observed that the left-quotients for languages of
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1867
  regular expressions can be calculated directly using the notion of
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1868
  \emph{derivatives of a regular expression}. We define
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1869
  this notion in Isabelle/HOL as follows:
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1870
  %
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1871
  \begin{center}
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  1872
  \begin{longtable}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1.5mm}}l@ {}}
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1873
  @{thm (lhs) deriv.simps(1)}  & @{text "\<equiv>"} & @{thm (rhs) deriv.simps(1)}\\
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1874
  @{thm (lhs) deriv.simps(2)}  & @{text "\<equiv>"} & @{thm (rhs) deriv.simps(2)}\\
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1875
  @{thm (lhs) deriv.simps(3)[where c'="d"]}  & @{text "\<equiv>"} & @{thm (rhs) deriv.simps(3)[where c'="d"]}\\
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1876
  @{thm (lhs) deriv.simps(4)[where ?r1.0="r\<^sub>1" and ?r2.0="r\<^sub>2"]}  
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1877
     & @{text "\<equiv>"} & @{thm (rhs) deriv.simps(4)[where ?r1.0="r\<^sub>1" and ?r2.0="r\<^sub>2"]}\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1878
  @{thm (lhs) deriv.simps(5)[where ?r1.0="r\<^sub>1" and ?r2.0="r\<^sub>2"]}  
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1879
     & @{text "\<equiv>"} & @{text "if"}~@{term "nullable r\<^sub>1"}~@{text "then"}~%
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1880
       @{term "Plus (Times (deriv c r\<^sub>1) r\<^sub>2) (deriv c r\<^sub>2)"}\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1881
     &             & \phantom{@{text "if"}~@{term "nullable r\<^sub>1"}~}@{text "else"}~%  
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1882
                    @{term "Times (deriv c r\<^sub>1) r\<^sub>2"}\\ 
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1883
  @{thm (lhs) deriv.simps(6)}  & @{text "\<equiv>"} & @{thm (rhs) deriv.simps(6)}\smallskip\\
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1884
  @{thm (lhs) derivs.simps(1)}  & @{text "\<equiv>"} & @{thm (rhs) derivs.simps(1)}\\
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  1885
  @{thm (lhs) derivs.simps(2)}  & @{text "\<equiv>"} & @{thm (rhs) derivs.simps(2)}
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  1886
  \end{longtable}
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1887
  \end{center}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1888
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1889
  \noindent
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1890
  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
  1891
  boolean function @{term "nullable r"} needed in the @{const Times}-case tests
197
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1892
  whether a regular expression can recognise the empty string. It can be defined as 
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1893
  follows.
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1894
  %
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1895
  \begin{center}
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1896
  \begin{longtable}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1.5mm}}l@ {}}
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1897
  @{thm (lhs) nullable.simps(1)}  & @{text "\<equiv>"} & @{thm (rhs) nullable.simps(1)}\\
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1898
  @{thm (lhs) nullable.simps(2)}  & @{text "\<equiv>"} & @{thm (rhs) nullable.simps(2)}\\
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1899
  @{thm (lhs) nullable.simps(3)}  & @{text "\<equiv>"} & @{thm (rhs) nullable.simps(3)}\\
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1900
  @{thm (lhs) nullable.simps(4)[where ?r1.0="r\<^sub>1" and ?r2.0="r\<^sub>2"]}  
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1901
     & @{text "\<equiv>"} & @{thm (rhs) nullable.simps(4)[where ?r1.0="r\<^sub>1" and ?r2.0="r\<^sub>2"]}\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1902
  @{thm (lhs) nullable.simps(5)[where ?r1.0="r\<^sub>1" and ?r2.0="r\<^sub>2"]}  
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1903
     & @{text "\<equiv>"} & @{thm (rhs) nullable.simps(5)[where ?r1.0="r\<^sub>1" and ?r2.0="r\<^sub>2"]}\\
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1904
  @{thm (lhs) nullable.simps(6)}  & @{text "\<equiv>"} & @{thm (rhs) nullable.simps(6)}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1905
  \end{longtable}
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1906
  \end{center}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1907
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1908
  \noindent
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  1909
  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
  1910
  one can easily show that left-quotients and derivatives of regular expressions 
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1911
  relate as follows (see for example~\cite{Sakarovitch09}):
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1912
  %
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1913
  \begin{equation}\label{Dersders}
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1914
  \mbox{\begin{tabular}{c}
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1915
  @{thm Deriv_deriv}\\
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1916
  @{thm Derivs_derivs}
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1917
  \end{tabular}}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1918
  \end{equation}
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1919
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1920
  \noindent
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
  1921
  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
  1922
  we can use \eqref{mhders} and \eqref{Dersders} in order to 
200
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
  1923
  establish that 
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
  1924
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
  1925
  \begin{center}
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
  1926
  @{term "x \<approx>(lang r) y"} \hspace{4mm}if and only if\hspace{4mm}
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1927
  @{term "lang (derivs x r) = lang (derivs y r)"}. 
200
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
  1928
  \end{center}
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
  1929
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
  1930
  \noindent
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
  1931
  holds and hence
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1932
  %
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1933
  \begin{equation}
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1934
  @{term "x \<approx>(lang r) y"}\hspace{4mm}\mbox{provided}\hspace{4mm}@{term "derivs x r = derivs y r"}
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1935
  \end{equation}
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1936
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1937
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1938
  \noindent
200
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
  1939
  This means the right-hand side (seen as a relation) refines the Myhill-Nerode
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
  1940
  Relation.  Consequently, we can use @{text
197
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1941
  "\<^raw:$\threesim$>\<^bsub>(\<lambda>x. ders x r)\<^esub>"} as a
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1942
  tagging-relation. However, in order to be useful for the second part of the
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
  1943
  Myhill-Nerode Theorem, we have to be able to establish that for the
197
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1944
  corresponding language there are only finitely many derivatives---thus
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1945
  ensuring that there are only finitely many equivalence
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1946
  classes. Unfortunately, this is not true in general. Sakarovitch gives an
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1947
  example where a regular expression has infinitely many derivatives
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1948
  w.r.t.~the language @{text "(ab)\<^sup>\<star> \<union> (ab)\<^sup>\<star>a"}, which is formally 
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1949
  written in our notation as \mbox{@{text "{[a,b]}\<^sup>\<star> \<union> ({[a,b]}\<^sup>\<star> \<cdot> {[a]})"}}
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  1950
  (see \cite[Page~141]{Sakarovitch09}).
197
cf1c17431dab a bit more polishing
urbanc
parents: 196
diff changeset
  1951
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1952
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1953
  What \citeN{Brzozowski64} established is that for every language there
199
11c3c302fa2e a little tuning
urbanc
parents: 198
diff changeset
  1954
  \emph{are} only finitely `dissimilar' derivatives for a regular
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1955
  expression. Two regular expressions are said to be \emph{similar} provided
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  1956
  they can be identified using the using the @{text "ACI"}-identities:
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1957
  %
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1958
  \begin{equation}\label{ACI}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1959
  \mbox{\begin{tabular}{cl}
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1960
  (@{text A}) & @{term "Plus (Plus r\<^sub>1 r\<^sub>2) r\<^sub>3"} $\equiv$ @{term "Plus r\<^sub>1 (Plus r\<^sub>2 r\<^sub>3)"}\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1961
  (@{text C}) & @{term "Plus r\<^sub>1 r\<^sub>2"} $\equiv$ @{term "Plus r\<^sub>2 r\<^sub>1"}\\
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1962
  (@{text I}) & @{term "Plus r r"} $\equiv$ @{term "r"}\\
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1963
  \end{tabular}}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1964
  \end{equation}
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1965
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1966
  \noindent
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1967
  Carrying this idea through, we must not consider the set of all derivatives,
199
11c3c302fa2e a little tuning
urbanc
parents: 198
diff changeset
  1968
  but the one modulo @{text "ACI"}.  In principle, this can be done formally, 
387
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
  1969
  but it is very painful in a theorem prover---since there is no
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
  1970
  direct characterisation of the set of dissimilar derivatives. Therefore
388
0da31edd95b9 new version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 387
diff changeset
  1971
  \citeN{CoquandSiles12} who follow through this idea had to spend extra effort and
387
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
  1972
  first formalise the quite intricate notion of \emph{inductively finite sets} in 
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
  1973
  order to formalise this property.
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1974
174
2b414a8a7132 more on the section about derivatives
urbanc
parents: 173
diff changeset
  1975
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1976
  Fortunately, there is a much simpler approach using \emph{partial
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1977
  derivatives}. They were introduced by \citeN{Antimirov95} and can be defined
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1978
  in Isabelle/HOL as follows:
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1979
  %
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
  1980
  \begin{center}
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1981
  \begin{longtable}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1.5mm}}l@ {}}
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1982
  @{thm (lhs) pderiv.simps(1)}  & @{text "\<equiv>"} & @{thm (rhs) pderiv.simps(1)}\\
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1983
  @{thm (lhs) pderiv.simps(2)}  & @{text "\<equiv>"} & @{thm (rhs) pderiv.simps(2)}\\
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1984
  @{thm (lhs) pderiv.simps(3)[where c'="d"]}  & @{text "\<equiv>"} & @{thm (rhs) pderiv.simps(3)[where c'="d"]}\\
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1985
  @{thm (lhs) pderiv.simps(4)[where ?r1.0="r\<^sub>1" and ?r2.0="r\<^sub>2"]}  
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1986
     & @{text "\<equiv>"} & @{thm (rhs) pderiv.simps(4)[where ?r1.0="r\<^sub>1" and ?r2.0="r\<^sub>2"]}\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1987
  @{thm (lhs) pderiv.simps(5)[where ?r1.0="r\<^sub>1" and ?r2.0="r\<^sub>2"]}  
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1988
     & @{text "\<equiv>"} & @{text "if"}~@{term "nullable r\<^sub>1"}~@{text "then"}~%
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1989
       @{term "(Timess (pderiv c r\<^sub>1) r\<^sub>2) \<union> (pderiv c r\<^sub>2)"}\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1990
     & & \phantom{@{text "if"}~@{term "nullable r\<^sub>1"}~}@{text "else"}~%  
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  1991
                    @{term "Timess (pderiv c r\<^sub>1) r\<^sub>2"}\\ 
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1992
  @{thm (lhs) pderiv.simps(6)}  & @{text "\<equiv>"} & @{thm (rhs) pderiv.simps(6)}\smallskip\\
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  1993
  @{thm (lhs) pderivs.simps(1)}  & @{text "\<equiv>"} & @{thm (rhs) pderivs.simps(1)}\\
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1994
  @{thm (lhs) pderivs.simps(2)}  & @{text "\<equiv>"} & @{text "\<Union> (pders s) ` (pder c r)"}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  1995
  \end{longtable}
175
edc642266a82 polished the introduction
urbanc
parents: 174
diff changeset
  1996
  \end{center}
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
  1997
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  1998
  \noindent
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  1999
  Again the last two clauses extend partial derivatives from characters to strings. 
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2000
  Unlike `simple' derivatives, the functions for partial derivatives return sets of regular
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2001
  expressions. In the @{const Times} and @{const Star} cases we therefore use the
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2002
  auxiliary definition
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2003
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2004
  \begin{center}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2005
  @{text "TIMESS rs r \<equiv> {TIMES r' r | r' \<in> rs}"}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2006
  \end{center}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2007
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2008
  \noindent
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2009
  in order to `sequence' a regular expression with a set of regular
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2010
  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
  2011
  derivatives w.r.t~the character @{text c}, then build the image of this set under the
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2012
  function @{term "pderivs s"} and finally `union up' all resulting sets. It will be
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  2013
  convenient to introduce for this the following abbreviation
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2014
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2015
  \begin{center}
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2016
  @{abbrev "pderivs_set s rs"}
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2017
  \end{center}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2018
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2019
  \noindent
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2020
  which simplifies the last clause of @{const "pderivs"} to
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2021
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2022
  \begin{center}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2023
  \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1.5mm}}l@ {}}
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2024
  @{thm (lhs) pderivs.simps(2)}  & @{text "\<equiv>"} & @{thm (rhs) pderivs.simps(2)}\\
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2025
  \end{tabular}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2026
  \end{center}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2027
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2028
  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
  2029
  taking the partial derivatives of the
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2030
  regular expressions in \eqref{ACI} gives us in each case
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2031
  equal sets.  \citeN{Antimirov95} showed a similar result to
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2032
  \eqref{Dersders} for partial derivatives, namely
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2033
  %
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  2034
  \begin{equation}\label{Derspders}
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2035
  \mbox{\begin{tabular}{lc}
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2036
  @{text "(i)"}  & @{thm Deriv_pderiv}\\
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2037
  @{text "(ii)"} & @{thm Derivs_pderivs}
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2038
  \end{tabular}}
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2039
  \end{equation} 
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2040
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2041
  \begin{proof}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2042
  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
  2043
  modify Antimirov's proof by performing an induction on @{text s} where we
194
5347d7556487 some typos
urbanc
parents: 193
diff changeset
  2044
  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
  2045
  induction hypothesis is
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2046
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2047
  \begin{center}
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2048
  @{text "(IH)"}\hspace{3mm}@{term "\<forall>r. Derivs s (lang r) = \<Union> (lang ` (pderivs s r))"}
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2049
  \end{center}
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2050
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2051
  \noindent
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2052
  With this we can establish
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2053
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2054
  \begin{center}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2055
  \begin{tabular}{r@ {\hspace{1.5mm}}c@ {\hspace{1.5mm}}ll}
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2056
  @{term "Derivs (c # s) (lang r)"} 
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2057
    & @{text "="} & @{term "Derivs s (Deriv c (lang r))"} & by def.\\
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2058
    & @{text "="} & @{term "Derivs s (\<Union> (lang ` (pderiv c r)))"} & by @{text "("}\ref{Derspders}@{text ".i)"}\\
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2059
    & @{text "="} & @{term "\<Union> ((Derivs s) ` (lang ` (pderiv c r)))"} & by def.~of @{text "Ders"}\\
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2060
    & @{text "="} & @{term "\<Union> (lang ` (\<Union> (pderivs s ` (pderiv c r))))"} & by IH\\
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2061
    & @{text "="} & @{term "\<Union> (lang ` (pderivs (c # s) r))"} & by def.\\
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2062
  \end{tabular}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2063
  \end{center}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2064
  
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2065
  \noindent
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  2066
  Note that in order to apply the induction hypothesis in the fourth equation, we
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  2067
  need the generalisation over all regular expressions @{text r}. The case for
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  2068
  the empty string is routine and omitted.\qed
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2069
  \end{proof}
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2070
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  2071
  \noindent
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  2072
  Taking \eqref{Dersders} and \eqref{Derspders} together gives the relationship 
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  2073
  between languages of derivatives and partial derivatives
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2074
  %
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  2075
  \begin{equation}
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  2076
  \mbox{\begin{tabular}{lc}
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2077
  @{text "(i)"}  & @{thm deriv_pderiv[symmetric]}\\
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2078
  @{text "(ii)"} & @{thm derivs_pderivs[symmetric]}
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  2079
  \end{tabular}}
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  2080
  \end{equation} 
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  2081
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  2082
  \noindent
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  2083
  These two properties confirm the observation made earlier 
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2084
  that by using sets, partial derivatives have the @{text "ACI"}-identities
190
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  2085
  of derivatives already built in. 
b73478aaf33e more on paper
urbanc
parents: 187
diff changeset
  2086
245
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  2087
  Antimirov also proved that for every language and every regular expression 
200
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
  2088
  there are only finitely many partial derivatives, whereby the set of partial
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2089
  derivatives of @{text r} w.r.t.~a language @{text A} is defined as
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2090
  %
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2091
  \begin{equation}\label{Pdersdef}
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2092
  @{thm pderivs_lang_def}
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2093
  \end{equation}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2094
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2095
  \begin{theorem}[\cite{Antimirov95}]\label{antimirov}
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2096
  For every language @{text A} and every regular expression @{text r}, 
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2097
  \mbox{@{thm finite_pderivs_lang}}.
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2098
  \end{theorem}
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2099
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2100
  \noindent
200
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
  2101
  Antimirov's proof first establishes this theorem for the language @{term UNIV1}, 
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2102
  which is the set of all non-empty strings. For this he proves:
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2103
  %
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2104
  \begin{equation}\label{pdersunivone}
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2105
  \mbox{\begin{tabular}{l}
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2106
  @{thm pderivs_lang_Zero}\\
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2107
  @{thm pderivs_lang_One}\\
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2108
  @{thm pderivs_lang_Atom}\\
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  2109
  @{thm pderivs_lang_Plus[where ?r1.0="r\<^sub>1" and ?r2.0="r\<^sub>2"]}\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  2110
  @{thm pderivs_lang_Times[where ?r1.0="r\<^sub>1" and ?r2.0="r\<^sub>2"]}\\
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2111
  @{thm pderivs_lang_Star}\\
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2112
  \end{tabular}}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2113
  \end{equation}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2114
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2115
  \noindent
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2116
  from which one can deduce by induction on @{text r} that
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2117
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2118
  \begin{center}
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2119
  @{thm finite_pderivs_lang_UNIV1}
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2120
  \end{center}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2121
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2122
  \noindent
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2123
  holds. Now Antimirov's theorem follows because
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2124
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2125
  \begin{center}
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2126
  @{thm pderivs_lang_UNIV}\\
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2127
  \end{center}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2128
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2129
  \noindent
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2130
  and for all languages @{text "A"}, @{term "pderivs_lang A r"} is a subset of
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2131
  @{term "pderivs_lang UNIV r"}.  Since we follow Antimirov's proof quite
199
11c3c302fa2e a little tuning
urbanc
parents: 198
diff changeset
  2132
  closely in our formalisation (only the last two cases of
200
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
  2133
  \eqref{pdersunivone} involve some non-routine induction arguments), we omit
199
11c3c302fa2e a little tuning
urbanc
parents: 198
diff changeset
  2134
  the details.
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2135
200
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
  2136
  Let us now return to our proof for the second direction in the Myhill-Nerode
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
  2137
  Theorem. The point of the above calculations is to use 
201
9fbf6d9f85ae added comments by Xingyuan
urbanc
parents: 200
diff changeset
  2138
  @{text "\<^raw:$\threesim$>\<^bsub>(\<lambda>x. pders x r)\<^esub>"} as tagging-relation. 
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2139
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2140
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  2141
  \begin{proof}[of Theorem~\ref{myhillnerodetwo} (second version)]
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2142
  Using \eqref{mhders}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2143
  and \eqref{Derspders} we can easily infer that
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2144
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2145
  \begin{center}
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2146
  @{term "x \<approx>(lang r) y"}\hspace{4mm}\mbox{provided}\hspace{4mm}@{term "pderivs x r = pderivs y r"}
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2147
  \end{center}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2148
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2149
  \noindent
201
9fbf6d9f85ae added comments by Xingyuan
urbanc
parents: 200
diff changeset
  2150
  which means the tagging-relation @{text "\<^raw:$\threesim$>\<^bsub>(\<lambda>x. pders x r)\<^esub>"} refines @{term "\<approx>(lang r)"}.
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2151
  So we know by Lemma~\ref{fintwo}, \mbox{@{term "finite (UNIV // (\<approx>(lang r)))"}} 
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2152
  holds if @{term "finite (UNIV // (=(\<lambda>x. pderivs x r)=))"}. In order to establish 
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2153
  the latter, we can use Lemma~\ref{finone} and show that the range of the 
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2154
  tagging-function \mbox{@{term "\<lambda>x. pderivs x r"}} is finite. For this recall Definition
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2155
  \ref{Pdersdef}, which gives us that 
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2156
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2157
  \begin{center}
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2158
  @{thm pderivs_lang_def[where A="UNIV", simplified]}
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2159
  \end{center}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2160
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2161
  \noindent
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2162
  Now the range of @{term "\<lambda>x. pderivs x r"} is a subset of @{term "Pow (pderivs_lang UNIV r)"},
200
204856ef5573 added an example for non-regularity and continuation lemma (the example does not yet work)
urbanc
parents: 199
diff changeset
  2163
  which we know is finite by Theorem~\ref{antimirov}. Consequently there 
245
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  2164
  are only finitely many equivalence classes of @{text "\<^raw:$\threesim$>\<^bsub>(\<lambda>x. pders x r)\<^esub>"}.
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  2165
  This relation refines @{term "\<approx>(lang r)"}, and therefore we can again conclude the 
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  2166
  second part of the Myhill-Nerode Theorem.\qed
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2167
  \end{proof}
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
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  2170
section {* Closure Properties of Regular Languages *}
39
a59473f0229d tuned a little bit the section about finite partitions
urbanc
parents: 37
diff changeset
  2171
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2172
text {*
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  2173
  \label{closure} 
187
9f46a9571e37 more on the derivatives section
urbanc
parents: 186
diff changeset
  2174
  \noindent
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2175
  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
  2176
  operations. Closure under union, concatenation and Kleene-star are trivial
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2177
  to establish given our definition of regularity (recall Definition~\ref{regular}).
240
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
  2178
  More interesting in our setting is the closure under complement, because it seems difficult
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2179
  to construct a regular expression for the complement language by direct
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2180
  means. However the existence of such a regular expression can now be easily
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
  2181
  proved using both parts of the Myhill-Nerode Theorem, since
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2182
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2183
  \begin{center}
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  2184
  @{term "s\<^sub>1 \<approx>A s\<^sub>2"} if and only if @{term "s\<^sub>1 \<approx>(-A) s\<^sub>2"}
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2185
  \end{center}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2186
  
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2187
  \noindent
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  2188
  holds for any strings @{text "s\<^sub>1"} and @{text
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  2189
  "s\<^sub>2"}. Therefore @{text A} and the complement language @{term "-A"}
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2190
  give rise to the same partitions. So if one is finite, the other is too, and
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2191
  vice versa. As noted earlier, our algorithm for solving equational systems
250
b1946e131ce8 small typo
urbanc
parents: 249
diff changeset
  2192
  actually calculates a regular expression for the complement language. 
b1946e131ce8 small typo
urbanc
parents: 249
diff changeset
  2193
  Calculating such a regular expression via
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2194
  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
  2195
  steps: regular expression @{text "\<Rightarrow>"} non-deterministic automaton @{text
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2196
  "\<Rightarrow>"} deterministic automaton @{text "\<Rightarrow>"} complement automaton @{text "\<Rightarrow>"}
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2197
  regular expression. Clearly not something you want to formalise in a theorem
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  2198
  prover if it is cumbersome to reason about automata.
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2199
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2200
  %Once closure under complement is established, closure under intersection
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2201
  %and set difference is also easy, because
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2202
  %
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2203
  %\begin{center}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2204
  %\begin{tabular}{c}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2205
  %@{term "A \<inter> B = - (- A \<union> - B)"}\\
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2206
  %@{term "A - B = - (- A \<union> B)"}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2207
  %\end{tabular}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2208
  %\end{center}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2209
  %
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2210
  %\noindent
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2211
  %Since all finite languages are regular, then by closure under complement also
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2212
  %all co-finite languages. 
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2213
  %
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2214
  %Closure of regular languages under reversal, that is
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2215
  %
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2216
  %\begin{center}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2217
  %@{text "A\<^bsup>-1\<^esup> \<equiv> {s\<^bsup>-1\<^esup> | s \<in> A}"}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2218
  %\end{center}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2219
  %
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2220
  %\noindent 
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2221
  %can be shown with the help of the following operation defined recursively over 
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2222
  %regular expressions
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2223
  %
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2224
  %\begin{center}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2225
  %\begin{tabular}{r@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2226
  %@{thm (lhs) Rev.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) Rev.simps(1)}\\
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2227
  %@{thm (lhs) Rev.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) Rev.simps(2)}\\
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2228
  %@{thm (lhs) Rev.simps(3)} & @{text "\<equiv>"} & @{thm (rhs) Rev.simps(3)}\\
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  2229
  %@{thm (lhs) Rev.simps(4)[where ?r1.0="r\<^sub>1" and ?r2.0="r\<^sub>2"]} & @{text "\<equiv>"} & 
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  2230
  %    @{thm (rhs) Rev.simps(4)[where ?r1.0="r\<^sub>1" and ?r2.0="r\<^sub>2"]}\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  2231
  %@{thm (lhs) Rev.simps(5)[where ?r1.0="r\<^sub>1" and ?r2.0="r\<^sub>2"]} & @{text "\<equiv>"} & 
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  2232
  %    @{thm (rhs) Rev.simps(5)[where ?r1.0="r\<^sub>1" and ?r2.0="r\<^sub>2"]}\\
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2233
  %@{thm (lhs) Rev.simps(6)} & @{text "\<equiv>"} & @{thm (rhs) Rev.simps(6)}\\
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2234
  %\end{tabular}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2235
  %\end{center}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2236
  %
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2237
  %\noindent
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2238
  %For this operation we can show
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2239
  %
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2240
  %\begin{center}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2241
  %@{text "(\<calL>(r))\<^bsup>-1\<^esup>"}~@{text "="}~@{thm (rhs) rev_lang}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2242
  %\end{center}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2243
  %
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2244
  %\noindent
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2245
  %from which closure under reversal of regular languages follows.
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2246
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2247
  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
  2248
  left-quotient. Define
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2249
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2250
  \begin{center}
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2251
  @{abbrev "Deriv_lang B A"}
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2252
  \end{center}
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2253
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2254
  \noindent
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2255
  and assume @{text B} is any language and @{text A} is regular, then @{term
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2256
  "Deriv_lang B A"} is regular. To see this consider the following argument
372
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
  2257
  using partial derivatives (which we used in Section~\ref{derivatives}): From @{text A} 
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
  2258
  being regular we know there exists
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2259
  a regular expression @{text r} such that @{term "A = lang r"}. We also know
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2260
  that @{term "pderivs_lang B r"} is finite for every language @{text B} and 
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2261
  regular expression @{text r} (recall Theorem~\ref{antimirov}). By definition 
245
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  2262
  and \eqref{Derspders} we have
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2263
  %
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2264
  \begin{equation}\label{eqq}
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2265
  @{term "Deriv_lang B (lang r) = (\<Union> (lang ` (pderivs_lang B r)))"}
193
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2266
  \end{equation}
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2267
2a5ac68db24b finished section about derivatives and closure properties
urbanc
parents: 190
diff changeset
  2268
  \noindent
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2269
  Since there are only finitely many regular expressions in @{term "pderivs_lang
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2270
  B r"}, we know by \eqref{uplus} that there exists a regular expression so that
203
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2271
  the right-hand side of \eqref{eqq} is equal to the language \mbox{@{term "lang (\<Uplus>(pderivs_lang B
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2272
  r))"}}. Thus the regular expression @{term "\<Uplus>(pderivs_lang B r)"} verifies that
5d724fe0e096 changes according to afp-submission
urbanc
parents: 201
diff changeset
  2273
  @{term "Deriv_lang B A"} is regular.
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2274
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2275
  Even more surprising is the fact given first by \citeN{Haines69} stating 
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2276
  that for \emph{every} language @{text A}, the language
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2277
  consisting of all (scattered) substrings of @{text A} is regular  (see also 
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  2278
  \cite{Shallit08,Gasarch09}). 
258
1abf8586ee6b added slides for a talk in St Andrews
urbanc
parents: 257
diff changeset
  2279
  A \emph{(scattered) substring} can be obtained
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2280
  by striking out zero or more characters from a string. This can be defined 
237
e02155fe8136 slight polishing to SUBSEQ
urbanc
parents: 233
diff changeset
  2281
  inductively in Isabelle/HOL by the following three rules:
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2282
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2283
  %\begin{center}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2284
  %@ {thm[mode=Axiom] emb0[where bs="x"]}\hspace{10mm} 
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2285
  %@ {thm[mode=Rule] emb1[where as="x" and b="c" and bs="y"]}\hspace{10mm} 
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2286
  %@ {thm[mode=Rule] emb2[where as="x" and a="c" and bs="y"]}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2287
  %\end{center}
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2288
  \begin{center}
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2289
  @{thm[mode=Axiom] emb0}\hspace{10mm} 
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2290
  @{thm[mode=Rule] emb1}\hspace{10mm} 
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2291
  @{thm[mode=Rule] emb2}
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2292
  \end{center}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2293
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2294
  \noindent
245
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  2295
  It is straightforward to prove that @{text "\<preceq>"} is a partial order. Now define the 
237
e02155fe8136 slight polishing to SUBSEQ
urbanc
parents: 233
diff changeset
  2296
  \emph{language of substrings} and \emph{superstrings} of a language @{text A} 
e02155fe8136 slight polishing to SUBSEQ
urbanc
parents: 233
diff changeset
  2297
  respectively as
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2298
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2299
  \begin{center}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2300
  \begin{tabular}{l}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2301
  @{thm SUBSEQ_def}\\
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2302
  @{thm SUPSEQ_def}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2303
  \end{tabular}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2304
  \end{center}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2305
  
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2306
  \noindent
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2307
  We like to establish
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2308
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2309
  \begin{theorem}[\cite{Haines69}]\label{subseqreg}
245
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  2310
  For every language @{text A}, the languages @{text "(i)"} @{term "SUBSEQ A"} and 
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  2311
  @{text "(ii)"} @{term "SUPSEQ A"}
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2312
  are regular.
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2313
  \end{theorem}
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2314
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2315
  \noindent
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2316
  Our proof follows the one given by \citeN[Pages 92--95]{Shallit08}, except that we use
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2317
  Higman's Lemma, which is already proved in the Isabelle/HOL library by
388
0da31edd95b9 new version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 387
diff changeset
  2318
  \citeN{Sternagel13}.
256
acbae3a11fb5 set -> language
urbanc
parents: 254
diff changeset
  2319
  Higman's Lemma allows us to infer that every language @{text A} of antichains, satisfying
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2320
  %
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2321
  \begin{equation}\label{higman}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2322
  @{text "\<forall>x, y \<in> A."}~@{term "x \<noteq> y \<longrightarrow> \<not>(x \<preceq> y) \<and> \<not>(y \<preceq> x)"}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2323
  \end{equation} 
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2324
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2325
  \noindent
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2326
  is finite.
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2327
247
087e6c255e33 some more polishing and a link to Haines
urbanc
parents: 245
diff changeset
  2328
  The first step in our proof of Theorem~\ref{subseqreg} is to establish the 
245
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  2329
  following simple properties for @{term SUPSEQ}
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2330
  %
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2331
  \begin{equation}\label{supseqprops}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2332
  \mbox{\begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  2333
  @{thm (lhs) SUPSEQ_simps(1)} & @{text "="} & @{thm (rhs) SUPSEQ_simps(1)}\\  
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  2334
  @{thm (lhs) SUPSEQ_simps(2)} & @{text "="} & @{thm (rhs) SUPSEQ_simps(2)}\\ 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  2335
  @{thm (lhs) SUPSEQ_atom} & @{text "="} & @{thm (rhs) SUPSEQ_atom}\\ 
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  2336
  @{thm (lhs) SUPSEQ_union} & @{text "="} & @{thm (rhs) SUPSEQ_union}\\
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  2337
  @{thm (lhs) SUPSEQ_conc} & @{text "="} & @{thm (rhs) SUPSEQ_conc}\\
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  2338
  @{thm (lhs) SUPSEQ_star} & @{text "="} & @{thm (rhs) SUPSEQ_star}
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2339
  \end{tabular}}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2340
  \end{equation}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2341
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2342
  \noindent
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2343
  whereby the last equation follows from the fact that @{term "A\<star>"} contains the
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2344
  empty string. With these properties at our disposal we can establish the lemma
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2345
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2346
  \begin{lemma}
237
e02155fe8136 slight polishing to SUBSEQ
urbanc
parents: 233
diff changeset
  2347
  If @{text A} is regular, then also @{term "SUPSEQ A"}.
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2348
  \end{lemma}
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2349
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2350
  \begin{proof}
239
13de6a49294e polished SUBSEQ
urbanc
parents: 238
diff changeset
  2351
  Since our alphabet is finite, we have a regular expression, written @{text ALL}, that
240
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
  2352
  matches every string. Using this regular expression we can inductively define
239
13de6a49294e polished SUBSEQ
urbanc
parents: 238
diff changeset
  2353
  the operation @{text "r\<up>"} 
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2354
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2355
  \begin{center}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2356
  \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2357
  @{thm (lhs) UP.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) UP.simps(1)}\\  
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2358
  @{thm (lhs) UP.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) UP.simps(2)}\\ 
239
13de6a49294e polished SUBSEQ
urbanc
parents: 238
diff changeset
  2359
  @{thm (lhs) UP.simps(3)} & @{text "\<equiv>"} & @{thm (rhs) UP.simps(3)}\\ 
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  2360
  @{thm (lhs) UP.simps(4)[where ?r1.0="r\<^sub>1" and ?r2.0="r\<^sub>2"]} & 
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  2361
     @{text "\<equiv>"} & @{thm (rhs) UP.simps(4)[where ?r1.0="r\<^sub>1" and ?r2.0="r\<^sub>2"]}\\
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  2362
  @{thm (lhs) UP.simps(5)[where ?r1.0="r\<^sub>1" and ?r2.0="r\<^sub>2"]} & 
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  2363
     @{text "\<equiv>"} & @{thm (rhs) UP.simps(5)[where ?r1.0="r\<^sub>1" and ?r2.0="r\<^sub>2"]}\\
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2364
  @{thm (lhs) UP.simps(6)} & @{text "\<equiv>"} & @{thm (rhs) UP.simps(6)}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2365
  \end{tabular}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2366
  \end{center}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2367
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2368
  \noindent
239
13de6a49294e polished SUBSEQ
urbanc
parents: 238
diff changeset
  2369
  and use \eqref{supseqprops} to establish that @{thm lang_UP} holds. This shows
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  2370
  that @{term "SUPSEQ A"} is regular, provided @{text A} is.\qed
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2371
  \end{proof}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2372
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2373
  \noindent
247
087e6c255e33 some more polishing and a link to Haines
urbanc
parents: 245
diff changeset
  2374
  Now we can prove the main lemma w.r.t.~@{const "SUPSEQ"}, namely
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2375
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2376
  \begin{lemma}\label{mset}
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2377
  For every language @{text A}, there exists a finite language @{text M} such that
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2378
  \begin{center}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2379
  \mbox{@{term "SUPSEQ M = SUPSEQ A"}}\;.
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2380
  \end{center}
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2381
  \end{lemma}
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2382
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2383
  \begin{proof}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2384
  For @{text M} we take the set of all minimal elements of @{text A}. An element @{text x} 
237
e02155fe8136 slight polishing to SUBSEQ
urbanc
parents: 233
diff changeset
  2385
  is said to be \emph{minimal} in @{text A} provided
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2386
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2387
  \begin{center}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2388
  @{thm minimal_def}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2389
  \end{center}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2390
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2391
  \noindent
239
13de6a49294e polished SUBSEQ
urbanc
parents: 238
diff changeset
  2392
  By Higman's Lemma \eqref{higman} we know
237
e02155fe8136 slight polishing to SUBSEQ
urbanc
parents: 233
diff changeset
  2393
  that @{term "M \<equiv> {x \<in> A. minimal x A}"} is finite, since every minimal element is incomparable, 
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2394
  except with itself.
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2395
  It is also straightforward to show that @{term "SUPSEQ M \<subseteq> SUPSEQ A"}. For
240
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
  2396
  the other direction we have  @{term "x \<in> SUPSEQ A"}. From this we obtain 
245
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  2397
  a @{text y} such that @{term "y \<in> A"} and @{term "y \<preceq> x"}. Since we have that
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2398
  the relation \mbox{@{term "{(y, x). y \<preceq> x \<and> x \<noteq> y}"}} is well-founded, there must
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2399
  be a minimal element @{text "z"} such that @{term "z \<in> A"} and @{term "z \<preceq> y"}, 
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2400
  and hence by transitivity also \mbox{@{term "z \<preceq> x"}} (here we deviate from the argument 
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2401
  given by \citeN{Shallit08}, because Isabelle/HOL provides already an extensive infrastructure
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2402
  for reasoning about well-foundedness). Since @{term "z"} is
237
e02155fe8136 slight polishing to SUBSEQ
urbanc
parents: 233
diff changeset
  2403
  minimal and an element in @{term A}, we also know that @{term z} is in @{term M}.
240
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
  2404
  From this together with \mbox{@{term "z \<preceq> x"}}, we can infer that @{term x} is in 
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  2405
  @{term "SUPSEQ M"}, as required.\qed
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2406
  \end{proof}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2407
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2408
  \noindent
247
087e6c255e33 some more polishing and a link to Haines
urbanc
parents: 245
diff changeset
  2409
  This lemma allows us to establish the second part of Theorem~\ref{subseqreg}.
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2410
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  2411
  \begin{proof}[of the Second Part of Theorem~\ref{subseqreg}]
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2412
  Given any language @{text A}, by Lemma~\ref{mset} we know there exists
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2413
  a finite, and thus regular, language @{text M}. We further have @{term "SUPSEQ M = SUPSEQ A"},
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  2414
  which establishes the second part.\qed    
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2415
  \end{proof}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2416
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2417
  \noindent
247
087e6c255e33 some more polishing and a link to Haines
urbanc
parents: 245
diff changeset
  2418
  In order to establish the first part of this theorem, we use the
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2419
  property proved by \citeN{Shallit08}, namely that
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2420
  %
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2421
  \begin{equation}\label{compl}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2422
  @{thm SUBSEQ_complement}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2423
  \end{equation}
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2424
237
e02155fe8136 slight polishing to SUBSEQ
urbanc
parents: 233
diff changeset
  2425
  \noindent
247
087e6c255e33 some more polishing and a link to Haines
urbanc
parents: 245
diff changeset
  2426
  holds. Now the first part of Theorem~\ref{subseqreg} is a simple consequence of the second part.
237
e02155fe8136 slight polishing to SUBSEQ
urbanc
parents: 233
diff changeset
  2427
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  2428
  \begin{proof}[of the First Part of Theorem~\ref{subseqreg}]
247
087e6c255e33 some more polishing and a link to Haines
urbanc
parents: 245
diff changeset
  2429
  By the second part, we know the right-hand side of \eqref{compl}
237
e02155fe8136 slight polishing to SUBSEQ
urbanc
parents: 233
diff changeset
  2430
  is regular, which means @{term "- SUBSEQ A"} is regular. But since
372
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
  2431
  we established already that regularity is preserved under complement (using Myhill-Nerode), 
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  2432
  also @{term "SUBSEQ A"} must be regular.\qed 
233
e2dc11e12e0b added section about SUBSEQ and SUPSEQ
urbanc
parents: 218
diff changeset
  2433
  \end{proof}
240
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
  2434
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
  2435
  Finally we like to show that the Myhill-Nerode Theorem is also convenient for establishing 
245
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  2436
  the non-regularity of languages. For this we use the following version of the Continuation
240
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
  2437
  Lemma (see for example~\cite{Rosenberg06}).
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
  2438
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2439
  \begin{lemma}[Continuation Lemma]
257
f512026d5d6e small change
urbanc
parents: 256
diff changeset
  2440
  If a language @{text A} is regular and a set of strings @{text B} is infinite,
240
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
  2441
  then there exist two distinct strings @{text x} and @{text y} in @{text B} 
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
  2442
  such that @{term "x \<approx>A y"}.
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2443
  \end{lemma}
240
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
  2444
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
  2445
  \noindent
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
  2446
  This lemma can be easily deduced from the Myhill-Nerode Theorem and the Pigeonhole
240
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
  2447
  Principle: Since @{text A} is regular, there can be only finitely many 
245
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  2448
  equivalence classes. Hence an infinite set must contain 
240
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
  2449
  at least two strings that are in the same equivalence class, that is
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
  2450
  they need to be related by the Myhill-Nerode Relation.
240
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
  2451
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
  2452
  Using this lemma, it is straightforward to establish that the language 
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  2453
  \mbox{@{text "A \<equiv> \<Union>\<^sub>n a\<^sup>n @ b\<^sup>n"}} is not regular (@{text "a\<^sup>n"} stands
245
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  2454
  for the strings consisting of @{text n} times the character a; similarly for
385
e5e32faa2446 updated to new Isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 384
diff changeset
  2455
  @{text "b\<^sup>n"}). For this consider the infinite set @{text "B \<equiv> \<Union>\<^sub>n a\<^sup>n"}.
240
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
  2456
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2457
  \begin{lemma}
247
087e6c255e33 some more polishing and a link to Haines
urbanc
parents: 245
diff changeset
  2458
  No two distinct strings in set @{text "B"} are Myhill-Nerode related by language @{text A}.
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2459
  \end{lemma} 
240
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
  2460
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
  2461
  \begin{proof}
252
8e2c497d699e clarified proof about non-regularity
urbanc
parents: 251
diff changeset
  2462
  After unfolding the definition of @{text "B"}, we need to establish that given @{term "i \<noteq> j"},
8e2c497d699e clarified proof about non-regularity
urbanc
parents: 251
diff changeset
  2463
  the strings @{text "a\<^sup>i"} and @{text "a\<^sup>j"} are not Myhill-Nerode related by @{text "A"}.
375
44c4450152e3 adapted to JAR
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 374
diff changeset
  2464
  That means we have to show that @{text "\<forall>z. a\<^sup>i @ z \<in> A = a\<^sup>j @ z \<in> A"} leads to 
252
8e2c497d699e clarified proof about non-regularity
urbanc
parents: 251
diff changeset
  2465
  a contradiction. Let us take @{text "b\<^sup>i"} for @{text "z"}. Then we know @{text "a\<^sup>i @ b\<^sup>i \<in> A"}.
8e2c497d699e clarified proof about non-regularity
urbanc
parents: 251
diff changeset
  2466
  But since @{term "i \<noteq> j"}, @{text "a\<^sup>j @ b\<^sup>i \<notin> A"}. Therefore  @{text "a\<^sup>i"} and @{text "a\<^sup>j"}
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  2467
  cannot be Myhill-Nerode related by @{text "A"}, and we are done.\qed
240
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
  2468
  \end{proof}
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
  2469
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
  2470
  \noindent
252
8e2c497d699e clarified proof about non-regularity
urbanc
parents: 251
diff changeset
  2471
  To conclude the proof of non-regularity for the language @{text A}, the
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
  2472
  Continuation Lemma and the lemma above lead to a contradiction assuming
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
  2473
  @{text A} is regular. Therefore the language @{text A} is not regular, as we
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
  2474
  wanted to show.
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2475
*}
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2476
117
22ba25b808c8 updated second direction
urbanc
parents: 116
diff changeset
  2477
240
17aa8c8fbe7d added section about non-regularity
urbanc
parents: 239
diff changeset
  2478
54
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
  2479
section {* Conclusion and Related Work *}
c19d2fc2cc69 a bit more on the paper
urbanc
parents: 53
diff changeset
  2480
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
  2481
text {*
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2482
  \noindent
112
62fdb4bf7239 more on the conclusion
urbanc
parents: 111
diff changeset
  2483
  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
  2484
  exists a regular expression that matches all of its strings. Regular
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  2485
  expressions can be conveniently defined as a datatype in theorem
145
099e20f25b25 corrected small typo
urbanc
parents: 143
diff changeset
  2486
  provers. For us it was therefore interesting to find out how far we can push
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  2487
  this point of view. But this question whether regular language theory can
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  2488
  be done without automata crops up also in non-theorem prover contexts. Recall 
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  2489
  Gasarch's comment cited in the Introduction. 
375
44c4450152e3 adapted to JAR
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 374
diff changeset
  2490
  We have established in Isabelle/HOL both directions 
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
  2491
  of the Myhill-Nerode Theorem.
132
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
  2492
  %
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2493
  \begin{theorem}[Myhill-Nerode Theorem]\mbox{}\\
132
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
  2494
  A language @{text A} is regular if and only if @{thm (rhs) Myhill_Nerode}.
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2495
  \end{theorem}
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2496
  
132
f77a7138f791 comments by Xingyuan
urbanc
parents: 131
diff changeset
  2497
  \noindent
186
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2498
  Having formalised this theorem means we pushed our point of view quite
07a269d9642b added more to the derivatives section
urbanc
parents: 185
diff changeset
  2499
  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
  2500
  regular---by establishing that it has infinitely many equivalence classes
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
  2501
  generated by the Myhill-Nerode Relation (this is usually the purpose of the
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2502
  Pumping Lemma).  We can also use it to establish the standard
375
44c4450152e3 adapted to JAR
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 374
diff changeset
  2503
  textbook results about closure properties of regular languages. The case of 
44c4450152e3 adapted to JAR
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 374
diff changeset
  2504
  closure under complement follows easily from the Myhill-Nerode Theorem.
383
c8cb6914f4c8 some more polishing
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 381
diff changeset
  2505
  So our answer to Gasarch is ``{\it yes we can''}!  
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2506
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2507
  %Our insistence on regular expressions for proving the Myhill-Nerode Theorem
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2508
  %arose from the problem of defining formally the regularity of a language.
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2509
  %In order to guarantee consistency,
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2510
  %formalisations in HOL can only extend the logic with definitions that introduce a new concept in
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2511
  %terms of already existing notions. A convenient definition for automata
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2512
  %(based on graphs) uses a polymorphic type for the state nodes. This allows
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2513
  %us to use the standard operation for disjoint union whenever we need to compose two
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2514
  %automata. Unfortunately, we cannot use such a polymorphic definition
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2515
  %in HOL as part of the definition for regularity of a language (a predicate
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2516
  %over sets of strings).  Consider for example the following attempt:
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2517
  %
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2518
  %\begin{center}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2519
  %@{text "is_regular A \<equiv> \<exists>M(\<alpha>). is_dfa (M) \<and> \<calL>(M) = A"}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2520
  %\end{center}
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2521
  %
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2522
  %\noindent
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2523
  %In this definifion, the definiens is polymorphic in the type of the automata
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2524
  %@{text "M"} (indicated by dependency on the type-variable @{text "\<alpha>"}), but the definiendum
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2525
  %@{text "is_regular"} is not. Such definitions are excluded from HOL, because
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2526
  %they can lead easily to inconsistencies (see \cite{PittsHOL4} for a simple
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2527
  %example). Also HOL does not contain type-quantifiers which would allow us to
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2528
  %get rid of the polymorphism by quantifying over the type-variable 
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2529
  %@{text "\<alpha>"}. Therefore when defining regularity in terms of automata, the
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2530
  %natural way out in HOL is to resort to state nodes with an identity, for
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2531
  %example a natural number. Unfortunatly, the consequence is that we have to
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2532
  %be careful when combining two automata so that there is no clash between two
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2533
  %such states. This makes formalisations quite fiddly and rather
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2534
  %unpleasant. Regular expressions proved much more convenient for reasoning in
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2535
  %HOL since they can be defined as inductive datatype and a reasoning
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2536
  %infrastructure comes for free. The definition of regularity in terms of
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2537
  %regular expressions poses no problem at all for HOL.  We showed in this
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2538
  %paper that they can be used for establishing the central result in regular
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2539
  %language theory---the Myhill-Nerode Theorem.
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2540
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2541
  While regular expressions are convenient, they have some limitations. One is
375
44c4450152e3 adapted to JAR
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 374
diff changeset
  2542
  that there are some regular expressions for which the smallest regular 
44c4450152e3 adapted to JAR
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 374
diff changeset
  2543
  expression for the complement language has a doubly-exponential blowup in size 
44c4450152e3 adapted to JAR
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 374
diff changeset
  2544
  as shown by \citeN{Gelade12}. 
44c4450152e3 adapted to JAR
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 374
diff changeset
  2545
  Another is
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2546
  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
  2547
  (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
  2548
  automata. On the other hand, efficient regular expression matching, without
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2549
  using automata, poses no problem as shown by \citeN{OwensReppyTuron09}.  For an
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2550
  implementation of a simple regular expression matcher, whose correctness has
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2551
  been formally established, we refer the reader to 
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2552
  \citeN{OwensSlind08}. In our opinion, their formalisation is considerably
245
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  2553
  slicker than for example the approach to regular expression matching taken
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2554
  by \citeN{Harper99} and by \citeN{Yi06}.
116
342983676c8f included comments by Chunhan
urbanc
parents: 115
diff changeset
  2555
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2556
  Our proof of the first direction is very much inspired by \emph{Brzozowski's
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2557
  algebraic method} \citeyear{Brzozowski64} used to convert a finite automaton to a regular expression. 
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2558
  The close connection can be seen by considering the
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2559
  equivalence classes as the states of the minimal automaton for the regular
201
9fbf6d9f85ae added comments by Xingyuan
urbanc
parents: 200
diff changeset
  2560
  language.  However there are some subtle differences. Because our equivalence 
248
47446f111550 one more itteration on the paper
urbanc
parents: 247
diff changeset
  2561
  classes (or correspondingly states) arise from the Myhill-Nerode Relation, the most natural
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2562
  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
  2563
  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
  2564
  characterised as the strings starting from that state leading to the
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2565
  terminal states.  The first choice has consequences about how the initial
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2566
  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
  2567
  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
  2568
  need to reverse the direction of Arden's Lemma. We have not found anything
249
061b32d78471 a final polishing before submitting later this week
urbanc
parents: 248
diff changeset
  2569
  in the `pencil-and-paper-reasoning' literature about our way of proving the 
061b32d78471 a final polishing before submitting later this week
urbanc
parents: 248
diff changeset
  2570
  first direction of the Myhill-Nerode Theorem, but it appears to be folklore.
112
62fdb4bf7239 more on the conclusion
urbanc
parents: 111
diff changeset
  2571
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  2572
  We presented two proofs for the second direction of the
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  2573
  Myhill-Nerode Theorem. One direct proof using tagging-functions and
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  2574
  another using partial derivatives. This part of our work is where
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  2575
  our method using regular expressions shines, because we can perform
386
92ca56c1a199 soem small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 385
diff changeset
  2576
  an induction on the structure of regular expressions. However, it is
381
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  2577
  also the direction where we had to spend most of the `conceptual'
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  2578
  time, as our first proof based on tagging-functions is new for
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  2579
  establishing the Myhill-Nerode Theorem. All standard proofs of this
99161cd17c0f added modified version adn answer to the reviewers
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 379
diff changeset
  2580
  direction proceed by arguments over automata.
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2581
  
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2582
  The indirect proof for the second direction arose from our interest in
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2583
  Brzozowski's derivatives for regular expression matching.  While \citeN{Brzozowski64} 
245
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  2584
  already established that there are only
196
fa8d33d13cb6 polishing of the closure section and conclusion
urbanc
parents: 194
diff changeset
  2585
  finitely many dissimilar derivatives for every regular expression, this
199
11c3c302fa2e a little tuning
urbanc
parents: 198
diff changeset
  2586
  result is not as straightforward to formalise in a theorem prover as one
11c3c302fa2e a little tuning
urbanc
parents: 198
diff changeset
  2587
  might wish. The reason is that the set of dissimilar derivatives is not
11c3c302fa2e a little tuning
urbanc
parents: 198
diff changeset
  2588
  defined inductively, but in terms of an ACI-equivalence relation. This
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2589
  difficulty prevented for example \citeN{KraussNipkow11} to prove termination of
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2590
  their equivalence checker for regular expressions. Their checker is based on Brzozowski's derivatives
199
11c3c302fa2e a little tuning
urbanc
parents: 198
diff changeset
  2591
  and for their argument the lack of a formal proof of termination is not
11c3c302fa2e a little tuning
urbanc
parents: 198
diff changeset
  2592
  crucial (it merely lets them ``sleep better'' \cite{KraussNipkow11}).  We
11c3c302fa2e a little tuning
urbanc
parents: 198
diff changeset
  2593
  expect that their development simplifies by using partial derivatives,
245
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  2594
  instead of derivatives, and that the termination of the algorithm can be
40b8d485ce8d polished a bit the journal paper
urbanc
parents: 242
diff changeset
  2595
  formally established (the main ingredient is
199
11c3c302fa2e a little tuning
urbanc
parents: 198
diff changeset
  2596
  Theorem~\ref{antimirov}). However, since partial derivatives use sets of
11c3c302fa2e a little tuning
urbanc
parents: 198
diff changeset
  2597
  regular expressions, one needs to carefully analyse whether the resulting
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2598
  algorithm is still executable. Given the infrastructure for
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2599
  executable sets  introduced by \citeN{Haftmann09} in Isabelle/HOL, it should.
199
11c3c302fa2e a little tuning
urbanc
parents: 198
diff changeset
  2600
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2601
  We started out by claiming that in a theorem prover it is easier to
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2602
  reason about regular expressions than about automata. Here are some
378
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2603
  numbers: Our formalisation of the Myhill-Nerode Theorem consists of
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2604
  780 lines of Isabelle/Isar code for the first direction and 460 for
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2605
  the second (the one based on tagging-functions), plus around 300
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2606
  lines of standard material about regular languages. The
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2607
  formalisation of derivatives and partial derivatives shown in
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2608
  Section~\ref{derivatives} consists of 390 lines of code.  The
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2609
  closure properties in Section~\ref{closure} (except
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2610
  Theorem~\ref{subseqreg}) can be established in 100 lines of
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2611
  code. The Continuation Lemma and the non-regularity of @{text "a\<^sup>n
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2612
  b\<^sup>n"} require 70 lines of code.  The algorithm for solving equational
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2613
  systems, which we used in the first direction, is conceptually
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2614
  relatively simple. Still the use of sets over which the algorithm
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2615
  operates means it is not as easy to formalise as one might
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2616
  wish. However, it seems sets cannot be avoided since the `input' of
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2617
  the algorithm consists of equivalence classes and we cannot see how
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2618
  to reformulate the theory so that we can use lists or
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2619
  matrices. Lists would be much easier to reason about, since we can
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2620
  define functions over them by recursion. For sets we have to use
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2621
  set-comprehensions, which is slightly unwieldy. Matrices would allow
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2622
  us to use the slick formalisation by \citeN{Nipkow11} of the
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2623
  Gauss-Jordan algorithm.
172
21ee3a852a02 more on the journal paper
urbanc
parents: 170
diff changeset
  2624
378
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2625
  % OLD
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2626
  %While our formalisation might appear large, it should be seen in the
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2627
  %context of the work done by \citeN{Constable00} who formalised the
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2628
  %Myhill-Nerode Theorem in Nuprl using automata. They write that their
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2629
  %four-member team would need something on the magnitude of 18 months
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2630
  %for their formalisation of the first eleven chapters of the textbook
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2631
  %by \citeN{HopcroftUllman69}, which includes the Myhill-Nerode
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2632
  %theorem. It is hard to gauge the size of a formalisation in Nurpl,
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2633
  %but from what is shown in the Nuprl Math Library about their
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2634
  %development it seems \emph{substantially} larger than ours. We
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2635
  %attribute this to our use of regular expressions, which meant we did
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2636
  %not need to `fight' the theorem prover. 
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2637
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2638
  %%% NEW
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2639
  While our formalisation might appear large, it should be seen in the
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2640
  context of the work done by \citeN{Constable00} who formalised the
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2641
  Myhill-Nerode Theorem in Nuprl using automata.  They choose to formalise the 
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2642
  this theorem, because it gives them state minimization of automata
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2643
  as a corollary. It is hard to gauge the size of a formalisation in Nurpl,
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2644
  but from what is shown in the Nuprl Math Library about this
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2645
  development it seems \emph{substantially} larger than ours. We
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2646
  attribute this to our use of regular expressions, which meant we did
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2647
  not need to `fight' the theorem prover. 
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2648
  %
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2649
  Recently,
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2650
  \citeN{LammichTuerk12} formalised Hopcroft's algorithm in
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2651
  Isabelle/HOL (in 7000 lines of code) using an automata library of
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2652
  27000 lines of code.  Also, \citeN{Filliatre97} reports that his
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2653
  formalisation in Coq of automata theory and Kleene's theorem is
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2654
  ``rather big''.  \citeN{Almeidaetal10} reported about another
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2655
  formalisation of regular languages in Coq. Their main result is the
218
28e98ede8599 added a few points
urbanc
parents: 217
diff changeset
  2656
  correctness of Mirkin's construction of an automaton from a regular
378
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2657
  expression using partial derivatives. This took approximately 10600
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2658
  lines of code.  \citeN{Braibant12} formalised a large part of
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2659
  regular language theory and Kleene algebras in Coq. While he is
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2660
  mainly interested in implementing decision procedures for Kleene
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2661
  algebras, his library includes a proof of the Myhill-Nerode
384
60bcf13adb77 comment by Chunhan
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 383
diff changeset
  2662
  theorem. He reckons that our Myhill-Nerode ``development is more concise'' than
378
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2663
  his one based on matrices \cite[Page 67]{Braibant12}.  He writes
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2664
  that there is no conceptual problems with formally reasoning about
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2665
  matrices for automata, but notes ``intrinsic difficult[ies]'' when
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2666
  working with matrices in Coq, which is the sort of `fighting' one
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2667
  would encounter also in other theorem provers.
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  2668
  In terms of time, the estimate for our formalisation is that we
198
b300f2c5d51d final(?) version of the paper
urbanc
parents: 197
diff changeset
  2669
  needed approximately 3 months and this included the time to find our proof
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2670
  arguments. Unlike \citeN{Constable00}, who were able to follow the Myhill-Nerode 
378
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2671
  proof by \citeN{HopcroftUllman69}, we had to find our own arguments.  
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2672
  So for us the formalisation was not the bottleneck. The conclusion we draw 
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2673
  from all these comparisons is that if one is interested in formalising
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2674
  results from regular language theory, not results from automata theory,
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2675
  then regular expressions are easier to work with formally. 
a0bcf886b8ef deleted utm-work from the repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 377
diff changeset
  2676
  The code of
372
2c56b20032a7 made changes and updates to the journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 350
diff changeset
  2677
  our formalisation \cite{myhillnerodeafp11} can be found in the Archive of Formal Proofs at 
386
92ca56c1a199 soem small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 385
diff changeset
  2678
  \mbox{\url{http://afp.sourceforge.net/entries/Myhill-Nerode.shtml}}.
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 160
diff changeset
  2679
  
386
92ca56c1a199 soem small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 385
diff changeset
  2680
  Our future work will focus on formalising the regular expression matchers
92ca56c1a199 soem small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 385
diff changeset
  2681
  developed by \citeN{Sulzmann12} which generate variable assignments for
387
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
  2682
  regular expression sub-matching. For this they use both, derivatives and
288637d9dcde more changes for final submission
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 386
diff changeset
  2683
  partial derivatives of regular expressions.\smallskip
386
92ca56c1a199 soem small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 385
diff changeset
  2684
162
e93760534354 added directory for journal version; took uptodate version of the theory files
urbanc
parents: 160
diff changeset
  2685
  \noindent
173
d371536861bc more on the introduction of the journal paper
urbanc
parents: 172
diff changeset
  2686
  {\bf Acknowledgements:}
242
093e45c44d91 polished proposal
urbanc
parents: 240
diff changeset
  2687
  We are grateful for the comments we received from Larry Paulson.  Tobias
247
087e6c255e33 some more polishing and a link to Haines
urbanc
parents: 245
diff changeset
  2688
  Nipkow made us aware of the properties in Theorem~\ref{subseqreg} and Tjark
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2689
  Weber helped us with proving them. Christian Sternagel provided us with a
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2690
  version of Higman's Lemma that applies to arbitrary, but finite alphabets. 
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  2691
350
8ce9a432680b made changes for another journal submission of the MN-paper
urbanc
parents: 348
diff changeset
  2692
  \bibliographystyle{acmtrans}
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  2693
  \bibliography{root}
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  2694
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2695
  %\newpage
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2696
  %\begin{appendix}
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2697
  %\section{Appendix$^\star$}
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  2698
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2699
  %\renewcommand{\thefootnote}{\mbox{$\star$}}
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2700
  %\footnotetext{If the reviewers deem more suitable, the authors are
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2701
  %prepared to drop material or move it to an electronic appendix.}
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  2702
  
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2703
  %\begin{proof}[of Lemma~\ref{arden}]
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2704
  %For the right-to-left direction we assume @{thm (rhs) reversed_Arden} and show
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2705
  %that @{thm (lhs) reversed_Arden} holds. From Property~\ref{langprops}@{text "(i)"} 
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2706
  %we have @{term "A\<star> = A \<cdot> A\<star> \<union> {[]}"},
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2707
  %which is equal to @{term "A\<star> = A\<star> \<cdot> A \<union> {[]}"}. Adding @{text B} to both 
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2708
  %sides gives @{term "B \<cdot> A\<star> = B \<cdot> (A\<star> \<cdot> A \<union> {[]})"}, whose right-hand side
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2709
  %is equal to @{term "(B \<cdot> A\<star>) \<cdot> A \<union> B"}. Applying the assumed equation 
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2710
  %completes this direction. 
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  2711
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2712
  %For the other direction we assume @{thm (lhs) reversed_Arden}. By a simple induction
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2713
  %on @{text n}, we can establish the property
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  2714
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2715
  %\begin{center}
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2716
  %@{text "("}@{text "*"}@{text ")"}\hspace{5mm} @{thm (concl) reversed_arden_helper}
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2717
  %\end{center}
348
bea94f1e6771 made changes for another journal submission of the MH-paper
urbanc
parents: 338
diff changeset
  2718
  
379
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2719
  %\noindent
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2720
  %Using this property we can show that @{term "B \<cdot> (A \<up> n) \<subseteq> X"} holds for
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2721
  %all @{text n}. From this we can infer @{term "B \<cdot> A\<star> \<subseteq> X"} using the definition
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2722
  %of @{text "\<star>"}.
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2723
  %For the inclusion in the other direction we assume a string @{text s}
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2724
  %with length @{text k} is an element in @{text X}. Since @{thm (prem 1) reversed_Arden}
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2725
  %we know by Property~\ref{langprops}@{text "(ii)"} that 
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2726
  %@{term "s \<notin> X \<cdot> (A \<up> Suc k)"} since its length is only @{text k}
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2727
  %(the strings in @{term "X \<cdot> (A \<up> Suc k)"} are all longer). 
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2728
  %From @{text "("}@{text "*"}@{text ")"} it follows then that
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2729
  %@{term s} must be an element in @{term "(\<Union>m\<le>k. B \<cdot> (A \<up> m))"}. This in turn
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2730
  %implies that @{term s} is in @{term "(\<Union>n. B \<cdot> (A \<up> n))"}. Using Property~\ref{langprops}@{text "(iii)"} 
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2731
  %this is equal to @{term "B \<cdot> A\<star>"}, as we needed to show.
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2732
  %\end{proof}
8c4b6fb43ebe polished more and updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 378
diff changeset
  2733
  % \end{appendix}
92
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
  2734
*}
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
  2735
a9ebc410a5c8 more on paper
urbanc
parents: 90
diff changeset
  2736
24
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
  2737
(*<*)
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
  2738
end
f72c82bf59e5 added paper
urbanc
parents:
diff changeset
  2739
(*>*)