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