218
+ − 1
(*<*)
+ − 2
theory Paper
+ − 3
imports
+ − 4
"../Lexer"
+ − 5
"../Simplifying"
280
+ − 6
"../Positions"
218
+ − 7
"~~/src/HOL/Library/LaTeXsugar"
+ − 8
begin
+ − 9
265
+ − 10
lemma Suc_0_fold:
+ − 11
"Suc 0 = 1"
+ − 12
by simp
+ − 13
+ − 14
+ − 15
218
+ − 16
declare [[show_question_marks = false]]
+ − 17
267
+ − 18
syntax (latex output)
274
+ − 19
"_Collect" :: "pttrn => bool => 'a set" ("(1{_ \<^latex>\<open>\\mbox{\\boldmath$\\mid$}\<close> _})")
267
+ − 20
"_CollectIn" :: "pttrn => 'a set => bool => 'a set" ("(1{_ \<in> _ |e _})")
+ − 21
273
+ − 22
syntax
+ − 23
"_Not_Ex" :: "idts \<Rightarrow> bool \<Rightarrow> bool" ("(3\<nexists>_.a./ _)" [0, 10] 10)
+ − 24
"_Not_Ex1" :: "pttrn \<Rightarrow> bool \<Rightarrow> bool" ("(3\<nexists>!_.a./ _)" [0, 10] 10)
+ − 25
267
+ − 26
218
+ − 27
abbreviation
265
+ − 28
"der_syn r c \<equiv> der c r"
218
+ − 29
+ − 30
abbreviation
265
+ − 31
"ders_syn r s \<equiv> ders s r"
+ − 32
+ − 33
+ − 34
abbreviation
+ − 35
"nprec v1 v2 \<equiv> \<not>(v1 :\<sqsubset>val v2)"
+ − 36
218
+ − 37
267
+ − 38
+ − 39
218
+ − 40
notation (latex output)
274
+ − 41
If ("(\<^latex>\<open>\\textrm{\<close>if\<^latex>\<open>}\<close> (_)/ \<^latex>\<open>\\textrm{\<close>then\<^latex>\<open>}\<close> (_)/ \<^latex>\<open>\\textrm{\<close>else\<^latex>\<open>}\<close> (_))" 10) and
+ − 42
Cons ("_\<^latex>\<open>\\mbox{$\\,$}\<close>::\<^latex>\<open>\\mbox{$\\,$}\<close>_" [75,73] 73) and
218
+ − 43
273
+ − 44
ZERO ("\<^bold>0" 81) and
+ − 45
ONE ("\<^bold>1" 81) and
218
+ − 46
CHAR ("_" [1000] 80) and
+ − 47
ALT ("_ + _" [77,77] 78) and
+ − 48
SEQ ("_ \<cdot> _" [77,77] 78) and
273
+ − 49
STAR ("_\<^sup>\<star>" [78] 78) and
218
+ − 50
265
+ − 51
val.Void ("Empty" 78) and
218
+ − 52
val.Char ("Char _" [1000] 78) and
+ − 53
val.Left ("Left _" [79] 78) and
+ − 54
val.Right ("Right _" [1000] 78) and
+ − 55
val.Seq ("Seq _ _" [79,79] 78) and
+ − 56
val.Stars ("Stars _" [79] 78) and
+ − 57
+ − 58
L ("L'(_')" [10] 78) and
272
+ − 59
LV ("LV _ _" [80,73] 78) and
218
+ − 60
der_syn ("_\\_" [79, 1000] 76) and
+ − 61
ders_syn ("_\\_" [79, 1000] 76) and
+ − 62
flat ("|_|" [75] 74) and
273
+ − 63
flats ("|_|" [72] 74) and
218
+ − 64
Sequ ("_ @ _" [78,77] 63) and
+ − 65
injval ("inj _ _ _" [79,77,79] 76) and
+ − 66
mkeps ("mkeps _" [79] 76) and
+ − 67
length ("len _" [73] 73) and
266
+ − 68
intlen ("len _" [73] 73) and
267
+ − 69
set ("_" [73] 73) and
218
+ − 70
267
+ − 71
Prf ("_ : _" [75,75] 75) and
218
+ − 72
Posix ("'(_, _') \<rightarrow> _" [63,75,75] 75) and
+ − 73
+ − 74
lexer ("lexer _ _" [78,78] 77) and
+ − 75
F_RIGHT ("F\<^bsub>Right\<^esub> _") and
+ − 76
F_LEFT ("F\<^bsub>Left\<^esub> _") and
+ − 77
F_ALT ("F\<^bsub>Alt\<^esub> _ _") and
+ − 78
F_SEQ1 ("F\<^bsub>Seq1\<^esub> _ _") and
+ − 79
F_SEQ2 ("F\<^bsub>Seq2\<^esub> _ _") and
+ − 80
F_SEQ ("F\<^bsub>Seq\<^esub> _ _") and
+ − 81
simp_SEQ ("simp\<^bsub>Seq\<^esub> _ _" [1000, 1000] 1) and
+ − 82
simp_ALT ("simp\<^bsub>Alt\<^esub> _ _" [1000, 1000] 1) and
+ − 83
slexer ("lexer\<^sup>+" 1000) and
+ − 84
274
+ − 85
at ("_\<^latex>\<open>\\mbox{$\\downharpoonleft$}\<close>\<^bsub>_\<^esub>") and
265
+ − 86
lex_list ("_ \<prec>\<^bsub>lex\<^esub> _") and
+ − 87
PosOrd ("_ \<prec>\<^bsub>_\<^esub> _" [77,77,77] 77) and
+ − 88
PosOrd_ex ("_ \<prec> _" [77,77] 77) and
274
+ − 89
PosOrd_ex_eq ("_ \<^latex>\<open>\\mbox{$\\preccurlyeq$}\<close> _" [77,77] 77) and
265
+ − 90
pflat_len ("\<parallel>_\<parallel>\<^bsub>_\<^esub>") and
274
+ − 91
nprec ("_ \<^latex>\<open>\\mbox{$\\not\\prec$}\<close> _" [77,77] 77) and
265
+ − 92
274
+ − 93
DUMMY ("\<^latex>\<open>\\underline{\\hspace{2mm}}\<close>")
268
+ − 94
218
+ − 95
+ − 96
definition
+ − 97
"match r s \<equiv> nullable (ders s r)"
+ − 98
267
+ − 99
268
+ − 100
lemma LV_STAR_ONE_empty:
+ − 101
shows "LV (STAR ONE) [] = {Stars []}"
+ − 102
by(auto simp add: LV_def elim: Prf.cases intro: Prf.intros)
267
+ − 103
+ − 104
+ − 105
218
+ − 106
(*
+ − 107
comments not implemented
+ − 108
272
+ − 109
p9. The condition "not exists s3 s4..." appears often enough (in particular in
218
+ − 110
the proof of Lemma 3) to warrant a definition.
+ − 111
+ − 112
*)
+ − 113
273
+ − 114
218
+ − 115
(*>*)
+ − 116
267
+ − 117
+ − 118
218
+ − 119
section {* Introduction *}
+ − 120
+ − 121
+ − 122
text {*
+ − 123
+ − 124
Brzozowski \cite{Brzozowski1964} introduced the notion of the {\em
+ − 125
derivative} @{term "der c r"} of a regular expression @{text r} w.r.t.\ a
+ − 126
character~@{text c}, and showed that it gave a simple solution to the
+ − 127
problem of matching a string @{term s} with a regular expression @{term r}:
+ − 128
if the derivative of @{term r} w.r.t.\ (in succession) all the characters of
+ − 129
the string matches the empty string, then @{term r} matches @{term s} (and
+ − 130
{\em vice versa}). The derivative has the property (which may almost be
+ − 131
regarded as its specification) that, for every string @{term s} and regular
+ − 132
expression @{term r} and character @{term c}, one has @{term "cs \<in> L(r)"} if
+ − 133
and only if \mbox{@{term "s \<in> L(der c r)"}}. The beauty of Brzozowski's
+ − 134
derivatives is that they are neatly expressible in any functional language,
+ − 135
and easily definable and reasoned about in theorem provers---the definitions
+ − 136
just consist of inductive datatypes and simple recursive functions. A
+ − 137
mechanised correctness proof of Brzozowski's matcher in for example HOL4
+ − 138
has been mentioned by Owens and Slind~\cite{Owens2008}. Another one in Isabelle/HOL is part
+ − 139
of the work by Krauss and Nipkow \cite{Krauss2011}. And another one in Coq is given
+ − 140
by Coquand and Siles \cite{Coquand2012}.
+ − 141
+ − 142
If a regular expression matches a string, then in general there is more than
+ − 143
one way of how the string is matched. There are two commonly used
+ − 144
disambiguation strategies to generate a unique answer: one is called GREEDY
+ − 145
matching \cite{Frisch2004} and the other is POSIX
268
+ − 146
matching~\cite{POSIX,Kuklewicz,OkuiSuzuki2010,Sulzmann2014,Vansummeren2006}. For example consider
218
+ − 147
the string @{term xy} and the regular expression \mbox{@{term "STAR (ALT
+ − 148
(ALT x y) xy)"}}. Either the string can be matched in two `iterations' by
+ − 149
the single letter-regular expressions @{term x} and @{term y}, or directly
+ − 150
in one iteration by @{term xy}. The first case corresponds to GREEDY
+ − 151
matching, which first matches with the left-most symbol and only matches the
+ − 152
next symbol in case of a mismatch (this is greedy in the sense of preferring
+ − 153
instant gratification to delayed repletion). The second case is POSIX
+ − 154
matching, which prefers the longest match.
+ − 155
268
+ − 156
In the context of lexing, where an input string needs to be split up
+ − 157
into a sequence of tokens, POSIX is the more natural disambiguation
+ − 158
strategy for what programmers consider basic syntactic building blocks
+ − 159
in their programs. These building blocks are often specified by some
+ − 160
regular expressions, say @{text "r\<^bsub>key\<^esub>"} and @{text
+ − 161
"r\<^bsub>id\<^esub>"} for recognising keywords and identifiers,
+ − 162
respectively. There are a few underlying (informal) rules behind
+ − 163
tokenising a string in a POSIX \cite{POSIX} fashion according to a
+ − 164
collection of regular expressions:
218
+ − 165
+ − 166
\begin{itemize}
265
+ − 167
\item[$\bullet$] \emph{The Longest Match Rule} (or \emph{``{M}aximal {M}unch {R}ule''}):
218
+ − 168
The longest initial substring matched by any regular expression is taken as
+ − 169
next token.\smallskip
+ − 170
+ − 171
\item[$\bullet$] \emph{Priority Rule:}
265
+ − 172
For a particular longest initial substring, the first (leftmost) regular expression
+ − 173
that can match determines the token.\smallskip
+ − 174
+ − 175
\item[$\bullet$] \emph{Star Rule:} A subexpression repeated by ${}^\star$ shall
+ − 176
not match an empty string unless this is the only match for the repetition.\smallskip
+ − 177
+ − 178
\item[$\bullet$] \emph{Empty String Rule:} An empty string shall be considered to
268
+ − 179
be longer than no match at all.
218
+ − 180
\end{itemize}
+ − 181
268
+ − 182
\noindent Consider for example a regular expression @{text
+ − 183
"r\<^bsub>key\<^esub>"} for recognising keywords such as @{text "if"},
+ − 184
@{text "then"} and so on; and @{text "r\<^bsub>id\<^esub>"}
218
+ − 185
recognising identifiers (say, a single character followed by
+ − 186
characters or numbers). Then we can form the regular expression
268
+ − 187
@{text "(r\<^bsub>key\<^esub> + r\<^bsub>id\<^esub>)\<^sup>\<star>"}
+ − 188
and use POSIX matching to tokenise strings, say @{text "iffoo"} and
+ − 189
@{text "if"}. For @{text "iffoo"} we obtain by the Longest Match Rule
+ − 190
a single identifier token, not a keyword followed by an
+ − 191
identifier. For @{text "if"} we obtain by the Priority Rule a keyword
+ − 192
token, not an identifier token---even if @{text "r\<^bsub>id\<^esub>"}
+ − 193
matches also. By the Star Rule we know @{text "(r\<^bsub>key\<^esub> +
+ − 194
r\<^bsub>id\<^esub>)\<^sup>\<star>"} matches @{text "iffoo"},
+ − 195
respectively @{text "if"}, in exactly one `iteration' of the star. The
273
+ − 196
Empty String Rule is for cases where, for example, the regular expression
+ − 197
@{text "(a\<^sup>\<star>)\<^sup>\<star>"} matches against the
268
+ − 198
string @{text "bc"}. Then the longest initial matched substring is the
+ − 199
empty string, which is matched by both the whole regular expression
272
+ − 200
and the parenthesised subexpression.
267
+ − 201
218
+ − 202
+ − 203
One limitation of Brzozowski's matcher is that it only generates a
+ − 204
YES/NO answer for whether a string is being matched by a regular
+ − 205
expression. Sulzmann and Lu~\cite{Sulzmann2014} extended this matcher
+ − 206
to allow generation not just of a YES/NO answer but of an actual
272
+ − 207
matching, called a [lexical] {\em value}. Assuming a regular
+ − 208
expression matches a string, values encode the information of
+ − 209
\emph{how} the string is matched by the regular expression---that is,
+ − 210
which part of the string is matched by which part of the regular
273
+ − 211
expression. For this consider again the string @{text "xy"} and
+ − 212
the regular expression \mbox{@{text "(x + (y + xy))\<^sup>\<star>"}}
+ − 213
(this time fully parenthesised). We can view this regular expression
+ − 214
as tree and if the string @{text xy} is matched by two Star
+ − 215
`iterations', then the @{text x} is matched by the left-most
+ − 216
alternative in this tree and the @{text y} by the right-left alternative. This
+ − 217
suggests to record this matching as
+ − 218
+ − 219
\begin{center}
+ − 220
@{term "Stars [Left(Char x), Right(Left(Char y))]"}
+ − 221
\end{center}
272
+ − 222
273
+ − 223
\noindent where @{const Stars}, @{text Left}, @{text Right} and @{text
+ − 224
Char} are constructors for values. @{text Stars} records how many
+ − 225
iterations were used; @{text Left}, respectively @{text Right}, which
275
+ − 226
alternative is used. This `tree view' leads naturally to the idea that
+ − 227
regular expressions act as types and values as inhabiting those types
+ − 228
(see, for example, \cite{HosoyaVouillonPierce2005}). The value for
+ − 229
the single `iteration', i.e.~the POSIX value, would look as follows
272
+ − 230
273
+ − 231
\begin{center}
+ − 232
@{term "Stars [Seq (Char x) (Char y)]"}
+ − 233
\end{center}
+ − 234
+ − 235
\noindent where @{const Stars} has only a single-element list for the
+ − 236
single iteration and @{const Seq} indicates that @{term xy} is matched
+ − 237
by a sequence regular expression, which we will in what follows
+ − 238
write more formally as @{term "SEQ x y"}.
272
+ − 239
218
+ − 240
272
+ − 241
Sulzmann and Lu give a simple algorithm to calculate a value that
+ − 242
appears to be the value associated with POSIX matching. The challenge
+ − 243
then is to specify that value, in an algorithm-independent fashion,
+ − 244
and to show that Sulzmann and Lu's derivative-based algorithm does
+ − 245
indeed calculate a value that is correct according to the
+ − 246
specification. The answer given by Sulzmann and Lu
+ − 247
\cite{Sulzmann2014} is to define a relation (called an ``order
+ − 248
relation'') on the set of values of @{term r}, and to show that (once
+ − 249
a string to be matched is chosen) there is a maximum element and that
+ − 250
it is computed by their derivative-based algorithm. This proof idea is
+ − 251
inspired by work of Frisch and Cardelli \cite{Frisch2004} on a GREEDY
+ − 252
regular expression matching algorithm. However, we were not able to
+ − 253
establish transitivity and totality for the ``order relation'' by
+ − 254
Sulzmann and Lu. There are some inherent problems with their approach
+ − 255
(of which some of the proofs are not published in
+ − 256
\cite{Sulzmann2014}); perhaps more importantly, we give in this paper
+ − 257
a simple inductive (and algorithm-independent) definition of what we
+ − 258
call being a {\em POSIX value} for a regular expression @{term r} and
+ − 259
a string @{term s}; we show that the algorithm by Sulzmann and Lu
+ − 260
computes such a value and that such a value is unique. Our proofs are
+ − 261
both done by hand and checked in Isabelle/HOL. The experience of
+ − 262
doing our proofs has been that this mechanical checking was absolutely
+ − 263
essential: this subject area has hidden snares. This was also noted by
+ − 264
Kuklewicz \cite{Kuklewicz} who found that nearly all POSIX matching
+ − 265
implementations are ``buggy'' \cite[Page 203]{Sulzmann2014} and by
+ − 266
Grathwohl et al \cite[Page 36]{CrashCourse2014} who wrote:
218
+ − 267
+ − 268
\begin{quote}
+ − 269
\it{}``The POSIX strategy is more complicated than the greedy because of
+ − 270
the dependence on information about the length of matched strings in the
+ − 271
various subexpressions.''
+ − 272
\end{quote}
+ − 273
+ − 274
%\footnote{The relation @{text "\<ge>\<^bsub>r\<^esub>"} defined by Sulzmann and Lu \cite{Sulzmann2014}
+ − 275
%is a relation on the
+ − 276
%values for the regular expression @{term r}; but it only holds between
+ − 277
%@{term "v\<^sub>1"} and @{term "v\<^sub>2"} in cases where @{term "v\<^sub>1"} and @{term "v\<^sub>2"} have
+ − 278
%the same flattening (underlying string). So a counterexample to totality is
+ − 279
%given by taking two values @{term "v\<^sub>1"} and @{term "v\<^sub>2"} for @{term r} that
+ − 280
%have different flattenings (see Section~\ref{posixsec}). A different
+ − 281
%relation @{text "\<ge>\<^bsub>r,s\<^esub>"} on the set of values for @{term r}
+ − 282
%with flattening @{term s} is definable by the same approach, and is indeed
+ − 283
%total; but that is not what Proposition 1 of \cite{Sulzmann2014} does.}
+ − 284
+ − 285
+ − 286
\noindent {\bf Contributions:} We have implemented in Isabelle/HOL the
+ − 287
derivative-based regular expression matching algorithm of
+ − 288
Sulzmann and Lu \cite{Sulzmann2014}. We have proved the correctness of this
+ − 289
algorithm according to our specification of what a POSIX value is (inspired
+ − 290
by work of Vansummeren \cite{Vansummeren2006}). Sulzmann
+ − 291
and Lu sketch in \cite{Sulzmann2014} an informal correctness proof: but to
+ − 292
us it contains unfillable gaps.\footnote{An extended version of
+ − 293
\cite{Sulzmann2014} is available at the website of its first author; this
+ − 294
extended version already includes remarks in the appendix that their
+ − 295
informal proof contains gaps, and possible fixes are not fully worked out.}
+ − 296
Our specification of a POSIX value consists of a simple inductive definition
+ − 297
that given a string and a regular expression uniquely determines this value.
267
+ − 298
We also show that our definition is equivalent to an ordering
268
+ − 299
of values based on positions by Okui and Suzuki \cite{OkuiSuzuki2010}.
218
+ − 300
Derivatives as calculated by Brzozowski's method are usually more complex
+ − 301
regular expressions than the initial one; various optimisations are
+ − 302
possible. We prove the correctness when simplifications of @{term "ALT ZERO
+ − 303
r"}, @{term "ALT r ZERO"}, @{term "SEQ ONE r"} and @{term "SEQ r ONE"} to
272
+ − 304
@{term r} are applied. We extend our results to ???
218
+ − 305
+ − 306
*}
+ − 307
+ − 308
section {* Preliminaries *}
+ − 309
273
+ − 310
text {* \noindent Strings in Isabelle/HOL are lists of characters with
+ − 311
the empty string being represented by the empty list, written @{term
+ − 312
"[]"}, and list-cons being written as @{term "DUMMY # DUMMY"}. Often
+ − 313
we use the usual bracket notation for lists also for strings; for
+ − 314
example a string consisting of just a single character @{term c} is
+ − 315
written @{term "[c]"}. We use the usual definitions for
+ − 316
\emph{prefixes} and \emph{strict prefixes} of strings. By using the
218
+ − 317
type @{type char} for characters we have a supply of finitely many
+ − 318
characters roughly corresponding to the ASCII character set. Regular
273
+ − 319
expressions are defined as usual as the elements of the following
+ − 320
inductive datatype:
218
+ − 321
+ − 322
\begin{center}
+ − 323
@{text "r :="}
+ − 324
@{const "ZERO"} $\mid$
+ − 325
@{const "ONE"} $\mid$
+ − 326
@{term "CHAR c"} $\mid$
+ − 327
@{term "ALT r\<^sub>1 r\<^sub>2"} $\mid$
+ − 328
@{term "SEQ r\<^sub>1 r\<^sub>2"} $\mid$
+ − 329
@{term "STAR r"}
+ − 330
\end{center}
+ − 331
+ − 332
\noindent where @{const ZERO} stands for the regular expression that does
+ − 333
not match any string, @{const ONE} for the regular expression that matches
+ − 334
only the empty string and @{term c} for matching a character literal. The
+ − 335
language of a regular expression is also defined as usual by the
+ − 336
recursive function @{term L} with the six clauses:
+ − 337
+ − 338
\begin{center}
267
+ − 339
\begin{tabular}{l@ {\hspace{4mm}}rcl}
273
+ − 340
\textit{(1)} & @{thm (lhs) L.simps(1)} & $\dn$ & @{thm (rhs) L.simps(1)}\\
+ − 341
\textit{(2)} & @{thm (lhs) L.simps(2)} & $\dn$ & @{thm (rhs) L.simps(2)}\\
+ − 342
\textit{(3)} & @{thm (lhs) L.simps(3)} & $\dn$ & @{thm (rhs) L.simps(3)}\\
+ − 343
\textit{(4)} & @{thm (lhs) L.simps(4)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ &
+ − 344
@{thm (rhs) L.simps(4)[of "r\<^sub>1" "r\<^sub>2"]}\\
+ − 345
\textit{(5)} & @{thm (lhs) L.simps(5)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ &
+ − 346
@{thm (rhs) L.simps(5)[of "r\<^sub>1" "r\<^sub>2"]}\\
+ − 347
\textit{(6)} & @{thm (lhs) L.simps(6)} & $\dn$ & @{thm (rhs) L.simps(6)}\\
218
+ − 348
\end{tabular}
+ − 349
\end{center}
+ − 350
273
+ − 351
\noindent In clause \textit{(4)} we use the operation @{term "DUMMY ;;
218
+ − 352
DUMMY"} for the concatenation of two languages (it is also list-append for
+ − 353
strings). We use the star-notation for regular expressions and for
+ − 354
languages (in the last clause above). The star for languages is defined
+ − 355
inductively by two clauses: @{text "(i)"} the empty string being in
+ − 356
the star of a language and @{text "(ii)"} if @{term "s\<^sub>1"} is in a
+ − 357
language and @{term "s\<^sub>2"} in the star of this language, then also @{term
+ − 358
"s\<^sub>1 @ s\<^sub>2"} is in the star of this language. It will also be convenient
+ − 359
to use the following notion of a \emph{semantic derivative} (or \emph{left
+ − 360
quotient}) of a language defined as
+ − 361
%
+ − 362
\begin{center}
+ − 363
@{thm Der_def}\;.
+ − 364
\end{center}
+ − 365
+ − 366
\noindent
+ − 367
For semantic derivatives we have the following equations (for example
+ − 368
mechanically proved in \cite{Krauss2011}):
+ − 369
%
+ − 370
\begin{equation}\label{SemDer}
+ − 371
\begin{array}{lcl}
+ − 372
@{thm (lhs) Der_null} & \dn & @{thm (rhs) Der_null}\\
+ − 373
@{thm (lhs) Der_empty} & \dn & @{thm (rhs) Der_empty}\\
+ − 374
@{thm (lhs) Der_char} & \dn & @{thm (rhs) Der_char}\\
+ − 375
@{thm (lhs) Der_union} & \dn & @{thm (rhs) Der_union}\\
+ − 376
@{thm (lhs) Der_Sequ} & \dn & @{thm (rhs) Der_Sequ}\\
+ − 377
@{thm (lhs) Der_star} & \dn & @{thm (rhs) Der_star}
+ − 378
\end{array}
+ − 379
\end{equation}
+ − 380
+ − 381
+ − 382
\noindent \emph{\Brz's derivatives} of regular expressions
+ − 383
\cite{Brzozowski1964} can be easily defined by two recursive functions:
+ − 384
the first is from regular expressions to booleans (implementing a test
+ − 385
when a regular expression can match the empty string), and the second
+ − 386
takes a regular expression and a character to a (derivative) regular
+ − 387
expression:
+ − 388
+ − 389
\begin{center}
+ − 390
\begin{tabular}{lcl}
+ − 391
@{thm (lhs) nullable.simps(1)} & $\dn$ & @{thm (rhs) nullable.simps(1)}\\
+ − 392
@{thm (lhs) nullable.simps(2)} & $\dn$ & @{thm (rhs) nullable.simps(2)}\\
+ − 393
@{thm (lhs) nullable.simps(3)} & $\dn$ & @{thm (rhs) nullable.simps(3)}\\
+ − 394
@{thm (lhs) nullable.simps(4)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) nullable.simps(4)[of "r\<^sub>1" "r\<^sub>2"]}\\
+ − 395
@{thm (lhs) nullable.simps(5)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) nullable.simps(5)[of "r\<^sub>1" "r\<^sub>2"]}\\
273
+ − 396
@{thm (lhs) nullable.simps(6)} & $\dn$ & @{thm (rhs) nullable.simps(6)}\medskip\\
218
+ − 397
273
+ − 398
% \end{tabular}
+ − 399
% \end{center}
+ − 400
+ − 401
% \begin{center}
+ − 402
% \begin{tabular}{lcl}
+ − 403
218
+ − 404
@{thm (lhs) der.simps(1)} & $\dn$ & @{thm (rhs) der.simps(1)}\\
+ − 405
@{thm (lhs) der.simps(2)} & $\dn$ & @{thm (rhs) der.simps(2)}\\
+ − 406
@{thm (lhs) der.simps(3)} & $\dn$ & @{thm (rhs) der.simps(3)}\\
+ − 407
@{thm (lhs) der.simps(4)[of c "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) der.simps(4)[of c "r\<^sub>1" "r\<^sub>2"]}\\
+ − 408
@{thm (lhs) der.simps(5)[of c "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) der.simps(5)[of c "r\<^sub>1" "r\<^sub>2"]}\\
+ − 409
@{thm (lhs) der.simps(6)} & $\dn$ & @{thm (rhs) der.simps(6)}
+ − 410
\end{tabular}
+ − 411
\end{center}
+ − 412
+ − 413
\noindent
+ − 414
We may extend this definition to give derivatives w.r.t.~strings:
+ − 415
+ − 416
\begin{center}
+ − 417
\begin{tabular}{lcl}
+ − 418
@{thm (lhs) ders.simps(1)} & $\dn$ & @{thm (rhs) ders.simps(1)}\\
+ − 419
@{thm (lhs) ders.simps(2)} & $\dn$ & @{thm (rhs) ders.simps(2)}\\
+ − 420
\end{tabular}
+ − 421
\end{center}
+ − 422
+ − 423
\noindent Given the equations in \eqref{SemDer}, it is a relatively easy
+ − 424
exercise in mechanical reasoning to establish that
+ − 425
+ − 426
\begin{proposition}\label{derprop}\mbox{}\\
+ − 427
\begin{tabular}{ll}
273
+ − 428
\textit{(1)} & @{thm (lhs) nullable_correctness} if and only if
218
+ − 429
@{thm (rhs) nullable_correctness}, and \\
273
+ − 430
\textit{(2)} & @{thm[mode=IfThen] der_correctness}.
218
+ − 431
\end{tabular}
+ − 432
\end{proposition}
+ − 433
+ − 434
\noindent With this in place it is also very routine to prove that the
+ − 435
regular expression matcher defined as
+ − 436
%
+ − 437
\begin{center}
+ − 438
@{thm match_def}
+ − 439
\end{center}
+ − 440
+ − 441
\noindent gives a positive answer if and only if @{term "s \<in> L r"}.
+ − 442
Consequently, this regular expression matching algorithm satisfies the
+ − 443
usual specification for regular expression matching. While the matcher
+ − 444
above calculates a provably correct YES/NO answer for whether a regular
+ − 445
expression matches a string or not, the novel idea of Sulzmann and Lu
+ − 446
\cite{Sulzmann2014} is to append another phase to this algorithm in order
+ − 447
to calculate a [lexical] value. We will explain the details next.
+ − 448
+ − 449
*}
+ − 450
+ − 451
section {* POSIX Regular Expression Matching\label{posixsec} *}
+ − 452
+ − 453
text {*
+ − 454
268
+ − 455
There have been many previous works that use values for encoding
+ − 456
\emph{how} a regular expression matches a string.
+ − 457
The clever idea by Sulzmann and Lu \cite{Sulzmann2014} is to
+ − 458
define a function on values that mirrors (but inverts) the
218
+ − 459
construction of the derivative on regular expressions. \emph{Values}
+ − 460
are defined as the inductive datatype
+ − 461
+ − 462
\begin{center}
+ − 463
@{text "v :="}
+ − 464
@{const "Void"} $\mid$
+ − 465
@{term "val.Char c"} $\mid$
+ − 466
@{term "Left v"} $\mid$
+ − 467
@{term "Right v"} $\mid$
+ − 468
@{term "Seq v\<^sub>1 v\<^sub>2"} $\mid$
+ − 469
@{term "Stars vs"}
+ − 470
\end{center}
+ − 471
+ − 472
\noindent where we use @{term vs} to stand for a list of
+ − 473
values. (This is similar to the approach taken by Frisch and
+ − 474
Cardelli for GREEDY matching \cite{Frisch2004}, and Sulzmann and Lu
+ − 475
for POSIX matching \cite{Sulzmann2014}). The string underlying a
+ − 476
value can be calculated by the @{const flat} function, written
+ − 477
@{term "flat DUMMY"} and defined as:
+ − 478
+ − 479
\begin{center}
+ − 480
\begin{tabular}[t]{lcl}
+ − 481
@{thm (lhs) flat.simps(1)} & $\dn$ & @{thm (rhs) flat.simps(1)}\\
+ − 482
@{thm (lhs) flat.simps(2)} & $\dn$ & @{thm (rhs) flat.simps(2)}\\
+ − 483
@{thm (lhs) flat.simps(3)} & $\dn$ & @{thm (rhs) flat.simps(3)}\\
+ − 484
@{thm (lhs) flat.simps(4)} & $\dn$ & @{thm (rhs) flat.simps(4)}
+ − 485
\end{tabular}\hspace{14mm}
+ − 486
\begin{tabular}[t]{lcl}
+ − 487
@{thm (lhs) flat.simps(5)[of "v\<^sub>1" "v\<^sub>2"]} & $\dn$ & @{thm (rhs) flat.simps(5)[of "v\<^sub>1" "v\<^sub>2"]}\\
+ − 488
@{thm (lhs) flat.simps(6)} & $\dn$ & @{thm (rhs) flat.simps(6)}\\
+ − 489
@{thm (lhs) flat.simps(7)} & $\dn$ & @{thm (rhs) flat.simps(7)}\\
+ − 490
\end{tabular}
+ − 491
\end{center}
+ − 492
273
+ − 493
\noindent We will sometimes refer to the underlying string of a
+ − 494
value as \emph{flattened value}. We will also overload our notation and
+ − 495
use @{term "flats vs"} for flattening a list of values and concatenating
+ − 496
the resulting strings.
+ − 497
+ − 498
Sulzmann and Lu define
+ − 499
inductively an \emph{inhabitation relation} that associates values to
+ − 500
regular expressions. We define this relation as
+ − 501
follows:\footnote{Note that the rule for @{term Stars} differs from
+ − 502
our earlier paper \cite{AusafDyckhoffUrban2016}. There we used the
+ − 503
original definition by Sulzmann and Lu which does not require that
+ − 504
the values @{term "v \<in> set vs"} flatten to a non-empty
+ − 505
string. The reason for introducing the more restricted version of
+ − 506
lexical values is convenience later on when reasoning about an
+ − 507
ordering relation for values.}
218
+ − 508
+ − 509
\begin{center}
280
+ − 510
\begin{tabular}{c@ {\hspace{12mm}}c}\label{prfintros}
218
+ − 511
\\[-8mm]
268
+ − 512
@{thm[mode=Axiom] Prf.intros(4)} &
218
+ − 513
@{thm[mode=Axiom] Prf.intros(5)[of "c"]}\\[4mm]
268
+ − 514
@{thm[mode=Rule] Prf.intros(2)[of "v\<^sub>1" "r\<^sub>1" "r\<^sub>2"]} &
218
+ − 515
@{thm[mode=Rule] Prf.intros(3)[of "v\<^sub>2" "r\<^sub>1" "r\<^sub>2"]}\\[4mm]
268
+ − 516
@{thm[mode=Rule] Prf.intros(1)[of "v\<^sub>1" "r\<^sub>1" "v\<^sub>2" "r\<^sub>2"]} &
266
+ − 517
@{thm[mode=Rule] Prf.intros(6)[of "vs"]}
218
+ − 518
\end{tabular}
+ − 519
\end{center}
+ − 520
268
+ − 521
\noindent where in the clause for @{const "Stars"} we use the
+ − 522
notation @{term "v \<in> set vs"} for indicating that @{text v} is a
+ − 523
member in the list @{text vs}. We require in this rule that every
+ − 524
value in @{term vs} flattens to a non-empty string. The idea is that
+ − 525
@{term "Stars"}-values satisfy the informal Star Rule (see Introduction)
+ − 526
where the $^\star$ does not match the empty string unless this is
+ − 527
the only match for the repetition. Note also that no values are
+ − 528
associated with the regular expression @{term ZERO}, and that the
+ − 529
only value associated with the regular expression @{term ONE} is
+ − 530
@{term Void}. It is routine to establish how values ``inhabiting''
+ − 531
a regular expression correspond to the language of a regular
+ − 532
expression, namely
218
+ − 533
269
+ − 534
\begin{proposition}\label{inhabs}
218
+ − 535
@{thm L_flat_Prf}
+ − 536
\end{proposition}
+ − 537
267
+ − 538
\noindent
268
+ − 539
Given a regular expression @{text r} and a string @{text s}, we define the
267
+ − 540
set of all \emph{Lexical Values} inhabited by @{text r} with the underlying string
268
+ − 541
being @{text s}:\footnote{Okui and Suzuki refer to our lexical values
+ − 542
as \emph{canonical values} in \cite{OkuiSuzuki2010}. The notion of \emph{non-problematic
273
+ − 543
values} by Cardelli and Frisch \cite{Frisch2004} is related, but not identical
268
+ − 544
to our lexical values.}
267
+ − 545
+ − 546
\begin{center}
+ − 547
@{thm LV_def}
+ − 548
\end{center}
+ − 549
268
+ − 550
\noindent The main property of @{term "LV r s"} is that it is alway finite.
+ − 551
+ − 552
\begin{proposition}
+ − 553
@{thm LV_finite}
+ − 554
\end{proposition}
267
+ − 555
268
+ − 556
\noindent This finiteness property does not hold in general if we
+ − 557
remove the side-condition about @{term "flat v \<noteq> []"} in the
+ − 558
@{term Stars}-rule above. For example using Sulzmann and Lu's
+ − 559
less restrictive definition, @{term "LV (STAR ONE) []"} would contain
+ − 560
infinitely many values, but according to our more restricted
273
+ − 561
definition only a single value, namely @{thm LV_STAR_ONE_empty}.
267
+ − 562
268
+ − 563
If a regular expression @{text r} matches a string @{text s}, then
+ − 564
generally the set @{term "LV r s"} is not just a singleton set. In
+ − 565
case of POSIX matching the problem is to calculate the unique lexical value
+ − 566
that satisfies the (informal) POSIX rules from the Introduction.
+ − 567
Graphically the POSIX value calculation algorithm by Sulzmann and Lu
+ − 568
can be illustrated by the picture in Figure~\ref{Sulz} where the
+ − 569
path from the left to the right involving @{term
+ − 570
derivatives}/@{const nullable} is the first phase of the algorithm
+ − 571
(calculating successive \Brz's derivatives) and @{const
+ − 572
mkeps}/@{text inj}, the path from right to left, the second
+ − 573
phase. This picture shows the steps required when a regular
+ − 574
expression, say @{text "r\<^sub>1"}, matches the string @{term
+ − 575
"[a,b,c]"}. We first build the three derivatives (according to
+ − 576
@{term a}, @{term b} and @{term c}). We then use @{const nullable}
+ − 577
to find out whether the resulting derivative regular expression
+ − 578
@{term "r\<^sub>4"} can match the empty string. If yes, we call the
+ − 579
function @{const mkeps} that produces a value @{term "v\<^sub>4"}
+ − 580
for how @{term "r\<^sub>4"} can match the empty string (taking into
+ − 581
account the POSIX constraints in case there are several ways). This
+ − 582
function is defined by the clauses:
218
+ − 583
+ − 584
\begin{figure}[t]
+ − 585
\begin{center}
+ − 586
\begin{tikzpicture}[scale=2,node distance=1.3cm,
+ − 587
every node/.style={minimum size=6mm}]
+ − 588
\node (r1) {@{term "r\<^sub>1"}};
+ − 589
\node (r2) [right=of r1]{@{term "r\<^sub>2"}};
+ − 590
\draw[->,line width=1mm](r1)--(r2) node[above,midway] {@{term "der a DUMMY"}};
+ − 591
\node (r3) [right=of r2]{@{term "r\<^sub>3"}};
+ − 592
\draw[->,line width=1mm](r2)--(r3) node[above,midway] {@{term "der b DUMMY"}};
+ − 593
\node (r4) [right=of r3]{@{term "r\<^sub>4"}};
+ − 594
\draw[->,line width=1mm](r3)--(r4) node[above,midway] {@{term "der c DUMMY"}};
+ − 595
\draw (r4) node[anchor=west] {\;\raisebox{3mm}{@{term nullable}}};
+ − 596
\node (v4) [below=of r4]{@{term "v\<^sub>4"}};
+ − 597
\draw[->,line width=1mm](r4) -- (v4);
+ − 598
\node (v3) [left=of v4] {@{term "v\<^sub>3"}};
+ − 599
\draw[->,line width=1mm](v4)--(v3) node[below,midway] {@{text "inj r\<^sub>3 c"}};
+ − 600
\node (v2) [left=of v3]{@{term "v\<^sub>2"}};
+ − 601
\draw[->,line width=1mm](v3)--(v2) node[below,midway] {@{text "inj r\<^sub>2 b"}};
+ − 602
\node (v1) [left=of v2] {@{term "v\<^sub>1"}};
+ − 603
\draw[->,line width=1mm](v2)--(v1) node[below,midway] {@{text "inj r\<^sub>1 a"}};
+ − 604
\draw (r4) node[anchor=north west] {\;\raisebox{-8mm}{@{term "mkeps"}}};
+ − 605
\end{tikzpicture}
+ − 606
\end{center}
+ − 607
\mbox{}\\[-13mm]
+ − 608
+ − 609
\caption{The two phases of the algorithm by Sulzmann \& Lu \cite{Sulzmann2014},
+ − 610
matching the string @{term "[a,b,c]"}. The first phase (the arrows from
+ − 611
left to right) is \Brz's matcher building successive derivatives. If the
+ − 612
last regular expression is @{term nullable}, then the functions of the
+ − 613
second phase are called (the top-down and right-to-left arrows): first
+ − 614
@{term mkeps} calculates a value @{term "v\<^sub>4"} witnessing
+ − 615
how the empty string has been recognised by @{term "r\<^sub>4"}. After
+ − 616
that the function @{term inj} ``injects back'' the characters of the string into
+ − 617
the values.
+ − 618
\label{Sulz}}
+ − 619
\end{figure}
+ − 620
+ − 621
\begin{center}
+ − 622
\begin{tabular}{lcl}
+ − 623
@{thm (lhs) mkeps.simps(1)} & $\dn$ & @{thm (rhs) mkeps.simps(1)}\\
+ − 624
@{thm (lhs) mkeps.simps(2)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) mkeps.simps(2)[of "r\<^sub>1" "r\<^sub>2"]}\\
+ − 625
@{thm (lhs) mkeps.simps(3)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) mkeps.simps(3)[of "r\<^sub>1" "r\<^sub>2"]}\\
+ − 626
@{thm (lhs) mkeps.simps(4)} & $\dn$ & @{thm (rhs) mkeps.simps(4)}\\
+ − 627
\end{tabular}
+ − 628
\end{center}
+ − 629
+ − 630
\noindent Note that this function needs only to be partially defined,
+ − 631
namely only for regular expressions that are nullable. In case @{const
+ − 632
nullable} fails, the string @{term "[a,b,c]"} cannot be matched by @{term
+ − 633
"r\<^sub>1"} and the null value @{term "None"} is returned. Note also how this function
+ − 634
makes some subtle choices leading to a POSIX value: for example if an
+ − 635
alternative regular expression, say @{term "ALT r\<^sub>1 r\<^sub>2"}, can
+ − 636
match the empty string and furthermore @{term "r\<^sub>1"} can match the
+ − 637
empty string, then we return a @{text Left}-value. The @{text
+ − 638
Right}-value will only be returned if @{term "r\<^sub>1"} cannot match the empty
+ − 639
string.
+ − 640
+ − 641
The most interesting idea from Sulzmann and Lu \cite{Sulzmann2014} is
+ − 642
the construction of a value for how @{term "r\<^sub>1"} can match the
+ − 643
string @{term "[a,b,c]"} from the value how the last derivative, @{term
+ − 644
"r\<^sub>4"} in Fig.~\ref{Sulz}, can match the empty string. Sulzmann and
+ − 645
Lu achieve this by stepwise ``injecting back'' the characters into the
+ − 646
values thus inverting the operation of building derivatives, but on the level
+ − 647
of values. The corresponding function, called @{term inj}, takes three
+ − 648
arguments, a regular expression, a character and a value. For example in
+ − 649
the first (or right-most) @{term inj}-step in Fig.~\ref{Sulz} the regular
+ − 650
expression @{term "r\<^sub>3"}, the character @{term c} from the last
+ − 651
derivative step and @{term "v\<^sub>4"}, which is the value corresponding
+ − 652
to the derivative regular expression @{term "r\<^sub>4"}. The result is
+ − 653
the new value @{term "v\<^sub>3"}. The final result of the algorithm is
+ − 654
the value @{term "v\<^sub>1"}. The @{term inj} function is defined by recursion on regular
+ − 655
expressions and by analysing the shape of values (corresponding to
+ − 656
the derivative regular expressions).
+ − 657
%
+ − 658
\begin{center}
+ − 659
\begin{tabular}{l@ {\hspace{5mm}}lcl}
273
+ − 660
\textit{(1)} & @{thm (lhs) injval.simps(1)} & $\dn$ & @{thm (rhs) injval.simps(1)}\\
+ − 661
\textit{(2)} & @{thm (lhs) injval.simps(2)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1"]} & $\dn$ &
218
+ − 662
@{thm (rhs) injval.simps(2)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1"]}\\
273
+ − 663
\textit{(3)} & @{thm (lhs) injval.simps(3)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]} & $\dn$ &
218
+ − 664
@{thm (rhs) injval.simps(3)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]}\\
273
+ − 665
\textit{(4)} & @{thm (lhs) injval.simps(4)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]} & $\dn$
218
+ − 666
& @{thm (rhs) injval.simps(4)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]}\\
273
+ − 667
\textit{(5)} & @{thm (lhs) injval.simps(5)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]} & $\dn$
218
+ − 668
& @{thm (rhs) injval.simps(5)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]}\\
273
+ − 669
\textit{(6)} & @{thm (lhs) injval.simps(6)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]} & $\dn$
218
+ − 670
& @{thm (rhs) injval.simps(6)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]}\\
273
+ − 671
\textit{(7)} & @{thm (lhs) injval.simps(7)[of "r" "c" "v" "vs"]} & $\dn$
218
+ − 672
& @{thm (rhs) injval.simps(7)[of "r" "c" "v" "vs"]}\\
+ − 673
\end{tabular}
+ − 674
\end{center}
+ − 675
+ − 676
\noindent To better understand what is going on in this definition it
+ − 677
might be instructive to look first at the three sequence cases (clauses
273
+ − 678
\textit{(4)} -- \textit{(6)}). In each case we need to construct an ``injected value'' for
218
+ − 679
@{term "SEQ r\<^sub>1 r\<^sub>2"}. This must be a value of the form @{term
+ − 680
"Seq DUMMY DUMMY"}\,. Recall the clause of the @{text derivative}-function
+ − 681
for sequence regular expressions:
+ − 682
+ − 683
\begin{center}
+ − 684
@{thm (lhs) der.simps(5)[of c "r\<^sub>1" "r\<^sub>2"]} $\dn$ @{thm (rhs) der.simps(5)[of c "r\<^sub>1" "r\<^sub>2"]}
+ − 685
\end{center}
+ − 686
+ − 687
\noindent Consider first the @{text "else"}-branch where the derivative is @{term
+ − 688
"SEQ (der c r\<^sub>1) r\<^sub>2"}. The corresponding value must therefore
+ − 689
be of the form @{term "Seq v\<^sub>1 v\<^sub>2"}, which matches the left-hand
273
+ − 690
side in clause~\textit{(4)} of @{term inj}. In the @{text "if"}-branch the derivative is an
218
+ − 691
alternative, namely @{term "ALT (SEQ (der c r\<^sub>1) r\<^sub>2) (der c
+ − 692
r\<^sub>2)"}. This means we either have to consider a @{text Left}- or
+ − 693
@{text Right}-value. In case of the @{text Left}-value we know further it
+ − 694
must be a value for a sequence regular expression. Therefore the pattern
273
+ − 695
we match in the clause \textit{(5)} is @{term "Left (Seq v\<^sub>1 v\<^sub>2)"},
+ − 696
while in \textit{(6)} it is just @{term "Right v\<^sub>2"}. One more interesting
+ − 697
point is in the right-hand side of clause \textit{(6)}: since in this case the
218
+ − 698
regular expression @{text "r\<^sub>1"} does not ``contribute'' to
+ − 699
matching the string, that means it only matches the empty string, we need to
+ − 700
call @{const mkeps} in order to construct a value for how @{term "r\<^sub>1"}
+ − 701
can match this empty string. A similar argument applies for why we can
273
+ − 702
expect in the left-hand side of clause \textit{(7)} that the value is of the form
218
+ − 703
@{term "Seq v (Stars vs)"}---the derivative of a star is @{term "SEQ (der c r)
+ − 704
(STAR r)"}. Finally, the reason for why we can ignore the second argument
273
+ − 705
in clause \textit{(1)} of @{term inj} is that it will only ever be called in cases
218
+ − 706
where @{term "c=d"}, but the usual linearity restrictions in patterns do
+ − 707
not allow us to build this constraint explicitly into our function
+ − 708
definition.\footnote{Sulzmann and Lu state this clause as @{thm (lhs)
+ − 709
injval.simps(1)[of "c" "c"]} $\dn$ @{thm (rhs) injval.simps(1)[of "c"]},
+ − 710
but our deviation is harmless.}
+ − 711
+ − 712
The idea of the @{term inj}-function to ``inject'' a character, say
+ − 713
@{term c}, into a value can be made precise by the first part of the
+ − 714
following lemma, which shows that the underlying string of an injected
+ − 715
value has a prepended character @{term c}; the second part shows that the
+ − 716
underlying string of an @{const mkeps}-value is always the empty string
+ − 717
(given the regular expression is nullable since otherwise @{text mkeps}
+ − 718
might not be defined).
+ − 719
+ − 720
\begin{lemma}\mbox{}\smallskip\\\label{Prf_injval_flat}
+ − 721
\begin{tabular}{ll}
+ − 722
(1) & @{thm[mode=IfThen] Prf_injval_flat}\\
+ − 723
(2) & @{thm[mode=IfThen] mkeps_flat}
+ − 724
\end{tabular}
+ − 725
\end{lemma}
+ − 726
+ − 727
\begin{proof}
+ − 728
Both properties are by routine inductions: the first one can, for example,
+ − 729
be proved by induction over the definition of @{term derivatives}; the second by
+ − 730
an induction on @{term r}. There are no interesting cases.\qed
+ − 731
\end{proof}
+ − 732
+ − 733
Having defined the @{const mkeps} and @{text inj} function we can extend
267
+ − 734
\Brz's matcher so that a value is constructed (assuming the
218
+ − 735
regular expression matches the string). The clauses of the Sulzmann and Lu lexer are
+ − 736
+ − 737
\begin{center}
+ − 738
\begin{tabular}{lcl}
+ − 739
@{thm (lhs) lexer.simps(1)} & $\dn$ & @{thm (rhs) lexer.simps(1)}\\
+ − 740
@{thm (lhs) lexer.simps(2)} & $\dn$ & @{text "case"} @{term "lexer (der c r) s"} @{text of}\\
+ − 741
& & \phantom{$|$} @{term "None"} @{text "\<Rightarrow>"} @{term None}\\
+ − 742
& & $|$ @{term "Some v"} @{text "\<Rightarrow>"} @{term "Some (injval r c v)"}
+ − 743
\end{tabular}
+ − 744
\end{center}
+ − 745
+ − 746
\noindent If the regular expression does not match the string, @{const None} is
+ − 747
returned. If the regular expression \emph{does}
+ − 748
match the string, then @{const Some} value is returned. One important
+ − 749
virtue of this algorithm is that it can be implemented with ease in any
+ − 750
functional programming language and also in Isabelle/HOL. In the remaining
+ − 751
part of this section we prove that this algorithm is correct.
+ − 752
267
+ − 753
The well-known idea of POSIX matching is informally defined by some
273
+ − 754
rules such as the Longest Match and Priority Rules (see
267
+ − 755
Introduction); as correctly argued in \cite{Sulzmann2014}, this
218
+ − 756
needs formal specification. Sulzmann and Lu define an ``ordering
267
+ − 757
relation'' between values and argue that there is a maximum value,
+ − 758
as given by the derivative-based algorithm. In contrast, we shall
+ − 759
introduce a simple inductive definition that specifies directly what
+ − 760
a \emph{POSIX value} is, incorporating the POSIX-specific choices
+ − 761
into the side-conditions of our rules. Our definition is inspired by
273
+ − 762
the matching relation given by Vansummeren~\cite{Vansummeren2006}.
+ − 763
The relation we define is ternary and
267
+ − 764
written as \mbox{@{term "s \<in> r \<rightarrow> v"}}, relating
+ − 765
strings, regular expressions and values; the inductive rules are given in
+ − 766
Figure~\ref{POSIXrules}.
+ − 767
We can prove that given a string @{term s} and regular expression @{term
+ − 768
r}, the POSIX value @{term v} is uniquely determined by @{term "s \<in> r \<rightarrow> v"}.
+ − 769
218
+ − 770
%
267
+ − 771
\begin{figure}[t]
218
+ − 772
\begin{center}
+ − 773
\begin{tabular}{c}
+ − 774
@{thm[mode=Axiom] Posix.intros(1)}@{text "P"}@{term "ONE"} \qquad
+ − 775
@{thm[mode=Axiom] Posix.intros(2)}@{text "P"}@{term "c"}\medskip\\
+ − 776
@{thm[mode=Rule] Posix.intros(3)[of "s" "r\<^sub>1" "v" "r\<^sub>2"]}@{text "P+L"}\qquad
+ − 777
@{thm[mode=Rule] Posix.intros(4)[of "s" "r\<^sub>2" "v" "r\<^sub>1"]}@{text "P+R"}\medskip\\
+ − 778
$\mprset{flushleft}
+ − 779
\inferrule
+ − 780
{@{thm (prem 1) Posix.intros(5)[of "s\<^sub>1" "r\<^sub>1" "v\<^sub>1" "s\<^sub>2" "r\<^sub>2" "v\<^sub>2"]} \qquad
+ − 781
@{thm (prem 2) Posix.intros(5)[of "s\<^sub>1" "r\<^sub>1" "v\<^sub>1" "s\<^sub>2" "r\<^sub>2" "v\<^sub>2"]} \\\\
+ − 782
@{thm (prem 3) Posix.intros(5)[of "s\<^sub>1" "r\<^sub>1" "v\<^sub>1" "s\<^sub>2" "r\<^sub>2" "v\<^sub>2"]}}
+ − 783
{@{thm (concl) Posix.intros(5)[of "s\<^sub>1" "r\<^sub>1" "v\<^sub>1" "s\<^sub>2" "r\<^sub>2" "v\<^sub>2"]}}$@{text "PS"}\\
+ − 784
@{thm[mode=Axiom] Posix.intros(7)}@{text "P[]"}\medskip\\
+ − 785
$\mprset{flushleft}
+ − 786
\inferrule
+ − 787
{@{thm (prem 1) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]} \qquad
+ − 788
@{thm (prem 2) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]} \qquad
+ − 789
@{thm (prem 3) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]} \\\\
+ − 790
@{thm (prem 4) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]}}
+ − 791
{@{thm (concl) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]}}$@{text "P\<star>"}
+ − 792
\end{tabular}
+ − 793
\end{center}
267
+ − 794
\caption{Our inductive definition of POSIX values.}\label{POSIXrules}
+ − 795
\end{figure}
218
+ − 796
267
+ − 797
218
+ − 798
+ − 799
\begin{theorem}\mbox{}\smallskip\\\label{posixdeterm}
+ − 800
\begin{tabular}{ll}
272
+ − 801
(1) & If @{thm (prem 1) Posix1(1)} then @{thm (concl)
218
+ − 802
Posix1(1)} and @{thm (concl) Posix1(2)}.\\
272
+ − 803
(2) & @{thm[mode=IfThen] Posix_determ(1)[of _ _ "v" "v'"]}
218
+ − 804
\end{tabular}
+ − 805
\end{theorem}
+ − 806
+ − 807
\begin{proof} Both by induction on the definition of @{term "s \<in> r \<rightarrow> v"}.
+ − 808
The second parts follows by a case analysis of @{term "s \<in> r \<rightarrow> v'"} and
+ − 809
the first part.\qed
+ − 810
\end{proof}
+ − 811
+ − 812
\noindent
267
+ − 813
We claim that our @{term "s \<in> r \<rightarrow> v"} relation captures the idea behind the four
218
+ − 814
informal POSIX rules shown in the Introduction: Consider for example the
+ − 815
rules @{text "P+L"} and @{text "P+R"} where the POSIX value for a string
+ − 816
and an alternative regular expression, that is @{term "(s, ALT r\<^sub>1 r\<^sub>2)"},
+ − 817
is specified---it is always a @{text "Left"}-value, \emph{except} when the
+ − 818
string to be matched is not in the language of @{term "r\<^sub>1"}; only then it
+ − 819
is a @{text Right}-value (see the side-condition in @{text "P+R"}).
+ − 820
Interesting is also the rule for sequence regular expressions (@{text
+ − 821
"PS"}). The first two premises state that @{term "v\<^sub>1"} and @{term "v\<^sub>2"}
+ − 822
are the POSIX values for @{term "(s\<^sub>1, r\<^sub>1)"} and @{term "(s\<^sub>2, r\<^sub>2)"}
+ − 823
respectively. Consider now the third premise and note that the POSIX value
+ − 824
of this rule should match the string \mbox{@{term "s\<^sub>1 @ s\<^sub>2"}}. According to the
272
+ − 825
Longest Match Rule, we want that the @{term "s\<^sub>1"} is the longest initial
218
+ − 826
split of \mbox{@{term "s\<^sub>1 @ s\<^sub>2"}} such that @{term "s\<^sub>2"} is still recognised
+ − 827
by @{term "r\<^sub>2"}. Let us assume, contrary to the third premise, that there
+ − 828
\emph{exist} an @{term "s\<^sub>3"} and @{term "s\<^sub>4"} such that @{term "s\<^sub>2"}
+ − 829
can be split up into a non-empty string @{term "s\<^sub>3"} and a possibly empty
+ − 830
string @{term "s\<^sub>4"}. Moreover the longer string @{term "s\<^sub>1 @ s\<^sub>3"} can be
+ − 831
matched by @{text "r\<^sub>1"} and the shorter @{term "s\<^sub>4"} can still be
+ − 832
matched by @{term "r\<^sub>2"}. In this case @{term "s\<^sub>1"} would \emph{not} be the
+ − 833
longest initial split of \mbox{@{term "s\<^sub>1 @ s\<^sub>2"}} and therefore @{term "Seq v\<^sub>1
+ − 834
v\<^sub>2"} cannot be a POSIX value for @{term "(s\<^sub>1 @ s\<^sub>2, SEQ r\<^sub>1 r\<^sub>2)"}.
272
+ − 835
The main point is that our side-condition ensures the Longest
+ − 836
Match Rule is satisfied.
218
+ − 837
+ − 838
A similar condition is imposed on the POSIX value in the @{text
+ − 839
"P\<star>"}-rule. Also there we want that @{term "s\<^sub>1"} is the longest initial
+ − 840
split of @{term "s\<^sub>1 @ s\<^sub>2"} and furthermore the corresponding value
+ − 841
@{term v} cannot be flattened to the empty string. In effect, we require
+ − 842
that in each ``iteration'' of the star, some non-empty substring needs to
+ − 843
be ``chipped'' away; only in case of the empty string we accept @{term
273
+ − 844
"Stars []"} as the POSIX value. Indeed we can show that our POSIX values
+ − 845
are lexical values which exclude those @{text Stars} that contain subvalues
267
+ − 846
that flatten to the empty string.
218
+ − 847
272
+ − 848
\begin{lemma}\label{LVposix}
268
+ − 849
@{thm [mode=IfThen] Posix_LV}
267
+ − 850
\end{lemma}
+ − 851
+ − 852
\begin{proof}
268
+ − 853
By routine induction on @{thm (prem 1) Posix_LV}.\qed
267
+ − 854
\end{proof}
+ − 855
+ − 856
\noindent
218
+ − 857
Next is the lemma that shows the function @{term "mkeps"} calculates
+ − 858
the POSIX value for the empty string and a nullable regular expression.
+ − 859
+ − 860
\begin{lemma}\label{lemmkeps}
+ − 861
@{thm[mode=IfThen] Posix_mkeps}
+ − 862
\end{lemma}
+ − 863
+ − 864
\begin{proof}
+ − 865
By routine induction on @{term r}.\qed
+ − 866
\end{proof}
+ − 867
+ − 868
\noindent
+ − 869
The central lemma for our POSIX relation is that the @{text inj}-function
+ − 870
preserves POSIX values.
+ − 871
+ − 872
\begin{lemma}\label{Posix2}
+ − 873
@{thm[mode=IfThen] Posix_injval}
+ − 874
\end{lemma}
+ − 875
+ − 876
\begin{proof}
+ − 877
By induction on @{text r}. We explain two cases.
+ − 878
+ − 879
\begin{itemize}
+ − 880
\item[$\bullet$] Case @{term "r = ALT r\<^sub>1 r\<^sub>2"}. There are
+ − 881
two subcases, namely @{text "(a)"} \mbox{@{term "v = Left v'"}} and @{term
+ − 882
"s \<in> der c r\<^sub>1 \<rightarrow> v'"}; and @{text "(b)"} @{term "v = Right v'"}, @{term
+ − 883
"s \<notin> L (der c r\<^sub>1)"} and @{term "s \<in> der c r\<^sub>2 \<rightarrow> v'"}. In @{text "(a)"} we
+ − 884
know @{term "s \<in> der c r\<^sub>1 \<rightarrow> v'"}, from which we can infer @{term "(c # s)
+ − 885
\<in> r\<^sub>1 \<rightarrow> injval r\<^sub>1 c v'"} by induction hypothesis and hence @{term "(c #
+ − 886
s) \<in> ALT r\<^sub>1 r\<^sub>2 \<rightarrow> injval (ALT r\<^sub>1 r\<^sub>2) c (Left v')"} as needed. Similarly
+ − 887
in subcase @{text "(b)"} where, however, in addition we have to use
272
+ − 888
Proposition~\ref{derprop}(2) in order to infer @{term "c # s \<notin> L r\<^sub>1"} from @{term
+ − 889
"s \<notin> L (der c r\<^sub>1)"}.\smallskip
218
+ − 890
+ − 891
\item[$\bullet$] Case @{term "r = SEQ r\<^sub>1 r\<^sub>2"}. There are three subcases:
+ − 892
+ − 893
\begin{quote}
+ − 894
\begin{description}
+ − 895
\item[@{text "(a)"}] @{term "v = Left (Seq v\<^sub>1 v\<^sub>2)"} and @{term "nullable r\<^sub>1"}
+ − 896
\item[@{text "(b)"}] @{term "v = Right v\<^sub>1"} and @{term "nullable r\<^sub>1"}
+ − 897
\item[@{text "(c)"}] @{term "v = Seq v\<^sub>1 v\<^sub>2"} and @{term "\<not> nullable r\<^sub>1"}
+ − 898
\end{description}
+ − 899
\end{quote}
+ − 900
+ − 901
\noindent For @{text "(a)"} we know @{term "s\<^sub>1 \<in> der c r\<^sub>1 \<rightarrow> v\<^sub>1"} and
+ − 902
@{term "s\<^sub>2 \<in> r\<^sub>2 \<rightarrow> v\<^sub>2"} as well as
+ − 903
%
+ − 904
\[@{term "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s\<^sub>2 \<and> s\<^sub>1 @ s\<^sub>3 \<in> L (der c r\<^sub>1) \<and> s\<^sub>4 \<in> L r\<^sub>2)"}\]
+ − 905
272
+ − 906
\noindent From the latter we can infer by Proposition~\ref{derprop}(2):
218
+ − 907
%
+ − 908
\[@{term "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s\<^sub>2 \<and> (c # s\<^sub>1) @ s\<^sub>3 \<in> L r\<^sub>1 \<and> s\<^sub>4 \<in> L r\<^sub>2)"}\]
+ − 909
+ − 910
\noindent We can use the induction hypothesis for @{text "r\<^sub>1"} to obtain
+ − 911
@{term "(c # s\<^sub>1) \<in> r\<^sub>1 \<rightarrow> injval r\<^sub>1 c v\<^sub>1"}. Putting this all together allows us to infer
+ − 912
@{term "((c # s\<^sub>1) @ s\<^sub>2) \<in> SEQ r\<^sub>1 r\<^sub>2 \<rightarrow> Seq (injval r\<^sub>1 c v\<^sub>1) v\<^sub>2"}. The case @{text "(c)"}
+ − 913
is similar.
+ − 914
+ − 915
For @{text "(b)"} we know @{term "s \<in> der c r\<^sub>2 \<rightarrow> v\<^sub>1"} and
+ − 916
@{term "s\<^sub>1 @ s\<^sub>2 \<notin> L (SEQ (der c r\<^sub>1) r\<^sub>2)"}. From the former
+ − 917
we have @{term "(c # s) \<in> r\<^sub>2 \<rightarrow> (injval r\<^sub>2 c v\<^sub>1)"} by induction hypothesis
+ − 918
for @{term "r\<^sub>2"}. From the latter we can infer
+ − 919
%
+ − 920
\[@{term "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = c # s \<and> s\<^sub>3 \<in> L r\<^sub>1 \<and> s\<^sub>4 \<in> L r\<^sub>2)"}\]
+ − 921
272
+ − 922
\noindent By Lemma~\ref{lemmkeps} we know @{term "[] \<in> r\<^sub>1 \<rightarrow> (mkeps r\<^sub>1)"}
218
+ − 923
holds. Putting this all together, we can conclude with @{term "(c #
+ − 924
s) \<in> SEQ r\<^sub>1 r\<^sub>2 \<rightarrow> Seq (mkeps r\<^sub>1) (injval r\<^sub>2 c v\<^sub>1)"}, as required.
+ − 925
+ − 926
Finally suppose @{term "r = STAR r\<^sub>1"}. This case is very similar to the
+ − 927
sequence case, except that we need to also ensure that @{term "flat (injval r\<^sub>1
+ − 928
c v\<^sub>1) \<noteq> []"}. This follows from @{term "(c # s\<^sub>1)
+ − 929
\<in> r\<^sub>1 \<rightarrow> injval r\<^sub>1 c v\<^sub>1"} (which in turn follows from @{term "s\<^sub>1 \<in> der c
+ − 930
r\<^sub>1 \<rightarrow> v\<^sub>1"} and the induction hypothesis).\qed
+ − 931
\end{itemize}
+ − 932
\end{proof}
+ − 933
+ − 934
\noindent
272
+ − 935
With Lemma~\ref{Posix2} in place, it is completely routine to establish
218
+ − 936
that the Sulzmann and Lu lexer satisfies our specification (returning
+ − 937
the null value @{term "None"} iff the string is not in the language of the regular expression,
+ − 938
and returning a unique POSIX value iff the string \emph{is} in the language):
+ − 939
+ − 940
\begin{theorem}\mbox{}\smallskip\\\label{lexercorrect}
+ − 941
\begin{tabular}{ll}
+ − 942
(1) & @{thm (lhs) lexer_correct_None} if and only if @{thm (rhs) lexer_correct_None}\\
+ − 943
(2) & @{thm (lhs) lexer_correct_Some} if and only if @{thm (rhs) lexer_correct_Some}\\
+ − 944
\end{tabular}
+ − 945
\end{theorem}
+ − 946
+ − 947
\begin{proof}
272
+ − 948
By induction on @{term s} using Lemma~\ref{lemmkeps} and \ref{Posix2}.\qed
218
+ − 949
\end{proof}
+ − 950
273
+ − 951
\noindent In \textit{(2)} we further know by Theorem~\ref{posixdeterm} that the
218
+ − 952
value returned by the lexer must be unique. A simple corollary
+ − 953
of our two theorems is:
+ − 954
+ − 955
\begin{corollary}\mbox{}\smallskip\\\label{lexercorrectcor}
+ − 956
\begin{tabular}{ll}
+ − 957
(1) & @{thm (lhs) lexer_correctness(2)} if and only if @{thm (rhs) lexer_correctness(2)}\\
+ − 958
(2) & @{thm (lhs) lexer_correctness(1)} if and only if @{thm (rhs) lexer_correctness(1)}\\
+ − 959
\end{tabular}
+ − 960
\end{corollary}
+ − 961
272
+ − 962
\noindent This concludes our correctness proof. Note that we have
+ − 963
not changed the algorithm of Sulzmann and Lu,\footnote{All
+ − 964
deviations we introduced are harmless.} but introduced our own
+ − 965
specification for what a correct result---a POSIX value---should be.
+ − 966
In the next section we show that our specification coincides with
+ − 967
another one given by Okui and Suzuki using a different technique.
218
+ − 968
+ − 969
*}
+ − 970
268
+ − 971
section {* Ordering of Values according to Okui and Suzuki*}
+ − 972
+ − 973
text {*
+ − 974
+ − 975
While in the previous section we have defined POSIX values directly
+ − 976
in terms of a ternary relation (see inference rules in Figure~\ref{POSIXrules}),
+ − 977
Sulzmann and Lu took a different approach in \cite{Sulzmann2014}:
+ − 978
they introduced an ordering for values and identified POSIX values
+ − 979
as the maximal elements. An extended version of \cite{Sulzmann2014}
+ − 980
is available at the website of its first author; this includes more
+ − 981
details of their proofs, but which are evidently not in final form
+ − 982
yet. Unfortunately, we were not able to verify claims that their
+ − 983
ordering has properties such as being transitive or having maximal
273
+ − 984
elements.
268
+ − 985
+ − 986
Okui and Suzuki \cite{OkuiSuzuki2010,OkuiSuzukiTech} described
269
+ − 987
another ordering of values, which they use to establish the
+ − 988
correctness of their automata-based algorithm for POSIX matching.
+ − 989
Their ordering resembles some aspects of the one given by Sulzmann
273
+ − 990
and Lu, but overall is quite different. To begin with, Okui and
+ − 991
Suzuki identify POSIX values as minimal, rather than maximal,
+ − 992
elements in their ordering. A more substantial difference is that
+ − 993
the ordering by Okui and Suzuki uses \emph{positions} in order to
+ − 994
identify and compare subvalues. Positions are lists of natural
+ − 995
numbers. This allows them to quite naturally formalise the Longest
+ − 996
Match and Priority rules of the informal POSIX standard. Consider
+ − 997
for example the value @{term v}
269
+ − 998
+ − 999
\begin{center}
+ − 1000
@{term "v == Stars [Seq (Char x) (Char y), Char z]"}
+ − 1001
\end{center}
+ − 1002
+ − 1003
\noindent
+ − 1004
At position @{text "[0,1]"} of this value is the
+ − 1005
subvalue @{text "Char y"} and at position @{text "[1]"} the
+ − 1006
subvalue @{term "Char z"}. At the `root' position, or empty list
273
+ − 1007
@{term "[]"}, is the whole value @{term v}. Positions such as @{text
+ − 1008
"[0,1,0]"} or @{text "[2]"} are outside of @{text
269
+ − 1009
v}. If it exists, the subvalue of @{term v} at a position @{text
+ − 1010
p}, written @{term "at v p"}, can be recursively defined by
268
+ − 1011
+ − 1012
\begin{center}
+ − 1013
\begin{tabular}{r@ {\hspace{0mm}}lcl}
+ − 1014
@{term v} & @{text "\<downharpoonleft>\<^bsub>[]\<^esub>"} & @{text "\<equiv>"}& @{thm (rhs) at.simps(1)}\\
+ − 1015
@{term "Left v"} & @{text "\<downharpoonleft>\<^bsub>0::ps\<^esub>"} & @{text "\<equiv>"}& @{thm (rhs) at.simps(2)}\\
+ − 1016
@{term "Right v"} & @{text "\<downharpoonleft>\<^bsub>1::ps\<^esub>"} & @{text "\<equiv>"} &
+ − 1017
@{thm (rhs) at.simps(3)[simplified Suc_0_fold]}\\
+ − 1018
@{term "Seq v\<^sub>1 v\<^sub>2"} & @{text "\<downharpoonleft>\<^bsub>0::ps\<^esub>"} & @{text "\<equiv>"} &
+ − 1019
@{thm (rhs) at.simps(4)[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]} \\
+ − 1020
@{term "Seq v\<^sub>1 v\<^sub>2"} & @{text "\<downharpoonleft>\<^bsub>1::ps\<^esub>"}
+ − 1021
& @{text "\<equiv>"} &
+ − 1022
@{thm (rhs) at.simps(5)[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2", simplified Suc_0_fold]} \\
+ − 1023
@{term "Stars vs"} & @{text "\<downharpoonleft>\<^bsub>n::ps\<^esub>"} & @{text "\<equiv>"}& @{thm (rhs) at.simps(6)}\\
+ − 1024
\end{tabular}
+ − 1025
\end{center}
+ − 1026
269
+ − 1027
\noindent In the last clause we use Isabelle's notation @{term "vs ! n"} for the
268
+ − 1028
@{text n}th element in a list. The set of positions inside a value @{text v},
273
+ − 1029
written @{term "Pos v"}, is given by
268
+ − 1030
+ − 1031
\begin{center}
+ − 1032
\begin{tabular}{lcl}
+ − 1033
@{thm (lhs) Pos.simps(1)} & @{text "\<equiv>"} & @{thm (rhs) Pos.simps(1)}\\
+ − 1034
@{thm (lhs) Pos.simps(2)} & @{text "\<equiv>"} & @{thm (rhs) Pos.simps(2)}\\
+ − 1035
@{thm (lhs) Pos.simps(3)} & @{text "\<equiv>"} & @{thm (rhs) Pos.simps(3)}\\
+ − 1036
@{thm (lhs) Pos.simps(4)} & @{text "\<equiv>"} & @{thm (rhs) Pos.simps(4)}\\
+ − 1037
@{thm (lhs) Pos.simps(5)[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}
+ − 1038
& @{text "\<equiv>"}
+ − 1039
& @{thm (rhs) Pos.simps(5)[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}\\
+ − 1040
@{thm (lhs) Pos_stars} & @{text "\<equiv>"} & @{thm (rhs) Pos_stars}\\
+ − 1041
\end{tabular}
+ − 1042
\end{center}
+ − 1043
+ − 1044
\noindent
273
+ − 1045
whereby @{text len} in the last clause stands for the length of a list. Clearly
268
+ − 1046
for every position inside a value there exists a subvalue at that position.
+ − 1047
+ − 1048
+ − 1049
To help understanding the ordering of Okui and Suzuki, consider again
+ − 1050
the earlier value
+ − 1051
@{text v} and compare it with the following @{text w}:
+ − 1052
+ − 1053
\begin{center}
+ − 1054
\begin{tabular}{l}
+ − 1055
@{term "v == Stars [Seq (Char x) (Char y), Char z]"}\\
+ − 1056
@{term "w == Stars [Char x, Char y, Char z]"}
+ − 1057
\end{tabular}
+ − 1058
\end{center}
+ − 1059
+ − 1060
\noindent Both values match the string @{text "xyz"}, that means if
273
+ − 1061
we flatten these values at their respective root position, we obtain
268
+ − 1062
@{text "xyz"}. However, at position @{text "[0]"}, @{text v} matches
+ − 1063
@{text xy} whereas @{text w} matches only the shorter @{text x}. So
+ − 1064
according to the Longest Match Rule, we should prefer @{text v},
+ − 1065
rather than @{text w} as POSIX value for string @{text xyz} (and
+ − 1066
corresponding regular expression). In order to
+ − 1067
formalise this idea, Okui and Suzuki introduce a measure for
+ − 1068
subvalues at position @{text p}, called the \emph{norm} of @{text v}
+ − 1069
at position @{text p}. We can define this measure in Isabelle as an
+ − 1070
integer as follows
+ − 1071
+ − 1072
\begin{center}
+ − 1073
@{thm pflat_len_def}
+ − 1074
\end{center}
+ − 1075
+ − 1076
\noindent where we take the length of the flattened value at
+ − 1077
position @{text p}, provided the position is inside @{text v}; if
272
+ − 1078
not, then the norm is @{text "-1"}. The default for outside
+ − 1079
positions is crucial for the POSIX requirement of preferring a
+ − 1080
@{text Left}-value over a @{text Right}-value (if they can match the
+ − 1081
same string---see the Priority Rule from the Introduction). For this
+ − 1082
consider
268
+ − 1083
+ − 1084
\begin{center}
+ − 1085
@{term "v == Left (Char x)"} \qquad and \qquad @{term "w == Right (Char x)"}
+ − 1086
\end{center}
+ − 1087
273
+ − 1088
\noindent Both values match @{text x}. At position @{text "[0]"}
272
+ − 1089
the norm of @{term v} is @{text 1} (the subvalue matches @{text x}),
+ − 1090
but the norm of @{text w} is @{text "-1"} (the position is outside
+ − 1091
@{text w} according to how we defined the `inside' positions of
+ − 1092
@{text Left}- and @{text Right}-values). Of course at position
+ − 1093
@{text "[1]"}, the norms @{term "pflat_len v [1]"} and @{term
+ − 1094
"pflat_len w [1]"} are reversed, but the point is that subvalues
+ − 1095
will be analysed according to lexicographically ordered
+ − 1096
positions. According to this ordering, the position @{text "[0]"}
273
+ − 1097
takes precedence over @{text "[1]"} and thus also @{text v} will be
+ − 1098
preferred over @{text w}. The lexicographic ordering of positions, written
272
+ − 1099
@{term "DUMMY \<sqsubset>lex DUMMY"}, can be conveniently formalised
+ − 1100
by three inference rules
268
+ − 1101
+ − 1102
\begin{center}
+ − 1103
\begin{tabular}{ccc}
+ − 1104
@{thm [mode=Axiom] lex_list.intros(1)}\hspace{1cm} &
+ − 1105
@{thm [mode=Rule] lex_list.intros(3)[where ?p1.0="p\<^sub>1" and ?p2.0="p\<^sub>2" and
+ − 1106
?ps1.0="ps\<^sub>1" and ?ps2.0="ps\<^sub>2"]}\hspace{1cm} &
+ − 1107
@{thm [mode=Rule] lex_list.intros(2)[where ?ps1.0="ps\<^sub>1" and ?ps2.0="ps\<^sub>2"]}
+ − 1108
\end{tabular}
+ − 1109
\end{center}
+ − 1110
272
+ − 1111
With the norm and lexicographic order in place,
268
+ − 1112
we can state the key definition of Okui and Suzuki
273
+ − 1113
\cite{OkuiSuzuki2010}: a value @{term "v\<^sub>1"} is \emph{smaller at position @{text p}} than
+ − 1114
@{term "v\<^sub>2"}, written @{term "v\<^sub>1 \<sqsubset>val p v\<^sub>2"},
+ − 1115
if and only if $(i)$ the norm at position @{text p} is
268
+ − 1116
greater in @{term "v\<^sub>1"} (that is the string @{term "flat (at v\<^sub>1 p)"} is longer
+ − 1117
than @{term "flat (at v\<^sub>2 p)"}) and $(ii)$ all subvalues at
+ − 1118
positions that are inside @{term "v\<^sub>1"} or @{term "v\<^sub>2"} and that are
+ − 1119
lexicographically smaller than @{text p}, we have the same norm, namely
+ − 1120
+ − 1121
\begin{center}
+ − 1122
\begin{tabular}{c}
+ − 1123
@{thm (lhs) PosOrd_def[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}
+ − 1124
@{text "\<equiv>"}
+ − 1125
$\begin{cases}
+ − 1126
(i) & @{term "pflat_len v\<^sub>1 p > pflat_len v\<^sub>2 p"} \quad\text{and}\smallskip \\
+ − 1127
(ii) & @{term "(\<forall>q \<in> Pos v\<^sub>1 \<union> Pos v\<^sub>2. q \<sqsubset>lex p --> pflat_len v\<^sub>1 q = pflat_len v\<^sub>2 q)"}
+ − 1128
\end{cases}$
+ − 1129
\end{tabular}
+ − 1130
\end{center}
+ − 1131
+ − 1132
\noindent The position @{text p} in this definition acts as the
+ − 1133
\emph{first distinct position} of @{text "v\<^sub>1"} and @{text
+ − 1134
"v\<^sub>2"}, where both values match strings of different length
+ − 1135
\cite{OkuiSuzuki2010}. Since at @{text p} the values @{text
272
+ − 1136
"v\<^sub>1"} and @{text "v\<^sub>2"} match different strings, the
268
+ − 1137
ordering is irreflexive. Derived from the definition above
+ − 1138
are the following two orderings:
+ − 1139
+ − 1140
\begin{center}
+ − 1141
\begin{tabular}{l}
+ − 1142
@{thm PosOrd_ex_def[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}\\
+ − 1143
@{thm PosOrd_ex_eq_def[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}
+ − 1144
\end{tabular}
+ − 1145
\end{center}
+ − 1146
272
+ − 1147
While we encountered a number of obstacles for establishing properties like
268
+ − 1148
transitivity for the ordering of Sulzmann and Lu (and which we failed
+ − 1149
to overcome), it is relatively straightforward to establish this
273
+ − 1150
property for the orderings
+ − 1151
@{term "DUMMY :\<sqsubset>val DUMMY"} and @{term "DUMMY :\<sqsubseteq>val DUMMY"}
+ − 1152
by Okui and Suzuki.
268
+ − 1153
+ − 1154
\begin{lemma}[Transitivity]\label{transitivity}
+ − 1155
@{thm [mode=IfThen] PosOrd_trans[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2" and ?v3.0="v\<^sub>3"]}
+ − 1156
\end{lemma}
+ − 1157
+ − 1158
\begin{proof} From the assumption we obtain two positions @{text p}
+ − 1159
and @{text q}, where the values @{text "v\<^sub>1"} and @{text
+ − 1160
"v\<^sub>2"} (respectively @{text "v\<^sub>2"} and @{text
+ − 1161
"v\<^sub>3"}) are `distinct'. Since @{text
+ − 1162
"\<prec>\<^bsub>lex\<^esub>"} is trichotomous, we need to consider
+ − 1163
three cases, namely @{term "p = q"}, @{term "p \<sqsubset>lex q"} and
273
+ − 1164
@{term "q \<sqsubset>lex p"}. Let us look at the first case. Clearly
+ − 1165
@{term "pflat_len v\<^sub>2 p < pflat_len v\<^sub>1 p"} and @{term
+ − 1166
"pflat_len v\<^sub>3 p < pflat_len v\<^sub>2 p"} imply @{term
+ − 1167
"pflat_len v\<^sub>3 p < pflat_len v\<^sub>1 p"}. It remains to show
+ − 1168
that for a @{term "p' \<in> Pos v\<^sub>1 \<union> Pos v\<^sub>3"}
+ − 1169
with @{term "p' \<sqsubset>lex p"} that @{term "pflat_len v\<^sub>1
+ − 1170
p' = pflat_len v\<^sub>3 p'"} holds. Suppose @{term "p' \<in> Pos
+ − 1171
v\<^sub>1"}, then we can infer from the first assumption that @{term
+ − 1172
"pflat_len v\<^sub>1 p' = pflat_len v\<^sub>2 p'"}. But this means
+ − 1173
that @{term "p'"} must be in @{term "Pos v\<^sub>2"} too (the norm
+ − 1174
cannot be @{text "-1"} given @{term "p' \<in> Pos v\<^sub>1"}).
+ − 1175
Hence we can use the second assumption and
+ − 1176
infer @{term "pflat_len v\<^sub>2 p' = pflat_len v\<^sub>3 p'"},
+ − 1177
which concludes this case with @{term "v\<^sub>1 :\<sqsubset>val
+ − 1178
v\<^sub>3"}. The reasoning in the other cases is similar.\qed
268
+ − 1179
\end{proof}
+ − 1180
273
+ − 1181
\noindent
+ − 1182
The proof for $\preccurlyeq$ is similar and omitted.
+ − 1183
It is also straightforward to show that @{text "\<prec>"} and
+ − 1184
$\preccurlyeq$ are partial orders. Okui and Suzuki furthermore show that they
+ − 1185
are linear orderings for lexical values \cite{OkuiSuzuki2010} of a given
+ − 1186
regular expression and given string, but we have not formalised this in Isabelle. It is
272
+ − 1187
not essential for our results. What we are going to show below is
273
+ − 1188
that for a given @{text r} and @{text s}, the orderings have a unique
269
+ − 1189
minimal element on the set @{term "LV r s"}, which is the POSIX value
273
+ − 1190
we defined in the previous section. We start with two properties that
+ − 1191
show how the length of a flattened value relates to the @{text "\<prec>"}-ordering.
+ − 1192
+ − 1193
\begin{proposition}\mbox{}\smallskip\\\label{ordlen}
+ − 1194
\begin{tabular}{@ {}ll}
+ − 1195
(1) &
+ − 1196
@{thm [mode=IfThen] PosOrd_shorterE[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}\\
+ − 1197
(2) &
+ − 1198
@{thm [mode=IfThen] PosOrd_shorterI[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}
+ − 1199
\end{tabular}
+ − 1200
\end{proposition}
+ − 1201
+ − 1202
\noindent Both properties follow from the definition of the ordering. Note that
+ − 1203
\textit{(2)} entails that a value, say @{term "v\<^sub>2"}, whose underlying
+ − 1204
string is a strict prefix of another flattened value, say @{term "v\<^sub>1"}, then
+ − 1205
@{term "v\<^sub>1"} must be smaller than @{term "v\<^sub>2"}. For our proofs it
+ − 1206
will be useful to have the following properties---in each case the underlying strings
+ − 1207
of the compared values are the same:
268
+ − 1208
273
+ − 1209
\begin{proposition}\mbox{}\smallskip\\\label{ordintros}
+ − 1210
\begin{tabular}{ll}
+ − 1211
\textit{(1)} &
+ − 1212
@{thm [mode=IfThen] PosOrd_Left_Right[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}\\
+ − 1213
\textit{(2)} & If
+ − 1214
@{thm (prem 1) PosOrd_Left_eq[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]} \;then\;
+ − 1215
@{thm (lhs) PosOrd_Left_eq[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]} \;iff\;
+ − 1216
@{thm (rhs) PosOrd_Left_eq[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}\\
+ − 1217
\textit{(3)} & If
+ − 1218
@{thm (prem 1) PosOrd_Right_eq[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]} \;then\;
+ − 1219
@{thm (lhs) PosOrd_Right_eq[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]} \;iff\;
+ − 1220
@{thm (rhs) PosOrd_Right_eq[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}\\
+ − 1221
\textit{(4)} & If
+ − 1222
@{thm (prem 1) PosOrd_Seq_eq[where ?v2.0="v\<^sub>2" and ?w2.0="w\<^sub>2"]} \;then\;
+ − 1223
@{thm (lhs) PosOrd_Seq_eq[where ?v2.0="v\<^sub>2" and ?w2.0="w\<^sub>2"]} \;iff\;
+ − 1224
@{thm (rhs) PosOrd_Seq_eq[where ?v2.0="v\<^sub>2" and ?w2.0="w\<^sub>2"]}\\
+ − 1225
\textit{(5)} & If
+ − 1226
@{thm (prem 2) PosOrd_SeqI1[simplified, where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2" and
+ − 1227
?w1.0="w\<^sub>1" and ?w2.0="w\<^sub>2"]} \;and\;
+ − 1228
@{thm (prem 1) PosOrd_SeqI1[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2" and
+ − 1229
?w1.0="w\<^sub>1" and ?w2.0="w\<^sub>2"]} \;then\;
+ − 1230
@{thm (concl) PosOrd_SeqI1[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2" and
+ − 1231
?w1.0="w\<^sub>1" and ?w2.0="w\<^sub>2"]}\\
+ − 1232
\textit{(6)} & If
+ − 1233
@{thm (prem 1) PosOrd_Stars_append_eq[where ?vs1.0="vs\<^sub>1" and ?vs2.0="vs\<^sub>2"]} \;then\;
+ − 1234
@{thm (lhs) PosOrd_Stars_append_eq[where ?vs1.0="vs\<^sub>1" and ?vs2.0="vs\<^sub>2"]} \;iff\;
+ − 1235
@{thm (rhs) PosOrd_Stars_append_eq[where ?vs1.0="vs\<^sub>1" and ?vs2.0="vs\<^sub>2"]}\\
+ − 1236
+ − 1237
\textit{(7)} & If
+ − 1238
@{thm (prem 2) PosOrd_StarsI[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2" and
+ − 1239
?vs1.0="vs\<^sub>1" and ?vs2.0="vs\<^sub>2"]} \;and\;
+ − 1240
@{thm (prem 1) PosOrd_StarsI[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2" and
+ − 1241
?vs1.0="vs\<^sub>1" and ?vs2.0="vs\<^sub>2"]} \;then\;
+ − 1242
@{thm (concl) PosOrd_StarsI[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2" and
+ − 1243
?vs1.0="vs\<^sub>1" and ?vs2.0="vs\<^sub>2"]}\\
+ − 1244
\end{tabular}
+ − 1245
\end{proposition}
268
+ − 1246
273
+ − 1247
\noindent One might prefer that statements \textit{(4)} and \textit{(5)}
+ − 1248
(respectively \textit{(6)} and \textit{(7)})
+ − 1249
are combined into a single \textit{iff}-statement (like the ones for @{text
+ − 1250
Left} and @{text Right}). Unfortunately this cannot be done easily: such
+ − 1251
a single statement would require an additional assumption about the
+ − 1252
two values @{term "Seq v\<^sub>1 v\<^sub>2"} and @{term "Seq w\<^sub>1 w\<^sub>2"}
+ − 1253
being inhabited by the same regular expression. The
+ − 1254
complexity of the proofs involved seems to not justify such a
+ − 1255
`cleaner' single statement. The statements given are just the properties that
275
+ − 1256
allow us to establish our theorems without any difficulty. The proofs
+ − 1257
for Proposition~\ref{ordintros} are routine.
268
+ − 1258
+ − 1259
273
+ − 1260
Next we establish how Okui and Suzuki's orderings relate to our
272
+ − 1261
definition of POSIX values. Given a @{text POSIX} value @{text "v\<^sub>1"}
268
+ − 1262
for @{text r} and @{text s}, then any other lexical value @{text
+ − 1263
"v\<^sub>2"} in @{term "LV r s"} is greater or equal than @{text
272
+ − 1264
"v\<^sub>1"}, namely:
268
+ − 1265
+ − 1266
272
+ − 1267
\begin{theorem}\label{orderone}
268
+ − 1268
@{thm [mode=IfThen] Posix_PosOrd[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}
+ − 1269
\end{theorem}
+ − 1270
270
+ − 1271
\begin{proof} By induction on our POSIX rules. By
272
+ − 1272
Theorem~\ref{posixdeterm} and the definition of @{const LV}, it is clear
270
+ − 1273
that @{text "v\<^sub>1"} and @{text "v\<^sub>2"} have the same
+ − 1274
underlying string @{term s}. The three base cases are
+ − 1275
straightforward: for example for @{term "v\<^sub>1 = Void"}, we have
+ − 1276
that @{term "v\<^sub>2 \<in> LV ONE []"} must also be of the form
+ − 1277
\mbox{@{term "v\<^sub>2 = Void"}}. Therefore we have @{term
+ − 1278
"v\<^sub>1 :\<sqsubseteq>val v\<^sub>2"}. The inductive cases for
272
+ − 1279
@{text r} being of the form @{term "ALT r\<^sub>1 r\<^sub>2"} and
+ − 1280
@{term "SEQ r\<^sub>1 r\<^sub>2"} are as follows:
269
+ − 1281
270
+ − 1282
+ − 1283
\begin{itemize}
268
+ − 1284
272
+ − 1285
\item[$\bullet$] Case @{text "P+L"} with @{term "s \<in> (ALT r\<^sub>1 r\<^sub>2)
273
+ − 1286
\<rightarrow> (Left w\<^sub>1)"}: In this case the value
+ − 1287
@{term "v\<^sub>2"} is either of the
270
+ − 1288
form @{term "Left w\<^sub>2"} or @{term "Right w\<^sub>2"}. In the
273
+ − 1289
latter case we can immediately conclude with \mbox{@{term "v\<^sub>1
+ − 1290
:\<sqsubseteq>val v\<^sub>2"}} since a @{text Left}-value with the
272
+ − 1291
same underlying string @{text s} is always smaller than a
273
+ − 1292
@{text Right}-value by Proposition~\ref{ordintros}\textit{(1)}.
+ − 1293
In the former case we have @{term "w\<^sub>2
270
+ − 1294
\<in> LV r\<^sub>1 s"} and can use the induction hypothesis to infer
+ − 1295
@{term "w\<^sub>1 :\<sqsubseteq>val w\<^sub>2"}. Because @{term
+ − 1296
"w\<^sub>1"} and @{term "w\<^sub>2"} have the same underlying string
+ − 1297
@{text s}, we can conclude with @{term "Left w\<^sub>1
273
+ − 1298
:\<sqsubseteq>val Left w\<^sub>2"} using
+ − 1299
Proposition~\ref{ordintros}\textit{(2)}.\smallskip
268
+ − 1300
272
+ − 1301
\item[$\bullet$] Case @{text "P+R"} with @{term "s \<in> (ALT r\<^sub>1 r\<^sub>2)
+ − 1302
\<rightarrow> (Right w\<^sub>1)"}: This case similar to the previous
+ − 1303
case, except that we additionally know @{term "s \<notin> L
270
+ − 1304
r\<^sub>1"}. This is needed when @{term "v\<^sub>2"} is of the form
273
+ − 1305
\mbox{@{term "Left w\<^sub>2"}}. Since \mbox{@{term "flat v\<^sub>2 = flat
270
+ − 1306
w\<^sub>2"} @{text "= s"}} and @{term "\<Turnstile> w\<^sub>2 :
272
+ − 1307
r\<^sub>1"}, we can derive a contradiction for \mbox{@{term "s \<notin> L
+ − 1308
r\<^sub>1"}} using
+ − 1309
Proposition~\ref{inhabs}. So also in this case \mbox{@{term "v\<^sub>1
270
+ − 1310
:\<sqsubseteq>val v\<^sub>2"}}.\smallskip
268
+ − 1311
273
+ − 1312
\item[$\bullet$] Case @{text "PS"} with @{term "(s\<^sub>1 @
+ − 1313
s\<^sub>2) \<in> (SEQ r\<^sub>1 r\<^sub>2) \<rightarrow> (Seq
+ − 1314
w\<^sub>1 w\<^sub>2)"}: We can assume @{term "v\<^sub>2 = Seq
+ − 1315
(u\<^sub>1) (u\<^sub>2)"} with @{term "\<Turnstile> u\<^sub>1 :
+ − 1316
r\<^sub>1"} and \mbox{@{term "\<Turnstile> u\<^sub>2 :
+ − 1317
r\<^sub>2"}}. We have @{term "s\<^sub>1 @ s\<^sub>2 = (flat
+ − 1318
u\<^sub>1) @ (flat u\<^sub>2)"}. By the side-condition of the
+ − 1319
@{text PS}-rule we know that either @{term "s\<^sub>1 = flat
+ − 1320
u\<^sub>1"} or that @{term "flat u\<^sub>1"} is a strict prefix of
+ − 1321
@{term "s\<^sub>1"}. In the latter case we can infer @{term
+ − 1322
"w\<^sub>1 :\<sqsubset>val u\<^sub>1"} by
+ − 1323
Proposition~\ref{ordlen}\textit{(2)} and from this @{term "v\<^sub>1
+ − 1324
:\<sqsubseteq>val v\<^sub>2"} by Proposition~\ref{ordintros}\textit{(5)}
+ − 1325
(as noted above @{term "v\<^sub>1"} and @{term "v\<^sub>2"} must have the
+ − 1326
same underlying string).
+ − 1327
In the former case we know
+ − 1328
@{term "u\<^sub>1 \<in> LV r\<^sub>1 s\<^sub>1"} and @{term
+ − 1329
"u\<^sub>2 \<in> LV r\<^sub>2 s\<^sub>2"}. With this we can use the
+ − 1330
induction hypotheses to infer @{term "w\<^sub>1 :\<sqsubseteq>val
+ − 1331
u\<^sub>1"} and @{term "w\<^sub>2 :\<sqsubseteq>val u\<^sub>2"}. By
+ − 1332
Proposition~\ref{ordintros}\textit{(4,5)} we can again infer
+ − 1333
@{term "v\<^sub>1 :\<sqsubseteq>val
+ − 1334
v\<^sub>2"}.
270
+ − 1335
268
+ − 1336
\end{itemize}
270
+ − 1337
272
+ − 1338
\noindent The case for @{text "P\<star>"} is similar to the @{text PS}-case and omitted.\qed
268
+ − 1339
\end{proof}
+ − 1340
272
+ − 1341
\noindent This theorem shows that our @{text POSIX} value for a
+ − 1342
regular expression @{text r} and string @{term s} is in fact a
273
+ − 1343
minimal element of the values in @{text "LV r s"}. By
+ − 1344
Proposition~\ref{ordlen}\textit{(2)} we also know that any value in
+ − 1345
@{text "LV s' r"}, with @{term "s'"} being a strict prefix, cannot be
+ − 1346
smaller than @{text "v\<^sub>1"}. The next theorem shows the
+ − 1347
opposite---namely any minimal element in @{term "LV r s"} must be a
+ − 1348
@{text POSIX} value. This can be established by induction on @{text
+ − 1349
r}, but the proof can be drastically simplified by using the fact
+ − 1350
from the previous section about the existence of a @{text POSIX} value
+ − 1351
whenever a string @{term "s \<in> L r"}.
+ − 1352
268
+ − 1353
+ − 1354
\begin{theorem}
272
+ − 1355
@{thm [mode=IfThen] PosOrd_Posix[where ?v1.0="v\<^sub>1"]}
268
+ − 1356
\end{theorem}
+ − 1357
272
+ − 1358
\begin{proof}
+ − 1359
If @{thm (prem 1) PosOrd_Posix[where ?v1.0="v\<^sub>1"]} then
+ − 1360
@{term "s \<in> L r"} by Proposition~\ref{inhabs}. Hence by Theorem~\ref{lexercorrect}(2)
+ − 1361
there exists a
+ − 1362
@{text POSIX} value @{term "v\<^sub>P"} with @{term "s \<in> r \<rightarrow> v\<^sub>P"}
273
+ − 1363
and by Lemma~\ref{LVposix} we also have \mbox{@{term "v\<^sub>P \<in> LV r s"}}.
272
+ − 1364
By Theorem~\ref{orderone} we therefore have
+ − 1365
@{term "v\<^sub>P :\<sqsubseteq>val v\<^sub>1"}. If @{term "v\<^sub>P = v\<^sub>1"} then
273
+ − 1366
we are done. Otherwise we have @{term "v\<^sub>P :\<sqsubset>val v\<^sub>1"}, which
+ − 1367
however contradicts the second assumption about @{term "v\<^sub>1"} being the smallest
+ − 1368
element in @{term "LV r s"}. So we are done in this case too.\qed
272
+ − 1369
\end{proof}
270
+ − 1370
273
+ − 1371
\noindent
+ − 1372
From this we can also show
+ − 1373
that if @{term "LV r s"} is non-empty (or equivalently @{term "s \<in> L r"}) then
+ − 1374
it has a unique minimal element:
+ − 1375
272
+ − 1376
\begin{corollary}
+ − 1377
@{thm [mode=IfThen] Least_existence1}
+ − 1378
\end{corollary}
270
+ − 1379
+ − 1380
+ − 1381
273
+ − 1382
\noindent To sum up, we have shown that the (unique) minimal elements
+ − 1383
of the ordering by Okui and Suzuki are exactly the @{text POSIX}
+ − 1384
values we defined inductively in Section~\ref{posixsec}. This provides
+ − 1385
an independent confirmation that our ternary relation formalise the
+ − 1386
informal POSIX rules.
+ − 1387
268
+ − 1388
*}
+ − 1389
272
+ − 1390
section {* Optimisations *}
218
+ − 1391
+ − 1392
text {*
+ − 1393
+ − 1394
Derivatives as calculated by \Brz's method are usually more complex
+ − 1395
regular expressions than the initial one; the result is that the
+ − 1396
derivative-based matching and lexing algorithms are often abysmally slow.
+ − 1397
However, various optimisations are possible, such as the simplifications
+ − 1398
of @{term "ALT ZERO r"}, @{term "ALT r ZERO"}, @{term "SEQ ONE r"} and
+ − 1399
@{term "SEQ r ONE"} to @{term r}. These simplifications can speed up the
+ − 1400
algorithms considerably, as noted in \cite{Sulzmann2014}. One of the
+ − 1401
advantages of having a simple specification and correctness proof is that
+ − 1402
the latter can be refined to prove the correctness of such simplification
+ − 1403
steps. While the simplification of regular expressions according to
+ − 1404
rules like
+ − 1405
+ − 1406
\begin{equation}\label{Simpl}
+ − 1407
\begin{array}{lcllcllcllcl}
+ − 1408
@{term "ALT ZERO r"} & @{text "\<Rightarrow>"} & @{term r} \hspace{8mm}%\\
+ − 1409
@{term "ALT r ZERO"} & @{text "\<Rightarrow>"} & @{term r} \hspace{8mm}%\\
+ − 1410
@{term "SEQ ONE r"} & @{text "\<Rightarrow>"} & @{term r} \hspace{8mm}%\\
+ − 1411
@{term "SEQ r ONE"} & @{text "\<Rightarrow>"} & @{term r}
+ − 1412
\end{array}
+ − 1413
\end{equation}
+ − 1414
+ − 1415
\noindent is well understood, there is an obstacle with the POSIX value
+ − 1416
calculation algorithm by Sulzmann and Lu: if we build a derivative regular
+ − 1417
expression and then simplify it, we will calculate a POSIX value for this
+ − 1418
simplified derivative regular expression, \emph{not} for the original (unsimplified)
+ − 1419
derivative regular expression. Sulzmann and Lu \cite{Sulzmann2014} overcome this obstacle by
+ − 1420
not just calculating a simplified regular expression, but also calculating
+ − 1421
a \emph{rectification function} that ``repairs'' the incorrect value.
+ − 1422
+ − 1423
The rectification functions can be (slightly clumsily) implemented in
+ − 1424
Isabelle/HOL as follows using some auxiliary functions:
+ − 1425
+ − 1426
\begin{center}
+ − 1427
\begin{tabular}{lcl}
+ − 1428
@{thm (lhs) F_RIGHT.simps(1)} & $\dn$ & @{text "Right (f v)"}\\
+ − 1429
@{thm (lhs) F_LEFT.simps(1)} & $\dn$ & @{text "Left (f v)"}\\
+ − 1430
+ − 1431
@{thm (lhs) F_ALT.simps(1)} & $\dn$ & @{text "Right (f\<^sub>2 v)"}\\
+ − 1432
@{thm (lhs) F_ALT.simps(2)} & $\dn$ & @{text "Left (f\<^sub>1 v)"}\\
+ − 1433
+ − 1434
@{thm (lhs) F_SEQ1.simps(1)} & $\dn$ & @{text "Seq (f\<^sub>1 ()) (f\<^sub>2 v)"}\\
+ − 1435
@{thm (lhs) F_SEQ2.simps(1)} & $\dn$ & @{text "Seq (f\<^sub>1 v) (f\<^sub>2 ())"}\\
+ − 1436
@{thm (lhs) F_SEQ.simps(1)} & $\dn$ & @{text "Seq (f\<^sub>1 v\<^sub>1) (f\<^sub>2 v\<^sub>2)"}\medskip\\
+ − 1437
%\end{tabular}
+ − 1438
%
+ − 1439
%\begin{tabular}{lcl}
+ − 1440
@{term "simp_ALT (ZERO, DUMMY) (r\<^sub>2, f\<^sub>2)"} & $\dn$ & @{term "(r\<^sub>2, F_RIGHT f\<^sub>2)"}\\
+ − 1441
@{term "simp_ALT (r\<^sub>1, f\<^sub>1) (ZERO, DUMMY)"} & $\dn$ & @{term "(r\<^sub>1, F_LEFT f\<^sub>1)"}\\
+ − 1442
@{term "simp_ALT (r\<^sub>1, f\<^sub>1) (r\<^sub>2, f\<^sub>2)"} & $\dn$ & @{term "(ALT r\<^sub>1 r\<^sub>2, F_ALT f\<^sub>1 f\<^sub>2)"}\\
+ − 1443
@{term "simp_SEQ (ONE, f\<^sub>1) (r\<^sub>2, f\<^sub>2)"} & $\dn$ & @{term "(r\<^sub>2, F_SEQ1 f\<^sub>1 f\<^sub>2)"}\\
+ − 1444
@{term "simp_SEQ (r\<^sub>1, f\<^sub>1) (ONE, f\<^sub>2)"} & $\dn$ & @{term "(r\<^sub>1, F_SEQ2 f\<^sub>1 f\<^sub>2)"}\\
+ − 1445
@{term "simp_SEQ (r\<^sub>1, f\<^sub>1) (r\<^sub>2, f\<^sub>2)"} & $\dn$ & @{term "(SEQ r\<^sub>1 r\<^sub>2, F_SEQ f\<^sub>1 f\<^sub>2)"}\\
+ − 1446
\end{tabular}
+ − 1447
\end{center}
+ − 1448
+ − 1449
\noindent
+ − 1450
The functions @{text "simp\<^bsub>Alt\<^esub>"} and @{text "simp\<^bsub>Seq\<^esub>"} encode the simplification rules
+ − 1451
in \eqref{Simpl} and compose the rectification functions (simplifications can occur
+ − 1452
deep inside the regular expression). The main simplification function is then
+ − 1453
+ − 1454
\begin{center}
+ − 1455
\begin{tabular}{lcl}
+ − 1456
@{term "simp (ALT r\<^sub>1 r\<^sub>2)"} & $\dn$ & @{term "simp_ALT (simp r\<^sub>1) (simp r\<^sub>2)"}\\
+ − 1457
@{term "simp (SEQ r\<^sub>1 r\<^sub>2)"} & $\dn$ & @{term "simp_SEQ (simp r\<^sub>1) (simp r\<^sub>2)"}\\
+ − 1458
@{term "simp r"} & $\dn$ & @{term "(r, id)"}\\
+ − 1459
\end{tabular}
+ − 1460
\end{center}
+ − 1461
+ − 1462
\noindent where @{term "id"} stands for the identity function. The
+ − 1463
function @{const simp} returns a simplified regular expression and a corresponding
+ − 1464
rectification function. Note that we do not simplify under stars: this
+ − 1465
seems to slow down the algorithm, rather than speed it up. The optimised
+ − 1466
lexer is then given by the clauses:
+ − 1467
+ − 1468
\begin{center}
+ − 1469
\begin{tabular}{lcl}
+ − 1470
@{thm (lhs) slexer.simps(1)} & $\dn$ & @{thm (rhs) slexer.simps(1)}\\
+ − 1471
@{thm (lhs) slexer.simps(2)} & $\dn$ &
+ − 1472
@{text "let (r\<^sub>s, f\<^sub>r) = simp (r "}$\backslash$@{text " c) in"}\\
+ − 1473
& & @{text "case"} @{term "slexer r\<^sub>s s"} @{text of}\\
+ − 1474
& & \phantom{$|$} @{term "None"} @{text "\<Rightarrow>"} @{term None}\\
+ − 1475
& & $|$ @{term "Some v"} @{text "\<Rightarrow>"} @{text "Some (inj r c (f\<^sub>r v))"}
+ − 1476
\end{tabular}
+ − 1477
\end{center}
+ − 1478
+ − 1479
\noindent
+ − 1480
In the second clause we first calculate the derivative @{term "der c r"}
+ − 1481
and then simplify the result. This gives us a simplified derivative
+ − 1482
@{text "r\<^sub>s"} and a rectification function @{text "f\<^sub>r"}. The lexer
+ − 1483
is then recursively called with the simplified derivative, but before
+ − 1484
we inject the character @{term c} into the value @{term v}, we need to rectify
+ − 1485
@{term v} (that is construct @{term "f\<^sub>r v"}). Before we can establish the correctness
+ − 1486
of @{term "slexer"}, we need to show that simplification preserves the language
+ − 1487
and simplification preserves our POSIX relation once the value is rectified
+ − 1488
(recall @{const "simp"} generates a (regular expression, rectification function) pair):
+ − 1489
+ − 1490
\begin{lemma}\mbox{}\smallskip\\\label{slexeraux}
+ − 1491
\begin{tabular}{ll}
+ − 1492
(1) & @{thm L_fst_simp[symmetric]}\\
+ − 1493
(2) & @{thm[mode=IfThen] Posix_simp}
+ − 1494
\end{tabular}
+ − 1495
\end{lemma}
+ − 1496
+ − 1497
\begin{proof} Both are by induction on @{text r}. There is no
+ − 1498
interesting case for the first statement. For the second statement,
+ − 1499
of interest are the @{term "r = ALT r\<^sub>1 r\<^sub>2"} and @{term "r = SEQ r\<^sub>1
+ − 1500
r\<^sub>2"} cases. In each case we have to analyse four subcases whether
+ − 1501
@{term "fst (simp r\<^sub>1)"} and @{term "fst (simp r\<^sub>2)"} equals @{const
+ − 1502
ZERO} (respectively @{const ONE}). For example for @{term "r = ALT
+ − 1503
r\<^sub>1 r\<^sub>2"}, consider the subcase @{term "fst (simp r\<^sub>1) = ZERO"} and
+ − 1504
@{term "fst (simp r\<^sub>2) \<noteq> ZERO"}. By assumption we know @{term "s \<in>
+ − 1505
fst (simp (ALT r\<^sub>1 r\<^sub>2)) \<rightarrow> v"}. From this we can infer @{term "s \<in> fst (simp r\<^sub>2) \<rightarrow> v"}
+ − 1506
and by IH also (*) @{term "s \<in> r\<^sub>2 \<rightarrow> (snd (simp r\<^sub>2) v)"}. Given @{term "fst (simp r\<^sub>1) = ZERO"}
+ − 1507
we know @{term "L (fst (simp r\<^sub>1)) = {}"}. By the first statement
+ − 1508
@{term "L r\<^sub>1"} is the empty set, meaning (**) @{term "s \<notin> L r\<^sub>1"}.
+ − 1509
Taking (*) and (**) together gives by the \mbox{@{text "P+R"}}-rule
+ − 1510
@{term "s \<in> ALT r\<^sub>1 r\<^sub>2 \<rightarrow> Right (snd (simp r\<^sub>2) v)"}. In turn this
+ − 1511
gives @{term "s \<in> ALT r\<^sub>1 r\<^sub>2 \<rightarrow> snd (simp (ALT r\<^sub>1 r\<^sub>2)) v"} as we need to show.
+ − 1512
The other cases are similar.\qed
+ − 1513
\end{proof}
+ − 1514
+ − 1515
\noindent We can now prove relatively straightforwardly that the
+ − 1516
optimised lexer produces the expected result:
+ − 1517
+ − 1518
\begin{theorem}
+ − 1519
@{thm slexer_correctness}
+ − 1520
\end{theorem}
+ − 1521
+ − 1522
\begin{proof} By induction on @{term s} generalising over @{term
+ − 1523
r}. The case @{term "[]"} is trivial. For the cons-case suppose the
+ − 1524
string is of the form @{term "c # s"}. By induction hypothesis we
+ − 1525
know @{term "slexer r s = lexer r s"} holds for all @{term r} (in
+ − 1526
particular for @{term "r"} being the derivative @{term "der c
+ − 1527
r"}). Let @{term "r\<^sub>s"} be the simplified derivative regular expression, that is @{term
+ − 1528
"fst (simp (der c r))"}, and @{term "f\<^sub>r"} be the rectification
+ − 1529
function, that is @{term "snd (simp (der c r))"}. We distinguish the cases
+ − 1530
whether (*) @{term "s \<in> L (der c r)"} or not. In the first case we
272
+ − 1531
have by Theorem~\ref{lexercorrect}(2) a value @{term "v"} so that @{term
218
+ − 1532
"lexer (der c r) s = Some v"} and @{term "s \<in> der c r \<rightarrow> v"} hold.
272
+ − 1533
By Lemma~\ref{slexeraux}(1) we can also infer from~(*) that @{term "s
+ − 1534
\<in> L r\<^sub>s"} holds. Hence we know by Theorem~\ref{lexercorrect}(2) that
218
+ − 1535
there exists a @{term "v'"} with @{term "lexer r\<^sub>s s = Some v'"} and
+ − 1536
@{term "s \<in> r\<^sub>s \<rightarrow> v'"}. From the latter we know by
272
+ − 1537
Lemma~\ref{slexeraux}(2) that @{term "s \<in> der c r \<rightarrow> (f\<^sub>r v')"} holds.
+ − 1538
By the uniqueness of the POSIX relation (Theorem~\ref{posixdeterm}) we
218
+ − 1539
can infer that @{term v} is equal to @{term "f\<^sub>r v'"}---that is the
+ − 1540
rectification function applied to @{term "v'"}
+ − 1541
produces the original @{term "v"}. Now the case follows by the
+ − 1542
definitions of @{const lexer} and @{const slexer}.
+ − 1543
+ − 1544
In the second case where @{term "s \<notin> L (der c r)"} we have that
272
+ − 1545
@{term "lexer (der c r) s = None"} by Theorem~\ref{lexercorrect}(1). We
+ − 1546
also know by Lemma~\ref{slexeraux}(1) that @{term "s \<notin> L r\<^sub>s"}. Hence
+ − 1547
@{term "lexer r\<^sub>s s = None"} by Theorem~\ref{lexercorrect}(1) and
218
+ − 1548
by IH then also @{term "slexer r\<^sub>s s = None"}. With this we can
+ − 1549
conclude in this case too.\qed
+ − 1550
+ − 1551
\end{proof}
272
+ − 1552
+ − 1553
*}
+ − 1554
+ − 1555
218
+ − 1556
+ − 1557
+ − 1558
(*<*)
+ − 1559
end
+ − 1560
(*>*)