--- a/thys/BitCoded.thy Fri Jan 07 22:25:26 2022 +0000
+++ b/thys/BitCoded.thy Fri Jan 07 22:28:23 2022 +0000
@@ -1,5 +1,5 @@
-theory BitCodedCT
+theory BitCoded
imports "Lexer"
begin
--- a/thys2/Journal/Paper.tex Fri Jan 07 22:25:26 2022 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,3369 +0,0 @@
-%
-\begin{isabellebody}%
-\setisabellecontext{Paper}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isatagproof
-%
-\endisatagproof
-{\isafoldproof}%
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isatagproof
-%
-\endisatagproof
-{\isafoldproof}%
-%
-\isadelimproof
-%
-\endisadelimproof
-%
-\isadelimdocument
-%
-\endisadelimdocument
-%
-\isatagdocument
-%
-\isamarkupsection{Core of the proof%
-}
-\isamarkuptrue%
-%
-\endisatagdocument
-{\isafolddocument}%
-%
-\isadelimdocument
-%
-\endisadelimdocument
-%
-\begin{isamarkuptext}%
-This paper builds on previous work by Ausaf and Urban using
-regular expression'd bit-coded derivatives to do lexing that
-is both fast and satisfies the POSIX specification.
-In their work, a bit-coded algorithm introduced by Sulzmann and Lu
-was formally verified in Isabelle, by a very clever use of
-flex function and retrieve to carefully mimic the way a value is
-built up by the injection funciton.
-
-In the previous work, Ausaf and Urban established the below equality:
-\begin{lemma}
-\isa{{\normalsize{}If\,}\ v\ {\isacharcolon}{\kern0pt}\ r{\isacharbackslash}{\kern0pt}s\ {\normalsize \,then\,}\ Some\ {\isacharparenleft}{\kern0pt}flex\ r\ id\ s\ v{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ decode\ {\isacharparenleft}{\kern0pt}retrieve\ {\isacharparenleft}{\kern0pt}r\mbox{$^\uparrow$}\mbox{$\bbslash$}s{\isacharparenright}{\kern0pt}\ v{\isacharparenright}{\kern0pt}\ r{\isachardot}{\kern0pt}}
-\end{lemma}
-
-This lemma establishes a link with the lexer without bit-codes.
-
-With it we get the correctness of bit-coded algorithm.
-\begin{lemma}
-\isa{lexer\mbox{$_b$}\ r\ s\ {\isacharequal}{\kern0pt}\ lexer\ r\ s}
-\end{lemma}
-
-However what is not certain is whether we can add simplification
-to the bit-coded algorithm, without breaking the correct lexing output.
-
-
-The reason that we do need to add a simplification phase
-after each derivative step of $\textit{blexer}$ is
-because it produces intermediate
-regular expressions that can grow exponentially.
-For example, the regular expression $(a+aa)^*$ after taking
-derivative against just 10 $a$s will have size 8192.
-
-%TODO: add figure for this?
-
-
-Therefore, we insert a simplification phase
-after each derivation step, as defined below:
-\begin{lemma}
-\isa{blexer{\isacharunderscore}{\kern0pt}simp\ r\ s\ {\isasymequiv}\ \textrm{if}\ nullable\mbox{$_b$}\ {\isacharparenleft}{\kern0pt}bders{\isacharunderscore}{\kern0pt}simp\ {\isacharparenleft}{\kern0pt}r\mbox{$^\uparrow$}{\isacharparenright}{\kern0pt}\ s{\isacharparenright}{\kern0pt}\ \textrm{then}\ decode\ {\isacharparenleft}{\kern0pt}mkeps\mbox{$_b$}\ {\isacharparenleft}{\kern0pt}bders{\isacharunderscore}{\kern0pt}simp\ {\isacharparenleft}{\kern0pt}r\mbox{$^\uparrow$}{\isacharparenright}{\kern0pt}\ s{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ r\ \textrm{else}\ None}
-\end{lemma}
-
-The simplification function is given as follows:
-
-\begin{center}
- \begin{tabular}{lcl}
- \isa{bsimp\ {\isacharparenleft}{\kern0pt}ASEQ\ bs\ r\isactrlsub {\isadigit{1}}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{bsimp{\isacharunderscore}{\kern0pt}ASEQ\ bs\ {\isacharparenleft}{\kern0pt}bsimp\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}bsimp\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}\\
- \isa{bsimp\ {\isacharparenleft}{\kern0pt}AALTs\ bs{\isadigit{1}}{\isachardot}{\kern0pt}{\isadigit{0}}\ rs{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{bsimp{\isacharunderscore}{\kern0pt}AALTs\ bs{\isadigit{1}}{\isachardot}{\kern0pt}{\isadigit{0}}\ {\isacharparenleft}{\kern0pt}distinctBy\ {\isacharparenleft}{\kern0pt}flts\ {\isacharparenleft}{\kern0pt}map\ bsimp\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ erase\ {\isasymemptyset}{\isacharparenright}{\kern0pt}}\\
- \isa{bsimp\ AZERO} & $\dn$ & \isa{AZERO}\\
-
-\end{tabular}
-\end{center}
-
-And the two helper functions are:
-\begin{center}
- \begin{tabular}{lcl}
- \isa{bsimp{\isacharunderscore}{\kern0pt}AALTs\ bs\isactrlsub {\isadigit{1}}\ {\isacharbrackleft}{\kern0pt}r{\isacharbrackright}{\kern0pt}} & $\dn$ & \isa{bsimp{\isacharunderscore}{\kern0pt}ASEQ\ bs\isactrlsub {\isadigit{1}}\ {\isacharparenleft}{\kern0pt}bsimp\ r{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}bsimp\ r{\isadigit{2}}{\isachardot}{\kern0pt}{\isadigit{0}}{\isacharparenright}{\kern0pt}}\\
- \isa{bsimp{\isacharunderscore}{\kern0pt}AALTs\ bs{\isadigit{1}}{\isachardot}{\kern0pt}{\isadigit{0}}\ {\isacharbrackleft}{\kern0pt}r{\isacharbrackright}{\kern0pt}} & $\dn$ & \isa{bsimp{\isacharunderscore}{\kern0pt}AALTs\ bs{\isadigit{1}}{\isachardot}{\kern0pt}{\isadigit{0}}\ {\isacharparenleft}{\kern0pt}distinctBy\ {\isacharparenleft}{\kern0pt}flts\ {\isacharparenleft}{\kern0pt}map\ bsimp\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ erase\ {\isasymemptyset}{\isacharparenright}{\kern0pt}}\\
- \isa{bsimp{\isacharunderscore}{\kern0pt}AALTs\ bs{\isadigit{1}}{\isachardot}{\kern0pt}{\isadigit{0}}\ {\isacharparenleft}{\kern0pt}v\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}vb\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}vc{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{AZERO}\\
-
-\end{tabular}
-\end{center}
-
-
-This might sound trivial in the case of producing a YES/NO answer,
-but once we require a lexing output to be produced (which is required
-in applications like compiler front-end, malicious attack domain extraction,
-etc.), it is not straightforward if we still extract what is needed according
-to the POSIX standard.
-
-
-
-
-
-By simplification, we mean specifically the following rules:
-
-\begin{center}
- \begin{tabular}{lcl}
- \isa{\mbox{}\inferrule{\mbox{}}{\mbox{ASEQ\ bs\ AZERO\ r\isactrlsub {\isadigit{2}}\ {\isasymleadsto}\ AZERO}}}\\
- \isa{\mbox{}\inferrule{\mbox{}}{\mbox{ASEQ\ bs\ r\isactrlsub {\isadigit{1}}\ AZERO\ {\isasymleadsto}\ AZERO}}}\\
- \isa{\mbox{}\inferrule{\mbox{}}{\mbox{ASEQ\ bs\ {\isacharparenleft}{\kern0pt}AONE\ bs\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isasymleadsto}\ fuse\ {\isacharparenleft}{\kern0pt}bs\ {\isacharat}{\kern0pt}\ bs\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ r\isactrlsub {\isadigit{1}}}}}\\
- \isa{\mbox{}\inferrule{\mbox{r\isactrlsub {\isadigit{1}}\ {\isasymleadsto}\ r\isactrlsub {\isadigit{2}}}}{\mbox{ASEQ\ bs\ r\isactrlsub {\isadigit{1}}\ r\isactrlsub {\isadigit{3}}\ {\isasymleadsto}\ ASEQ\ bs\ r\isactrlsub {\isadigit{2}}\ r\isactrlsub {\isadigit{3}}}}}\\
- \isa{\mbox{}\inferrule{\mbox{r\isactrlsub {\isadigit{3}}\ {\isasymleadsto}\ r\isactrlsub {\isadigit{4}}}}{\mbox{ASEQ\ bs\ r\isactrlsub {\isadigit{1}}\ r\isactrlsub {\isadigit{3}}\ {\isasymleadsto}\ ASEQ\ bs\ r\isactrlsub {\isadigit{1}}\ r\isactrlsub {\isadigit{4}}}}}\\
- \isa{\mbox{}\inferrule{\mbox{r\ {\isasymleadsto}\ r{\isacharprime}{\kern0pt}}}{\mbox{AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs\isactrlsub {\isadigit{1}}\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}r{\isacharbrackright}{\kern0pt}\ {\isacharat}{\kern0pt}\ rs\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isasymleadsto}\ AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs\isactrlsub {\isadigit{1}}\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}r{\isacharprime}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isacharat}{\kern0pt}\ rs\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}}}\\
- \isa{\mbox{}\inferrule{\mbox{}}{\mbox{AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs\isactrlsub a\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}AZERO{\isacharbrackright}{\kern0pt}\ {\isacharat}{\kern0pt}\ rs\isactrlsub b{\isacharparenright}{\kern0pt}\ {\isasymleadsto}\ AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs\isactrlsub a\ {\isacharat}{\kern0pt}\ rs\isactrlsub b{\isacharparenright}{\kern0pt}}}}\\
- \isa{\mbox{}\inferrule{\mbox{}}{\mbox{AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs\isactrlsub a\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}AALTs\ bs\isactrlsub {\isadigit{1}}\ rs\isactrlsub {\isadigit{1}}{\isacharbrackright}{\kern0pt}\ {\isacharat}{\kern0pt}\ rs\isactrlsub b{\isacharparenright}{\kern0pt}\ {\isasymleadsto}\ AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs\isactrlsub a\ {\isacharat}{\kern0pt}\ map\ {\isacharparenleft}{\kern0pt}fuse\ bs\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ rs\isactrlsub {\isadigit{1}}\ {\isacharat}{\kern0pt}\ rs\isactrlsub b{\isacharparenright}{\kern0pt}}}}\\
- \isa{\mbox{}\inferrule{\mbox{}}{\mbox{AALTs\ {\isacharparenleft}{\kern0pt}bs\ {\isacharat}{\kern0pt}\ bs\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ rs\ {\isasymleadsto}\ AALTs\ bs\ {\isacharparenleft}{\kern0pt}map\ {\isacharparenleft}{\kern0pt}fuse\ bs\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}}}}\\
- \isa{\mbox{}\inferrule{\mbox{}}{\mbox{AALTs\ bs\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isasymleadsto}\ AZERO}}}\\
- \isa{\mbox{}\inferrule{\mbox{}}{\mbox{AALTs\ bs\ {\isacharbrackleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharbrackright}{\kern0pt}\ {\isasymleadsto}\ fuse\ bs\ r\isactrlsub {\isadigit{1}}}}}\\
- \isa{\mbox{}\inferrule{\mbox{a\isactrlsub {\isadigit{1}}\mbox{$^\downarrow$}\ {\isacharequal}{\kern0pt}\ a\isactrlsub {\isadigit{2}}\mbox{$^\downarrow$}}}{\mbox{AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs\isactrlsub a\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}a\isactrlsub {\isadigit{1}}{\isacharbrackright}{\kern0pt}\ {\isacharat}{\kern0pt}\ rs\isactrlsub b\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}a\isactrlsub {\isadigit{2}}{\isacharbrackright}{\kern0pt}\ {\isacharat}{\kern0pt}\ rs\isactrlsub c{\isacharparenright}{\kern0pt}\ {\isasymleadsto}\ AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs\isactrlsub a\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}a\isactrlsub {\isadigit{1}}{\isacharbrackright}{\kern0pt}\ {\isacharat}{\kern0pt}\ rs\isactrlsub b\ {\isacharat}{\kern0pt}\ rs\isactrlsub c{\isacharparenright}{\kern0pt}}}}\\
-
- \end{tabular}
-\end{center}
-
-
-And these can be made compact by the following simplification function:
-
-where the function $\mathit{bsimp_AALTs}$
-
-The core idea of the proof is that two regular expressions,
-if "isomorphic" up to a finite number of rewrite steps, will
-remain "isomorphic" when we take the same sequence of
-derivatives on both of them.
-This can be expressed by the following rewrite relation lemma:
-\begin{lemma}
-\isa{{\isacharparenleft}{\kern0pt}r\mbox{$\bbslash$}s{\isacharparenright}{\kern0pt}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ bders{\isacharunderscore}{\kern0pt}simp\ r\ s}
-\end{lemma}
-
-This isomorphic relation implies a property that leads to the
-correctness result:
-if two (nullable) regular expressions are "rewritable" in many steps
-from one another,
-then a call to function $\textit{bmkeps}$ gives the same
-bit-sequence :
-\begin{lemma}
-\isa{{\normalsize{}If\,}\ \mbox{r{\isadigit{1}}{\isachardot}{\kern0pt}{\isadigit{0}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ r{\isadigit{2}}{\isachardot}{\kern0pt}{\isadigit{0}}}\ {\normalsize \,and\,}\ \mbox{nullable\mbox{$_b$}\ r{\isadigit{1}}{\isachardot}{\kern0pt}{\isadigit{0}}}\ {\normalsize \,then\,}\ mkeps\mbox{$_b$}\ r{\isadigit{1}}{\isachardot}{\kern0pt}{\isadigit{0}}\ {\isacharequal}{\kern0pt}\ mkeps\mbox{$_b$}\ r{\isadigit{2}}{\isachardot}{\kern0pt}{\isadigit{0}}{\isachardot}{\kern0pt}}
-\end{lemma}
-
-Given the same bit-sequence, the decode function
-will give out the same value, which is the output
-of both lexers:
-\begin{lemma}
-\isa{lexer\mbox{$_b$}\ r\ s\ {\isasymequiv}\ \textrm{if}\ nullable\mbox{$_b$}\ {\isacharparenleft}{\kern0pt}r\mbox{$^\uparrow$}\mbox{$\bbslash$}s{\isacharparenright}{\kern0pt}\ \textrm{then}\ decode\ {\isacharparenleft}{\kern0pt}mkeps\mbox{$_b$}\ {\isacharparenleft}{\kern0pt}r\mbox{$^\uparrow$}\mbox{$\bbslash$}s{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ r\ \textrm{else}\ None}
-\end{lemma}
-
-\begin{lemma}
-\isa{blexer{\isacharunderscore}{\kern0pt}simp\ r\ s\ {\isasymequiv}\ \textrm{if}\ nullable\mbox{$_b$}\ {\isacharparenleft}{\kern0pt}bders{\isacharunderscore}{\kern0pt}simp\ {\isacharparenleft}{\kern0pt}r\mbox{$^\uparrow$}{\isacharparenright}{\kern0pt}\ s{\isacharparenright}{\kern0pt}\ \textrm{then}\ decode\ {\isacharparenleft}{\kern0pt}mkeps\mbox{$_b$}\ {\isacharparenleft}{\kern0pt}bders{\isacharunderscore}{\kern0pt}simp\ {\isacharparenleft}{\kern0pt}r\mbox{$^\uparrow$}{\isacharparenright}{\kern0pt}\ s{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ r\ \textrm{else}\ None}
-\end{lemma}
-
-And that yields the correctness result:
-\begin{lemma}
-\isa{lexer\ r\ s\ {\isacharequal}{\kern0pt}\ blexer{\isacharunderscore}{\kern0pt}simp\ r\ s}
-\end{lemma}
-
-The nice thing about the above%
-\end{isamarkuptext}\isamarkuptrue%
-%
-\isadelimdocument
-%
-\endisadelimdocument
-%
-\isatagdocument
-%
-\isamarkupsection{Additional Simp Rules?%
-}
-\isamarkuptrue%
-%
-\endisatagdocument
-{\isafolddocument}%
-%
-\isadelimdocument
-%
-\endisadelimdocument
-%
-\begin{isamarkuptext}%
-One question someone would ask is:
-can we add more "atomic" simplification/rewriting rules,
-so the simplification is even more aggressive, making
-the intermediate results smaller, and therefore more space-efficient?
-For example, one might want to do open up alternatives who is a child
-of a sequence:
-
-\begin{center}
- \begin{tabular}{lcl}
- \isa{ASEQ\ bs\ {\isacharparenleft}{\kern0pt}AALTs\ bs{\isadigit{1}}\ rs{\isacharparenright}{\kern0pt}\ r\ {\isasymleadsto}{\isacharquery}{\kern0pt}\ AALTs\ {\isacharparenleft}{\kern0pt}bs\ {\isacharat}{\kern0pt}\ bs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}map\ {\isacharparenleft}{\kern0pt}{\isasymlambda}r{\isacharprime}{\kern0pt}{\isachardot}{\kern0pt}\ ASEQ\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ r{\isacharprime}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}}\\
- \end{tabular}
-\end{center}
-
-This rule allows us to simplify \mbox{\isa{a\ {\isacharplus}{\kern0pt}\ b\ {\isasymcdot}\ c\ {\isacharplus}{\kern0pt}\ a\ {\isasymcdot}\ c}}%
-\end{isamarkuptext}\isamarkuptrue%
-%
-\isadelimdocument
-%
-\endisadelimdocument
-%
-\isatagdocument
-%
-\isamarkupsection{Introduction%
-}
-\isamarkuptrue%
-%
-\endisatagdocument
-{\isafolddocument}%
-%
-\isadelimdocument
-%
-\endisadelimdocument
-%
-\begin{isamarkuptext}%
-Brzozowski \cite{Brzozowski1964} introduced the notion of the {\em
-derivative} \isa{r{\isacharbackslash}{\kern0pt}c} of a regular expression \isa{r} w.r.t.\
-a character~\isa{c}, and showed that it gave a simple solution to the
-problem of matching a string \isa{s} with a regular expression \isa{r}: if the derivative of \isa{r} w.r.t.\ (in succession) all the
-characters of the string matches the empty string, then \isa{r}
-matches \isa{s} (and {\em vice versa}). The derivative has the
-property (which may almost be regarded as its specification) that, for
-every string \isa{s} and regular expression \isa{r} and character
-\isa{c}, one has \isa{cs\ {\isasymin}\ L{\isacharparenleft}{\kern0pt}r{\isacharparenright}{\kern0pt}} if and only if \mbox{\isa{s\ {\isasymin}\ L{\isacharparenleft}{\kern0pt}r{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}}}.
-The beauty of Brzozowski's derivatives is that
-they are neatly expressible in any functional language, and easily
-definable and reasoned about in theorem provers---the definitions just
-consist of inductive datatypes and simple recursive functions. A
-mechanised correctness proof of Brzozowski's matcher in for example HOL4
-has been mentioned by Owens and Slind~\cite{Owens2008}. Another one in
-Isabelle/HOL is part of the work by Krauss and Nipkow \cite{Krauss2011}.
-And another one in Coq is given by Coquand and Siles \cite{Coquand2012}.
-
-If a regular expression matches a string, then in general there is more
-than one way of how the string is matched. There are two commonly used
-disambiguation strategies to generate a unique answer: one is called
-GREEDY matching \cite{Frisch2004} and the other is POSIX
-matching~\cite{POSIX,Kuklewicz,OkuiSuzuki2010,Sulzmann2014,Vansummeren2006}.
-For example consider the string \isa{xy} and the regular expression
-\mbox{\isa{{\isacharparenleft}{\kern0pt}x\ {\isacharplus}{\kern0pt}\ y\ {\isacharplus}{\kern0pt}\ xy{\isacharparenright}{\kern0pt}\isactrlsup {\isasymstar}}}. Either the string can be
-matched in two `iterations' by the single letter-regular expressions
-\isa{x} and \isa{y}, or directly in one iteration by \isa{xy}. The
-first case corresponds to GREEDY matching, which first matches with the
-left-most symbol and only matches the next symbol in case of a mismatch
-(this is greedy in the sense of preferring instant gratification to
-delayed repletion). The second case is POSIX matching, which prefers the
-longest match.
-
-In the context of lexing, where an input string needs to be split up
-into a sequence of tokens, POSIX is the more natural disambiguation
-strategy for what programmers consider basic syntactic building blocks
-in their programs. These building blocks are often specified by some
-regular expressions, say \isa{r\isactrlbsub key\isactrlesub } and \isa{r\isactrlbsub id\isactrlesub } for recognising keywords and identifiers,
-respectively. There are a few underlying (informal) rules behind
-tokenising a string in a POSIX \cite{POSIX} fashion:
-
-\begin{itemize}
-\item[$\bullet$] \emph{The Longest Match Rule} (or \emph{``{M}aximal {M}unch {R}ule''}):
-The longest initial substring matched by any regular expression is taken as
-next token.\smallskip
-
-\item[$\bullet$] \emph{Priority Rule:}
-For a particular longest initial substring, the first (leftmost) regular expression
-that can match determines the token.\smallskip
-
-\item[$\bullet$] \emph{Star Rule:} A subexpression repeated by ${}^\star$ shall
-not match an empty string unless this is the only match for the repetition.\smallskip
-
-\item[$\bullet$] \emph{Empty String Rule:} An empty string shall be considered to
-be longer than no match at all.
-\end{itemize}
-
-\noindent Consider for example a regular expression \isa{r\isactrlbsub key\isactrlesub } for recognising keywords such as \isa{if},
-\isa{then} and so on; and \isa{r\isactrlbsub id\isactrlesub }
-recognising identifiers (say, a single character followed by
-characters or numbers). Then we can form the regular expression
-\isa{{\isacharparenleft}{\kern0pt}r\isactrlbsub key\isactrlesub \ {\isacharplus}{\kern0pt}\ r\isactrlbsub id\isactrlesub {\isacharparenright}{\kern0pt}\isactrlsup {\isasymstar}}
-and use POSIX matching to tokenise strings, say \isa{iffoo} and
-\isa{if}. For \isa{iffoo} we obtain by the Longest Match Rule
-a single identifier token, not a keyword followed by an
-identifier. For \isa{if} we obtain by the Priority Rule a keyword
-token, not an identifier token---even if \isa{r\isactrlbsub id\isactrlesub }
-matches also. By the Star Rule we know \isa{{\isacharparenleft}{\kern0pt}r\isactrlbsub key\isactrlesub \ {\isacharplus}{\kern0pt}\ r\isactrlbsub id\isactrlesub {\isacharparenright}{\kern0pt}\isactrlsup {\isasymstar}} matches \isa{iffoo},
-respectively \isa{if}, in exactly one `iteration' of the star. The
-Empty String Rule is for cases where, for example, the regular expression
-\isa{{\isacharparenleft}{\kern0pt}a\isactrlsup {\isasymstar}{\isacharparenright}{\kern0pt}\isactrlsup {\isasymstar}} matches against the
-string \isa{bc}. Then the longest initial matched substring is the
-empty string, which is matched by both the whole regular expression
-and the parenthesised subexpression.
-
-
-One limitation of Brzozowski's matcher is that it only generates a
-YES/NO answer for whether a string is being matched by a regular
-expression. Sulzmann and Lu~\cite{Sulzmann2014} extended this matcher
-to allow generation not just of a YES/NO answer but of an actual
-matching, called a [lexical] {\em value}. Assuming a regular
-expression matches a string, values encode the information of
-\emph{how} the string is matched by the regular expression---that is,
-which part of the string is matched by which part of the regular
-expression. For this consider again the string \isa{xy} and
-the regular expression \mbox{\isa{{\isacharparenleft}{\kern0pt}x\ {\isacharplus}{\kern0pt}\ {\isacharparenleft}{\kern0pt}y\ {\isacharplus}{\kern0pt}\ xy{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isactrlsup {\isasymstar}}}
-(this time fully parenthesised). We can view this regular expression
-as tree and if the string \isa{xy} is matched by two Star
-`iterations', then the \isa{x} is matched by the left-most
-alternative in this tree and the \isa{y} by the right-left alternative. This
-suggests to record this matching as
-
-\begin{center}
-\isa{Stars\ {\isacharbrackleft}{\kern0pt}Left\ {\isacharparenleft}{\kern0pt}Char\ x{\isacharparenright}{\kern0pt}{\isacharcomma}{\kern0pt}\ Right\ {\isacharparenleft}{\kern0pt}Left\ {\isacharparenleft}{\kern0pt}Char\ y{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharbrackright}{\kern0pt}}
-\end{center}
-
-\noindent where \isa{Stars}, \isa{Left}, \isa{Right} and \isa{Char} are constructors for values. \isa{Stars} records how many
-iterations were used; \isa{Left}, respectively \isa{Right}, which
-alternative is used. This `tree view' leads naturally to the idea that
-regular expressions act as types and values as inhabiting those types
-(see, for example, \cite{HosoyaVouillonPierce2005}). The value for
-matching \isa{xy} in a single `iteration', i.e.~the POSIX value,
-would look as follows
-
-\begin{center}
-\isa{Stars\ {\isacharbrackleft}{\kern0pt}Seq\ {\isacharparenleft}{\kern0pt}Char\ x{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}Char\ y{\isacharparenright}{\kern0pt}{\isacharbrackright}{\kern0pt}}
-\end{center}
-
-\noindent where \isa{Stars} has only a single-element list for the
-single iteration and \isa{Seq} indicates that \isa{xy} is matched
-by a sequence regular expression.
-
-%, which we will in what follows
-%write more formally as \isa{x\ {\isasymcdot}\ y}.
-
-
-Sulzmann and Lu give a simple algorithm to calculate a value that
-appears to be the value associated with POSIX matching. The challenge
-then is to specify that value, in an algorithm-independent fashion,
-and to show that Sulzmann and Lu's derivative-based algorithm does
-indeed calculate a value that is correct according to the
-specification. The answer given by Sulzmann and Lu
-\cite{Sulzmann2014} is to define a relation (called an ``order
-relation'') on the set of values of \isa{r}, and to show that (once
-a string to be matched is chosen) there is a maximum element and that
-it is computed by their derivative-based algorithm. This proof idea is
-inspired by work of Frisch and Cardelli \cite{Frisch2004} on a GREEDY
-regular expression matching algorithm. However, we were not able to
-establish transitivity and totality for the ``order relation'' by
-Sulzmann and Lu. There are some inherent problems with their approach
-(of which some of the proofs are not published in
-\cite{Sulzmann2014}); perhaps more importantly, we give in this paper
-a simple inductive (and algorithm-independent) definition of what we
-call being a {\em POSIX value} for a regular expression \isa{r} and
-a string \isa{s}; we show that the algorithm by Sulzmann and Lu
-computes such a value and that such a value is unique. Our proofs are
-both done by hand and checked in Isabelle/HOL. The experience of
-doing our proofs has been that this mechanical checking was absolutely
-essential: this subject area has hidden snares. This was also noted by
-Kuklewicz \cite{Kuklewicz} who found that nearly all POSIX matching
-implementations are ``buggy'' \cite[Page 203]{Sulzmann2014} and by
-Grathwohl et al \cite[Page 36]{CrashCourse2014} who wrote:
-
-\begin{quote}
-\it{}``The POSIX strategy is more complicated than the greedy because of
-the dependence on information about the length of matched strings in the
-various subexpressions.''
-\end{quote}
-
-
-
-\noindent {\bf Contributions:} We have implemented in Isabelle/HOL the
-derivative-based regular expression matching algorithm of
-Sulzmann and Lu \cite{Sulzmann2014}. We have proved the correctness of this
-algorithm according to our specification of what a POSIX value is (inspired
-by work of Vansummeren \cite{Vansummeren2006}). Sulzmann
-and Lu sketch in \cite{Sulzmann2014} an informal correctness proof: but to
-us it contains unfillable gaps.\footnote{An extended version of
-\cite{Sulzmann2014} is available at the website of its first author; this
-extended version already includes remarks in the appendix that their
-informal proof contains gaps, and possible fixes are not fully worked out.}
-Our specification of a POSIX value consists of a simple inductive definition
-that given a string and a regular expression uniquely determines this value.
-We also show that our definition is equivalent to an ordering
-of values based on positions by Okui and Suzuki \cite{OkuiSuzuki2010}.
-
-%Derivatives as calculated by Brzozowski's method are usually more complex
-%regular expressions than the initial one; various optimisations are
-%possible. We prove the correctness when simplifications of \isa{\isactrlbold {\isadigit{0}}\ {\isacharplus}{\kern0pt}\ r},
-%\isa{r\ {\isacharplus}{\kern0pt}\ \isactrlbold {\isadigit{0}}}, \isa{\isactrlbold {\isadigit{1}}\ {\isasymcdot}\ r} and \isa{r\ {\isasymcdot}\ \isactrlbold {\isadigit{1}}} to
-%\isa{r} are applied.
-
-We extend our results to ??? Bitcoded version??%
-\end{isamarkuptext}\isamarkuptrue%
-%
-\isadelimdocument
-%
-\endisadelimdocument
-%
-\isatagdocument
-%
-\isamarkupsection{Preliminaries%
-}
-\isamarkuptrue%
-%
-\endisatagdocument
-{\isafolddocument}%
-%
-\isadelimdocument
-%
-\endisadelimdocument
-%
-\begin{isamarkuptext}%
-\noindent Strings in Isabelle/HOL are lists of characters with
-the empty string being represented by the empty list, written \isa{{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}}, and list-cons being written as \isa{\underline{\hspace{2mm}}\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}\underline{\hspace{2mm}}}. Often
-we use the usual bracket notation for lists also for strings; for
-example a string consisting of just a single character \isa{c} is
-written \isa{{\isacharbrackleft}{\kern0pt}c{\isacharbrackright}{\kern0pt}}. We use the usual definitions for
-\emph{prefixes} and \emph{strict prefixes} of strings. By using the
-type \isa{char} for characters we have a supply of finitely many
-characters roughly corresponding to the ASCII character set. Regular
-expressions are defined as usual as the elements of the following
-inductive datatype:
-
- \begin{center}
- \isa{r\ {\isacharcolon}{\kern0pt}{\isacharequal}{\kern0pt}}
- \isa{\isactrlbold {\isadigit{0}}} $\mid$
- \isa{\isactrlbold {\isadigit{1}}} $\mid$
- \isa{c} $\mid$
- \isa{r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}} $\mid$
- \isa{r\isactrlsub {\isadigit{1}}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}} $\mid$
- \isa{r\isactrlsup {\isasymstar}}
- \end{center}
-
- \noindent where \isa{\isactrlbold {\isadigit{0}}} stands for the regular expression that does
- not match any string, \isa{\isactrlbold {\isadigit{1}}} for the regular expression that matches
- only the empty string and \isa{c} for matching a character literal. The
- language of a regular expression is also defined as usual by the
- recursive function \isa{L} with the six clauses:
-
- \begin{center}
- \begin{tabular}{l@ {\hspace{4mm}}rcl}
- \textit{(1)} & \isa{L{\isacharparenleft}{\kern0pt}\isactrlbold {\isadigit{0}}{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{{\isasymemptyset}}\\
- \textit{(2)} & \isa{L{\isacharparenleft}{\kern0pt}\isactrlbold {\isadigit{1}}{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{{\isacharbraceleft}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}}\\
- \textit{(3)} & \isa{L{\isacharparenleft}{\kern0pt}c{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{{\isacharbraceleft}{\kern0pt}{\isacharbrackleft}{\kern0pt}c{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}}\\
- \textit{(4)} & \isa{L{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}} & $\dn$ &
- \isa{L{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharat}{\kern0pt}\ L{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}\\
- \textit{(5)} & \isa{L{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}} & $\dn$ &
- \isa{L{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isasymunion}\ L{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}\\
- \textit{(6)} & \isa{L{\isacharparenleft}{\kern0pt}r\isactrlsup {\isasymstar}{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{{\isacharparenleft}{\kern0pt}L{\isacharparenleft}{\kern0pt}r{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isasymstar}}\\
- \end{tabular}
- \end{center}
-
- \noindent In clause \textit{(4)} we use the operation \isa{\underline{\hspace{2mm}}\ {\isacharat}{\kern0pt}\ \underline{\hspace{2mm}}} for the concatenation of two languages (it is also list-append for
- strings). We use the star-notation for regular expressions and for
- languages (in the last clause above). The star for languages is defined
- inductively by two clauses: \isa{{\isacharparenleft}{\kern0pt}i{\isacharparenright}{\kern0pt}} the empty string being in
- the star of a language and \isa{{\isacharparenleft}{\kern0pt}ii{\isacharparenright}{\kern0pt}} if \isa{s\isactrlsub {\isadigit{1}}} is in a
- language and \isa{s\isactrlsub {\isadigit{2}}} in the star of this language, then also \isa{s\isactrlsub {\isadigit{1}}\ {\isacharat}{\kern0pt}\ s\isactrlsub {\isadigit{2}}} is in the star of this language. It will also be convenient
- to use the following notion of a \emph{semantic derivative} (or \emph{left
- quotient}) of a language defined as
- %
- \begin{center}
- \isa{Der\ c\ A\ {\isasymequiv}\ {\isacharbraceleft}{\kern0pt}s\ \mbox{\boldmath$\mid$}\ c\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}s\ {\isasymin}\ A{\isacharbraceright}{\kern0pt}}\;.
- \end{center}
-
- \noindent
- For semantic derivatives we have the following equations (for example
- mechanically proved in \cite{Krauss2011}):
- %
- \begin{equation}\label{SemDer}
- \begin{array}{lcl}
- \isa{Der\ c\ {\isasymemptyset}} & \dn & \isa{{\isasymemptyset}}\\
- \isa{Der\ c\ {\isacharbraceleft}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}} & \dn & \isa{{\isasymemptyset}}\\
- \isa{Der\ c\ {\isacharbraceleft}{\kern0pt}{\isacharbrackleft}{\kern0pt}d{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}} & \dn & \isa{\textrm{if}\ c\ {\isacharequal}{\kern0pt}\ d\ \textrm{then}\ {\isacharbraceleft}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ \textrm{else}\ {\isasymemptyset}}\\
- \isa{Der\ c\ {\isacharparenleft}{\kern0pt}A\ {\isasymunion}\ B{\isacharparenright}{\kern0pt}} & \dn & \isa{Der\ c\ A\ {\isasymunion}\ Der\ c\ B}\\
- \isa{Der\ c\ {\isacharparenleft}{\kern0pt}A\ {\isacharat}{\kern0pt}\ B{\isacharparenright}{\kern0pt}} & \dn & \isa{{\isacharparenleft}{\kern0pt}Der\ c\ A\ {\isacharat}{\kern0pt}\ B{\isacharparenright}{\kern0pt}\ {\isasymunion}\ {\isacharparenleft}{\kern0pt}\textrm{if}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isasymin}\ A\ \textrm{then}\ Der\ c\ B\ \textrm{else}\ {\isasymemptyset}{\isacharparenright}{\kern0pt}}\\
- \isa{Der\ c\ {\isacharparenleft}{\kern0pt}A{\isasymstar}{\isacharparenright}{\kern0pt}} & \dn & \isa{Der\ c\ A\ {\isacharat}{\kern0pt}\ A{\isasymstar}}
- \end{array}
- \end{equation}
-
-
- \noindent \emph{\Brz's derivatives} of regular expressions
- \cite{Brzozowski1964} can be easily defined by two recursive functions:
- the first is from regular expressions to booleans (implementing a test
- when a regular expression can match the empty string), and the second
- takes a regular expression and a character to a (derivative) regular
- expression:
-
- \begin{center}
- \begin{tabular}{lcl}
- \isa{nullable\ {\isacharparenleft}{\kern0pt}\isactrlbold {\isadigit{0}}{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{False}\\
- \isa{nullable\ {\isacharparenleft}{\kern0pt}\isactrlbold {\isadigit{1}}{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{True}\\
- \isa{nullable\ {\isacharparenleft}{\kern0pt}c{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{False}\\
- \isa{nullable\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{nullable\ r\isactrlsub {\isadigit{1}}\ {\isasymor}\ nullable\ r\isactrlsub {\isadigit{2}}}\\
- \isa{nullable\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{nullable\ r\isactrlsub {\isadigit{1}}\ {\isasymand}\ nullable\ r\isactrlsub {\isadigit{2}}}\\
- \isa{nullable\ {\isacharparenleft}{\kern0pt}r\isactrlsup {\isasymstar}{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{True}\medskip\\
-
-% \end{tabular}
-% \end{center}
-
-% \begin{center}
-% \begin{tabular}{lcl}
-
- \isa{\isactrlbold {\isadigit{0}}{\isacharbackslash}{\kern0pt}c} & $\dn$ & \isa{\isactrlbold {\isadigit{0}}}\\
- \isa{\isactrlbold {\isadigit{1}}{\isacharbackslash}{\kern0pt}c} & $\dn$ & \isa{\isactrlbold {\isadigit{0}}}\\
- \isa{d{\isacharbackslash}{\kern0pt}c} & $\dn$ & \isa{\textrm{if}\ c\ {\isacharequal}{\kern0pt}\ d\ \textrm{then}\ \isactrlbold {\isadigit{1}}\ \textrm{else}\ \isactrlbold {\isadigit{0}}}\\
- \isa{{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharbackslash}{\kern0pt}c} & $\dn$ & \isa{{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isacharplus}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}}\\
- \isa{{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharbackslash}{\kern0pt}c} & $\dn$ & \isa{\textrm{if}\ nullable\ r\isactrlsub {\isadigit{1}}\ \textrm{then}\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}\ {\isacharplus}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ \textrm{else}\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}}\\
- \isa{{\isacharparenleft}{\kern0pt}r\isactrlsup {\isasymstar}{\isacharparenright}{\kern0pt}{\isacharbackslash}{\kern0pt}c} & $\dn$ & \isa{{\isacharparenleft}{\kern0pt}r{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isasymcdot}\ r\isactrlsup {\isasymstar}}
- \end{tabular}
- \end{center}
-
- \noindent
- We may extend this definition to give derivatives w.r.t.~strings:
-
- \begin{center}
- \begin{tabular}{lcl}
- \isa{r{\isacharbackslash}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}} & $\dn$ & \isa{r}\\
- \isa{r{\isacharbackslash}{\kern0pt}{\isacharparenleft}{\kern0pt}c\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}s{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{{\isacharparenleft}{\kern0pt}r{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}{\isacharbackslash}{\kern0pt}s}\\
- \end{tabular}
- \end{center}
-
- \noindent Given the equations in \eqref{SemDer}, it is a relatively easy
- exercise in mechanical reasoning to establish that
-
- \begin{proposition}\label{derprop}\mbox{}\\
- \begin{tabular}{ll}
- \textit{(1)} & \isa{nullable\ r} if and only if
- \isa{{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isasymin}\ L{\isacharparenleft}{\kern0pt}r{\isacharparenright}{\kern0pt}}, and \\
- \textit{(2)} & \isa{L{\isacharparenleft}{\kern0pt}r{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ Der\ c\ {\isacharparenleft}{\kern0pt}L{\isacharparenleft}{\kern0pt}r{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}}.
- \end{tabular}
- \end{proposition}
-
- \noindent With this in place it is also very routine to prove that the
- regular expression matcher defined as
- %
- \begin{center}
- \isa{match\ r\ s\ {\isasymequiv}\ nullable\ {\isacharparenleft}{\kern0pt}r{\isacharbackslash}{\kern0pt}s{\isacharparenright}{\kern0pt}}
- \end{center}
-
- \noindent gives a positive answer if and only if \isa{s\ {\isasymin}\ L{\isacharparenleft}{\kern0pt}r{\isacharparenright}{\kern0pt}}.
- Consequently, this regular expression matching algorithm satisfies the
- usual specification for regular expression matching. While the matcher
- above calculates a provably correct YES/NO answer for whether a regular
- expression matches a string or not, the novel idea of Sulzmann and Lu
- \cite{Sulzmann2014} is to append another phase to this algorithm in order
- to calculate a [lexical] value. We will explain the details next.%
-\end{isamarkuptext}\isamarkuptrue%
-%
-\isadelimdocument
-%
-\endisadelimdocument
-%
-\isatagdocument
-%
-\isamarkupsection{POSIX Regular Expression Matching\label{posixsec}%
-}
-\isamarkuptrue%
-%
-\endisatagdocument
-{\isafolddocument}%
-%
-\isadelimdocument
-%
-\endisadelimdocument
-%
-\begin{isamarkuptext}%
-There have been many previous works that use values for encoding
- \emph{how} a regular expression matches a string.
- The clever idea by Sulzmann and Lu \cite{Sulzmann2014} is to
- define a function on values that mirrors (but inverts) the
- construction of the derivative on regular expressions. \emph{Values}
- are defined as the inductive datatype
-
- \begin{center}
- \isa{v\ {\isacharcolon}{\kern0pt}{\isacharequal}{\kern0pt}}
- \isa{Empty} $\mid$
- \isa{Char\ c} $\mid$
- \isa{Left\ v} $\mid$
- \isa{Right\ v} $\mid$
- \isa{Seq\ v\isactrlsub {\isadigit{1}}\ v\isactrlsub {\isadigit{2}}} $\mid$
- \isa{Stars\ vs}
- \end{center}
-
- \noindent where we use \isa{vs} to stand for a list of
- values. (This is similar to the approach taken by Frisch and
- Cardelli for GREEDY matching \cite{Frisch2004}, and Sulzmann and Lu
- for POSIX matching \cite{Sulzmann2014}). The string underlying a
- value can be calculated by the \isa{flat} function, written
- \isa{{\isacharbar}{\kern0pt}\underline{\hspace{2mm}}{\isacharbar}{\kern0pt}} and defined as:
-
- \begin{center}
- \begin{tabular}[t]{lcl}
- \isa{{\isacharbar}{\kern0pt}Empty{\isacharbar}{\kern0pt}} & $\dn$ & \isa{{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}}\\
- \isa{{\isacharbar}{\kern0pt}Char\ c{\isacharbar}{\kern0pt}} & $\dn$ & \isa{{\isacharbrackleft}{\kern0pt}c{\isacharbrackright}{\kern0pt}}\\
- \isa{{\isacharbar}{\kern0pt}Left\ v{\isacharbar}{\kern0pt}} & $\dn$ & \isa{{\isacharbar}{\kern0pt}v{\isacharbar}{\kern0pt}}\\
- \isa{{\isacharbar}{\kern0pt}Right\ v{\isacharbar}{\kern0pt}} & $\dn$ & \isa{{\isacharbar}{\kern0pt}v{\isacharbar}{\kern0pt}}
- \end{tabular}\hspace{14mm}
- \begin{tabular}[t]{lcl}
- \isa{{\isacharbar}{\kern0pt}Seq\ v\isactrlsub {\isadigit{1}}\ v\isactrlsub {\isadigit{2}}{\isacharbar}{\kern0pt}} & $\dn$ & \isa{{\isacharbar}{\kern0pt}v\isactrlsub {\isadigit{1}}{\isacharbar}{\kern0pt}\ {\isacharat}{\kern0pt}\ {\isacharbar}{\kern0pt}v\isactrlsub {\isadigit{2}}{\isacharbar}{\kern0pt}}\\
- \isa{{\isacharbar}{\kern0pt}Stars\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbar}{\kern0pt}} & $\dn$ & \isa{{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}}\\
- \isa{{\isacharbar}{\kern0pt}Stars\ {\isacharparenleft}{\kern0pt}v\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}vs{\isacharparenright}{\kern0pt}{\isacharbar}{\kern0pt}} & $\dn$ & \isa{{\isacharbar}{\kern0pt}v{\isacharbar}{\kern0pt}\ {\isacharat}{\kern0pt}\ {\isacharbar}{\kern0pt}Stars\ vs{\isacharbar}{\kern0pt}}\\
- \end{tabular}
- \end{center}
-
- \noindent We will sometimes refer to the underlying string of a
- value as \emph{flattened value}. We will also overload our notation and
- use \isa{{\isacharbar}{\kern0pt}vs{\isacharbar}{\kern0pt}} for flattening a list of values and concatenating
- the resulting strings.
-
- Sulzmann and Lu define
- inductively an \emph{inhabitation relation} that associates values to
- regular expressions. We define this relation as
- follows:\footnote{Note that the rule for \isa{Stars} differs from
- our earlier paper \cite{AusafDyckhoffUrban2016}. There we used the
- original definition by Sulzmann and Lu which does not require that
- the values \isa{v\ {\isasymin}\ vs} flatten to a non-empty
- string. The reason for introducing the more restricted version of
- lexical values is convenience later on when reasoning about an
- ordering relation for values.}
-
- \begin{center}
- \begin{tabular}{c@ {\hspace{12mm}}c}\label{prfintros}
- \\[-8mm]
- \isa{\mbox{}\inferrule{\mbox{}}{\mbox{Empty\ {\isacharcolon}{\kern0pt}\ \isactrlbold {\isadigit{1}}}}} &
- \isa{\mbox{}\inferrule{\mbox{}}{\mbox{Char\ c\ {\isacharcolon}{\kern0pt}\ c}}}\\[4mm]
- \isa{\mbox{}\inferrule{\mbox{v\isactrlsub {\isadigit{1}}\ {\isacharcolon}{\kern0pt}\ r\isactrlsub {\isadigit{1}}}}{\mbox{Left\ v\isactrlsub {\isadigit{1}}\ {\isacharcolon}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}}}} &
- \isa{\mbox{}\inferrule{\mbox{v\isactrlsub {\isadigit{2}}\ {\isacharcolon}{\kern0pt}\ r\isactrlsub {\isadigit{1}}}}{\mbox{Right\ v\isactrlsub {\isadigit{2}}\ {\isacharcolon}{\kern0pt}\ r\isactrlsub {\isadigit{2}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{1}}}}}\\[4mm]
- \isa{\mbox{}\inferrule{\mbox{v\isactrlsub {\isadigit{1}}\ {\isacharcolon}{\kern0pt}\ r\isactrlsub {\isadigit{1}}}\\\ \mbox{v\isactrlsub {\isadigit{2}}\ {\isacharcolon}{\kern0pt}\ r\isactrlsub {\isadigit{2}}}}{\mbox{Seq\ v\isactrlsub {\isadigit{1}}\ v\isactrlsub {\isadigit{2}}\ {\isacharcolon}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}}}} &
- \isa{\mbox{}\inferrule{\mbox{{\isasymforall}v{\isasymin}vs{\isachardot}{\kern0pt}\ v\ {\isacharcolon}{\kern0pt}\ r\ {\isasymand}\ {\isacharbar}{\kern0pt}v{\isacharbar}{\kern0pt}\ {\isasymnoteq}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}}}{\mbox{Stars\ vs\ {\isacharcolon}{\kern0pt}\ r\isactrlsup {\isasymstar}}}}
- \end{tabular}
- \end{center}
-
- \noindent where in the clause for \isa{Stars} we use the
- notation \isa{v\ {\isasymin}\ vs} for indicating that \isa{v} is a
- member in the list \isa{vs}. We require in this rule that every
- value in \isa{vs} flattens to a non-empty string. The idea is that
- \isa{Stars}-values satisfy the informal Star Rule (see Introduction)
- where the $^\star$ does not match the empty string unless this is
- the only match for the repetition. Note also that no values are
- associated with the regular expression \isa{\isactrlbold {\isadigit{0}}}, and that the
- only value associated with the regular expression \isa{\isactrlbold {\isadigit{1}}} is
- \isa{Empty}. It is routine to establish how values ``inhabiting''
- a regular expression correspond to the language of a regular
- expression, namely
-
- \begin{proposition}\label{inhabs}
- \isa{L{\isacharparenleft}{\kern0pt}r{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharbraceleft}{\kern0pt}{\isacharbar}{\kern0pt}v{\isacharbar}{\kern0pt}\ \mbox{\boldmath$\mid$}\ v\ {\isacharcolon}{\kern0pt}\ r{\isacharbraceright}{\kern0pt}}
- \end{proposition}
-
- \noindent
- Given a regular expression \isa{r} and a string \isa{s}, we define the
- set of all \emph{Lexical Values} inhabited by \isa{r} with the underlying string
- being \isa{s}:\footnote{Okui and Suzuki refer to our lexical values
- as \emph{canonical values} in \cite{OkuiSuzuki2010}. The notion of \emph{non-problematic
- values} by Cardelli and Frisch \cite{Frisch2004} is related, but not identical
- to our lexical values.}
-
- \begin{center}
- \isa{LV\ r\ s\ {\isasymequiv}\ {\isacharbraceleft}{\kern0pt}v\ \mbox{\boldmath$\mid$}\ v\ {\isacharcolon}{\kern0pt}\ r\ {\isasymand}\ {\isacharbar}{\kern0pt}v{\isacharbar}{\kern0pt}\ {\isacharequal}{\kern0pt}\ s{\isacharbraceright}{\kern0pt}}
- \end{center}
-
- \noindent The main property of \isa{LV\ r\ s} is that it is alway finite.
-
- \begin{proposition}
- \isa{finite\ {\isacharparenleft}{\kern0pt}LV\ r\ s{\isacharparenright}{\kern0pt}}
- \end{proposition}
-
- \noindent This finiteness property does not hold in general if we
- remove the side-condition about \isa{{\isacharbar}{\kern0pt}v{\isacharbar}{\kern0pt}\ {\isasymnoteq}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}} in the
- \isa{Stars}-rule above. For example using Sulzmann and Lu's
- less restrictive definition, \isa{LV\ {\isacharparenleft}{\kern0pt}\isactrlbold {\isadigit{1}}\isactrlsup {\isasymstar}{\isacharparenright}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}} would contain
- infinitely many values, but according to our more restricted
- definition only a single value, namely \isa{LV\ {\isacharparenleft}{\kern0pt}\isactrlbold {\isadigit{1}}\isactrlsup {\isasymstar}{\isacharparenright}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharbraceleft}{\kern0pt}Stars\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}}.
-
- If a regular expression \isa{r} matches a string \isa{s}, then
- generally the set \isa{LV\ r\ s} is not just a singleton set. In
- case of POSIX matching the problem is to calculate the unique lexical value
- that satisfies the (informal) POSIX rules from the Introduction.
- Graphically the POSIX value calculation algorithm by Sulzmann and Lu
- can be illustrated by the picture in Figure~\ref{Sulz} where the
- path from the left to the right involving \isa{derivatives}/\isa{nullable} is the first phase of the algorithm
- (calculating successive \Brz's derivatives) and \isa{mkeps}/\isa{inj}, the path from right to left, the second
- phase. This picture shows the steps required when a regular
- expression, say \isa{r\isactrlsub {\isadigit{1}}}, matches the string \isa{{\isacharbrackleft}{\kern0pt}a{\isacharcomma}{\kern0pt}\ b{\isacharcomma}{\kern0pt}\ c{\isacharbrackright}{\kern0pt}}. We first build the three derivatives (according to
- \isa{a}, \isa{b} and \isa{c}). We then use \isa{nullable}
- to find out whether the resulting derivative regular expression
- \isa{r\isactrlsub {\isadigit{4}}} can match the empty string. If yes, we call the
- function \isa{mkeps} that produces a value \isa{v\isactrlsub {\isadigit{4}}}
- for how \isa{r\isactrlsub {\isadigit{4}}} can match the empty string (taking into
- account the POSIX constraints in case there are several ways). This
- function is defined by the clauses:
-
-\begin{figure}[t]
-\begin{center}
-\begin{tikzpicture}[scale=2,node distance=1.3cm,
- every node/.style={minimum size=6mm}]
-\node (r1) {\isa{r\isactrlsub {\isadigit{1}}}};
-\node (r2) [right=of r1]{\isa{r\isactrlsub {\isadigit{2}}}};
-\draw[->,line width=1mm](r1)--(r2) node[above,midway] {\isa{\underline{\hspace{2mm}}{\isacharbackslash}{\kern0pt}a}};
-\node (r3) [right=of r2]{\isa{r\isactrlsub {\isadigit{3}}}};
-\draw[->,line width=1mm](r2)--(r3) node[above,midway] {\isa{\underline{\hspace{2mm}}{\isacharbackslash}{\kern0pt}b}};
-\node (r4) [right=of r3]{\isa{r\isactrlsub {\isadigit{4}}}};
-\draw[->,line width=1mm](r3)--(r4) node[above,midway] {\isa{\underline{\hspace{2mm}}{\isacharbackslash}{\kern0pt}c}};
-\draw (r4) node[anchor=west] {\;\raisebox{3mm}{\isa{nullable}}};
-\node (v4) [below=of r4]{\isa{v\isactrlsub {\isadigit{4}}}};
-\draw[->,line width=1mm](r4) -- (v4);
-\node (v3) [left=of v4] {\isa{v\isactrlsub {\isadigit{3}}}};
-\draw[->,line width=1mm](v4)--(v3) node[below,midway] {\isa{inj\ r\isactrlsub {\isadigit{3}}\ c}};
-\node (v2) [left=of v3]{\isa{v\isactrlsub {\isadigit{2}}}};
-\draw[->,line width=1mm](v3)--(v2) node[below,midway] {\isa{inj\ r\isactrlsub {\isadigit{2}}\ b}};
-\node (v1) [left=of v2] {\isa{v\isactrlsub {\isadigit{1}}}};
-\draw[->,line width=1mm](v2)--(v1) node[below,midway] {\isa{inj\ r\isactrlsub {\isadigit{1}}\ a}};
-\draw (r4) node[anchor=north west] {\;\raisebox{-8mm}{\isa{mkeps}}};
-\end{tikzpicture}
-\end{center}
-\mbox{}\\[-13mm]
-
-\caption{The two phases of the algorithm by Sulzmann \& Lu \cite{Sulzmann2014},
-matching the string \isa{{\isacharbrackleft}{\kern0pt}a{\isacharcomma}{\kern0pt}\ b{\isacharcomma}{\kern0pt}\ c{\isacharbrackright}{\kern0pt}}. The first phase (the arrows from
-left to right) is \Brz's matcher building successive derivatives. If the
-last regular expression is \isa{nullable}, then the functions of the
-second phase are called (the top-down and right-to-left arrows): first
-\isa{mkeps} calculates a value \isa{v\isactrlsub {\isadigit{4}}} witnessing
-how the empty string has been recognised by \isa{r\isactrlsub {\isadigit{4}}}. After
-that the function \isa{inj} ``injects back'' the characters of the string into
-the values.
-\label{Sulz}}
-\end{figure}
-
- \begin{center}
- \begin{tabular}{lcl}
- \isa{mkeps\ \isactrlbold {\isadigit{1}}} & $\dn$ & \isa{Empty}\\
- \isa{mkeps\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{Seq\ {\isacharparenleft}{\kern0pt}mkeps\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}mkeps\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}\\
- \isa{mkeps\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{\textrm{if}\ nullable\ r\isactrlsub {\isadigit{1}}\ \textrm{then}\ Left\ {\isacharparenleft}{\kern0pt}mkeps\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ \textrm{else}\ Right\ {\isacharparenleft}{\kern0pt}mkeps\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}\\
- \isa{mkeps\ {\isacharparenleft}{\kern0pt}r\isactrlsup {\isasymstar}{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{Stars\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}}\\
- \end{tabular}
- \end{center}
-
- \noindent Note that this function needs only to be partially defined,
- namely only for regular expressions that are nullable. In case \isa{nullable} fails, the string \isa{{\isacharbrackleft}{\kern0pt}a{\isacharcomma}{\kern0pt}\ b{\isacharcomma}{\kern0pt}\ c{\isacharbrackright}{\kern0pt}} cannot be matched by \isa{r\isactrlsub {\isadigit{1}}} and the null value \isa{None} is returned. Note also how this function
- makes some subtle choices leading to a POSIX value: for example if an
- alternative regular expression, say \isa{r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}}, can
- match the empty string and furthermore \isa{r\isactrlsub {\isadigit{1}}} can match the
- empty string, then we return a \isa{Left}-value. The \isa{Right}-value will only be returned if \isa{r\isactrlsub {\isadigit{1}}} cannot match the empty
- string.
-
- The most interesting idea from Sulzmann and Lu \cite{Sulzmann2014} is
- the construction of a value for how \isa{r\isactrlsub {\isadigit{1}}} can match the
- string \isa{{\isacharbrackleft}{\kern0pt}a{\isacharcomma}{\kern0pt}\ b{\isacharcomma}{\kern0pt}\ c{\isacharbrackright}{\kern0pt}} from the value how the last derivative, \isa{r\isactrlsub {\isadigit{4}}} in Fig.~\ref{Sulz}, can match the empty string. Sulzmann and
- Lu achieve this by stepwise ``injecting back'' the characters into the
- values thus inverting the operation of building derivatives, but on the level
- of values. The corresponding function, called \isa{inj}, takes three
- arguments, a regular expression, a character and a value. For example in
- the first (or right-most) \isa{inj}-step in Fig.~\ref{Sulz} the regular
- expression \isa{r\isactrlsub {\isadigit{3}}}, the character \isa{c} from the last
- derivative step and \isa{v\isactrlsub {\isadigit{4}}}, which is the value corresponding
- to the derivative regular expression \isa{r\isactrlsub {\isadigit{4}}}. The result is
- the new value \isa{v\isactrlsub {\isadigit{3}}}. The final result of the algorithm is
- the value \isa{v\isactrlsub {\isadigit{1}}}. The \isa{inj} function is defined by recursion on regular
- expressions and by analysing the shape of values (corresponding to
- the derivative regular expressions).
- %
- \begin{center}
- \begin{tabular}{l@ {\hspace{5mm}}lcl}
- \textit{(1)} & \isa{inj\ d\ c\ {\isacharparenleft}{\kern0pt}Empty{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{Char\ d}\\
- \textit{(2)} & \isa{inj\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ c\ {\isacharparenleft}{\kern0pt}Left\ v\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}} & $\dn$ &
- \isa{Left\ {\isacharparenleft}{\kern0pt}inj\ r\isactrlsub {\isadigit{1}}\ c\ v\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}}\\
- \textit{(3)} & \isa{inj\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ c\ {\isacharparenleft}{\kern0pt}Right\ v\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}} & $\dn$ &
- \isa{Right\ {\isacharparenleft}{\kern0pt}inj\ r\isactrlsub {\isadigit{2}}\ c\ v\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}\\
- \textit{(4)} & \isa{inj\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ c\ {\isacharparenleft}{\kern0pt}Seq\ v\isactrlsub {\isadigit{1}}\ v\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}} & $\dn$
- & \isa{Seq\ {\isacharparenleft}{\kern0pt}inj\ r\isactrlsub {\isadigit{1}}\ c\ v\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ v\isactrlsub {\isadigit{2}}}\\
- \textit{(5)} & \isa{inj\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ c\ {\isacharparenleft}{\kern0pt}Left\ {\isacharparenleft}{\kern0pt}Seq\ v\isactrlsub {\isadigit{1}}\ v\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}} & $\dn$
- & \isa{Seq\ {\isacharparenleft}{\kern0pt}inj\ r\isactrlsub {\isadigit{1}}\ c\ v\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ v\isactrlsub {\isadigit{2}}}\\
- \textit{(6)} & \isa{inj\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ c\ {\isacharparenleft}{\kern0pt}Right\ v\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}} & $\dn$
- & \isa{Seq\ {\isacharparenleft}{\kern0pt}mkeps\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}inj\ r\isactrlsub {\isadigit{2}}\ c\ v\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}\\
- \textit{(7)} & \isa{inj\ {\isacharparenleft}{\kern0pt}r\isactrlsup {\isasymstar}{\isacharparenright}{\kern0pt}\ c\ {\isacharparenleft}{\kern0pt}Seq\ v\ {\isacharparenleft}{\kern0pt}Stars\ vs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}} & $\dn$
- & \isa{Stars\ {\isacharparenleft}{\kern0pt}inj\ r\ c\ v\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}vs{\isacharparenright}{\kern0pt}}\\
- \end{tabular}
- \end{center}
-
- \noindent To better understand what is going on in this definition it
- might be instructive to look first at the three sequence cases (clauses
- \textit{(4)} -- \textit{(6)}). In each case we need to construct an ``injected value'' for
- \isa{r\isactrlsub {\isadigit{1}}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}}. This must be a value of the form \isa{Seq\ \underline{\hspace{2mm}}\ \underline{\hspace{2mm}}}\,. Recall the clause of the \isa{derivative}-function
- for sequence regular expressions:
-
- \begin{center}
- \isa{{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharbackslash}{\kern0pt}c} $\dn$ \isa{\textrm{if}\ nullable\ r\isactrlsub {\isadigit{1}}\ \textrm{then}\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}\ {\isacharplus}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ \textrm{else}\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}}
- \end{center}
-
- \noindent Consider first the \isa{else}-branch where the derivative is \isa{{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}}. The corresponding value must therefore
- be of the form \isa{Seq\ v\isactrlsub {\isadigit{1}}\ v\isactrlsub {\isadigit{2}}}, which matches the left-hand
- side in clause~\textit{(4)} of \isa{inj}. In the \isa{if}-branch the derivative is an
- alternative, namely \isa{{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}\ {\isacharplus}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}}. This means we either have to consider a \isa{Left}- or
- \isa{Right}-value. In case of the \isa{Left}-value we know further it
- must be a value for a sequence regular expression. Therefore the pattern
- we match in the clause \textit{(5)} is \isa{Left\ {\isacharparenleft}{\kern0pt}Seq\ v\isactrlsub {\isadigit{1}}\ v\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}},
- while in \textit{(6)} it is just \isa{Right\ v\isactrlsub {\isadigit{2}}}. One more interesting
- point is in the right-hand side of clause \textit{(6)}: since in this case the
- regular expression \isa{r\isactrlsub {\isadigit{1}}} does not ``contribute'' to
- matching the string, that means it only matches the empty string, we need to
- call \isa{mkeps} in order to construct a value for how \isa{r\isactrlsub {\isadigit{1}}}
- can match this empty string. A similar argument applies for why we can
- expect in the left-hand side of clause \textit{(7)} that the value is of the form
- \isa{Seq\ v\ {\isacharparenleft}{\kern0pt}Stars\ vs{\isacharparenright}{\kern0pt}}---the derivative of a star is \isa{{\isacharparenleft}{\kern0pt}r{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isasymcdot}\ r\isactrlsup {\isasymstar}}. Finally, the reason for why we can ignore the second argument
- in clause \textit{(1)} of \isa{inj} is that it will only ever be called in cases
- where \isa{c\ {\isacharequal}{\kern0pt}\ d}, but the usual linearity restrictions in patterns do
- not allow us to build this constraint explicitly into our function
- definition.\footnote{Sulzmann and Lu state this clause as \isa{inj\ c\ c\ {\isacharparenleft}{\kern0pt}Empty{\isacharparenright}{\kern0pt}} $\dn$ \isa{Char\ c},
- but our deviation is harmless.}
-
- The idea of the \isa{inj}-function to ``inject'' a character, say
- \isa{c}, into a value can be made precise by the first part of the
- following lemma, which shows that the underlying string of an injected
- value has a prepended character \isa{c}; the second part shows that
- the underlying string of an \isa{mkeps}-value is always the empty
- string (given the regular expression is nullable since otherwise
- \isa{mkeps} might not be defined).
-
- \begin{lemma}\mbox{}\smallskip\\\label{Prf_injval_flat}
- \begin{tabular}{ll}
- (1) & \isa{{\normalsize{}If\,}\ v\ {\isacharcolon}{\kern0pt}\ r{\isacharbackslash}{\kern0pt}c\ {\normalsize \,then\,}\ {\isacharbar}{\kern0pt}inj\ r\ c\ v{\isacharbar}{\kern0pt}\ {\isacharequal}{\kern0pt}\ c\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}{\isacharbar}{\kern0pt}v{\isacharbar}{\kern0pt}{\isachardot}{\kern0pt}}\\
- (2) & \isa{{\normalsize{}If\,}\ nullable\ r\ {\normalsize \,then\,}\ {\isacharbar}{\kern0pt}mkeps\ r{\isacharbar}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isachardot}{\kern0pt}}
- \end{tabular}
- \end{lemma}
-
- \begin{proof}
- Both properties are by routine inductions: the first one can, for example,
- be proved by induction over the definition of \isa{derivatives}; the second by
- an induction on \isa{r}. There are no interesting cases.\qed
- \end{proof}
-
- Having defined the \isa{mkeps} and \isa{inj} function we can extend
- \Brz's matcher so that a value is constructed (assuming the
- regular expression matches the string). The clauses of the Sulzmann and Lu lexer are
-
- \begin{center}
- \begin{tabular}{lcl}
- \isa{lexer\ r\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}} & $\dn$ & \isa{\textrm{if}\ nullable\ r\ \textrm{then}\ Some\ {\isacharparenleft}{\kern0pt}mkeps\ r{\isacharparenright}{\kern0pt}\ \textrm{else}\ None}\\
- \isa{lexer\ r\ {\isacharparenleft}{\kern0pt}c\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}s{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{case} \isa{lexer\ {\isacharparenleft}{\kern0pt}r{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ s} \isa{of}\\
- & & \phantom{$|$} \isa{None} \isa{{\isasymRightarrow}} \isa{None}\\
- & & $|$ \isa{Some\ v} \isa{{\isasymRightarrow}} \isa{Some\ {\isacharparenleft}{\kern0pt}inj\ r\ c\ v{\isacharparenright}{\kern0pt}}
- \end{tabular}
- \end{center}
-
- \noindent If the regular expression does not match the string, \isa{None} is
- returned. If the regular expression \emph{does}
- match the string, then \isa{Some} value is returned. One important
- virtue of this algorithm is that it can be implemented with ease in any
- functional programming language and also in Isabelle/HOL. In the remaining
- part of this section we prove that this algorithm is correct.
-
- The well-known idea of POSIX matching is informally defined by some
- rules such as the Longest Match and Priority Rules (see
- Introduction); as correctly argued in \cite{Sulzmann2014}, this
- needs formal specification. Sulzmann and Lu define an ``ordering
- relation'' between values and argue that there is a maximum value,
- as given by the derivative-based algorithm. In contrast, we shall
- introduce a simple inductive definition that specifies directly what
- a \emph{POSIX value} is, incorporating the POSIX-specific choices
- into the side-conditions of our rules. Our definition is inspired by
- the matching relation given by Vansummeren~\cite{Vansummeren2006}.
- The relation we define is ternary and
- written as \mbox{\isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v}}, relating
- strings, regular expressions and values; the inductive rules are given in
- Figure~\ref{POSIXrules}.
- We can prove that given a string \isa{s} and regular expression \isa{r}, the POSIX value \isa{v} is uniquely determined by \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v}.
-
- %
- \begin{figure}[t]
- \begin{center}
- \begin{tabular}{c}
- \isa{\mbox{}\inferrule{\mbox{}}{\mbox{{\isacharparenleft}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharcomma}{\kern0pt}\ \isactrlbold {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ Empty}}}\isa{P}\isa{\isactrlbold {\isadigit{1}}} \qquad
- \isa{\mbox{}\inferrule{\mbox{}}{\mbox{{\isacharparenleft}{\kern0pt}{\isacharbrackleft}{\kern0pt}c{\isacharbrackright}{\kern0pt}{\isacharcomma}{\kern0pt}\ c{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ Char\ c}}}\isa{P}\isa{c}\medskip\\
- \isa{\mbox{}\inferrule{\mbox{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v}}{\mbox{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ Left\ v}}}\isa{P{\isacharplus}{\kern0pt}L}\qquad
- \isa{\mbox{}\inferrule{\mbox{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v}\\\ \mbox{s\ {\isasymnotin}\ L{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}}}{\mbox{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ Right\ v}}}\isa{P{\isacharplus}{\kern0pt}R}\medskip\\
- $\mprset{flushleft}
- \inferrule
- {\isa{{\isacharparenleft}{\kern0pt}s\isactrlsub {\isadigit{1}}{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v\isactrlsub {\isadigit{1}}} \qquad
- \isa{{\isacharparenleft}{\kern0pt}s\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v\isactrlsub {\isadigit{2}}} \\\\
- \isa{{\isasymnexists}s\isactrlsub {\isadigit{3}}\ s\isactrlsub {\isadigit{4}}{\isachardot}{\kern0pt}a{\isachardot}{\kern0pt}\ s\isactrlsub {\isadigit{3}}\ {\isasymnoteq}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isasymand}\ s\isactrlsub {\isadigit{3}}\ {\isacharat}{\kern0pt}\ s\isactrlsub {\isadigit{4}}\ {\isacharequal}{\kern0pt}\ s\isactrlsub {\isadigit{2}}\ {\isasymand}\ s\isactrlsub {\isadigit{1}}\ {\isacharat}{\kern0pt}\ s\isactrlsub {\isadigit{3}}\ {\isasymin}\ L{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isasymand}\ s\isactrlsub {\isadigit{4}}\ {\isasymin}\ L{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}}
- {\isa{{\isacharparenleft}{\kern0pt}s\isactrlsub {\isadigit{1}}\ {\isacharat}{\kern0pt}\ s\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ Seq\ v\isactrlsub {\isadigit{1}}\ v\isactrlsub {\isadigit{2}}}}$\isa{PS}\\
- \isa{\mbox{}\inferrule{\mbox{}}{\mbox{{\isacharparenleft}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharcomma}{\kern0pt}\ r\isactrlsup {\isasymstar}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ Stars\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}}}}\isa{P{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}}\medskip\\
- $\mprset{flushleft}
- \inferrule
- {\isa{{\isacharparenleft}{\kern0pt}s\isactrlsub {\isadigit{1}}{\isacharcomma}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v} \qquad
- \isa{{\isacharparenleft}{\kern0pt}s\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ r\isactrlsup {\isasymstar}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ Stars\ vs} \qquad
- \isa{{\isacharbar}{\kern0pt}v{\isacharbar}{\kern0pt}\ {\isasymnoteq}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}} \\\\
- \isa{{\isasymnexists}s\isactrlsub {\isadigit{3}}\ s\isactrlsub {\isadigit{4}}{\isachardot}{\kern0pt}a{\isachardot}{\kern0pt}\ s\isactrlsub {\isadigit{3}}\ {\isasymnoteq}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isasymand}\ s\isactrlsub {\isadigit{3}}\ {\isacharat}{\kern0pt}\ s\isactrlsub {\isadigit{4}}\ {\isacharequal}{\kern0pt}\ s\isactrlsub {\isadigit{2}}\ {\isasymand}\ s\isactrlsub {\isadigit{1}}\ {\isacharat}{\kern0pt}\ s\isactrlsub {\isadigit{3}}\ {\isasymin}\ L{\isacharparenleft}{\kern0pt}r{\isacharparenright}{\kern0pt}\ {\isasymand}\ s\isactrlsub {\isadigit{4}}\ {\isasymin}\ L{\isacharparenleft}{\kern0pt}r\isactrlsup {\isasymstar}{\isacharparenright}{\kern0pt}}}
- {\isa{{\isacharparenleft}{\kern0pt}s\isactrlsub {\isadigit{1}}\ {\isacharat}{\kern0pt}\ s\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ r\isactrlsup {\isasymstar}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ Stars\ {\isacharparenleft}{\kern0pt}v\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}vs{\isacharparenright}{\kern0pt}}}$\isa{P{\isasymstar}}
- \end{tabular}
- \end{center}
- \caption{Our inductive definition of POSIX values.}\label{POSIXrules}
- \end{figure}
-
-
-
- \begin{theorem}\mbox{}\smallskip\\\label{posixdeterm}
- \begin{tabular}{ll}
- (1) & If \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v} then \isa{s\ {\isasymin}\ L{\isacharparenleft}{\kern0pt}r{\isacharparenright}{\kern0pt}} and \isa{{\isacharbar}{\kern0pt}v{\isacharbar}{\kern0pt}\ {\isacharequal}{\kern0pt}\ s}.\\
- (2) & \isa{{\normalsize{}If\,}\ \mbox{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v}\ {\normalsize \,and\,}\ \mbox{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v{\isacharprime}{\kern0pt}}\ {\normalsize \,then\,}\ v\ {\isacharequal}{\kern0pt}\ v{\isacharprime}{\kern0pt}{\isachardot}{\kern0pt}}
- \end{tabular}
- \end{theorem}
-
- \begin{proof} Both by induction on the definition of \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v}.
- The second parts follows by a case analysis of \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v{\isacharprime}{\kern0pt}} and
- the first part.\qed
- \end{proof}
-
- \noindent
- We claim that our \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v} relation captures the idea behind the four
- informal POSIX rules shown in the Introduction: Consider for example the
- rules \isa{P{\isacharplus}{\kern0pt}L} and \isa{P{\isacharplus}{\kern0pt}R} where the POSIX value for a string
- and an alternative regular expression, that is \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}},
- is specified---it is always a \isa{Left}-value, \emph{except} when the
- string to be matched is not in the language of \isa{r\isactrlsub {\isadigit{1}}}; only then it
- is a \isa{Right}-value (see the side-condition in \isa{P{\isacharplus}{\kern0pt}R}).
- Interesting is also the rule for sequence regular expressions (\isa{PS}). The first two premises state that \isa{v\isactrlsub {\isadigit{1}}} and \isa{v\isactrlsub {\isadigit{2}}}
- are the POSIX values for \isa{{\isacharparenleft}{\kern0pt}s\isactrlsub {\isadigit{1}}{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}} and \isa{{\isacharparenleft}{\kern0pt}s\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}
- respectively. Consider now the third premise and note that the POSIX value
- of this rule should match the string \mbox{\isa{s\isactrlsub {\isadigit{1}}\ {\isacharat}{\kern0pt}\ s\isactrlsub {\isadigit{2}}}}. According to the
- Longest Match Rule, we want that the \isa{s\isactrlsub {\isadigit{1}}} is the longest initial
- split of \mbox{\isa{s\isactrlsub {\isadigit{1}}\ {\isacharat}{\kern0pt}\ s\isactrlsub {\isadigit{2}}}} such that \isa{s\isactrlsub {\isadigit{2}}} is still recognised
- by \isa{r\isactrlsub {\isadigit{2}}}. Let us assume, contrary to the third premise, that there
- \emph{exist} an \isa{s\isactrlsub {\isadigit{3}}} and \isa{s\isactrlsub {\isadigit{4}}} such that \isa{s\isactrlsub {\isadigit{2}}}
- can be split up into a non-empty string \isa{s\isactrlsub {\isadigit{3}}} and a possibly empty
- string \isa{s\isactrlsub {\isadigit{4}}}. Moreover the longer string \isa{s\isactrlsub {\isadigit{1}}\ {\isacharat}{\kern0pt}\ s\isactrlsub {\isadigit{3}}} can be
- matched by \isa{r\isactrlsub {\isadigit{1}}} and the shorter \isa{s\isactrlsub {\isadigit{4}}} can still be
- matched by \isa{r\isactrlsub {\isadigit{2}}}. In this case \isa{s\isactrlsub {\isadigit{1}}} would \emph{not} be the
- longest initial split of \mbox{\isa{s\isactrlsub {\isadigit{1}}\ {\isacharat}{\kern0pt}\ s\isactrlsub {\isadigit{2}}}} and therefore \isa{Seq\ v\isactrlsub {\isadigit{1}}\ v\isactrlsub {\isadigit{2}}} cannot be a POSIX value for \isa{{\isacharparenleft}{\kern0pt}s\isactrlsub {\isadigit{1}}\ {\isacharat}{\kern0pt}\ s\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}.
- The main point is that our side-condition ensures the Longest
- Match Rule is satisfied.
-
- A similar condition is imposed on the POSIX value in the \isa{P{\isasymstar}}-rule. Also there we want that \isa{s\isactrlsub {\isadigit{1}}} is the longest initial
- split of \isa{s\isactrlsub {\isadigit{1}}\ {\isacharat}{\kern0pt}\ s\isactrlsub {\isadigit{2}}} and furthermore the corresponding value
- \isa{v} cannot be flattened to the empty string. In effect, we require
- that in each ``iteration'' of the star, some non-empty substring needs to
- be ``chipped'' away; only in case of the empty string we accept \isa{Stars\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}} as the POSIX value. Indeed we can show that our POSIX values
- are lexical values which exclude those \isa{Stars} that contain subvalues
- that flatten to the empty string.
-
- \begin{lemma}\label{LVposix}
- \isa{{\normalsize{}If\,}\ {\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v\ {\normalsize \,then\,}\ v\ {\isasymin}\ LV\ r\ s{\isachardot}{\kern0pt}}
- \end{lemma}
-
- \begin{proof}
- By routine induction on \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v}.\qed
- \end{proof}
-
- \noindent
- Next is the lemma that shows the function \isa{mkeps} calculates
- the POSIX value for the empty string and a nullable regular expression.
-
- \begin{lemma}\label{lemmkeps}
- \isa{{\normalsize{}If\,}\ nullable\ r\ {\normalsize \,then\,}\ {\isacharparenleft}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharcomma}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ mkeps\ r{\isachardot}{\kern0pt}}
- \end{lemma}
-
- \begin{proof}
- By routine induction on \isa{r}.\qed
- \end{proof}
-
- \noindent
- The central lemma for our POSIX relation is that the \isa{inj}-function
- preserves POSIX values.
-
- \begin{lemma}\label{Posix2}
- \isa{{\normalsize{}If\,}\ {\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v\ {\normalsize \,then\,}\ {\isacharparenleft}{\kern0pt}c\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}s{\isacharcomma}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ inj\ r\ c\ v{\isachardot}{\kern0pt}}
- \end{lemma}
-
- \begin{proof}
- By induction on \isa{r}. We explain two cases.
-
- \begin{itemize}
- \item[$\bullet$] Case \isa{r\ {\isacharequal}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}}. There are
- two subcases, namely \isa{{\isacharparenleft}{\kern0pt}a{\isacharparenright}{\kern0pt}} \mbox{\isa{v\ {\isacharequal}{\kern0pt}\ Left\ v{\isacharprime}{\kern0pt}}} and \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v{\isacharprime}{\kern0pt}}; and \isa{{\isacharparenleft}{\kern0pt}b{\isacharparenright}{\kern0pt}} \isa{v\ {\isacharequal}{\kern0pt}\ Right\ v{\isacharprime}{\kern0pt}}, \isa{s\ {\isasymnotin}\ L{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}} and \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v{\isacharprime}{\kern0pt}}. In \isa{{\isacharparenleft}{\kern0pt}a{\isacharparenright}{\kern0pt}} we
- know \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v{\isacharprime}{\kern0pt}}, from which we can infer \isa{{\isacharparenleft}{\kern0pt}c\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}s{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ inj\ r\isactrlsub {\isadigit{1}}\ c\ v{\isacharprime}{\kern0pt}} by induction hypothesis and hence \isa{{\isacharparenleft}{\kern0pt}c\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}s{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ inj\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ c\ {\isacharparenleft}{\kern0pt}Left\ v{\isacharprime}{\kern0pt}{\isacharparenright}{\kern0pt}} as needed. Similarly
- in subcase \isa{{\isacharparenleft}{\kern0pt}b{\isacharparenright}{\kern0pt}} where, however, in addition we have to use
- Proposition~\ref{derprop}(2) in order to infer \isa{c\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}s\ {\isasymnotin}\ L{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}} from \isa{s\ {\isasymnotin}\ L{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}}.\smallskip
-
- \item[$\bullet$] Case \isa{r\ {\isacharequal}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}}. There are three subcases:
-
- \begin{quote}
- \begin{description}
- \item[\isa{{\isacharparenleft}{\kern0pt}a{\isacharparenright}{\kern0pt}}] \isa{v\ {\isacharequal}{\kern0pt}\ Left\ {\isacharparenleft}{\kern0pt}Seq\ v\isactrlsub {\isadigit{1}}\ v\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}} and \isa{nullable\ r\isactrlsub {\isadigit{1}}}
- \item[\isa{{\isacharparenleft}{\kern0pt}b{\isacharparenright}{\kern0pt}}] \isa{v\ {\isacharequal}{\kern0pt}\ Right\ v\isactrlsub {\isadigit{1}}} and \isa{nullable\ r\isactrlsub {\isadigit{1}}}
- \item[\isa{{\isacharparenleft}{\kern0pt}c{\isacharparenright}{\kern0pt}}] \isa{v\ {\isacharequal}{\kern0pt}\ Seq\ v\isactrlsub {\isadigit{1}}\ v\isactrlsub {\isadigit{2}}} and \isa{{\isasymnot}\ nullable\ r\isactrlsub {\isadigit{1}}}
- \end{description}
- \end{quote}
-
- \noindent For \isa{{\isacharparenleft}{\kern0pt}a{\isacharparenright}{\kern0pt}} we know \isa{{\isacharparenleft}{\kern0pt}s\isactrlsub {\isadigit{1}}{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v\isactrlsub {\isadigit{1}}} and
- \isa{{\isacharparenleft}{\kern0pt}s\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v\isactrlsub {\isadigit{2}}} as well as
- %
- \[\isa{{\isasymnexists}s\isactrlsub {\isadigit{3}}\ s\isactrlsub {\isadigit{4}}{\isachardot}{\kern0pt}a{\isachardot}{\kern0pt}\ s\isactrlsub {\isadigit{3}}\ {\isasymnoteq}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isasymand}\ s\isactrlsub {\isadigit{3}}\ {\isacharat}{\kern0pt}\ s\isactrlsub {\isadigit{4}}\ {\isacharequal}{\kern0pt}\ s\isactrlsub {\isadigit{2}}\ {\isasymand}\ s\isactrlsub {\isadigit{1}}\ {\isacharat}{\kern0pt}\ s\isactrlsub {\isadigit{3}}\ {\isasymin}\ L{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isasymand}\ s\isactrlsub {\isadigit{4}}\ {\isasymin}\ L{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}\]
-
- \noindent From the latter we can infer by Proposition~\ref{derprop}(2):
- %
- \[\isa{{\isasymnexists}s\isactrlsub {\isadigit{3}}\ s\isactrlsub {\isadigit{4}}{\isachardot}{\kern0pt}a{\isachardot}{\kern0pt}\ s\isactrlsub {\isadigit{3}}\ {\isasymnoteq}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isasymand}\ s\isactrlsub {\isadigit{3}}\ {\isacharat}{\kern0pt}\ s\isactrlsub {\isadigit{4}}\ {\isacharequal}{\kern0pt}\ s\isactrlsub {\isadigit{2}}\ {\isasymand}\ c\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}s\isactrlsub {\isadigit{1}}\ {\isacharat}{\kern0pt}\ s\isactrlsub {\isadigit{3}}\ {\isasymin}\ L{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isasymand}\ s\isactrlsub {\isadigit{4}}\ {\isasymin}\ L{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}\]
-
- \noindent We can use the induction hypothesis for \isa{r\isactrlsub {\isadigit{1}}} to obtain
- \isa{{\isacharparenleft}{\kern0pt}c\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}s\isactrlsub {\isadigit{1}}{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ inj\ r\isactrlsub {\isadigit{1}}\ c\ v\isactrlsub {\isadigit{1}}}. Putting this all together allows us to infer
- \isa{{\isacharparenleft}{\kern0pt}c\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}s\isactrlsub {\isadigit{1}}\ {\isacharat}{\kern0pt}\ s\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ Seq\ {\isacharparenleft}{\kern0pt}inj\ r\isactrlsub {\isadigit{1}}\ c\ v\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ v\isactrlsub {\isadigit{2}}}. The case \isa{{\isacharparenleft}{\kern0pt}c{\isacharparenright}{\kern0pt}}
- is similar.
-
- For \isa{{\isacharparenleft}{\kern0pt}b{\isacharparenright}{\kern0pt}} we know \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v\isactrlsub {\isadigit{1}}} and
- \isa{s\isactrlsub {\isadigit{1}}\ {\isacharat}{\kern0pt}\ s\isactrlsub {\isadigit{2}}\ {\isasymnotin}\ L{\isacharparenleft}{\kern0pt}{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}. From the former
- we have \isa{{\isacharparenleft}{\kern0pt}c\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}s{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ inj\ r\isactrlsub {\isadigit{2}}\ c\ v\isactrlsub {\isadigit{1}}} by induction hypothesis
- for \isa{r\isactrlsub {\isadigit{2}}}. From the latter we can infer
- %
- \[\isa{{\isasymnexists}s\isactrlsub {\isadigit{3}}\ s\isactrlsub {\isadigit{4}}{\isachardot}{\kern0pt}a{\isachardot}{\kern0pt}\ s\isactrlsub {\isadigit{3}}\ {\isasymnoteq}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isasymand}\ s\isactrlsub {\isadigit{3}}\ {\isacharat}{\kern0pt}\ s\isactrlsub {\isadigit{4}}\ {\isacharequal}{\kern0pt}\ c\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}s\ {\isasymand}\ s\isactrlsub {\isadigit{3}}\ {\isasymin}\ L{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isasymand}\ s\isactrlsub {\isadigit{4}}\ {\isasymin}\ L{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}\]
-
- \noindent By Lemma~\ref{lemmkeps} we know \isa{{\isacharparenleft}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ mkeps\ r\isactrlsub {\isadigit{1}}}
- holds. Putting this all together, we can conclude with \isa{{\isacharparenleft}{\kern0pt}c\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}s{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ Seq\ {\isacharparenleft}{\kern0pt}mkeps\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}inj\ r\isactrlsub {\isadigit{2}}\ c\ v\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}}, as required.
-
- Finally suppose \isa{r\ {\isacharequal}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\isactrlsup {\isasymstar}}. This case is very similar to the
- sequence case, except that we need to also ensure that \isa{{\isacharbar}{\kern0pt}inj\ r\isactrlsub {\isadigit{1}}\ c\ v\isactrlsub {\isadigit{1}}{\isacharbar}{\kern0pt}\ {\isasymnoteq}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}}. This follows from \isa{{\isacharparenleft}{\kern0pt}c\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}s\isactrlsub {\isadigit{1}}{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ inj\ r\isactrlsub {\isadigit{1}}\ c\ v\isactrlsub {\isadigit{1}}} (which in turn follows from \isa{{\isacharparenleft}{\kern0pt}s\isactrlsub {\isadigit{1}}{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v\isactrlsub {\isadigit{1}}} and the induction hypothesis).\qed
- \end{itemize}
- \end{proof}
-
- \noindent
- With Lemma~\ref{Posix2} in place, it is completely routine to establish
- that the Sulzmann and Lu lexer satisfies our specification (returning
- the null value \isa{None} iff the string is not in the language of the regular expression,
- and returning a unique POSIX value iff the string \emph{is} in the language):
-
- \begin{theorem}\mbox{}\smallskip\\\label{lexercorrect}
- \begin{tabular}{ll}
- (1) & \isa{s\ {\isasymnotin}\ L{\isacharparenleft}{\kern0pt}r{\isacharparenright}{\kern0pt}} if and only if \isa{lexer\ r\ s\ {\isacharequal}{\kern0pt}\ None}\\
- (2) & \isa{s\ {\isasymin}\ L{\isacharparenleft}{\kern0pt}r{\isacharparenright}{\kern0pt}} if and only if \isa{{\isasymexists}v{\isachardot}{\kern0pt}\ lexer\ r\ s\ {\isacharequal}{\kern0pt}\ Some\ v\ {\isasymand}\ {\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v}\\
- \end{tabular}
- \end{theorem}
-
- \begin{proof}
- By induction on \isa{s} using Lemma~\ref{lemmkeps} and \ref{Posix2}.\qed
- \end{proof}
-
- \noindent In \textit{(2)} we further know by Theorem~\ref{posixdeterm} that the
- value returned by the lexer must be unique. A simple corollary
- of our two theorems is:
-
- \begin{corollary}\mbox{}\smallskip\\\label{lexercorrectcor}
- \begin{tabular}{ll}
- (1) & \isa{lexer\ r\ s\ {\isacharequal}{\kern0pt}\ None} if and only if \isa{{\isasymnexists}v{\isachardot}{\kern0pt}a{\isachardot}{\kern0pt}\ {\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v}\\
- (2) & \isa{lexer\ r\ s\ {\isacharequal}{\kern0pt}\ Some\ v} if and only if \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v}\\
- \end{tabular}
- \end{corollary}
-
- \noindent This concludes our correctness proof. Note that we have
- not changed the algorithm of Sulzmann and Lu,\footnote{All
- deviations we introduced are harmless.} but introduced our own
- specification for what a correct result---a POSIX value---should be.
- In the next section we show that our specification coincides with
- another one given by Okui and Suzuki using a different technique.%
-\end{isamarkuptext}\isamarkuptrue%
-%
-\isadelimdocument
-%
-\endisadelimdocument
-%
-\isatagdocument
-%
-\isamarkupsection{Ordering of Values according to Okui and Suzuki%
-}
-\isamarkuptrue%
-%
-\endisatagdocument
-{\isafolddocument}%
-%
-\isadelimdocument
-%
-\endisadelimdocument
-%
-\begin{isamarkuptext}%
-While in the previous section we have defined POSIX values directly
- in terms of a ternary relation (see inference rules in Figure~\ref{POSIXrules}),
- Sulzmann and Lu took a different approach in \cite{Sulzmann2014}:
- they introduced an ordering for values and identified POSIX values
- as the maximal elements. An extended version of \cite{Sulzmann2014}
- is available at the website of its first author; this includes more
- details of their proofs, but which are evidently not in final form
- yet. Unfortunately, we were not able to verify claims that their
- ordering has properties such as being transitive or having maximal
- elements.
-
- Okui and Suzuki \cite{OkuiSuzuki2010,OkuiSuzukiTech} described
- another ordering of values, which they use to establish the
- correctness of their automata-based algorithm for POSIX matching.
- Their ordering resembles some aspects of the one given by Sulzmann
- and Lu, but overall is quite different. To begin with, Okui and
- Suzuki identify POSIX values as minimal, rather than maximal,
- elements in their ordering. A more substantial difference is that
- the ordering by Okui and Suzuki uses \emph{positions} in order to
- identify and compare subvalues. Positions are lists of natural
- numbers. This allows them to quite naturally formalise the Longest
- Match and Priority rules of the informal POSIX standard. Consider
- for example the value \isa{v}
-
- \begin{center}
- \isa{v\ {\isasymequiv}\ Stars\ {\isacharbrackleft}{\kern0pt}Seq\ {\isacharparenleft}{\kern0pt}Char\ x{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}Char\ y{\isacharparenright}{\kern0pt}{\isacharcomma}{\kern0pt}\ Char\ z{\isacharbrackright}{\kern0pt}}
- \end{center}
-
- \noindent
- At position \isa{{\isacharbrackleft}{\kern0pt}{\isadigit{0}}{\isacharcomma}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}} of this value is the
- subvalue \isa{Char\ y} and at position \isa{{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}} the
- subvalue \isa{Char\ z}. At the `root' position, or empty list
- \isa{{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}}, is the whole value \isa{v}. Positions such as \isa{{\isacharbrackleft}{\kern0pt}{\isadigit{0}}{\isacharcomma}{\kern0pt}{\isadigit{1}}{\isacharcomma}{\kern0pt}{\isadigit{0}}{\isacharbrackright}{\kern0pt}} or \isa{{\isacharbrackleft}{\kern0pt}{\isadigit{2}}{\isacharbrackright}{\kern0pt}} are outside of \isa{v}. If it exists, the subvalue of \isa{v} at a position \isa{p}, written \isa{v\mbox{$\downharpoonleft$}\isactrlbsub p\isactrlesub }, can be recursively defined by
-
- \begin{center}
- \begin{tabular}{r@ {\hspace{0mm}}lcl}
- \isa{v} & \isa{{\isasymdownharpoonleft}\isactrlbsub {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\isactrlesub } & \isa{{\isasymequiv}}& \isa{v}\\
- \isa{Left\ v} & \isa{{\isasymdownharpoonleft}\isactrlbsub {\isadigit{0}}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}ps\isactrlesub } & \isa{{\isasymequiv}}& \isa{v\mbox{$\downharpoonleft$}\isactrlbsub ps\isactrlesub }\\
- \isa{Right\ v} & \isa{{\isasymdownharpoonleft}\isactrlbsub {\isadigit{1}}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}ps\isactrlesub } & \isa{{\isasymequiv}} &
- \isa{v\mbox{$\downharpoonleft$}\isactrlbsub ps\isactrlesub }\\
- \isa{Seq\ v\isactrlsub {\isadigit{1}}\ v\isactrlsub {\isadigit{2}}} & \isa{{\isasymdownharpoonleft}\isactrlbsub {\isadigit{0}}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}ps\isactrlesub } & \isa{{\isasymequiv}} &
- \isa{v\isactrlsub {\isadigit{1}}\mbox{$\downharpoonleft$}\isactrlbsub ps\isactrlesub } \\
- \isa{Seq\ v\isactrlsub {\isadigit{1}}\ v\isactrlsub {\isadigit{2}}} & \isa{{\isasymdownharpoonleft}\isactrlbsub {\isadigit{1}}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}ps\isactrlesub }
- & \isa{{\isasymequiv}} &
- \isa{v\isactrlsub {\isadigit{2}}\mbox{$\downharpoonleft$}\isactrlbsub ps\isactrlesub } \\
- \isa{Stars\ vs} & \isa{{\isasymdownharpoonleft}\isactrlbsub n{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}ps\isactrlesub } & \isa{{\isasymequiv}}& \isa{vs\ensuremath{_{[\mathit{n}]}}\mbox{$\downharpoonleft$}\isactrlbsub ps\isactrlesub }\\
- \end{tabular}
- \end{center}
-
- \noindent In the last clause we use Isabelle's notation \isa{vs\ensuremath{_{[\mathit{n}]}}} for the
- \isa{n}th element in a list. The set of positions inside a value \isa{v},
- written \isa{Pos\ v}, is given by
-
- \begin{center}
- \begin{tabular}{lcl}
- \isa{Pos\ {\isacharparenleft}{\kern0pt}Empty{\isacharparenright}{\kern0pt}} & \isa{{\isasymequiv}} & \isa{{\isacharbraceleft}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}}\\
- \isa{Pos\ {\isacharparenleft}{\kern0pt}Char\ c{\isacharparenright}{\kern0pt}} & \isa{{\isasymequiv}} & \isa{{\isacharbraceleft}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}}\\
- \isa{Pos\ {\isacharparenleft}{\kern0pt}Left\ v{\isacharparenright}{\kern0pt}} & \isa{{\isasymequiv}} & \isa{{\isacharbraceleft}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isasymunion}\ {\isacharbraceleft}{\kern0pt}{\isadigit{0}}\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}ps\ \mbox{\boldmath$\mid$}\ ps\ {\isasymin}\ Pos\ v{\isacharbraceright}{\kern0pt}}\\
- \isa{Pos\ {\isacharparenleft}{\kern0pt}Right\ v{\isacharparenright}{\kern0pt}} & \isa{{\isasymequiv}} & \isa{{\isacharbraceleft}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isasymunion}\ {\isacharbraceleft}{\kern0pt}{\isadigit{1}}\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}ps\ \mbox{\boldmath$\mid$}\ ps\ {\isasymin}\ Pos\ v{\isacharbraceright}{\kern0pt}}\\
- \isa{Pos\ {\isacharparenleft}{\kern0pt}Seq\ v\isactrlsub {\isadigit{1}}\ v\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}
- & \isa{{\isasymequiv}}
- & \isa{{\isacharbraceleft}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isasymunion}\ {\isacharbraceleft}{\kern0pt}{\isadigit{0}}\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}ps\ \mbox{\boldmath$\mid$}\ ps\ {\isasymin}\ Pos\ v\isactrlsub {\isadigit{1}}{\isacharbraceright}{\kern0pt}\ {\isasymunion}\ {\isacharbraceleft}{\kern0pt}{\isadigit{1}}\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}ps\ \mbox{\boldmath$\mid$}\ ps\ {\isasymin}\ Pos\ v\isactrlsub {\isadigit{2}}{\isacharbraceright}{\kern0pt}}\\
- \isa{Pos\ {\isacharparenleft}{\kern0pt}Stars\ vs{\isacharparenright}{\kern0pt}} & \isa{{\isasymequiv}} & \isa{{\isacharbraceleft}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isasymunion}\ {\isacharparenleft}{\kern0pt}{\isasymUnion}n\ {\isacharless}{\kern0pt}\ len\ vs\ {\isacharbraceleft}{\kern0pt}n\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}ps\ \mbox{\boldmath$\mid$}\ ps\ {\isasymin}\ Pos\ vs\ensuremath{_{[\mathit{n}]}}{\isacharbraceright}{\kern0pt}{\isacharparenright}{\kern0pt}}\\
- \end{tabular}
- \end{center}
-
- \noindent
- whereby \isa{len} in the last clause stands for the length of a list. Clearly
- for every position inside a value there exists a subvalue at that position.
-
-
- To help understanding the ordering of Okui and Suzuki, consider again
- the earlier value
- \isa{v} and compare it with the following \isa{w}:
-
- \begin{center}
- \begin{tabular}{l}
- \isa{v\ {\isasymequiv}\ Stars\ {\isacharbrackleft}{\kern0pt}Seq\ {\isacharparenleft}{\kern0pt}Char\ x{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}Char\ y{\isacharparenright}{\kern0pt}{\isacharcomma}{\kern0pt}\ Char\ z{\isacharbrackright}{\kern0pt}}\\
- \isa{w\ {\isasymequiv}\ Stars\ {\isacharbrackleft}{\kern0pt}Char\ x{\isacharcomma}{\kern0pt}\ Char\ y{\isacharcomma}{\kern0pt}\ Char\ z{\isacharbrackright}{\kern0pt}}
- \end{tabular}
- \end{center}
-
- \noindent Both values match the string \isa{xyz}, that means if
- we flatten these values at their respective root position, we obtain
- \isa{xyz}. However, at position \isa{{\isacharbrackleft}{\kern0pt}{\isadigit{0}}{\isacharbrackright}{\kern0pt}}, \isa{v} matches
- \isa{xy} whereas \isa{w} matches only the shorter \isa{x}. So
- according to the Longest Match Rule, we should prefer \isa{v},
- rather than \isa{w} as POSIX value for string \isa{xyz} (and
- corresponding regular expression). In order to
- formalise this idea, Okui and Suzuki introduce a measure for
- subvalues at position \isa{p}, called the \emph{norm} of \isa{v}
- at position \isa{p}. We can define this measure in Isabelle as an
- integer as follows
-
- \begin{center}
- \isa{{\isasymparallel}v{\isasymparallel}\isactrlbsub p\isactrlesub \ {\isasymequiv}\ \textrm{if}\ p\ {\isasymin}\ Pos\ v\ \textrm{then}\ len\ {\isacharbar}{\kern0pt}v\mbox{$\downharpoonleft$}\isactrlbsub p\isactrlesub {\isacharbar}{\kern0pt}\ \textrm{else}\ {\isacharminus}{\kern0pt}\ {\isadigit{1}}}
- \end{center}
-
- \noindent where we take the length of the flattened value at
- position \isa{p}, provided the position is inside \isa{v}; if
- not, then the norm is \isa{{\isacharminus}{\kern0pt}{\isadigit{1}}}. The default for outside
- positions is crucial for the POSIX requirement of preferring a
- \isa{Left}-value over a \isa{Right}-value (if they can match the
- same string---see the Priority Rule from the Introduction). For this
- consider
-
- \begin{center}
- \isa{v\ {\isasymequiv}\ Left\ {\isacharparenleft}{\kern0pt}Char\ x{\isacharparenright}{\kern0pt}} \qquad and \qquad \isa{w\ {\isasymequiv}\ Right\ {\isacharparenleft}{\kern0pt}Char\ x{\isacharparenright}{\kern0pt}}
- \end{center}
-
- \noindent Both values match \isa{x}. At position \isa{{\isacharbrackleft}{\kern0pt}{\isadigit{0}}{\isacharbrackright}{\kern0pt}}
- the norm of \isa{v} is \isa{{\isadigit{1}}} (the subvalue matches \isa{x}),
- but the norm of \isa{w} is \isa{{\isacharminus}{\kern0pt}{\isadigit{1}}} (the position is outside
- \isa{w} according to how we defined the `inside' positions of
- \isa{Left}- and \isa{Right}-values). Of course at position
- \isa{{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}}, the norms \isa{{\isasymparallel}v{\isasymparallel}\isactrlbsub {\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isactrlesub } and \isa{{\isasymparallel}w{\isasymparallel}\isactrlbsub {\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isactrlesub } are reversed, but the point is that subvalues
- will be analysed according to lexicographically ordered
- positions. According to this ordering, the position \isa{{\isacharbrackleft}{\kern0pt}{\isadigit{0}}{\isacharbrackright}{\kern0pt}}
- takes precedence over \isa{{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}} and thus also \isa{v} will be
- preferred over \isa{w}. The lexicographic ordering of positions, written
- \isa{\underline{\hspace{2mm}}\ {\isasymprec}\isactrlbsub lex\isactrlesub \ \underline{\hspace{2mm}}}, can be conveniently formalised
- by three inference rules
-
- \begin{center}
- \begin{tabular}{ccc}
- \isa{\mbox{}\inferrule{\mbox{}}{\mbox{{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isasymprec}\isactrlbsub lex\isactrlesub \ p\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}ps}}}\hspace{1cm} &
- \isa{\mbox{}\inferrule{\mbox{p\isactrlsub {\isadigit{1}}\ {\isacharless}{\kern0pt}\ p\isactrlsub {\isadigit{2}}}}{\mbox{p\isactrlsub {\isadigit{1}}\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}ps\isactrlsub {\isadigit{1}}\ {\isasymprec}\isactrlbsub lex\isactrlesub \ p\isactrlsub {\isadigit{2}}\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}ps\isactrlsub {\isadigit{2}}}}}\hspace{1cm} &
- \isa{\mbox{}\inferrule{\mbox{ps\isactrlsub {\isadigit{1}}\ {\isasymprec}\isactrlbsub lex\isactrlesub \ ps\isactrlsub {\isadigit{2}}}}{\mbox{p\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}ps\isactrlsub {\isadigit{1}}\ {\isasymprec}\isactrlbsub lex\isactrlesub \ p\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}ps\isactrlsub {\isadigit{2}}}}}
- \end{tabular}
- \end{center}
-
- With the norm and lexicographic order in place,
- we can state the key definition of Okui and Suzuki
- \cite{OkuiSuzuki2010}: a value \isa{v\isactrlsub {\isadigit{1}}} is \emph{smaller at position \isa{p}} than
- \isa{v\isactrlsub {\isadigit{2}}}, written \isa{v\isactrlsub {\isadigit{1}}\ {\isasymprec}\isactrlbsub p\isactrlesub \ v\isactrlsub {\isadigit{2}}},
- if and only if $(i)$ the norm at position \isa{p} is
- greater in \isa{v\isactrlsub {\isadigit{1}}} (that is the string \isa{{\isacharbar}{\kern0pt}v\isactrlsub {\isadigit{1}}\mbox{$\downharpoonleft$}\isactrlbsub p\isactrlesub {\isacharbar}{\kern0pt}} is longer
- than \isa{{\isacharbar}{\kern0pt}v\isactrlsub {\isadigit{2}}\mbox{$\downharpoonleft$}\isactrlbsub p\isactrlesub {\isacharbar}{\kern0pt}}) and $(ii)$ all subvalues at
- positions that are inside \isa{v\isactrlsub {\isadigit{1}}} or \isa{v\isactrlsub {\isadigit{2}}} and that are
- lexicographically smaller than \isa{p}, we have the same norm, namely
-
- \begin{center}
- \begin{tabular}{c}
- \isa{v\isactrlsub {\isadigit{1}}\ {\isasymprec}\isactrlbsub p\isactrlesub \ v\isactrlsub {\isadigit{2}}}
- \isa{{\isasymequiv}}
- $\begin{cases}
- (i) & \isa{{\isasymparallel}v\isactrlsub {\isadigit{2}}{\isasymparallel}\isactrlbsub p\isactrlesub \ {\isacharless}{\kern0pt}\ {\isasymparallel}v\isactrlsub {\isadigit{1}}{\isasymparallel}\isactrlbsub p\isactrlesub } \quad\text{and}\smallskip \\
- (ii) & \isa{{\isasymforall}q{\isasymin}Pos\ v\isactrlsub {\isadigit{1}}\ {\isasymunion}\ Pos\ v\isactrlsub {\isadigit{2}}{\isachardot}{\kern0pt}\ q\ {\isasymprec}\isactrlbsub lex\isactrlesub \ p\ {\isasymlongrightarrow}\ {\isasymparallel}v\isactrlsub {\isadigit{1}}{\isasymparallel}\isactrlbsub q\isactrlesub \ {\isacharequal}{\kern0pt}\ {\isasymparallel}v\isactrlsub {\isadigit{2}}{\isasymparallel}\isactrlbsub q\isactrlesub }
- \end{cases}$
- \end{tabular}
- \end{center}
-
- \noindent The position \isa{p} in this definition acts as the
- \emph{first distinct position} of \isa{v\isactrlsub {\isadigit{1}}} and \isa{v\isactrlsub {\isadigit{2}}}, where both values match strings of different length
- \cite{OkuiSuzuki2010}. Since at \isa{p} the values \isa{v\isactrlsub {\isadigit{1}}} and \isa{v\isactrlsub {\isadigit{2}}} match different strings, the
- ordering is irreflexive. Derived from the definition above
- are the following two orderings:
-
- \begin{center}
- \begin{tabular}{l}
- \isa{v\isactrlsub {\isadigit{1}}\ {\isasymprec}\ v\isactrlsub {\isadigit{2}}\ {\isasymequiv}\ {\isasymexists}p{\isachardot}{\kern0pt}\ v\isactrlsub {\isadigit{1}}\ {\isasymprec}\isactrlbsub p\isactrlesub \ v\isactrlsub {\isadigit{2}}}\\
- \isa{v\isactrlsub {\isadigit{1}}\ \mbox{$\preccurlyeq$}\ v\isactrlsub {\isadigit{2}}\ {\isasymequiv}\ v\isactrlsub {\isadigit{1}}\ {\isasymprec}\ v\isactrlsub {\isadigit{2}}\ {\isasymor}\ v\isactrlsub {\isadigit{1}}\ {\isacharequal}{\kern0pt}\ v\isactrlsub {\isadigit{2}}}
- \end{tabular}
- \end{center}
-
- While we encountered a number of obstacles for establishing properties like
- transitivity for the ordering of Sulzmann and Lu (and which we failed
- to overcome), it is relatively straightforward to establish this
- property for the orderings
- \isa{\underline{\hspace{2mm}}\ {\isasymprec}\ \underline{\hspace{2mm}}} and \isa{\underline{\hspace{2mm}}\ \mbox{$\preccurlyeq$}\ \underline{\hspace{2mm}}}
- by Okui and Suzuki.
-
- \begin{lemma}[Transitivity]\label{transitivity}
- \isa{{\normalsize{}If\,}\ \mbox{v\isactrlsub {\isadigit{1}}\ {\isasymprec}\ v\isactrlsub {\isadigit{2}}}\ {\normalsize \,and\,}\ \mbox{v\isactrlsub {\isadigit{2}}\ {\isasymprec}\ v\isactrlsub {\isadigit{3}}}\ {\normalsize \,then\,}\ v\isactrlsub {\isadigit{1}}\ {\isasymprec}\ v\isactrlsub {\isadigit{3}}{\isachardot}{\kern0pt}}
- \end{lemma}
-
- \begin{proof} From the assumption we obtain two positions \isa{p}
- and \isa{q}, where the values \isa{v\isactrlsub {\isadigit{1}}} and \isa{v\isactrlsub {\isadigit{2}}} (respectively \isa{v\isactrlsub {\isadigit{2}}} and \isa{v\isactrlsub {\isadigit{3}}}) are `distinct'. Since \isa{{\isasymprec}\isactrlbsub lex\isactrlesub } is trichotomous, we need to consider
- three cases, namely \isa{p\ {\isacharequal}{\kern0pt}\ q}, \isa{p\ {\isasymprec}\isactrlbsub lex\isactrlesub \ q} and
- \isa{q\ {\isasymprec}\isactrlbsub lex\isactrlesub \ p}. Let us look at the first case. Clearly
- \isa{{\isasymparallel}v\isactrlsub {\isadigit{2}}{\isasymparallel}\isactrlbsub p\isactrlesub \ {\isacharless}{\kern0pt}\ {\isasymparallel}v\isactrlsub {\isadigit{1}}{\isasymparallel}\isactrlbsub p\isactrlesub } and \isa{{\isasymparallel}v\isactrlsub {\isadigit{3}}{\isasymparallel}\isactrlbsub p\isactrlesub \ {\isacharless}{\kern0pt}\ {\isasymparallel}v\isactrlsub {\isadigit{2}}{\isasymparallel}\isactrlbsub p\isactrlesub } imply \isa{{\isasymparallel}v\isactrlsub {\isadigit{3}}{\isasymparallel}\isactrlbsub p\isactrlesub \ {\isacharless}{\kern0pt}\ {\isasymparallel}v\isactrlsub {\isadigit{1}}{\isasymparallel}\isactrlbsub p\isactrlesub }. It remains to show
- that for a \isa{p{\isacharprime}{\kern0pt}\ {\isasymin}\ Pos\ v\isactrlsub {\isadigit{1}}\ {\isasymunion}\ Pos\ v\isactrlsub {\isadigit{3}}}
- with \isa{p{\isacharprime}{\kern0pt}\ {\isasymprec}\isactrlbsub lex\isactrlesub \ p} that \isa{{\isasymparallel}v\isactrlsub {\isadigit{1}}{\isasymparallel}\isactrlbsub p{\isacharprime}{\kern0pt}\isactrlesub \ {\isacharequal}{\kern0pt}\ {\isasymparallel}v\isactrlsub {\isadigit{3}}{\isasymparallel}\isactrlbsub p{\isacharprime}{\kern0pt}\isactrlesub } holds. Suppose \isa{p{\isacharprime}{\kern0pt}\ {\isasymin}\ Pos\ v\isactrlsub {\isadigit{1}}}, then we can infer from the first assumption that \isa{{\isasymparallel}v\isactrlsub {\isadigit{1}}{\isasymparallel}\isactrlbsub p{\isacharprime}{\kern0pt}\isactrlesub \ {\isacharequal}{\kern0pt}\ {\isasymparallel}v\isactrlsub {\isadigit{2}}{\isasymparallel}\isactrlbsub p{\isacharprime}{\kern0pt}\isactrlesub }. But this means
- that \isa{p{\isacharprime}{\kern0pt}} must be in \isa{Pos\ v\isactrlsub {\isadigit{2}}} too (the norm
- cannot be \isa{{\isacharminus}{\kern0pt}{\isadigit{1}}} given \isa{p{\isacharprime}{\kern0pt}\ {\isasymin}\ Pos\ v\isactrlsub {\isadigit{1}}}).
- Hence we can use the second assumption and
- infer \isa{{\isasymparallel}v\isactrlsub {\isadigit{2}}{\isasymparallel}\isactrlbsub p{\isacharprime}{\kern0pt}\isactrlesub \ {\isacharequal}{\kern0pt}\ {\isasymparallel}v\isactrlsub {\isadigit{3}}{\isasymparallel}\isactrlbsub p{\isacharprime}{\kern0pt}\isactrlesub },
- which concludes this case with \isa{v\isactrlsub {\isadigit{1}}\ {\isasymprec}\ v\isactrlsub {\isadigit{3}}}. The reasoning in the other cases is similar.\qed
- \end{proof}
-
- \noindent
- The proof for $\preccurlyeq$ is similar and omitted.
- It is also straightforward to show that \isa{{\isasymprec}} and
- $\preccurlyeq$ are partial orders. Okui and Suzuki furthermore show that they
- are linear orderings for lexical values \cite{OkuiSuzuki2010} of a given
- regular expression and given string, but we have not formalised this in Isabelle. It is
- not essential for our results. What we are going to show below is
- that for a given \isa{r} and \isa{s}, the orderings have a unique
- minimal element on the set \isa{LV\ r\ s}, which is the POSIX value
- we defined in the previous section. We start with two properties that
- show how the length of a flattened value relates to the \isa{{\isasymprec}}-ordering.
-
- \begin{proposition}\mbox{}\smallskip\\\label{ordlen}
- \begin{tabular}{@ {}ll}
- (1) &
- \isa{{\normalsize{}If\,}\ v\isactrlsub {\isadigit{1}}\ {\isasymprec}\ v\isactrlsub {\isadigit{2}}\ {\normalsize \,then\,}\ len\ {\isacharbar}{\kern0pt}v\isactrlsub {\isadigit{2}}{\isacharbar}{\kern0pt}\ {\isasymle}\ len\ {\isacharbar}{\kern0pt}v\isactrlsub {\isadigit{1}}{\isacharbar}{\kern0pt}{\isachardot}{\kern0pt}}\\
- (2) &
- \isa{{\normalsize{}If\,}\ len\ {\isacharbar}{\kern0pt}v\isactrlsub {\isadigit{2}}{\isacharbar}{\kern0pt}\ {\isacharless}{\kern0pt}\ len\ {\isacharbar}{\kern0pt}v\isactrlsub {\isadigit{1}}{\isacharbar}{\kern0pt}\ {\normalsize \,then\,}\ v\isactrlsub {\isadigit{1}}\ {\isasymprec}\ v\isactrlsub {\isadigit{2}}{\isachardot}{\kern0pt}}
- \end{tabular}
- \end{proposition}
-
- \noindent Both properties follow from the definition of the ordering. Note that
- \textit{(2)} entails that a value, say \isa{v\isactrlsub {\isadigit{2}}}, whose underlying
- string is a strict prefix of another flattened value, say \isa{v\isactrlsub {\isadigit{1}}}, then
- \isa{v\isactrlsub {\isadigit{1}}} must be smaller than \isa{v\isactrlsub {\isadigit{2}}}. For our proofs it
- will be useful to have the following properties---in each case the underlying strings
- of the compared values are the same:
-
- \begin{proposition}\mbox{}\smallskip\\\label{ordintros}
- \begin{tabular}{ll}
- \textit{(1)} &
- \isa{{\normalsize{}If\,}\ {\isacharbar}{\kern0pt}v\isactrlsub {\isadigit{1}}{\isacharbar}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharbar}{\kern0pt}v\isactrlsub {\isadigit{2}}{\isacharbar}{\kern0pt}\ {\normalsize \,then\,}\ Left\ v\isactrlsub {\isadigit{1}}\ {\isasymprec}\ Right\ v\isactrlsub {\isadigit{2}}{\isachardot}{\kern0pt}}\\
- \textit{(2)} & If
- \isa{{\isacharbar}{\kern0pt}v\isactrlsub {\isadigit{1}}{\isacharbar}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharbar}{\kern0pt}v\isactrlsub {\isadigit{2}}{\isacharbar}{\kern0pt}} \;then\;
- \isa{Left\ v\isactrlsub {\isadigit{1}}\ {\isasymprec}\ Left\ v\isactrlsub {\isadigit{2}}} \;iff\;
- \isa{v\isactrlsub {\isadigit{1}}\ {\isasymprec}\ v\isactrlsub {\isadigit{2}}}\\
- \textit{(3)} & If
- \isa{{\isacharbar}{\kern0pt}v\isactrlsub {\isadigit{1}}{\isacharbar}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharbar}{\kern0pt}v\isactrlsub {\isadigit{2}}{\isacharbar}{\kern0pt}} \;then\;
- \isa{Right\ v\isactrlsub {\isadigit{1}}\ {\isasymprec}\ Right\ v\isactrlsub {\isadigit{2}}} \;iff\;
- \isa{v\isactrlsub {\isadigit{1}}\ {\isasymprec}\ v\isactrlsub {\isadigit{2}}}\\
- \textit{(4)} & If
- \isa{{\isacharbar}{\kern0pt}v\isactrlsub {\isadigit{2}}{\isacharbar}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharbar}{\kern0pt}w\isactrlsub {\isadigit{2}}{\isacharbar}{\kern0pt}} \;then\;
- \isa{Seq\ v\ v\isactrlsub {\isadigit{2}}\ {\isasymprec}\ Seq\ v\ w\isactrlsub {\isadigit{2}}} \;iff\;
- \isa{v\isactrlsub {\isadigit{2}}\ {\isasymprec}\ w\isactrlsub {\isadigit{2}}}\\
- \textit{(5)} & If
- \isa{{\isacharbar}{\kern0pt}v\isactrlsub {\isadigit{1}}{\isacharbar}{\kern0pt}\ {\isacharat}{\kern0pt}\ {\isacharbar}{\kern0pt}v\isactrlsub {\isadigit{2}}{\isacharbar}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharbar}{\kern0pt}w\isactrlsub {\isadigit{1}}{\isacharbar}{\kern0pt}\ {\isacharat}{\kern0pt}\ {\isacharbar}{\kern0pt}w\isactrlsub {\isadigit{2}}{\isacharbar}{\kern0pt}} \;and\;
- \isa{v\isactrlsub {\isadigit{1}}\ {\isasymprec}\ w\isactrlsub {\isadigit{1}}} \;then\;
- \isa{Seq\ v\isactrlsub {\isadigit{1}}\ v\isactrlsub {\isadigit{2}}\ {\isasymprec}\ Seq\ w\isactrlsub {\isadigit{1}}\ w\isactrlsub {\isadigit{2}}}\\
- \textit{(6)} & If
- \isa{{\isacharbar}{\kern0pt}vs\isactrlsub {\isadigit{1}}{\isacharbar}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharbar}{\kern0pt}vs\isactrlsub {\isadigit{2}}{\isacharbar}{\kern0pt}} \;then\;
- \isa{Stars\ {\isacharparenleft}{\kern0pt}vs\ {\isacharat}{\kern0pt}\ vs\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isasymprec}\ Stars\ {\isacharparenleft}{\kern0pt}vs\ {\isacharat}{\kern0pt}\ vs\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}} \;iff\;
- \isa{Stars\ vs\isactrlsub {\isadigit{1}}\ {\isasymprec}\ Stars\ vs\isactrlsub {\isadigit{2}}}\\
-
- \textit{(7)} & If
- \isa{{\isacharbar}{\kern0pt}v\isactrlsub {\isadigit{1}}\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}vs\isactrlsub {\isadigit{1}}{\isacharbar}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharbar}{\kern0pt}v\isactrlsub {\isadigit{2}}\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}vs\isactrlsub {\isadigit{2}}{\isacharbar}{\kern0pt}} \;and\;
- \isa{v\isactrlsub {\isadigit{1}}\ {\isasymprec}\ v\isactrlsub {\isadigit{2}}} \;then\;
- \isa{Stars\ {\isacharparenleft}{\kern0pt}v\isactrlsub {\isadigit{1}}\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}vs\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isasymprec}\ Stars\ {\isacharparenleft}{\kern0pt}v\isactrlsub {\isadigit{2}}\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}vs\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}\\
- \end{tabular}
- \end{proposition}
-
- \noindent One might prefer that statements \textit{(4)} and \textit{(5)}
- (respectively \textit{(6)} and \textit{(7)})
- are combined into a single \textit{iff}-statement (like the ones for \isa{Left} and \isa{Right}). Unfortunately this cannot be done easily: such
- a single statement would require an additional assumption about the
- two values \isa{Seq\ v\isactrlsub {\isadigit{1}}\ v\isactrlsub {\isadigit{2}}} and \isa{Seq\ w\isactrlsub {\isadigit{1}}\ w\isactrlsub {\isadigit{2}}}
- being inhabited by the same regular expression. The
- complexity of the proofs involved seems to not justify such a
- `cleaner' single statement. The statements given are just the properties that
- allow us to establish our theorems without any difficulty. The proofs
- for Proposition~\ref{ordintros} are routine.
-
-
- Next we establish how Okui and Suzuki's orderings relate to our
- definition of POSIX values. Given a \isa{POSIX} value \isa{v\isactrlsub {\isadigit{1}}}
- for \isa{r} and \isa{s}, then any other lexical value \isa{v\isactrlsub {\isadigit{2}}} in \isa{LV\ r\ s} is greater or equal than \isa{v\isactrlsub {\isadigit{1}}}, namely:
-
-
- \begin{theorem}\label{orderone}
- \isa{{\normalsize{}If\,}\ \mbox{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v\isactrlsub {\isadigit{1}}}\ {\normalsize \,and\,}\ \mbox{v\isactrlsub {\isadigit{2}}\ {\isasymin}\ LV\ r\ s}\ {\normalsize \,then\,}\ v\isactrlsub {\isadigit{1}}\ \mbox{$\preccurlyeq$}\ v\isactrlsub {\isadigit{2}}{\isachardot}{\kern0pt}}
- \end{theorem}
-
- \begin{proof} By induction on our POSIX rules. By
- Theorem~\ref{posixdeterm} and the definition of \isa{LV}, it is clear
- that \isa{v\isactrlsub {\isadigit{1}}} and \isa{v\isactrlsub {\isadigit{2}}} have the same
- underlying string \isa{s}. The three base cases are
- straightforward: for example for \isa{v\isactrlsub {\isadigit{1}}\ {\isacharequal}{\kern0pt}\ Empty}, we have
- that \isa{v\isactrlsub {\isadigit{2}}\ {\isasymin}\ LV\ \isactrlbold {\isadigit{1}}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}} must also be of the form
- \mbox{\isa{v\isactrlsub {\isadigit{2}}\ {\isacharequal}{\kern0pt}\ Empty}}. Therefore we have \isa{v\isactrlsub {\isadigit{1}}\ \mbox{$\preccurlyeq$}\ v\isactrlsub {\isadigit{2}}}. The inductive cases for
- \isa{r} being of the form \isa{r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}} and
- \isa{r\isactrlsub {\isadigit{1}}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}} are as follows:
-
-
- \begin{itemize}
-
- \item[$\bullet$] Case \isa{P{\isacharplus}{\kern0pt}L} with \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ Left\ w\isactrlsub {\isadigit{1}}}: In this case the value
- \isa{v\isactrlsub {\isadigit{2}}} is either of the
- form \isa{Left\ w\isactrlsub {\isadigit{2}}} or \isa{Right\ w\isactrlsub {\isadigit{2}}}. In the
- latter case we can immediately conclude with \mbox{\isa{v\isactrlsub {\isadigit{1}}\ \mbox{$\preccurlyeq$}\ v\isactrlsub {\isadigit{2}}}} since a \isa{Left}-value with the
- same underlying string \isa{s} is always smaller than a
- \isa{Right}-value by Proposition~\ref{ordintros}\textit{(1)}.
- In the former case we have \isa{w\isactrlsub {\isadigit{2}}\ {\isasymin}\ LV\ r\isactrlsub {\isadigit{1}}\ s} and can use the induction hypothesis to infer
- \isa{w\isactrlsub {\isadigit{1}}\ \mbox{$\preccurlyeq$}\ w\isactrlsub {\isadigit{2}}}. Because \isa{w\isactrlsub {\isadigit{1}}} and \isa{w\isactrlsub {\isadigit{2}}} have the same underlying string
- \isa{s}, we can conclude with \isa{Left\ w\isactrlsub {\isadigit{1}}\ \mbox{$\preccurlyeq$}\ Left\ w\isactrlsub {\isadigit{2}}} using
- Proposition~\ref{ordintros}\textit{(2)}.\smallskip
-
- \item[$\bullet$] Case \isa{P{\isacharplus}{\kern0pt}R} with \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ Right\ w\isactrlsub {\isadigit{1}}}: This case similar to the previous
- case, except that we additionally know \isa{s\ {\isasymnotin}\ L{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}}. This is needed when \isa{v\isactrlsub {\isadigit{2}}} is of the form
- \mbox{\isa{Left\ w\isactrlsub {\isadigit{2}}}}. Since \mbox{\isa{{\isacharbar}{\kern0pt}v\isactrlsub {\isadigit{2}}{\isacharbar}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharbar}{\kern0pt}w\isactrlsub {\isadigit{2}}{\isacharbar}{\kern0pt}} \isa{{\isacharequal}{\kern0pt}\ s}} and \isa{w\isactrlsub {\isadigit{2}}\ {\isacharcolon}{\kern0pt}\ r\isactrlsub {\isadigit{1}}}, we can derive a contradiction for \mbox{\isa{s\ {\isasymnotin}\ L{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}}} using
- Proposition~\ref{inhabs}. So also in this case \mbox{\isa{v\isactrlsub {\isadigit{1}}\ \mbox{$\preccurlyeq$}\ v\isactrlsub {\isadigit{2}}}}.\smallskip
-
- \item[$\bullet$] Case \isa{PS} with \isa{{\isacharparenleft}{\kern0pt}s\isactrlsub {\isadigit{1}}\ {\isacharat}{\kern0pt}\ s\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ Seq\ w\isactrlsub {\isadigit{1}}\ w\isactrlsub {\isadigit{2}}}: We can assume \isa{v\isactrlsub {\isadigit{2}}\ {\isacharequal}{\kern0pt}\ Seq\ u\isactrlsub {\isadigit{1}}\ u\isactrlsub {\isadigit{2}}} with \isa{u\isactrlsub {\isadigit{1}}\ {\isacharcolon}{\kern0pt}\ r\isactrlsub {\isadigit{1}}} and \mbox{\isa{u\isactrlsub {\isadigit{2}}\ {\isacharcolon}{\kern0pt}\ r\isactrlsub {\isadigit{2}}}}. We have \isa{s\isactrlsub {\isadigit{1}}\ {\isacharat}{\kern0pt}\ s\isactrlsub {\isadigit{2}}\ {\isacharequal}{\kern0pt}\ {\isacharbar}{\kern0pt}u\isactrlsub {\isadigit{1}}{\isacharbar}{\kern0pt}\ {\isacharat}{\kern0pt}\ {\isacharbar}{\kern0pt}u\isactrlsub {\isadigit{2}}{\isacharbar}{\kern0pt}}. By the side-condition of the
- \isa{PS}-rule we know that either \isa{s\isactrlsub {\isadigit{1}}\ {\isacharequal}{\kern0pt}\ {\isacharbar}{\kern0pt}u\isactrlsub {\isadigit{1}}{\isacharbar}{\kern0pt}} or that \isa{{\isacharbar}{\kern0pt}u\isactrlsub {\isadigit{1}}{\isacharbar}{\kern0pt}} is a strict prefix of
- \isa{s\isactrlsub {\isadigit{1}}}. In the latter case we can infer \isa{w\isactrlsub {\isadigit{1}}\ {\isasymprec}\ u\isactrlsub {\isadigit{1}}} by
- Proposition~\ref{ordlen}\textit{(2)} and from this \isa{v\isactrlsub {\isadigit{1}}\ \mbox{$\preccurlyeq$}\ v\isactrlsub {\isadigit{2}}} by Proposition~\ref{ordintros}\textit{(5)}
- (as noted above \isa{v\isactrlsub {\isadigit{1}}} and \isa{v\isactrlsub {\isadigit{2}}} must have the
- same underlying string).
- In the former case we know
- \isa{u\isactrlsub {\isadigit{1}}\ {\isasymin}\ LV\ r\isactrlsub {\isadigit{1}}\ s\isactrlsub {\isadigit{1}}} and \isa{u\isactrlsub {\isadigit{2}}\ {\isasymin}\ LV\ r\isactrlsub {\isadigit{2}}\ s\isactrlsub {\isadigit{2}}}. With this we can use the
- induction hypotheses to infer \isa{w\isactrlsub {\isadigit{1}}\ \mbox{$\preccurlyeq$}\ u\isactrlsub {\isadigit{1}}} and \isa{w\isactrlsub {\isadigit{2}}\ \mbox{$\preccurlyeq$}\ u\isactrlsub {\isadigit{2}}}. By
- Proposition~\ref{ordintros}\textit{(4,5)} we can again infer
- \isa{v\isactrlsub {\isadigit{1}}\ \mbox{$\preccurlyeq$}\ v\isactrlsub {\isadigit{2}}}.
-
- \end{itemize}
-
- \noindent The case for \isa{P{\isasymstar}} is similar to the \isa{PS}-case and omitted.\qed
- \end{proof}
-
- \noindent This theorem shows that our \isa{POSIX} value for a
- regular expression \isa{r} and string \isa{s} is in fact a
- minimal element of the values in \isa{LV\ r\ s}. By
- Proposition~\ref{ordlen}\textit{(2)} we also know that any value in
- \isa{LV\ r\ s{\isacharprime}{\kern0pt}}, with \isa{s{\isacharprime}{\kern0pt}} being a strict prefix, cannot be
- smaller than \isa{v\isactrlsub {\isadigit{1}}}. The next theorem shows the
- opposite---namely any minimal element in \isa{LV\ r\ s} must be a
- \isa{POSIX} value. This can be established by induction on \isa{r}, but the proof can be drastically simplified by using the fact
- from the previous section about the existence of a \isa{POSIX} value
- whenever a string \isa{s\ {\isasymin}\ L{\isacharparenleft}{\kern0pt}r{\isacharparenright}{\kern0pt}}.
-
-
- \begin{theorem}
- \isa{{\normalsize{}If\,}\ \mbox{v\isactrlsub {\isadigit{1}}\ {\isasymin}\ LV\ r\ s}\ {\normalsize \,and\,}\ \mbox{{\isasymforall}v\isactrlsub {\isadigit{2}}{\isasymin}LV\ r\ s{\isachardot}{\kern0pt}\ v\isactrlsub {\isadigit{2}}\ \mbox{$\not\prec$}\ v\isactrlsub {\isadigit{1}}}\ {\normalsize \,then\,}\ {\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v\isactrlsub {\isadigit{1}}{\isachardot}{\kern0pt}}
- \end{theorem}
-
- \begin{proof}
- If \isa{v\isactrlsub {\isadigit{1}}\ {\isasymin}\ LV\ r\ s} then
- \isa{s\ {\isasymin}\ L{\isacharparenleft}{\kern0pt}r{\isacharparenright}{\kern0pt}} by Proposition~\ref{inhabs}. Hence by Theorem~\ref{lexercorrect}(2)
- there exists a
- \isa{POSIX} value \isa{v\isactrlsub P} with \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v\isactrlsub P}
- and by Lemma~\ref{LVposix} we also have \mbox{\isa{v\isactrlsub P\ {\isasymin}\ LV\ r\ s}}.
- By Theorem~\ref{orderone} we therefore have
- \isa{v\isactrlsub P\ \mbox{$\preccurlyeq$}\ v\isactrlsub {\isadigit{1}}}. If \isa{v\isactrlsub P\ {\isacharequal}{\kern0pt}\ v\isactrlsub {\isadigit{1}}} then
- we are done. Otherwise we have \isa{v\isactrlsub P\ {\isasymprec}\ v\isactrlsub {\isadigit{1}}}, which
- however contradicts the second assumption about \isa{v\isactrlsub {\isadigit{1}}} being the smallest
- element in \isa{LV\ r\ s}. So we are done in this case too.\qed
- \end{proof}
-
- \noindent
- From this we can also show
- that if \isa{LV\ r\ s} is non-empty (or equivalently \isa{s\ {\isasymin}\ L{\isacharparenleft}{\kern0pt}r{\isacharparenright}{\kern0pt}}) then
- it has a unique minimal element:
-
- \begin{corollary}
- \isa{{\normalsize{}If\,}\ LV\ r\ s\ {\isasymnoteq}\ {\isasymemptyset}\ {\normalsize \,then\,}\ {\isasymexists}{\isacharbang}{\kern0pt}vmin{\isachardot}{\kern0pt}\ vmin\ {\isasymin}\ LV\ r\ s\ {\isasymand}\ {\isacharparenleft}{\kern0pt}{\isasymforall}v{\isasymin}LV\ r\ s{\isachardot}{\kern0pt}\ vmin\ \mbox{$\preccurlyeq$}\ v{\isacharparenright}{\kern0pt}{\isachardot}{\kern0pt}}
- \end{corollary}
-
-
-
- \noindent To sum up, we have shown that the (unique) minimal elements
- of the ordering by Okui and Suzuki are exactly the \isa{POSIX}
- values we defined inductively in Section~\ref{posixsec}. This provides
- an independent confirmation that our ternary relation formalises the
- informal POSIX rules.%
-\end{isamarkuptext}\isamarkuptrue%
-%
-\isadelimdocument
-%
-\endisadelimdocument
-%
-\isatagdocument
-%
-\isamarkupsection{Bitcoded Lexing%
-}
-\isamarkuptrue%
-%
-\endisatagdocument
-{\isafolddocument}%
-%
-\isadelimdocument
-%
-\endisadelimdocument
-%
-\begin{isamarkuptext}%
-Incremental calculation of the value. To simplify the proof we first define the function
-\isa{flex} which calculates the ``iterated'' injection function. With this we can
-rewrite the lexer as
-
-\begin{center}
-\isa{lexer\ r\ s\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}\textrm{if}\ nullable\ {\isacharparenleft}{\kern0pt}r{\isacharbackslash}{\kern0pt}s{\isacharparenright}{\kern0pt}\ \textrm{then}\ Some\ {\isacharparenleft}{\kern0pt}flex\ r\ id\ s\ {\isacharparenleft}{\kern0pt}mkeps\ {\isacharparenleft}{\kern0pt}r{\isacharbackslash}{\kern0pt}s{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ \textrm{else}\ None{\isacharparenright}{\kern0pt}}
-\end{center}%
-\end{isamarkuptext}\isamarkuptrue%
-%
-\isadelimdocument
-%
-\endisadelimdocument
-%
-\isatagdocument
-%
-\isamarkupsection{Optimisations%
-}
-\isamarkuptrue%
-%
-\endisatagdocument
-{\isafolddocument}%
-%
-\isadelimdocument
-%
-\endisadelimdocument
-%
-\begin{isamarkuptext}%
-Derivatives as calculated by \Brz's method are usually more complex
- regular expressions than the initial one; the result is that the
- derivative-based matching and lexing algorithms are often abysmally slow.
- However, various optimisations are possible, such as the simplifications
- of \isa{\isactrlbold {\isadigit{0}}\ {\isacharplus}{\kern0pt}\ r}, \isa{r\ {\isacharplus}{\kern0pt}\ \isactrlbold {\isadigit{0}}}, \isa{\isactrlbold {\isadigit{1}}\ {\isasymcdot}\ r} and
- \isa{r\ {\isasymcdot}\ \isactrlbold {\isadigit{1}}} to \isa{r}. These simplifications can speed up the
- algorithms considerably, as noted in \cite{Sulzmann2014}. One of the
- advantages of having a simple specification and correctness proof is that
- the latter can be refined to prove the correctness of such simplification
- steps. While the simplification of regular expressions according to
- rules like
-
- \begin{equation}\label{Simpl}
- \begin{array}{lcllcllcllcl}
- \isa{\isactrlbold {\isadigit{0}}\ {\isacharplus}{\kern0pt}\ r} & \isa{{\isasymRightarrow}} & \isa{r} \hspace{8mm}%\\
- \isa{r\ {\isacharplus}{\kern0pt}\ \isactrlbold {\isadigit{0}}} & \isa{{\isasymRightarrow}} & \isa{r} \hspace{8mm}%\\
- \isa{\isactrlbold {\isadigit{1}}\ {\isasymcdot}\ r} & \isa{{\isasymRightarrow}} & \isa{r} \hspace{8mm}%\\
- \isa{r\ {\isasymcdot}\ \isactrlbold {\isadigit{1}}} & \isa{{\isasymRightarrow}} & \isa{r}
- \end{array}
- \end{equation}
-
- \noindent is well understood, there is an obstacle with the POSIX value
- calculation algorithm by Sulzmann and Lu: if we build a derivative regular
- expression and then simplify it, we will calculate a POSIX value for this
- simplified derivative regular expression, \emph{not} for the original (unsimplified)
- derivative regular expression. Sulzmann and Lu \cite{Sulzmann2014} overcome this obstacle by
- not just calculating a simplified regular expression, but also calculating
- a \emph{rectification function} that ``repairs'' the incorrect value.
-
- The rectification functions can be (slightly clumsily) implemented in
- Isabelle/HOL as follows using some auxiliary functions:
-
- \begin{center}
- \begin{tabular}{lcl}
- \isa{F\isactrlbsub Right\isactrlesub \ f\ v} & $\dn$ & \isa{Right\ {\isacharparenleft}{\kern0pt}f\ v{\isacharparenright}{\kern0pt}}\\
- \isa{F\isactrlbsub Left\isactrlesub \ f\ v} & $\dn$ & \isa{Left\ {\isacharparenleft}{\kern0pt}f\ v{\isacharparenright}{\kern0pt}}\\
-
- \isa{F\isactrlbsub Alt\isactrlesub \ f\isactrlsub {\isadigit{1}}\ f\isactrlsub {\isadigit{2}}\ {\isacharparenleft}{\kern0pt}Right\ v{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{Right\ {\isacharparenleft}{\kern0pt}f\isactrlsub {\isadigit{2}}\ v{\isacharparenright}{\kern0pt}}\\
- \isa{F\isactrlbsub Alt\isactrlesub \ f\isactrlsub {\isadigit{1}}\ f\isactrlsub {\isadigit{2}}\ {\isacharparenleft}{\kern0pt}Left\ v{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{Left\ {\isacharparenleft}{\kern0pt}f\isactrlsub {\isadigit{1}}\ v{\isacharparenright}{\kern0pt}}\\
-
- \isa{F\isactrlbsub Seq{\isadigit{1}}\isactrlesub \ f\isactrlsub {\isadigit{1}}\ f\isactrlsub {\isadigit{2}}\ v} & $\dn$ & \isa{Seq\ {\isacharparenleft}{\kern0pt}f\isactrlsub {\isadigit{1}}\ {\isacharparenleft}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}f\isactrlsub {\isadigit{2}}\ v{\isacharparenright}{\kern0pt}}\\
- \isa{F\isactrlbsub Seq{\isadigit{2}}\isactrlesub \ f\isactrlsub {\isadigit{1}}\ f\isactrlsub {\isadigit{2}}\ v} & $\dn$ & \isa{Seq\ {\isacharparenleft}{\kern0pt}f\isactrlsub {\isadigit{1}}\ v{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}f\isactrlsub {\isadigit{2}}\ {\isacharparenleft}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}}\\
- \isa{F\isactrlbsub Seq\isactrlesub \ f\isactrlsub {\isadigit{1}}\ f\isactrlsub {\isadigit{2}}\ {\isacharparenleft}{\kern0pt}Seq\ v\isactrlsub {\isadigit{1}}\ v\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{Seq\ {\isacharparenleft}{\kern0pt}f\isactrlsub {\isadigit{1}}\ v\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}f\isactrlsub {\isadigit{2}}\ v\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}\medskip\\
- %\end{tabular}
- %
- %\begin{tabular}{lcl}
- \isa{simp\isactrlbsub Alt\isactrlesub \ {\isacharparenleft}{\kern0pt}\isactrlbold {\isadigit{0}}{\isacharcomma}{\kern0pt}\ \underline{\hspace{2mm}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ f\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ F\isactrlbsub Right\isactrlesub \ f\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}\\
- \isa{simp\isactrlbsub Alt\isactrlesub \ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharcomma}{\kern0pt}\ f\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}\isactrlbold {\isadigit{0}}{\isacharcomma}{\kern0pt}\ \underline{\hspace{2mm}}{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharcomma}{\kern0pt}\ F\isactrlbsub Left\isactrlesub \ f\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}}\\
- \isa{simp\isactrlbsub Alt\isactrlesub \ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharcomma}{\kern0pt}\ f\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ f\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ F\isactrlbsub Alt\isactrlesub \ f\isactrlsub {\isadigit{1}}\ f\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}\\
- \isa{simp\isactrlbsub Seq\isactrlesub \ {\isacharparenleft}{\kern0pt}\isactrlbold {\isadigit{1}}{\isacharcomma}{\kern0pt}\ f\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ f\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ F\isactrlbsub Seq{\isadigit{1}}\isactrlesub \ f\isactrlsub {\isadigit{1}}\ f\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}\\
- \isa{simp\isactrlbsub Seq\isactrlesub \ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharcomma}{\kern0pt}\ f\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}\isactrlbold {\isadigit{1}}{\isacharcomma}{\kern0pt}\ f\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharcomma}{\kern0pt}\ F\isactrlbsub Seq{\isadigit{2}}\isactrlesub \ f\isactrlsub {\isadigit{1}}\ f\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}\\
- \isa{simp\isactrlbsub Seq\isactrlesub \ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharcomma}{\kern0pt}\ f\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ f\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ F\isactrlbsub Seq\isactrlesub \ f\isactrlsub {\isadigit{1}}\ f\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}\\
- \end{tabular}
- \end{center}
-
- \noindent
- The functions \isa{simp\isactrlbsub Alt\isactrlesub } and \isa{simp\isactrlbsub Seq\isactrlesub } encode the simplification rules
- in \eqref{Simpl} and compose the rectification functions (simplifications can occur
- deep inside the regular expression). The main simplification function is then
-
- \begin{center}
- \begin{tabular}{lcl}
- \isa{simp\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{simp\isactrlbsub Alt\isactrlesub \ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}\\
- \isa{simp\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}} & $\dn$ & \isa{simp\isactrlbsub Seq\isactrlesub \ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}}\\
- \isa{simp\ r} & $\dn$ & \isa{{\isacharparenleft}{\kern0pt}r{\isacharcomma}{\kern0pt}\ id{\isacharparenright}{\kern0pt}}\\
- \end{tabular}
- \end{center}
-
- \noindent where \isa{id} stands for the identity function. The
- function \isa{simp} returns a simplified regular expression and a corresponding
- rectification function. Note that we do not simplify under stars: this
- seems to slow down the algorithm, rather than speed it up. The optimised
- lexer is then given by the clauses:
-
- \begin{center}
- \begin{tabular}{lcl}
- \isa{lexer\isactrlsup {\isacharplus}{\kern0pt}\ r\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}} & $\dn$ & \isa{\textrm{if}\ nullable\ r\ \textrm{then}\ Some\ {\isacharparenleft}{\kern0pt}mkeps\ r{\isacharparenright}{\kern0pt}\ \textrm{else}\ None}\\
- \isa{lexer\isactrlsup {\isacharplus}{\kern0pt}\ r\ {\isacharparenleft}{\kern0pt}c\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}s{\isacharparenright}{\kern0pt}} & $\dn$ &
- \isa{let\ {\isacharparenleft}{\kern0pt}r\isactrlsub s{\isacharcomma}{\kern0pt}\ f\isactrlsub r{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ simp\ {\isacharparenleft}{\kern0pt}r}$\backslash$\isa{c{\isacharparenright}{\kern0pt}\ in}\\
- & & \isa{case} \isa{lexer\isactrlsup {\isacharplus}{\kern0pt}\ r\isactrlsub s\ s} \isa{of}\\
- & & \phantom{$|$} \isa{None} \isa{{\isasymRightarrow}} \isa{None}\\
- & & $|$ \isa{Some\ v} \isa{{\isasymRightarrow}} \isa{Some\ {\isacharparenleft}{\kern0pt}inj\ r\ c\ {\isacharparenleft}{\kern0pt}f\isactrlsub r\ v{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}}
- \end{tabular}
- \end{center}
-
- \noindent
- In the second clause we first calculate the derivative \isa{r{\isacharbackslash}{\kern0pt}c}
- and then simpli
-
-text \isa{\ \ Incremental\ calculation\ of\ the\ value{\isachardot}{\kern0pt}\ To\ simplify\ the\ proof\ we\ first\ define\ the\ function\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}const\ flex{\isacharbraceright}{\kern0pt}\ which\ calculates\ the\ {\isacharbackquote}{\kern0pt}{\isacharbackquote}{\kern0pt}iterated{\isacharprime}{\kern0pt}{\isacharprime}{\kern0pt}\ injection\ function{\isachardot}{\kern0pt}\ With\ this\ we\ can\ rewrite\ the\ lexer\ as\ \ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}center{\isacharbraceright}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ lexer{\isacharunderscore}{\kern0pt}flex{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}center{\isacharbraceright}{\kern0pt}\ \ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}center{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}{\isacharbraceleft}{\kern0pt}lcl{\isacharbraceright}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ code{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ code{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ code{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ code{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ code{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ code{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ code{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ code{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ code{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{5}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ {\isachardoublequote}{\kern0pt}v\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}\ {\isachardoublequote}{\kern0pt}v\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ code{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{5}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ {\isachardoublequote}{\kern0pt}v\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}\ {\isachardoublequote}{\kern0pt}v\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ code{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{6}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ code{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{6}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ code{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{7}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ code{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{7}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}center{\isacharbraceright}{\kern0pt}\ \ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}center{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}{\isacharbraceleft}{\kern0pt}lcl{\isacharbraceright}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ areg{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}{\isacharequal}{\kern0pt}{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}AZERO{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}mid{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}AONE\ bs{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}mid{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}ACHAR\ bs\ c{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}mid{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}AALT\ bs\ r{\isadigit{1}}\ r{\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}mid{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}ASEQ\ bs\ r\isactrlsub {\isadigit{1}}\ r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}mid{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}ASTAR\ bs\ r{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}center{\isacharbraceright}{\kern0pt}\ \ \ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}center{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}{\isacharbraceleft}{\kern0pt}lcl{\isacharbraceright}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ intern{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ intern{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ intern{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ intern{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ intern{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ intern{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ intern{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ intern{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ intern{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{5}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ intern{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{5}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ intern{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{6}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ intern{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{6}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}center{\isacharbraceright}{\kern0pt}\ \ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}center{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}{\isacharbraceleft}{\kern0pt}lcl{\isacharbraceright}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ erase{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ erase{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ erase{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ bs{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ erase{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ bs{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ erase{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ bs{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ erase{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ bs{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ erase{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ bs\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ erase{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ bs\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ erase{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{5}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ bs\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ erase{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{5}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ bs\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ erase{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{6}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ bs{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ erase{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{6}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ bs{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}center{\isacharbraceright}{\kern0pt}\ \ Some\ simple\ facts\ about\ erase\ \ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}lemma{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}mbox{\isacharbraceleft}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ erase{\isacharunderscore}{\kern0pt}bder{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ erase{\isacharunderscore}{\kern0pt}intern{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}lemma{\isacharbraceright}{\kern0pt}\ \ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}center{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}{\isacharbraceleft}{\kern0pt}lcl{\isacharbraceright}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ bnullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ bnullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ bnullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ bnullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ bnullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ bnullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ bnullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ bs\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ bnullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ bs\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ bnullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{5}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ bs\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ bnullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{5}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ bs\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ bnullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{6}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ bnullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{6}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}medskip{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ \ {\isacharpercent}{\kern0pt}\ \ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}\ {\isacharpercent}{\kern0pt}\ \ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}center{\isacharbraceright}{\kern0pt}\ \ {\isacharpercent}{\kern0pt}\ \ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}center{\isacharbraceright}{\kern0pt}\ {\isacharpercent}{\kern0pt}\ \ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}{\isacharbraceleft}{\kern0pt}lcl{\isacharbraceright}{\kern0pt}\ \ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ bder{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ bder{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ bder{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ bder{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ bder{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ bder{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ bder{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ bs\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ bder{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ bs\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ bder{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{5}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ bs\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ bder{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{5}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ bs\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ bder{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{6}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ bder{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{6}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}center{\isacharbraceright}{\kern0pt}\ \ \ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}center{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}{\isacharbraceleft}{\kern0pt}lcl{\isacharbraceright}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ bmkeps{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ bmkeps{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ bmkeps{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ bs\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ bmkeps{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ bs\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ bmkeps{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ bs\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ bmkeps{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}of\ bs\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}\ {\isachardoublequote}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ bmkeps{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ bmkeps{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}medskip{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}center{\isacharbraceright}{\kern0pt}\ \ \ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharbrackleft}{\kern0pt}mode{\isacharequal}{\kern0pt}IfThen{\isacharbrackright}{\kern0pt}\ bder{\isacharunderscore}{\kern0pt}retrieve{\isacharbraceright}{\kern0pt}\ \ By\ induction\ on\ {\isasymopen}r{\isasymclose}\ \ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}theorem{\isacharbraceright}{\kern0pt}{\isacharbrackleft}{\kern0pt}Main\ Lemma{\isacharbrackright}{\kern0pt}{\isacharbackslash}{\kern0pt}mbox{\isacharbraceleft}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharbrackleft}{\kern0pt}mode{\isacharequal}{\kern0pt}IfThen{\isacharbrackright}{\kern0pt}\ MAIN{\isacharunderscore}{\kern0pt}decode{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}theorem{\isacharbraceright}{\kern0pt}\ \ {\isacharbackslash}{\kern0pt}noindent\ Definition\ of\ the\ bitcoded\ lexer\ \ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ blexer{\isacharunderscore}{\kern0pt}def{\isacharbraceright}{\kern0pt}\ \ \ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}theorem{\isacharbraceright}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ blexer{\isacharunderscore}{\kern0pt}correctness{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}theorem{\isacharbraceright}{\kern0pt}\ \ }
-
-section \isa{Optimisations}
-
-text \isa{\ \ Derivatives\ as\ calculated\ by\ {\isacharbackslash}{\kern0pt}Brz{\isacharprime}{\kern0pt}s\ method\ are\ usually\ more\ complex\ regular\ expressions\ than\ the\ initial\ one{\isacharsemicolon}{\kern0pt}\ the\ result\ is\ that\ the\ derivative{\isacharminus}{\kern0pt}based\ matching\ and\ lexing\ algorithms\ are\ often\ abysmally\ slow{\isachardot}{\kern0pt}\ However{\isacharcomma}{\kern0pt}\ various\ optimisations\ are\ possible{\isacharcomma}{\kern0pt}\ such\ as\ the\ simplifications\ of\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}ALT\ ZERO\ r{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharcomma}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}ALT\ r\ ZERO{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharcomma}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}SEQ\ ONE\ r{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ and\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}SEQ\ r\ ONE{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ to\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ r{\isacharbraceright}{\kern0pt}{\isachardot}{\kern0pt}\ These\ simplifications\ can\ speed\ up\ the\ algorithms\ considerably{\isacharcomma}{\kern0pt}\ as\ noted\ in\ {\isacharbackslash}{\kern0pt}cite{\isacharbraceleft}{\kern0pt}Sulzmann{\isadigit{2}}{\isadigit{0}}{\isadigit{1}}{\isadigit{4}}{\isacharbraceright}{\kern0pt}{\isachardot}{\kern0pt}\ One\ of\ the\ advantages\ of\ having\ a\ simple\ specification\ and\ correctness\ proof\ is\ that\ the\ latter\ can\ be\ refined\ to\ prove\ the\ correctness\ of\ such\ simplification\ steps{\isachardot}{\kern0pt}\ While\ the\ simplification\ of\ regular\ expressions\ according\ to\ rules\ like\ \ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}equation{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}label{\isacharbraceleft}{\kern0pt}Simpl{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}array{\isacharbraceright}{\kern0pt}{\isacharbraceleft}{\kern0pt}lcllcllcllcl{\isacharbraceright}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}ALT\ ZERO\ r{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isasymopen}{\isasymRightarrow}{\isasymclose}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ r{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}hspace{\isacharbraceleft}{\kern0pt}{\isadigit{8}}mm{\isacharbraceright}{\kern0pt}{\isacharpercent}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}ALT\ r\ ZERO{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isasymopen}{\isasymRightarrow}{\isasymclose}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ r{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}hspace{\isacharbraceleft}{\kern0pt}{\isadigit{8}}mm{\isacharbraceright}{\kern0pt}{\isacharpercent}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}SEQ\ ONE\ r{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ \ {\isacharampersand}{\kern0pt}\ {\isasymopen}{\isasymRightarrow}{\isasymclose}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ r{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}hspace{\isacharbraceleft}{\kern0pt}{\isadigit{8}}mm{\isacharbraceright}{\kern0pt}{\isacharpercent}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}SEQ\ r\ ONE{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ \ {\isacharampersand}{\kern0pt}\ {\isasymopen}{\isasymRightarrow}{\isasymclose}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ r{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}array{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}equation{\isacharbraceright}{\kern0pt}\ \ {\isacharbackslash}{\kern0pt}noindent\ is\ well\ understood{\isacharcomma}{\kern0pt}\ there\ is\ an\ obstacle\ with\ the\ POSIX\ value\ calculation\ algorithm\ by\ Sulzmann\ and\ Lu{\isacharcolon}{\kern0pt}\ if\ we\ build\ a\ derivative\ regular\ expression\ and\ then\ simplify\ it{\isacharcomma}{\kern0pt}\ we\ will\ calculate\ a\ POSIX\ value\ for\ this\ simplified\ derivative\ regular\ expression{\isacharcomma}{\kern0pt}\ {\isacharbackslash}{\kern0pt}emph{\isacharbraceleft}{\kern0pt}not{\isacharbraceright}{\kern0pt}\ for\ the\ original\ {\isacharparenleft}{\kern0pt}unsimplified{\isacharparenright}{\kern0pt}\ derivative\ regular\ expression{\isachardot}{\kern0pt}\ Sulzmann\ and\ Lu\ {\isacharbackslash}{\kern0pt}cite{\isacharbraceleft}{\kern0pt}Sulzmann{\isadigit{2}}{\isadigit{0}}{\isadigit{1}}{\isadigit{4}}{\isacharbraceright}{\kern0pt}\ overcome\ this\ obstacle\ by\ not\ just\ calculating\ a\ simplified\ regular\ expression{\isacharcomma}{\kern0pt}\ but\ also\ calculating\ a\ {\isacharbackslash}{\kern0pt}emph{\isacharbraceleft}{\kern0pt}rectification\ function{\isacharbraceright}{\kern0pt}\ that\ {\isacharbackquote}{\kern0pt}{\isacharbackquote}{\kern0pt}repairs{\isacharprime}{\kern0pt}{\isacharprime}{\kern0pt}\ the\ incorrect\ value{\isachardot}{\kern0pt}\ \ The\ rectification\ functions\ can\ be\ {\isacharparenleft}{\kern0pt}slightly\ clumsily{\isacharparenright}{\kern0pt}\ implemented\ \ in\ Isabelle{\isacharslash}{\kern0pt}HOL\ as\ follows\ using\ some\ auxiliary\ functions{\isacharcolon}{\kern0pt}\ \ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}center{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}{\isacharbraceleft}{\kern0pt}lcl{\isacharbraceright}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ F{\isacharunderscore}{\kern0pt}RIGHT{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isasymopen}Right\ {\isacharparenleft}{\kern0pt}f\ v{\isacharparenright}{\kern0pt}{\isasymclose}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ F{\isacharunderscore}{\kern0pt}LEFT{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isasymopen}Left\ {\isacharparenleft}{\kern0pt}f\ v{\isacharparenright}{\kern0pt}{\isasymclose}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ \ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ F{\isacharunderscore}{\kern0pt}ALT{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isasymopen}Right\ {\isacharparenleft}{\kern0pt}f\isactrlsub {\isadigit{2}}\ v{\isacharparenright}{\kern0pt}{\isasymclose}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ F{\isacharunderscore}{\kern0pt}ALT{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isasymopen}Left\ {\isacharparenleft}{\kern0pt}f\isactrlsub {\isadigit{1}}\ v{\isacharparenright}{\kern0pt}{\isasymclose}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ \ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ F{\isacharunderscore}{\kern0pt}SEQ{\isadigit{1}}{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isasymopen}Seq\ {\isacharparenleft}{\kern0pt}f\isactrlsub {\isadigit{1}}\ {\isacharparenleft}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}f\isactrlsub {\isadigit{2}}\ v{\isacharparenright}{\kern0pt}{\isasymclose}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ F{\isacharunderscore}{\kern0pt}SEQ{\isadigit{2}}{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isasymopen}Seq\ {\isacharparenleft}{\kern0pt}f\isactrlsub {\isadigit{1}}\ v{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}f\isactrlsub {\isadigit{2}}\ {\isacharparenleft}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isasymclose}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ F{\isacharunderscore}{\kern0pt}SEQ{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isasymopen}Seq\ {\isacharparenleft}{\kern0pt}f\isactrlsub {\isadigit{1}}\ v\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}f\isactrlsub {\isadigit{2}}\ v\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isasymclose}{\isacharbackslash}{\kern0pt}medskip{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharpercent}{\kern0pt}{\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}\ {\isacharpercent}{\kern0pt}\ {\isacharpercent}{\kern0pt}{\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}{\isacharbraceleft}{\kern0pt}lcl{\isacharbraceright}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}simp{\isacharunderscore}{\kern0pt}ALT\ {\isacharparenleft}{\kern0pt}ZERO{\isacharcomma}{\kern0pt}\ DUMMY{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ f\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ F{\isacharunderscore}{\kern0pt}RIGHT\ f\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}simp{\isacharunderscore}{\kern0pt}ALT\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharcomma}{\kern0pt}\ f\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}ZERO{\isacharcomma}{\kern0pt}\ DUMMY{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharcomma}{\kern0pt}\ F{\isacharunderscore}{\kern0pt}LEFT\ f\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}simp{\isacharunderscore}{\kern0pt}ALT\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharcomma}{\kern0pt}\ f\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ f\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}{\isacharparenleft}{\kern0pt}ALT\ r\isactrlsub {\isadigit{1}}\ r\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ F{\isacharunderscore}{\kern0pt}ALT\ f\isactrlsub {\isadigit{1}}\ f\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}simp{\isacharunderscore}{\kern0pt}SEQ\ {\isacharparenleft}{\kern0pt}ONE{\isacharcomma}{\kern0pt}\ f\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ f\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ F{\isacharunderscore}{\kern0pt}SEQ{\isadigit{1}}\ f\isactrlsub {\isadigit{1}}\ f\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}simp{\isacharunderscore}{\kern0pt}SEQ\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharcomma}{\kern0pt}\ f\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}ONE{\isacharcomma}{\kern0pt}\ f\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharcomma}{\kern0pt}\ F{\isacharunderscore}{\kern0pt}SEQ{\isadigit{2}}\ f\isactrlsub {\isadigit{1}}\ f\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}simp{\isacharunderscore}{\kern0pt}SEQ\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharcomma}{\kern0pt}\ f\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ f\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}{\isacharparenleft}{\kern0pt}SEQ\ r\isactrlsub {\isadigit{1}}\ r\isactrlsub {\isadigit{2}}{\isacharcomma}{\kern0pt}\ F{\isacharunderscore}{\kern0pt}SEQ\ f\isactrlsub {\isadigit{1}}\ f\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}center{\isacharbraceright}{\kern0pt}\ \ {\isacharbackslash}{\kern0pt}noindent\ The\ functions\ {\isasymopen}simp\isactrlbsub Alt\isactrlesub {\isasymclose}\ and\ {\isasymopen}simp\isactrlbsub Seq\isactrlesub {\isasymclose}\ encode\ the\ simplification\ rules\ in\ {\isacharbackslash}{\kern0pt}eqref{\isacharbraceleft}{\kern0pt}Simpl{\isacharbraceright}{\kern0pt}\ and\ compose\ the\ rectification\ functions\ {\isacharparenleft}{\kern0pt}simplifications\ can\ occur\ deep\ inside\ the\ regular\ expression{\isacharparenright}{\kern0pt}{\isachardot}{\kern0pt}\ The\ main\ simplification\ function\ is\ then\ \ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}center{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}{\isacharbraceleft}{\kern0pt}lcl{\isacharbraceright}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}simp\ {\isacharparenleft}{\kern0pt}ALT\ r\isactrlsub {\isadigit{1}}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}simp{\isacharunderscore}{\kern0pt}ALT\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}simp\ {\isacharparenleft}{\kern0pt}SEQ\ r\isactrlsub {\isadigit{1}}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}simp{\isacharunderscore}{\kern0pt}SEQ\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}simp\ r{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}{\isacharparenleft}{\kern0pt}r{\isacharcomma}{\kern0pt}\ id{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}center{\isacharbraceright}{\kern0pt}\ \ {\isacharbackslash}{\kern0pt}noindent\ where\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}id{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ stands\ for\ the\ identity\ function{\isachardot}{\kern0pt}\ The\ function\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}const\ simp{\isacharbraceright}{\kern0pt}\ returns\ a\ simplified\ regular\ expression\ and\ a\ corresponding\ rectification\ function{\isachardot}{\kern0pt}\ Note\ that\ we\ do\ not\ simplify\ under\ stars{\isacharcolon}{\kern0pt}\ this\ seems\ to\ slow\ down\ the\ algorithm{\isacharcomma}{\kern0pt}\ rather\ than\ speed\ it\ up{\isachardot}{\kern0pt}\ The\ optimised\ lexer\ is\ then\ given\ by\ the\ clauses{\isacharcolon}{\kern0pt}\ \ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}center{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}{\isacharbraceleft}{\kern0pt}lcl{\isacharbraceright}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ slexer{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}rhs{\isacharparenright}{\kern0pt}\ slexer{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ {\isacharparenleft}{\kern0pt}lhs{\isacharparenright}{\kern0pt}\ slexer{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}dn{\isachardollar}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isasymopen}let\ {\isacharparenleft}{\kern0pt}r\isactrlsub s{\isacharcomma}{\kern0pt}\ f\isactrlsub r{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ simp\ {\isacharparenleft}{\kern0pt}r\ {\isasymclose}{\isachardollar}{\kern0pt}{\isacharbackslash}{\kern0pt}backslash{\isachardollar}{\kern0pt}{\isasymopen}\ c{\isacharparenright}{\kern0pt}\ in{\isasymclose}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isasymopen}case{\isasymclose}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}slexer\ r\isactrlsub s\ s{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isasymopen}of{\isasymclose}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharbackslash}{\kern0pt}phantom{\isacharbraceleft}{\kern0pt}{\isachardollar}{\kern0pt}{\isacharbar}{\kern0pt}{\isachardollar}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}None{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ \ {\isasymopen}{\isasymRightarrow}{\isasymclose}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ None{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isachardollar}{\kern0pt}{\isacharbar}{\kern0pt}{\isachardollar}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}Some\ v{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isasymopen}{\isasymRightarrow}{\isasymclose}\ {\isasymopen}Some\ {\isacharparenleft}{\kern0pt}inj\ r\ c\ {\isacharparenleft}{\kern0pt}f\isactrlsub r\ v{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isasymclose}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}center{\isacharbraceright}{\kern0pt}\ \ {\isacharbackslash}{\kern0pt}noindent\ In\ the\ second\ clause\ we\ first\ calculate\ the\ derivative\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}der\ c\ r{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ and\ then\ simplify\ the\ result{\isachardot}{\kern0pt}\ This\ gives\ us\ a\ simplified\ derivative\ {\isasymopen}r\isactrlsub s{\isasymclose}\ and\ a\ rectification\ function\ {\isasymopen}f\isactrlsub r{\isasymclose}{\isachardot}{\kern0pt}\ The\ lexer\ is\ then\ recursively\ called\ with\ the\ simplified\ derivative{\isacharcomma}{\kern0pt}\ but\ before\ we\ inject\ the\ character\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ c{\isacharbraceright}{\kern0pt}\ into\ the\ value\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ v{\isacharbraceright}{\kern0pt}{\isacharcomma}{\kern0pt}\ we\ need\ to\ rectify\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ v{\isacharbraceright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}that\ is\ construct\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}f\isactrlsub r\ v{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardot}{\kern0pt}\ Before\ we\ can\ establish\ the\ correctness\ of\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}slexer{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharcomma}{\kern0pt}\ we\ need\ to\ show\ that\ simplification\ preserves\ the\ language\ and\ simplification\ preserves\ our\ POSIX\ relation\ once\ the\ value\ is\ rectified\ {\isacharparenleft}{\kern0pt}recall\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}const\ {\isachardoublequote}{\kern0pt}simp{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ generates\ a\ {\isacharparenleft}{\kern0pt}regular\ expression{\isacharcomma}{\kern0pt}\ rectification\ function{\isacharparenright}{\kern0pt}\ pair{\isacharparenright}{\kern0pt}{\isacharcolon}{\kern0pt}\ \ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}lemma{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}mbox{\isacharbraceleft}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}smallskip{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}label{\isacharbraceleft}{\kern0pt}slexeraux{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}{\isacharbraceleft}{\kern0pt}ll{\isacharbraceright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ L{\isacharunderscore}{\kern0pt}fst{\isacharunderscore}{\kern0pt}simp{\isacharbrackleft}{\kern0pt}symmetric{\isacharbrackright}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharbackslash}{\kern0pt}{\isacharbackslash}{\kern0pt}\ {\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharampersand}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm{\isacharbrackleft}{\kern0pt}mode{\isacharequal}{\kern0pt}IfThen{\isacharbrackright}{\kern0pt}\ Posix{\isacharunderscore}{\kern0pt}simp{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}tabular{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}lemma{\isacharbraceright}{\kern0pt}\ \ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}proof{\isacharbraceright}{\kern0pt}\ Both\ are\ by\ induction\ on\ {\isasymopen}r{\isasymclose}{\isachardot}{\kern0pt}\ There\ is\ no\ interesting\ case\ for\ the\ first\ statement{\isachardot}{\kern0pt}\ For\ the\ second\ statement{\isacharcomma}{\kern0pt}\ of\ interest\ are\ the\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}r\ {\isacharequal}{\kern0pt}\ ALT\ r\isactrlsub {\isadigit{1}}\ r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ and\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}r\ {\isacharequal}{\kern0pt}\ SEQ\ r\isactrlsub {\isadigit{1}}\ r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ cases{\isachardot}{\kern0pt}\ In\ each\ case\ we\ have\ to\ analyse\ four\ subcases\ whether\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}fst\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ and\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}fst\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ equals\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}const\ ZERO{\isacharbraceright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}respectively\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}const\ ONE{\isacharbraceright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardot}{\kern0pt}\ For\ example\ for\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}r\ {\isacharequal}{\kern0pt}\ ALT\ r\isactrlsub {\isadigit{1}}\ r\isactrlsub {\isadigit{2}}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharcomma}{\kern0pt}\ consider\ the\ subcase\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}fst\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ ZERO{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ and\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}fst\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isasymnoteq}\ ZERO{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isachardot}{\kern0pt}\ By\ assumption\ we\ know\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}s\ {\isasymin}\ fst\ {\isacharparenleft}{\kern0pt}simp\ {\isacharparenleft}{\kern0pt}ALT\ r\isactrlsub {\isadigit{1}}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isachardot}{\kern0pt}\ From\ this\ we\ can\ infer\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}s\ {\isasymin}\ fst\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ and\ by\ IH\ also\ {\isacharparenleft}{\kern0pt}{\isacharasterisk}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}s\ {\isasymin}\ r\isactrlsub {\isadigit{2}}\ {\isasymrightarrow}\ {\isacharparenleft}{\kern0pt}snd\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ v{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isachardot}{\kern0pt}\ Given\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}fst\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ ZERO{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ we\ know\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}L\ {\isacharparenleft}{\kern0pt}fst\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharbraceleft}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isachardot}{\kern0pt}\ By\ the\ first\ statement\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}L\ r\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ is\ the\ empty\ set{\isacharcomma}{\kern0pt}\ meaning\ {\isacharparenleft}{\kern0pt}{\isacharasterisk}{\kern0pt}{\isacharasterisk}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}s\ {\isasymnotin}\ L\ r\isactrlsub {\isadigit{1}}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isachardot}{\kern0pt}\ Taking\ {\isacharparenleft}{\kern0pt}{\isacharasterisk}{\kern0pt}{\isacharparenright}{\kern0pt}\ and\ {\isacharparenleft}{\kern0pt}{\isacharasterisk}{\kern0pt}{\isacharasterisk}{\kern0pt}{\isacharparenright}{\kern0pt}\ together\ gives\ by\ the\ {\isacharbackslash}{\kern0pt}mbox{\isacharbraceleft}{\kern0pt}{\isasymopen}P{\isacharplus}{\kern0pt}R{\isasymclose}{\isacharbraceright}{\kern0pt}{\isacharminus}{\kern0pt}rule\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}s\ {\isasymin}\ ALT\ r\isactrlsub {\isadigit{1}}\ r\isactrlsub {\isadigit{2}}\ {\isasymrightarrow}\ Right\ {\isacharparenleft}{\kern0pt}snd\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ v{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isachardot}{\kern0pt}\ In\ turn\ this\ gives\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}s\ {\isasymin}\ ALT\ r\isactrlsub {\isadigit{1}}\ r\isactrlsub {\isadigit{2}}\ {\isasymrightarrow}\ snd\ {\isacharparenleft}{\kern0pt}simp\ {\isacharparenleft}{\kern0pt}ALT\ r\isactrlsub {\isadigit{1}}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ v{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ as\ we\ need\ to\ show{\isachardot}{\kern0pt}\ The\ other\ cases\ are\ similar{\isachardot}{\kern0pt}{\isacharbackslash}{\kern0pt}qed\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}proof{\isacharbraceright}{\kern0pt}\ \ {\isacharbackslash}{\kern0pt}noindent\ We\ can\ now\ prove\ relatively\ straightforwardly\ that\ the\ optimised\ lexer\ produces\ the\ expected\ result{\isacharcolon}{\kern0pt}\ \ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}theorem{\isacharbraceright}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}thm\ slexer{\isacharunderscore}{\kern0pt}correctness{\isacharbraceright}{\kern0pt}\ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}theorem{\isacharbraceright}{\kern0pt}\ \ {\isacharbackslash}{\kern0pt}begin{\isacharbraceleft}{\kern0pt}proof{\isacharbraceright}{\kern0pt}\ By\ induction\ on\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ s{\isacharbraceright}{\kern0pt}\ generalising\ over\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ r{\isacharbraceright}{\kern0pt}{\isachardot}{\kern0pt}\ The\ case\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ is\ trivial{\isachardot}{\kern0pt}\ For\ the\ cons{\isacharminus}{\kern0pt}case\ suppose\ the\ string\ is\ of\ the\ form\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}c\ {\isacharhash}{\kern0pt}\ s{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isachardot}{\kern0pt}\ By\ induction\ hypothesis\ we\ know\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}slexer\ r\ s\ {\isacharequal}{\kern0pt}\ lexer\ r\ s{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ holds\ for\ all\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ r{\isacharbraceright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}in\ particular\ for\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}r{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ being\ the\ derivative\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}der\ c\ r{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardot}{\kern0pt}\ Let\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}r\isactrlsub s{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ be\ the\ simplified\ derivative\ regular\ expression{\isacharcomma}{\kern0pt}\ that\ is\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}fst\ {\isacharparenleft}{\kern0pt}simp\ {\isacharparenleft}{\kern0pt}der\ c\ r{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharcomma}{\kern0pt}\ and\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}f\isactrlsub r{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ be\ the\ rectification\ function{\isacharcomma}{\kern0pt}\ that\ is\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}snd\ {\isacharparenleft}{\kern0pt}simp\ {\isacharparenleft}{\kern0pt}der\ c\ r{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isachardot}{\kern0pt}\ \ We\ distinguish\ the\ cases\ whether\ {\isacharparenleft}{\kern0pt}{\isacharasterisk}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}s\ {\isasymin}\ L\ {\isacharparenleft}{\kern0pt}der\ c\ r{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ or\ not{\isachardot}{\kern0pt}\ In\ the\ first\ case\ we\ have\ by\ Theorem{\isachartilde}{\kern0pt}{\isacharbackslash}{\kern0pt}ref{\isacharbraceleft}{\kern0pt}lexercorrect{\isacharbraceright}{\kern0pt}{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}\ a\ value\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}v{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ so\ that\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}lexer\ {\isacharparenleft}{\kern0pt}der\ c\ r{\isacharparenright}{\kern0pt}\ s\ {\isacharequal}{\kern0pt}\ Some\ v{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ and\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}s\ {\isasymin}\ der\ c\ r\ {\isasymrightarrow}\ v{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ hold{\isachardot}{\kern0pt}\ By\ Lemma{\isachartilde}{\kern0pt}{\isacharbackslash}{\kern0pt}ref{\isacharbraceleft}{\kern0pt}slexeraux{\isacharbraceright}{\kern0pt}{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}\ we\ can\ also\ infer\ from{\isachartilde}{\kern0pt}{\isacharparenleft}{\kern0pt}{\isacharasterisk}{\kern0pt}{\isacharparenright}{\kern0pt}\ that\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}s\ {\isasymin}\ L\ r\isactrlsub s{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ holds{\isachardot}{\kern0pt}\ \ Hence\ we\ know\ by\ Theorem{\isachartilde}{\kern0pt}{\isacharbackslash}{\kern0pt}ref{\isacharbraceleft}{\kern0pt}lexercorrect{\isacharbraceright}{\kern0pt}{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}\ that\ there\ exists\ a\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}v{\isacharprime}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ with\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}lexer\ r\isactrlsub s\ s\ {\isacharequal}{\kern0pt}\ Some\ v{\isacharprime}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ and\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}s\ {\isasymin}\ r\isactrlsub s\ {\isasymrightarrow}\ v{\isacharprime}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isachardot}{\kern0pt}\ From\ the\ latter\ we\ know\ by\ Lemma{\isachartilde}{\kern0pt}{\isacharbackslash}{\kern0pt}ref{\isacharbraceleft}{\kern0pt}slexeraux{\isacharbraceright}{\kern0pt}{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}\ that\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}s\ {\isasymin}\ der\ c\ r\ {\isasymrightarrow}\ {\isacharparenleft}{\kern0pt}f\isactrlsub r\ v{\isacharprime}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ holds{\isachardot}{\kern0pt}\ By\ the\ uniqueness\ of\ the\ POSIX\ relation\ {\isacharparenleft}{\kern0pt}Theorem{\isachartilde}{\kern0pt}{\isacharbackslash}{\kern0pt}ref{\isacharbraceleft}{\kern0pt}posixdeterm{\isacharbraceright}{\kern0pt}{\isacharparenright}{\kern0pt}\ we\ can\ infer\ that\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ v{\isacharbraceright}{\kern0pt}\ is\ equal\ to\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}f\isactrlsub r\ v{\isacharprime}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharminus}{\kern0pt}{\isacharminus}{\kern0pt}{\isacharminus}{\kern0pt}that\ is\ the\ rectification\ function\ applied\ to\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}v{\isacharprime}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ produces\ the\ original\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}v{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isachardot}{\kern0pt}\ \ Now\ the\ case\ follows\ by\ the\ definitions\ of\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}const\ lexer{\isacharbraceright}{\kern0pt}\ and\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}const\ slexer{\isacharbraceright}{\kern0pt}{\isachardot}{\kern0pt}\ \ In\ the\ second\ case\ where\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}s\ {\isasymnotin}\ L\ {\isacharparenleft}{\kern0pt}der\ c\ r{\isacharparenright}{\kern0pt}{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ we\ have\ that\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}lexer\ {\isacharparenleft}{\kern0pt}der\ c\ r{\isacharparenright}{\kern0pt}\ s\ {\isacharequal}{\kern0pt}\ None{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ by\ Theorem{\isachartilde}{\kern0pt}{\isacharbackslash}{\kern0pt}ref{\isacharbraceleft}{\kern0pt}lexercorrect{\isacharbraceright}{\kern0pt}{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isachardot}{\kern0pt}\ \ We\ also\ know\ by\ Lemma{\isachartilde}{\kern0pt}{\isacharbackslash}{\kern0pt}ref{\isacharbraceleft}{\kern0pt}slexeraux{\isacharbraceright}{\kern0pt}{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}\ that\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}s\ {\isasymnotin}\ L\ r\isactrlsub s{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isachardot}{\kern0pt}\ Hence\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}lexer\ r\isactrlsub s\ s\ {\isacharequal}{\kern0pt}\ None{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}\ by\ Theorem{\isachartilde}{\kern0pt}{\isacharbackslash}{\kern0pt}ref{\isacharbraceleft}{\kern0pt}lexercorrect{\isacharbraceright}{\kern0pt}{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}\ and\ by\ IH\ then\ also\ {\isacharat}{\kern0pt}{\isacharbraceleft}{\kern0pt}term\ {\isachardoublequote}{\kern0pt}slexer\ r\isactrlsub s\ s\ {\isacharequal}{\kern0pt}\ None{\isachardoublequote}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isachardot}{\kern0pt}\ With\ this\ we\ can\ conclude\ in\ this\ case\ too{\isachardot}{\kern0pt}{\isacharbackslash}{\kern0pt}qed\ \ {\isacharbackslash}{\kern0pt}end{\isacharbraceleft}{\kern0pt}proof{\isacharbraceright}{\kern0pt}\ \ }
-fy the result. This gives us a simplified derivative
- \isa{r\isactrlsub s} and a rectification function \isa{f\isactrlsub r}. The lexer
- is then recursively called with the simplified derivative, but before
- we inject the character \isa{c} into the value \isa{v}, we need to rectify
- \isa{v} (that is construct \isa{f\isactrlsub r\ v}). Before we can establish the correctness
- of \isa{lexer\isactrlsup {\isacharplus}{\kern0pt}}, we need to show that simplification preserves the language
- and simplification preserves our POSIX relation once the value is rectified
- (recall \isa{simp} generates a (regular expression, rectification function) pair):
-
- \begin{lemma}\mbox{}\smallskip\\\label{slexeraux}
- \begin{tabular}{ll}
- (1) & \isa{L{\isacharparenleft}{\kern0pt}fst\ {\isacharparenleft}{\kern0pt}simp\ r{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ L{\isacharparenleft}{\kern0pt}r{\isacharparenright}{\kern0pt}}\\
- (2) & \isa{{\normalsize{}If\,}\ {\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ fst\ {\isacharparenleft}{\kern0pt}simp\ r{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v\ {\normalsize \,then\,}\ {\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ snd\ {\isacharparenleft}{\kern0pt}simp\ r{\isacharparenright}{\kern0pt}\ v{\isachardot}{\kern0pt}}
- \end{tabular}
- \end{lemma}
-
- \begin{proof} Both are by induction on \isa{r}. There is no
- interesting case for the first statement. For the second statement,
- of interest are the \isa{r\ {\isacharequal}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}} and \isa{r\ {\isacharequal}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isasymcdot}\ r\isactrlsub {\isadigit{2}}} cases. In each case we have to analyse four subcases whether
- \isa{fst\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}} and \isa{fst\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}} equals \isa{\isactrlbold {\isadigit{0}}} (respectively \isa{\isactrlbold {\isadigit{1}}}). For example for \isa{r\ {\isacharequal}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}}, consider the subcase \isa{fst\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ \isactrlbold {\isadigit{0}}} and
- \isa{fst\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isasymnoteq}\ \isactrlbold {\isadigit{0}}}. By assumption we know \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ fst\ {\isacharparenleft}{\kern0pt}simp\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v}. From this we can infer \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ fst\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v}
- and by IH also (*) \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ snd\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ v}. Given \isa{fst\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ \isactrlbold {\isadigit{0}}}
- we know \isa{L{\isacharparenleft}{\kern0pt}fst\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isasymemptyset}}. By the first statement
- \isa{L{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}} is the empty set, meaning (**) \isa{s\ {\isasymnotin}\ L{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}{\isacharparenright}{\kern0pt}}.
- Taking (*) and (**) together gives by the \mbox{\isa{P{\isacharplus}{\kern0pt}R}}-rule
- \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ Right\ {\isacharparenleft}{\kern0pt}snd\ {\isacharparenleft}{\kern0pt}simp\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ v{\isacharparenright}{\kern0pt}}. In turn this
- gives \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ snd\ {\isacharparenleft}{\kern0pt}simp\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\ {\isacharplus}{\kern0pt}\ r\isactrlsub {\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ v} as we need to show.
- The other cases are similar.\qed
- \end{proof}
-
- \noindent We can now prove relatively straightforwardly that the
- optimised lexer produces the expected result:
-
- \begin{theorem}
- \isa{lexer\isactrlsup {\isacharplus}{\kern0pt}\ r\ s\ {\isacharequal}{\kern0pt}\ lexer\ r\ s}
- \end{theorem}
-
- \begin{proof} By induction on \isa{s} generalising over \isa{r}. The case \isa{{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}} is trivial. For the cons-case suppose the
- string is of the form \isa{c\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}s}. By induction hypothesis we
- know \isa{lexer\isactrlsup {\isacharplus}{\kern0pt}\ r\ s\ {\isacharequal}{\kern0pt}\ lexer\ r\ s} holds for all \isa{r} (in
- particular for \isa{r} being the derivative \isa{r{\isacharbackslash}{\kern0pt}c}). Let \isa{r\isactrlsub s} be the simplified derivative regular expression, that is \isa{fst\ {\isacharparenleft}{\kern0pt}simp\ {\isacharparenleft}{\kern0pt}r{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}}, and \isa{f\isactrlsub r} be the rectification
- function, that is \isa{snd\ {\isacharparenleft}{\kern0pt}simp\ {\isacharparenleft}{\kern0pt}r{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}}. We distinguish the cases
- whether (*) \isa{s\ {\isasymin}\ L{\isacharparenleft}{\kern0pt}r{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}} or not. In the first case we
- have by Theorem~\ref{lexercorrect}(2) a value \isa{v} so that \isa{lexer\ {\isacharparenleft}{\kern0pt}r{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ s\ {\isacharequal}{\kern0pt}\ Some\ v} and \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v} hold.
- By Lemma~\ref{slexeraux}(1) we can also infer from~(*) that \isa{s\ {\isasymin}\ L{\isacharparenleft}{\kern0pt}r\isactrlsub s{\isacharparenright}{\kern0pt}} holds. Hence we know by Theorem~\ref{lexercorrect}(2) that
- there exists a \isa{v{\isacharprime}{\kern0pt}} with \isa{lexer\ r\isactrlsub s\ s\ {\isacharequal}{\kern0pt}\ Some\ v{\isacharprime}{\kern0pt}} and
- \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r\isactrlsub s{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ v{\isacharprime}{\kern0pt}}. From the latter we know by
- Lemma~\ref{slexeraux}(2) that \isa{{\isacharparenleft}{\kern0pt}s{\isacharcomma}{\kern0pt}\ r{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isasymrightarrow}\ f\isactrlsub r\ v{\isacharprime}{\kern0pt}} holds.
- By the uniqueness of the POSIX relation (Theorem~\ref{posixdeterm}) we
- can infer that \isa{v} is equal to \isa{f\isactrlsub r\ v{\isacharprime}{\kern0pt}}---that is the
- rectification function applied to \isa{v{\isacharprime}{\kern0pt}}
- produces the original \isa{v}. Now the case follows by the
- definitions of \isa{lexer} and \isa{lexer\isactrlsup {\isacharplus}{\kern0pt}}.
-
- In the second case where \isa{s\ {\isasymnotin}\ L{\isacharparenleft}{\kern0pt}r{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}} we have that
- \isa{lexer\ {\isacharparenleft}{\kern0pt}r{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ s\ {\isacharequal}{\kern0pt}\ None} by Theorem~\ref{lexercorrect}(1). We
- also know by Lemma~\ref{slexeraux}(1) that \isa{s\ {\isasymnotin}\ L{\isacharparenleft}{\kern0pt}r\isactrlsub s{\isacharparenright}{\kern0pt}}. Hence
- \isa{lexer\ r\isactrlsub s\ s\ {\isacharequal}{\kern0pt}\ None} by Theorem~\ref{lexercorrect}(1) and
- by IH then also \isa{lexer\isactrlsup {\isacharplus}{\kern0pt}\ r\isactrlsub s\ s\ {\isacharequal}{\kern0pt}\ None}. With this we can
- conclude in this case too.\qed
-
- \end{proof}%
-\end{isamarkuptext}\isamarkuptrue%
-%
-\isadelimdocument
-%
-\endisadelimdocument
-%
-\isatagdocument
-%
-\isamarkupsection{HERE%
-}
-\isamarkuptrue%
-%
-\endisatagdocument
-{\isafolddocument}%
-%
-\isadelimdocument
-%
-\endisadelimdocument
-%
-\begin{isamarkuptext}%
-\begin{lemma}
- \isa{{\normalsize{}If\,}\ v\ {\isacharcolon}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r\mbox{$^\downarrow$}{\isacharparenright}{\kern0pt}{\isacharbackslash}{\kern0pt}c\ {\normalsize \,then\,}\ retrieve\ {\isacharparenleft}{\kern0pt}r\mbox{$\bbslash$}c{\isacharparenright}{\kern0pt}\ v\ {\isacharequal}{\kern0pt}\ retrieve\ r\ {\isacharparenleft}{\kern0pt}inj\ {\isacharparenleft}{\kern0pt}r\mbox{$^\downarrow$}{\isacharparenright}{\kern0pt}\ c\ v{\isacharparenright}{\kern0pt}{\isachardot}{\kern0pt}}
- \end{lemma}
-
- \begin{proof}
- By induction on the definition of \isa{r\mbox{$^\downarrow$}}. The cases for rule 1) and 2) are
- straightforward as \isa{\isactrlbold {\isadigit{0}}{\isacharbackslash}{\kern0pt}c} and \isa{\isactrlbold {\isadigit{1}}{\isacharbackslash}{\kern0pt}c} are both equal to
- \isa{\isactrlbold {\isadigit{0}}}. This means \isa{v\ {\isacharcolon}{\kern0pt}\ \isactrlbold {\isadigit{0}}} cannot hold. Similarly in case of rule 3)
- where \isa{r} is of the form \isa{ACHAR\ d} with \isa{c\ {\isacharequal}{\kern0pt}\ d}. Then by assumption
- we know \isa{v\ {\isacharcolon}{\kern0pt}\ \isactrlbold {\isadigit{1}}}, which implies \isa{v\ {\isacharequal}{\kern0pt}\ Empty}. The equation follows by
- simplification of left- and right-hand side. In case \isa{c\ {\isasymnoteq}\ d} we have again
- \isa{v\ {\isacharcolon}{\kern0pt}\ \isactrlbold {\isadigit{0}}}, which cannot hold.
-
- For rule 4a) we have again \isa{v\ {\isacharcolon}{\kern0pt}\ \isactrlbold {\isadigit{0}}}. The property holds by IH for rule 4b).
- The induction hypothesis is
- \[
- \isa{retrieve\ {\isacharparenleft}{\kern0pt}r\mbox{$\bbslash$}c{\isacharparenright}{\kern0pt}\ v\ {\isacharequal}{\kern0pt}\ retrieve\ r\ {\isacharparenleft}{\kern0pt}inj\ {\isacharparenleft}{\kern0pt}r\mbox{$^\downarrow$}{\isacharparenright}{\kern0pt}\ c\ v{\isacharparenright}{\kern0pt}}
- \]
- which is what left- and right-hand side simplify to. The slightly more interesting case
- is for 4c). By assumption we have
- \isa{v\ {\isacharcolon}{\kern0pt}\ {\isacharparenleft}{\kern0pt}{\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\mbox{$^\downarrow$}{\isacharparenright}{\kern0pt}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isacharplus}{\kern0pt}\ {\isacharparenleft}{\kern0pt}{\isacharparenleft}{\kern0pt}{\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\mbox{$^\downarrow$}{\isacharparenright}{\kern0pt}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}}. This means we
- have either (*) \isa{v{\isadigit{1}}\ {\isacharcolon}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{1}}\mbox{$^\downarrow$}{\isacharparenright}{\kern0pt}{\isacharbackslash}{\kern0pt}c} with \isa{v\ {\isacharequal}{\kern0pt}\ Left\ v{\isadigit{1}}} or
- (**) \isa{v{\isadigit{2}}\ {\isacharcolon}{\kern0pt}\ {\isacharparenleft}{\kern0pt}{\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}r\isactrlsub {\isadigit{2}}\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\mbox{$^\downarrow$}{\isacharparenright}{\kern0pt}{\isacharbackslash}{\kern0pt}c} with \isa{v\ {\isacharequal}{\kern0pt}\ Right\ v{\isadigit{2}}}.
- The former case is straightforward by simplification. The second case is \ldots TBD.
-
- Rule 5) TBD.
-
- Finally for rule 6) the reasoning is as follows: By assumption we have
- \isa{v\ {\isacharcolon}{\kern0pt}\ {\isacharparenleft}{\kern0pt}{\isacharparenleft}{\kern0pt}r\mbox{$^\downarrow$}{\isacharparenright}{\kern0pt}{\isacharbackslash}{\kern0pt}c{\isacharparenright}{\kern0pt}\ {\isasymcdot}\ {\isacharparenleft}{\kern0pt}r\mbox{$^\downarrow$}{\isacharparenright}{\kern0pt}\isactrlsup {\isasymstar}}. This means we also have
- \isa{v\ {\isacharequal}{\kern0pt}\ Seq\ v{\isadigit{1}}\ v{\isadigit{2}}}, \isa{v{\isadigit{1}}\ {\isacharcolon}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r\mbox{$^\downarrow$}{\isacharparenright}{\kern0pt}{\isacharbackslash}{\kern0pt}c} and \isa{v{\isadigit{2}}\ {\isacharequal}{\kern0pt}\ Stars\ vs}.
- We want to prove
- \begin{align}
- & \isa{retrieve\ {\isacharparenleft}{\kern0pt}ASEQ\ bs\ {\isacharparenleft}{\kern0pt}fuse\ {\isacharbrackleft}{\kern0pt}Z{\isacharbrackright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r\mbox{$\bbslash$}c{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}ASTAR\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ r{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ v}\\
- &= \isa{retrieve\ {\isacharparenleft}{\kern0pt}ASTAR\ bs\ r{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}inj\ {\isacharparenleft}{\kern0pt}{\isacharparenleft}{\kern0pt}r\mbox{$^\downarrow$}{\isacharparenright}{\kern0pt}\isactrlsup {\isasymstar}{\isacharparenright}{\kern0pt}\ c\ v{\isacharparenright}{\kern0pt}}
- \end{align}
- The right-hand side \isa{inj}-expression is equal to
- \isa{Stars\ {\isacharparenleft}{\kern0pt}inj\ {\isacharparenleft}{\kern0pt}r\mbox{$^\downarrow$}{\isacharparenright}{\kern0pt}\ c\ v{\isadigit{1}}\mbox{$\,$}{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\mbox{$\,$}vs{\isacharparenright}{\kern0pt}}, which means the \isa{retrieve}-expression
- simplifies to
- \[
- \isa{bs\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}Z{\isacharbrackright}{\kern0pt}\ {\isacharat}{\kern0pt}\ retrieve\ r\ {\isacharparenleft}{\kern0pt}inj\ {\isacharparenleft}{\kern0pt}r\mbox{$^\downarrow$}{\isacharparenright}{\kern0pt}\ c\ v{\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharat}{\kern0pt}\ retrieve\ {\isacharparenleft}{\kern0pt}ASTAR\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}Stars\ vs{\isacharparenright}{\kern0pt}}
- \]
- The left-hand side (3) above simplifies to
- \[
- \isa{bs\ {\isacharat}{\kern0pt}\ retrieve\ {\isacharparenleft}{\kern0pt}fuse\ {\isacharbrackleft}{\kern0pt}Z{\isacharbrackright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r\mbox{$\bbslash$}c{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ v{\isadigit{1}}\ {\isacharat}{\kern0pt}\ retrieve\ {\isacharparenleft}{\kern0pt}ASTAR\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}Stars\ vs{\isacharparenright}{\kern0pt}}
- \]
- We can move out the \isa{fuse\ {\isacharbrackleft}{\kern0pt}Z{\isacharbrackright}{\kern0pt}} and then use the IH to show that left-hand side
- and right-hand side are equal. This completes the proof.
- \end{proof}
-
-
-
- \bibliographystyle{plain}
- \bibliography{root}%
-\end{isamarkuptext}\isamarkuptrue%
-%
-\isadelimtheory
-%
-\endisadelimtheory
-%
-\isatagtheory
-%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-\isanewline
-%
-\endisadelimtheory
-\isanewline
-\isanewline
-%
-\end{isabellebody}%
-\endinput
-%:%file=~/Dropbox/Workspace/journalpaper/lexing/thys2/Journal/Paper.thy%:%
-%:%50=134%:%
-%:%62=136%:%
-%:%63=137%:%
-%:%64=138%:%
-%:%65=139%:%
-%:%66=140%:%
-%:%67=141%:%
-%:%68=142%:%
-%:%69=143%:%
-%:%70=144%:%
-%:%71=145%:%
-%:%72=146%:%
-%:%73=147%:%
-%:%74=148%:%
-%:%75=149%:%
-%:%76=150%:%
-%:%77=151%:%
-%:%78=152%:%
-%:%79=153%:%
-%:%80=154%:%
-%:%81=155%:%
-%:%82=156%:%
-%:%83=157%:%
-%:%84=158%:%
-%:%85=159%:%
-%:%86=160%:%
-%:%87=161%:%
-%:%88=162%:%
-%:%89=163%:%
-%:%90=164%:%
-%:%91=165%:%
-%:%92=166%:%
-%:%93=167%:%
-%:%94=168%:%
-%:%95=169%:%
-%:%96=170%:%
-%:%97=171%:%
-%:%98=172%:%
-%:%99=173%:%
-%:%100=174%:%
-%:%101=175%:%
-%:%102=176%:%
-%:%103=177%:%
-%:%104=178%:%
-%:%105=179%:%
-%:%106=180%:%
-%:%107=181%:%
-%:%108=182%:%
-%:%109=183%:%
-%:%110=184%:%
-%:%111=185%:%
-%:%112=186%:%
-%:%113=187%:%
-%:%114=188%:%
-%:%115=189%:%
-%:%116=190%:%
-%:%117=191%:%
-%:%118=192%:%
-%:%119=193%:%
-%:%120=194%:%
-%:%121=195%:%
-%:%122=196%:%
-%:%123=197%:%
-%:%124=198%:%
-%:%125=199%:%
-%:%126=200%:%
-%:%127=201%:%
-%:%128=202%:%
-%:%129=203%:%
-%:%130=204%:%
-%:%131=205%:%
-%:%132=206%:%
-%:%133=207%:%
-%:%134=208%:%
-%:%135=209%:%
-%:%136=210%:%
-%:%137=211%:%
-%:%138=212%:%
-%:%139=213%:%
-%:%140=214%:%
-%:%141=215%:%
-%:%142=216%:%
-%:%143=217%:%
-%:%144=218%:%
-%:%145=219%:%
-%:%146=220%:%
-%:%147=221%:%
-%:%148=222%:%
-%:%149=223%:%
-%:%150=224%:%
-%:%151=225%:%
-%:%152=226%:%
-%:%153=227%:%
-%:%154=228%:%
-%:%155=229%:%
-%:%156=230%:%
-%:%157=231%:%
-%:%158=232%:%
-%:%159=233%:%
-%:%160=234%:%
-%:%161=235%:%
-%:%162=236%:%
-%:%163=237%:%
-%:%164=238%:%
-%:%165=239%:%
-%:%166=240%:%
-%:%167=241%:%
-%:%168=242%:%
-%:%169=243%:%
-%:%170=244%:%
-%:%171=245%:%
-%:%172=246%:%
-%:%173=247%:%
-%:%174=248%:%
-%:%175=249%:%
-%:%176=250%:%
-%:%177=251%:%
-%:%178=252%:%
-%:%179=253%:%
-%:%180=254%:%
-%:%181=255%:%
-%:%182=256%:%
-%:%183=257%:%
-%:%184=258%:%
-%:%185=259%:%
-%:%186=260%:%
-%:%187=261%:%
-%:%188=262%:%
-%:%189=263%:%
-%:%190=264%:%
-%:%191=265%:%
-%:%192=266%:%
-%:%193=267%:%
-%:%194=268%:%
-%:%203=272%:%
-%:%215=276%:%
-%:%216=277%:%
-%:%217=278%:%
-%:%218=279%:%
-%:%219=280%:%
-%:%220=281%:%
-%:%221=282%:%
-%:%222=283%:%
-%:%223=284%:%
-%:%224=285%:%
-%:%225=286%:%
-%:%226=287%:%
-%:%227=288%:%
-%:%228=289%:%
-%:%237=297%:%
-%:%249=303%:%
-%:%250=304%:%
-%:%251=305%:%
-%:%252=306%:%
-%:%252=307%:%
-%:%253=308%:%
-%:%254=309%:%
-%:%255=310%:%
-%:%256=311%:%
-%:%257=312%:%
-%:%258=313%:%
-%:%259=314%:%
-%:%260=315%:%
-%:%261=316%:%
-%:%262=317%:%
-%:%263=318%:%
-%:%264=319%:%
-%:%265=320%:%
-%:%266=321%:%
-%:%267=322%:%
-%:%268=323%:%
-%:%269=324%:%
-%:%270=325%:%
-%:%271=326%:%
-%:%272=327%:%
-%:%273=328%:%
-%:%274=329%:%
-%:%275=330%:%
-%:%276=331%:%
-%:%277=332%:%
-%:%278=333%:%
-%:%279=334%:%
-%:%280=335%:%
-%:%281=336%:%
-%:%282=337%:%
-%:%283=338%:%
-%:%284=339%:%
-%:%285=340%:%
-%:%286=341%:%
-%:%287=342%:%
-%:%288=343%:%
-%:%289=344%:%
-%:%290=345%:%
-%:%291=346%:%
-%:%292=347%:%
-%:%293=348%:%
-%:%294=349%:%
-%:%295=350%:%
-%:%296=351%:%
-%:%297=352%:%
-%:%298=353%:%
-%:%299=354%:%
-%:%300=355%:%
-%:%301=356%:%
-%:%302=357%:%
-%:%303=358%:%
-%:%304=359%:%
-%:%305=360%:%
-%:%306=361%:%
-%:%307=362%:%
-%:%308=363%:%
-%:%309=364%:%
-%:%310=365%:%
-%:%311=366%:%
-%:%312=367%:%
-%:%313=368%:%
-%:%314=369%:%
-%:%315=370%:%
-%:%316=371%:%
-%:%316=372%:%
-%:%317=373%:%
-%:%318=374%:%
-%:%319=375%:%
-%:%320=376%:%
-%:%321=377%:%
-%:%322=378%:%
-%:%323=379%:%
-%:%324=380%:%
-%:%325=381%:%
-%:%326=382%:%
-%:%327=383%:%
-%:%328=384%:%
-%:%329=385%:%
-%:%330=386%:%
-%:%331=387%:%
-%:%332=388%:%
-%:%333=389%:%
-%:%334=390%:%
-%:%335=391%:%
-%:%336=392%:%
-%:%337=393%:%
-%:%338=394%:%
-%:%339=395%:%
-%:%340=396%:%
-%:%341=397%:%
-%:%342=398%:%
-%:%343=399%:%
-%:%344=400%:%
-%:%345=401%:%
-%:%346=402%:%
-%:%347=403%:%
-%:%348=404%:%
-%:%349=405%:%
-%:%350=406%:%
-%:%351=407%:%
-%:%352=408%:%
-%:%353=409%:%
-%:%354=410%:%
-%:%355=411%:%
-%:%356=412%:%
-%:%357=413%:%
-%:%358=414%:%
-%:%359=415%:%
-%:%360=416%:%
-%:%361=417%:%
-%:%362=418%:%
-%:%363=419%:%
-%:%364=420%:%
-%:%365=421%:%
-%:%366=422%:%
-%:%367=423%:%
-%:%368=424%:%
-%:%369=425%:%
-%:%370=426%:%
-%:%371=427%:%
-%:%372=428%:%
-%:%373=429%:%
-%:%374=430%:%
-%:%375=431%:%
-%:%376=432%:%
-%:%377=433%:%
-%:%378=434%:%
-%:%379=435%:%
-%:%380=436%:%
-%:%381=437%:%
-%:%382=438%:%
-%:%383=439%:%
-%:%384=440%:%
-%:%385=441%:%
-%:%386=442%:%
-%:%387=443%:%
-%:%388=444%:%
-%:%389=445%:%
-%:%390=446%:%
-%:%391=447%:%
-%:%392=448%:%
-%:%393=449%:%
-%:%394=450%:%
-%:%395=451%:%
-%:%396=452%:%
-%:%397=453%:%
-%:%398=454%:%
-%:%399=455%:%
-%:%400=456%:%
-%:%401=457%:%
-%:%402=458%:%
-%:%403=459%:%
-%:%404=460%:%
-%:%405=461%:%
-%:%406=462%:%
-%:%407=463%:%
-%:%408=464%:%
-%:%409=465%:%
-%:%410=466%:%
-%:%411=467%:%
-%:%412=468%:%
-%:%413=469%:%
-%:%414=470%:%
-%:%415=471%:%
-%:%416=472%:%
-%:%417=473%:%
-%:%418=474%:%
-%:%419=475%:%
-%:%420=476%:%
-%:%421=477%:%
-%:%430=484%:%
-%:%442=486%:%
-%:%443=487%:%
-%:%443=488%:%
-%:%444=489%:%
-%:%445=490%:%
-%:%446=491%:%
-%:%447=492%:%
-%:%448=493%:%
-%:%449=494%:%
-%:%450=495%:%
-%:%451=496%:%
-%:%452=497%:%
-%:%453=498%:%
-%:%454=499%:%
-%:%455=500%:%
-%:%456=501%:%
-%:%457=502%:%
-%:%458=503%:%
-%:%459=504%:%
-%:%460=505%:%
-%:%461=506%:%
-%:%462=507%:%
-%:%463=508%:%
-%:%464=509%:%
-%:%465=510%:%
-%:%466=511%:%
-%:%467=512%:%
-%:%468=513%:%
-%:%469=514%:%
-%:%470=515%:%
-%:%471=516%:%
-%:%472=517%:%
-%:%473=518%:%
-%:%474=519%:%
-%:%475=520%:%
-%:%476=521%:%
-%:%477=522%:%
-%:%478=523%:%
-%:%479=524%:%
-%:%480=525%:%
-%:%481=526%:%
-%:%482=527%:%
-%:%482=528%:%
-%:%483=529%:%
-%:%484=530%:%
-%:%485=531%:%
-%:%486=532%:%
-%:%487=533%:%
-%:%487=534%:%
-%:%488=535%:%
-%:%489=536%:%
-%:%490=537%:%
-%:%491=538%:%
-%:%492=539%:%
-%:%493=540%:%
-%:%494=541%:%
-%:%495=542%:%
-%:%496=543%:%
-%:%497=544%:%
-%:%498=545%:%
-%:%499=546%:%
-%:%500=547%:%
-%:%501=548%:%
-%:%502=549%:%
-%:%503=550%:%
-%:%504=551%:%
-%:%505=552%:%
-%:%506=553%:%
-%:%507=554%:%
-%:%508=555%:%
-%:%509=556%:%
-%:%510=557%:%
-%:%511=558%:%
-%:%512=559%:%
-%:%513=560%:%
-%:%514=561%:%
-%:%515=562%:%
-%:%516=563%:%
-%:%517=564%:%
-%:%518=565%:%
-%:%519=566%:%
-%:%520=567%:%
-%:%521=568%:%
-%:%522=569%:%
-%:%523=570%:%
-%:%524=571%:%
-%:%525=572%:%
-%:%526=573%:%
-%:%527=574%:%
-%:%528=575%:%
-%:%529=576%:%
-%:%530=577%:%
-%:%531=578%:%
-%:%532=579%:%
-%:%533=580%:%
-%:%534=581%:%
-%:%535=582%:%
-%:%536=583%:%
-%:%537=584%:%
-%:%538=585%:%
-%:%539=586%:%
-%:%540=587%:%
-%:%541=588%:%
-%:%542=589%:%
-%:%543=590%:%
-%:%544=591%:%
-%:%545=592%:%
-%:%546=593%:%
-%:%547=594%:%
-%:%548=595%:%
-%:%549=596%:%
-%:%550=597%:%
-%:%551=598%:%
-%:%552=599%:%
-%:%553=600%:%
-%:%554=601%:%
-%:%555=602%:%
-%:%556=603%:%
-%:%557=604%:%
-%:%558=605%:%
-%:%559=606%:%
-%:%560=607%:%
-%:%561=608%:%
-%:%562=609%:%
-%:%563=610%:%
-%:%564=611%:%
-%:%565=612%:%
-%:%566=613%:%
-%:%567=614%:%
-%:%568=615%:%
-%:%569=616%:%
-%:%570=617%:%
-%:%571=618%:%
-%:%572=619%:%
-%:%573=620%:%
-%:%574=621%:%
-%:%575=622%:%
-%:%576=623%:%
-%:%585=627%:%
-%:%597=631%:%
-%:%598=632%:%
-%:%599=633%:%
-%:%600=634%:%
-%:%601=635%:%
-%:%602=636%:%
-%:%603=637%:%
-%:%604=638%:%
-%:%605=639%:%
-%:%606=640%:%
-%:%607=641%:%
-%:%608=642%:%
-%:%609=643%:%
-%:%610=644%:%
-%:%611=645%:%
-%:%612=646%:%
-%:%613=647%:%
-%:%614=648%:%
-%:%615=649%:%
-%:%616=650%:%
-%:%617=651%:%
-%:%618=652%:%
-%:%619=653%:%
-%:%620=654%:%
-%:%621=655%:%
-%:%622=656%:%
-%:%623=657%:%
-%:%624=658%:%
-%:%625=659%:%
-%:%626=660%:%
-%:%627=661%:%
-%:%628=662%:%
-%:%629=663%:%
-%:%630=664%:%
-%:%631=665%:%
-%:%632=666%:%
-%:%633=667%:%
-%:%634=668%:%
-%:%635=669%:%
-%:%636=670%:%
-%:%637=671%:%
-%:%638=672%:%
-%:%639=673%:%
-%:%640=674%:%
-%:%641=675%:%
-%:%642=676%:%
-%:%643=677%:%
-%:%644=678%:%
-%:%645=679%:%
-%:%646=680%:%
-%:%647=681%:%
-%:%648=682%:%
-%:%649=683%:%
-%:%650=684%:%
-%:%651=685%:%
-%:%652=686%:%
-%:%653=687%:%
-%:%654=688%:%
-%:%655=689%:%
-%:%656=690%:%
-%:%657=691%:%
-%:%658=692%:%
-%:%659=693%:%
-%:%660=694%:%
-%:%661=695%:%
-%:%662=696%:%
-%:%663=697%:%
-%:%664=698%:%
-%:%665=699%:%
-%:%666=700%:%
-%:%667=701%:%
-%:%668=702%:%
-%:%669=703%:%
-%:%670=704%:%
-%:%671=705%:%
-%:%672=706%:%
-%:%673=707%:%
-%:%674=708%:%
-%:%675=709%:%
-%:%676=710%:%
-%:%677=711%:%
-%:%678=712%:%
-%:%679=713%:%
-%:%680=714%:%
-%:%681=715%:%
-%:%682=716%:%
-%:%683=717%:%
-%:%684=718%:%
-%:%685=719%:%
-%:%686=720%:%
-%:%687=721%:%
-%:%688=722%:%
-%:%689=723%:%
-%:%690=724%:%
-%:%691=725%:%
-%:%692=726%:%
-%:%693=727%:%
-%:%694=728%:%
-%:%695=729%:%
-%:%696=730%:%
-%:%697=731%:%
-%:%698=732%:%
-%:%699=733%:%
-%:%700=734%:%
-%:%701=735%:%
-%:%702=736%:%
-%:%703=737%:%
-%:%704=738%:%
-%:%705=739%:%
-%:%706=740%:%
-%:%707=741%:%
-%:%708=742%:%
-%:%709=743%:%
-%:%710=744%:%
-%:%711=745%:%
-%:%711=746%:%
-%:%712=747%:%
-%:%712=748%:%
-%:%713=749%:%
-%:%714=750%:%
-%:%714=751%:%
-%:%715=752%:%
-%:%716=753%:%
-%:%717=754%:%
-%:%718=755%:%
-%:%719=756%:%
-%:%720=757%:%
-%:%721=758%:%
-%:%722=759%:%
-%:%723=760%:%
-%:%724=761%:%
-%:%725=762%:%
-%:%726=763%:%
-%:%727=764%:%
-%:%728=765%:%
-%:%729=766%:%
-%:%730=767%:%
-%:%731=768%:%
-%:%732=769%:%
-%:%733=770%:%
-%:%734=771%:%
-%:%735=772%:%
-%:%736=773%:%
-%:%737=774%:%
-%:%738=775%:%
-%:%739=776%:%
-%:%740=777%:%
-%:%741=778%:%
-%:%742=779%:%
-%:%743=780%:%
-%:%744=781%:%
-%:%745=782%:%
-%:%746=783%:%
-%:%747=784%:%
-%:%748=785%:%
-%:%749=786%:%
-%:%750=787%:%
-%:%751=788%:%
-%:%752=789%:%
-%:%753=790%:%
-%:%754=791%:%
-%:%755=792%:%
-%:%756=793%:%
-%:%757=794%:%
-%:%758=795%:%
-%:%759=796%:%
-%:%760=797%:%
-%:%761=798%:%
-%:%762=799%:%
-%:%763=800%:%
-%:%764=801%:%
-%:%765=802%:%
-%:%766=803%:%
-%:%767=804%:%
-%:%768=805%:%
-%:%769=806%:%
-%:%770=807%:%
-%:%770=808%:%
-%:%770=809%:%
-%:%771=810%:%
-%:%772=811%:%
-%:%773=812%:%
-%:%774=813%:%
-%:%775=814%:%
-%:%776=815%:%
-%:%777=816%:%
-%:%778=817%:%
-%:%779=818%:%
-%:%779=819%:%
-%:%780=820%:%
-%:%781=821%:%
-%:%782=822%:%
-%:%783=823%:%
-%:%784=824%:%
-%:%785=825%:%
-%:%786=826%:%
-%:%787=827%:%
-%:%788=828%:%
-%:%789=829%:%
-%:%790=830%:%
-%:%791=831%:%
-%:%792=832%:%
-%:%793=833%:%
-%:%794=834%:%
-%:%795=835%:%
-%:%796=836%:%
-%:%797=837%:%
-%:%798=838%:%
-%:%799=839%:%
-%:%800=840%:%
-%:%801=841%:%
-%:%802=842%:%
-%:%803=843%:%
-%:%804=844%:%
-%:%805=845%:%
-%:%806=846%:%
-%:%807=847%:%
-%:%808=848%:%
-%:%809=849%:%
-%:%810=850%:%
-%:%811=851%:%
-%:%812=852%:%
-%:%813=853%:%
-%:%814=854%:%
-%:%814=855%:%
-%:%815=856%:%
-%:%816=857%:%
-%:%817=858%:%
-%:%818=859%:%
-%:%819=860%:%
-%:%820=861%:%
-%:%821=862%:%
-%:%821=863%:%
-%:%822=864%:%
-%:%823=865%:%
-%:%824=866%:%
-%:%824=867%:%
-%:%825=868%:%
-%:%826=869%:%
-%:%827=870%:%
-%:%828=871%:%
-%:%829=872%:%
-%:%830=873%:%
-%:%831=874%:%
-%:%832=875%:%
-%:%833=876%:%
-%:%834=877%:%
-%:%835=878%:%
-%:%835=879%:%
-%:%836=880%:%
-%:%837=881%:%
-%:%838=882%:%
-%:%839=883%:%
-%:%839=884%:%
-%:%840=885%:%
-%:%841=886%:%
-%:%842=887%:%
-%:%843=888%:%
-%:%844=889%:%
-%:%845=890%:%
-%:%846=891%:%
-%:%847=892%:%
-%:%848=893%:%
-%:%849=894%:%
-%:%850=895%:%
-%:%851=896%:%
-%:%852=897%:%
-%:%853=898%:%
-%:%854=899%:%
-%:%855=900%:%
-%:%856=901%:%
-%:%857=902%:%
-%:%858=903%:%
-%:%859=904%:%
-%:%860=905%:%
-%:%861=906%:%
-%:%862=907%:%
-%:%863=908%:%
-%:%864=909%:%
-%:%865=910%:%
-%:%866=911%:%
-%:%867=912%:%
-%:%868=913%:%
-%:%869=914%:%
-%:%870=915%:%
-%:%871=916%:%
-%:%872=917%:%
-%:%873=918%:%
-%:%874=919%:%
-%:%875=920%:%
-%:%876=921%:%
-%:%877=922%:%
-%:%878=923%:%
-%:%879=924%:%
-%:%880=925%:%
-%:%881=926%:%
-%:%882=927%:%
-%:%883=928%:%
-%:%884=929%:%
-%:%885=930%:%
-%:%886=931%:%
-%:%887=932%:%
-%:%888=933%:%
-%:%889=934%:%
-%:%890=935%:%
-%:%891=936%:%
-%:%892=937%:%
-%:%893=938%:%
-%:%894=939%:%
-%:%895=940%:%
-%:%896=941%:%
-%:%897=942%:%
-%:%897=943%:%
-%:%898=944%:%
-%:%899=945%:%
-%:%900=946%:%
-%:%901=947%:%
-%:%902=948%:%
-%:%903=949%:%
-%:%904=950%:%
-%:%905=951%:%
-%:%906=952%:%
-%:%907=953%:%
-%:%908=954%:%
-%:%909=955%:%
-%:%910=956%:%
-%:%911=957%:%
-%:%912=958%:%
-%:%913=959%:%
-%:%914=960%:%
-%:%915=961%:%
-%:%916=962%:%
-%:%917=963%:%
-%:%918=964%:%
-%:%919=965%:%
-%:%920=966%:%
-%:%921=967%:%
-%:%922=968%:%
-%:%923=969%:%
-%:%924=970%:%
-%:%925=971%:%
-%:%926=972%:%
-%:%927=973%:%
-%:%928=974%:%
-%:%929=975%:%
-%:%930=976%:%
-%:%930=977%:%
-%:%931=978%:%
-%:%932=979%:%
-%:%933=980%:%
-%:%934=981%:%
-%:%935=982%:%
-%:%936=983%:%
-%:%937=984%:%
-%:%938=985%:%
-%:%939=986%:%
-%:%940=987%:%
-%:%941=988%:%
-%:%942=989%:%
-%:%943=990%:%
-%:%944=991%:%
-%:%945=992%:%
-%:%946=993%:%
-%:%947=994%:%
-%:%948=995%:%
-%:%949=996%:%
-%:%950=997%:%
-%:%951=998%:%
-%:%952=999%:%
-%:%953=1000%:%
-%:%954=1001%:%
-%:%955=1002%:%
-%:%956=1003%:%
-%:%957=1004%:%
-%:%958=1005%:%
-%:%959=1006%:%
-%:%960=1007%:%
-%:%960=1008%:%
-%:%961=1009%:%
-%:%962=1010%:%
-%:%963=1011%:%
-%:%964=1012%:%
-%:%965=1013%:%
-%:%966=1014%:%
-%:%967=1015%:%
-%:%968=1016%:%
-%:%968=1017%:%
-%:%969=1018%:%
-%:%970=1019%:%
-%:%971=1020%:%
-%:%972=1021%:%
-%:%973=1022%:%
-%:%974=1023%:%
-%:%975=1024%:%
-%:%976=1025%:%
-%:%977=1026%:%
-%:%978=1027%:%
-%:%979=1028%:%
-%:%980=1029%:%
-%:%981=1030%:%
-%:%982=1031%:%
-%:%983=1032%:%
-%:%984=1033%:%
-%:%985=1034%:%
-%:%986=1035%:%
-%:%987=1036%:%
-%:%988=1037%:%
-%:%989=1038%:%
-%:%990=1039%:%
-%:%991=1040%:%
-%:%992=1041%:%
-%:%993=1042%:%
-%:%994=1043%:%
-%:%995=1044%:%
-%:%996=1045%:%
-%:%997=1046%:%
-%:%998=1047%:%
-%:%999=1048%:%
-%:%1000=1049%:%
-%:%1001=1050%:%
-%:%1002=1051%:%
-%:%1003=1052%:%
-%:%1004=1053%:%
-%:%1005=1054%:%
-%:%1005=1055%:%
-%:%1005=1056%:%
-%:%1006=1057%:%
-%:%1006=1058%:%
-%:%1006=1059%:%
-%:%1007=1060%:%
-%:%1008=1061%:%
-%:%1008=1062%:%
-%:%1009=1063%:%
-%:%1010=1064%:%
-%:%1011=1065%:%
-%:%1012=1066%:%
-%:%1013=1067%:%
-%:%1014=1068%:%
-%:%1015=1069%:%
-%:%1016=1070%:%
-%:%1017=1071%:%
-%:%1018=1072%:%
-%:%1019=1073%:%
-%:%1020=1074%:%
-%:%1021=1075%:%
-%:%1022=1076%:%
-%:%1023=1077%:%
-%:%1024=1078%:%
-%:%1025=1079%:%
-%:%1026=1080%:%
-%:%1027=1081%:%
-%:%1028=1082%:%
-%:%1029=1083%:%
-%:%1030=1084%:%
-%:%1031=1085%:%
-%:%1032=1086%:%
-%:%1033=1087%:%
-%:%1034=1088%:%
-%:%1035=1089%:%
-%:%1036=1090%:%
-%:%1037=1091%:%
-%:%1038=1092%:%
-%:%1039=1093%:%
-%:%1040=1094%:%
-%:%1041=1095%:%
-%:%1042=1096%:%
-%:%1042=1097%:%
-%:%1043=1098%:%
-%:%1044=1099%:%
-%:%1045=1100%:%
-%:%1045=1101%:%
-%:%1045=1102%:%
-%:%1045=1103%:%
-%:%1046=1104%:%
-%:%1047=1105%:%
-%:%1048=1106%:%
-%:%1049=1107%:%
-%:%1050=1108%:%
-%:%1051=1109%:%
-%:%1052=1110%:%
-%:%1053=1111%:%
-%:%1054=1112%:%
-%:%1055=1113%:%
-%:%1056=1114%:%
-%:%1057=1115%:%
-%:%1058=1116%:%
-%:%1059=1117%:%
-%:%1060=1118%:%
-%:%1061=1119%:%
-%:%1062=1120%:%
-%:%1063=1121%:%
-%:%1064=1122%:%
-%:%1065=1123%:%
-%:%1066=1124%:%
-%:%1067=1125%:%
-%:%1068=1126%:%
-%:%1069=1127%:%
-%:%1070=1128%:%
-%:%1071=1129%:%
-%:%1072=1130%:%
-%:%1073=1131%:%
-%:%1074=1132%:%
-%:%1075=1133%:%
-%:%1076=1134%:%
-%:%1077=1135%:%
-%:%1078=1136%:%
-%:%1079=1137%:%
-%:%1080=1138%:%
-%:%1081=1139%:%
-%:%1082=1140%:%
-%:%1091=1144%:%
-%:%1103=1148%:%
-%:%1104=1149%:%
-%:%1105=1150%:%
-%:%1106=1151%:%
-%:%1107=1152%:%
-%:%1108=1153%:%
-%:%1109=1154%:%
-%:%1110=1155%:%
-%:%1111=1156%:%
-%:%1112=1157%:%
-%:%1113=1158%:%
-%:%1114=1159%:%
-%:%1115=1160%:%
-%:%1116=1161%:%
-%:%1117=1162%:%
-%:%1118=1163%:%
-%:%1119=1164%:%
-%:%1120=1165%:%
-%:%1121=1166%:%
-%:%1122=1167%:%
-%:%1123=1168%:%
-%:%1124=1169%:%
-%:%1125=1170%:%
-%:%1126=1171%:%
-%:%1127=1172%:%
-%:%1128=1173%:%
-%:%1129=1174%:%
-%:%1130=1175%:%
-%:%1131=1176%:%
-%:%1132=1177%:%
-%:%1133=1178%:%
-%:%1134=1179%:%
-%:%1135=1180%:%
-%:%1136=1181%:%
-%:%1137=1182%:%
-%:%1138=1183%:%
-%:%1139=1184%:%
-%:%1140=1185%:%
-%:%1141=1186%:%
-%:%1142=1187%:%
-%:%1143=1188%:%
-%:%1144=1189%:%
-%:%1145=1190%:%
-%:%1146=1191%:%
-%:%1147=1192%:%
-%:%1148=1193%:%
-%:%1149=1194%:%
-%:%1150=1195%:%
-%:%1151=1196%:%
-%:%1152=1197%:%
-%:%1153=1198%:%
-%:%1154=1199%:%
-%:%1155=1200%:%
-%:%1156=1201%:%
-%:%1157=1202%:%
-%:%1158=1203%:%
-%:%1159=1204%:%
-%:%1160=1205%:%
-%:%1161=1206%:%
-%:%1162=1207%:%
-%:%1163=1208%:%
-%:%1164=1209%:%
-%:%1165=1210%:%
-%:%1166=1211%:%
-%:%1167=1212%:%
-%:%1168=1213%:%
-%:%1169=1214%:%
-%:%1170=1215%:%
-%:%1171=1216%:%
-%:%1172=1217%:%
-%:%1173=1218%:%
-%:%1174=1219%:%
-%:%1175=1220%:%
-%:%1176=1221%:%
-%:%1177=1222%:%
-%:%1178=1223%:%
-%:%1179=1224%:%
-%:%1180=1225%:%
-%:%1181=1226%:%
-%:%1182=1227%:%
-%:%1183=1228%:%
-%:%1184=1229%:%
-%:%1185=1230%:%
-%:%1186=1231%:%
-%:%1187=1232%:%
-%:%1188=1233%:%
-%:%1189=1234%:%
-%:%1190=1235%:%
-%:%1191=1236%:%
-%:%1192=1237%:%
-%:%1193=1238%:%
-%:%1194=1239%:%
-%:%1195=1240%:%
-%:%1196=1241%:%
-%:%1197=1242%:%
-%:%1198=1243%:%
-%:%1199=1244%:%
-%:%1200=1245%:%
-%:%1201=1246%:%
-%:%1202=1247%:%
-%:%1203=1248%:%
-%:%1204=1249%:%
-%:%1205=1250%:%
-%:%1206=1251%:%
-%:%1207=1252%:%
-%:%1208=1253%:%
-%:%1209=1254%:%
-%:%1210=1255%:%
-%:%1211=1256%:%
-%:%1212=1257%:%
-%:%1213=1258%:%
-%:%1214=1259%:%
-%:%1215=1260%:%
-%:%1216=1261%:%
-%:%1217=1262%:%
-%:%1218=1263%:%
-%:%1218=1264%:%
-%:%1219=1265%:%
-%:%1220=1266%:%
-%:%1221=1267%:%
-%:%1222=1268%:%
-%:%1223=1269%:%
-%:%1224=1270%:%
-%:%1225=1271%:%
-%:%1226=1272%:%
-%:%1227=1273%:%
-%:%1228=1274%:%
-%:%1229=1275%:%
-%:%1229=1276%:%
-%:%1230=1277%:%
-%:%1231=1278%:%
-%:%1232=1279%:%
-%:%1233=1280%:%
-%:%1234=1281%:%
-%:%1235=1282%:%
-%:%1236=1283%:%
-%:%1237=1284%:%
-%:%1238=1285%:%
-%:%1239=1286%:%
-%:%1240=1287%:%
-%:%1241=1288%:%
-%:%1242=1289%:%
-%:%1243=1290%:%
-%:%1244=1291%:%
-%:%1245=1292%:%
-%:%1246=1293%:%
-%:%1247=1294%:%
-%:%1248=1295%:%
-%:%1249=1296%:%
-%:%1250=1297%:%
-%:%1251=1298%:%
-%:%1252=1299%:%
-%:%1253=1300%:%
-%:%1254=1301%:%
-%:%1255=1302%:%
-%:%1256=1303%:%
-%:%1257=1304%:%
-%:%1258=1305%:%
-%:%1259=1306%:%
-%:%1260=1307%:%
-%:%1261=1308%:%
-%:%1262=1309%:%
-%:%1263=1310%:%
-%:%1264=1311%:%
-%:%1265=1312%:%
-%:%1266=1313%:%
-%:%1267=1314%:%
-%:%1268=1315%:%
-%:%1269=1316%:%
-%:%1270=1317%:%
-%:%1271=1318%:%
-%:%1272=1319%:%
-%:%1273=1320%:%
-%:%1274=1321%:%
-%:%1275=1322%:%
-%:%1276=1323%:%
-%:%1277=1324%:%
-%:%1278=1325%:%
-%:%1279=1326%:%
-%:%1280=1327%:%
-%:%1281=1328%:%
-%:%1282=1329%:%
-%:%1283=1330%:%
-%:%1283=1331%:%
-%:%1283=1332%:%
-%:%1284=1333%:%
-%:%1285=1334%:%
-%:%1285=1335%:%
-%:%1285=1336%:%
-%:%1285=1337%:%
-%:%1286=1338%:%
-%:%1287=1339%:%
-%:%1288=1340%:%
-%:%1289=1341%:%
-%:%1290=1342%:%
-%:%1290=1343%:%
-%:%1291=1344%:%
-%:%1292=1345%:%
-%:%1293=1346%:%
-%:%1294=1347%:%
-%:%1295=1348%:%
-%:%1296=1349%:%
-%:%1297=1350%:%
-%:%1298=1351%:%
-%:%1299=1352%:%
-%:%1300=1353%:%
-%:%1301=1354%:%
-%:%1302=1355%:%
-%:%1303=1356%:%
-%:%1304=1357%:%
-%:%1305=1358%:%
-%:%1306=1359%:%
-%:%1307=1360%:%
-%:%1308=1361%:%
-%:%1309=1362%:%
-%:%1310=1363%:%
-%:%1311=1364%:%
-%:%1312=1365%:%
-%:%1313=1366%:%
-%:%1314=1367%:%
-%:%1315=1368%:%
-%:%1316=1369%:%
-%:%1317=1370%:%
-%:%1318=1371%:%
-%:%1319=1372%:%
-%:%1320=1373%:%
-%:%1321=1374%:%
-%:%1322=1375%:%
-%:%1323=1376%:%
-%:%1324=1377%:%
-%:%1325=1378%:%
-%:%1326=1379%:%
-%:%1327=1380%:%
-%:%1328=1381%:%
-%:%1329=1382%:%
-%:%1330=1383%:%
-%:%1331=1384%:%
-%:%1332=1385%:%
-%:%1333=1386%:%
-%:%1334=1387%:%
-%:%1335=1388%:%
-%:%1336=1389%:%
-%:%1337=1390%:%
-%:%1338=1391%:%
-%:%1338=1392%:%
-%:%1339=1393%:%
-%:%1339=1394%:%
-%:%1340=1395%:%
-%:%1340=1396%:%
-%:%1341=1397%:%
-%:%1342=1398%:%
-%:%1343=1399%:%
-%:%1344=1400%:%
-%:%1345=1401%:%
-%:%1346=1402%:%
-%:%1347=1403%:%
-%:%1347=1404%:%
-%:%1348=1405%:%
-%:%1348=1406%:%
-%:%1349=1407%:%
-%:%1349=1408%:%
-%:%1350=1409%:%
-%:%1351=1410%:%
-%:%1352=1411%:%
-%:%1353=1412%:%
-%:%1354=1413%:%
-%:%1355=1414%:%
-%:%1356=1415%:%
-%:%1357=1416%:%
-%:%1358=1417%:%
-%:%1359=1418%:%
-%:%1360=1419%:%
-%:%1361=1420%:%
-%:%1362=1421%:%
-%:%1363=1422%:%
-%:%1364=1423%:%
-%:%1365=1424%:%
-%:%1366=1425%:%
-%:%1367=1426%:%
-%:%1368=1427%:%
-%:%1369=1428%:%
-%:%1370=1429%:%
-%:%1371=1430%:%
-%:%1372=1431%:%
-%:%1373=1432%:%
-%:%1374=1433%:%
-%:%1375=1434%:%
-%:%1376=1435%:%
-%:%1377=1436%:%
-%:%1378=1437%:%
-%:%1379=1438%:%
-%:%1380=1439%:%
-%:%1380=1440%:%
-%:%1381=1441%:%
-%:%1382=1442%:%
-%:%1383=1443%:%
-%:%1384=1444%:%
-%:%1385=1445%:%
-%:%1386=1446%:%
-%:%1387=1447%:%
-%:%1387=1448%:%
-%:%1388=1449%:%
-%:%1389=1450%:%
-%:%1390=1451%:%
-%:%1390=1452%:%
-%:%1391=1453%:%
-%:%1392=1454%:%
-%:%1393=1455%:%
-%:%1393=1456%:%
-%:%1394=1457%:%
-%:%1394=1458%:%
-%:%1395=1459%:%
-%:%1395=1460%:%
-%:%1396=1461%:%
-%:%1397=1462%:%
-%:%1398=1463%:%
-%:%1398=1464%:%
-%:%1399=1465%:%
-%:%1399=1466%:%
-%:%1400=1467%:%
-%:%1400=1468%:%
-%:%1400=1469%:%
-%:%1400=1470%:%
-%:%1401=1471%:%
-%:%1401=1472%:%
-%:%1402=1473%:%
-%:%1403=1474%:%
-%:%1403=1476%:%
-%:%1403=1477%:%
-%:%1403=1478%:%
-%:%1403=1479%:%
-%:%1403=1480%:%
-%:%1404=1481%:%
-%:%1404=1482%:%
-%:%1405=1483%:%
-%:%1405=1484%:%
-%:%1406=1485%:%
-%:%1406=1486%:%
-%:%1407=1487%:%
-%:%1408=1488%:%
-%:%1409=1489%:%
-%:%1410=1490%:%
-%:%1410=1491%:%
-%:%1411=1492%:%
-%:%1411=1493%:%
-%:%1412=1494%:%
-%:%1413=1495%:%
-%:%1413=1496%:%
-%:%1414=1497%:%
-%:%1415=1498%:%
-%:%1416=1499%:%
-%:%1417=1500%:%
-%:%1418=1501%:%
-%:%1419=1502%:%
-%:%1420=1503%:%
-%:%1421=1504%:%
-%:%1422=1505%:%
-%:%1423=1506%:%
-%:%1424=1507%:%
-%:%1425=1508%:%
-%:%1426=1509%:%
-%:%1427=1510%:%
-%:%1428=1511%:%
-%:%1429=1512%:%
-%:%1430=1513%:%
-%:%1431=1514%:%
-%:%1432=1515%:%
-%:%1433=1516%:%
-%:%1434=1517%:%
-%:%1435=1518%:%
-%:%1436=1519%:%
-%:%1437=1520%:%
-%:%1438=1521%:%
-%:%1439=1522%:%
-%:%1440=1523%:%
-%:%1441=1524%:%
-%:%1442=1525%:%
-%:%1443=1526%:%
-%:%1444=1527%:%
-%:%1445=1528%:%
-%:%1446=1529%:%
-%:%1447=1530%:%
-%:%1448=1531%:%
-%:%1449=1532%:%
-%:%1450=1533%:%
-%:%1451=1534%:%
-%:%1452=1535%:%
-%:%1453=1536%:%
-%:%1454=1537%:%
-%:%1455=1538%:%
-%:%1456=1539%:%
-%:%1457=1540%:%
-%:%1458=1541%:%
-%:%1459=1542%:%
-%:%1460=1543%:%
-%:%1461=1544%:%
-%:%1462=1545%:%
-%:%1463=1546%:%
-%:%1464=1547%:%
-%:%1473=1551%:%
-%:%1485=1558%:%
-%:%1486=1559%:%
-%:%1487=1560%:%
-%:%1488=1561%:%
-%:%1489=1562%:%
-%:%1490=1563%:%
-%:%1491=1564%:%
-%:%1500=1569%:%
-%:%1512=1573%:%
-%:%1513=1574%:%
-%:%1514=1575%:%
-%:%1515=1576%:%
-%:%1516=1577%:%
-%:%1517=1578%:%
-%:%1518=1579%:%
-%:%1519=1580%:%
-%:%1520=1581%:%
-%:%1521=1582%:%
-%:%1522=1583%:%
-%:%1523=1584%:%
-%:%1524=1585%:%
-%:%1525=1586%:%
-%:%1526=1587%:%
-%:%1527=1588%:%
-%:%1528=1589%:%
-%:%1529=1590%:%
-%:%1530=1591%:%
-%:%1531=1592%:%
-%:%1532=1593%:%
-%:%1533=1594%:%
-%:%1534=1595%:%
-%:%1535=1596%:%
-%:%1536=1597%:%
-%:%1537=1598%:%
-%:%1538=1599%:%
-%:%1539=1600%:%
-%:%1540=1601%:%
-%:%1541=1602%:%
-%:%1542=1603%:%
-%:%1543=1604%:%
-%:%1544=1605%:%
-%:%1545=1606%:%
-%:%1546=1607%:%
-%:%1547=1608%:%
-%:%1548=1609%:%
-%:%1549=1610%:%
-%:%1550=1611%:%
-%:%1551=1612%:%
-%:%1552=1613%:%
-%:%1553=1614%:%
-%:%1554=1615%:%
-%:%1555=1616%:%
-%:%1556=1617%:%
-%:%1557=1618%:%
-%:%1558=1619%:%
-%:%1559=1620%:%
-%:%1560=1621%:%
-%:%1561=1622%:%
-%:%1562=1623%:%
-%:%1563=1624%:%
-%:%1564=1625%:%
-%:%1565=1626%:%
-%:%1566=1627%:%
-%:%1567=1628%:%
-%:%1568=1629%:%
-%:%1569=1630%:%
-%:%1570=1631%:%
-%:%1571=1632%:%
-%:%1572=1633%:%
-%:%1573=1634%:%
-%:%1574=1635%:%
-%:%1575=1636%:%
-%:%1576=1637%:%
-%:%1577=1638%:%
-%:%1578=1639%:%
-%:%1579=1640%:%
-%:%1580=1641%:%
-%:%1581=1642%:%
-%:%1582=1643%:%
-%:%1583=1644%:%
-%:%1584=1645%:%
-%:%1585=1646%:%
-%:%1586=1647%:%
-%:%1587=1648%:%
-%:%1588=1649%:%
-%:%1589=1650%:%
-%:%1590=1651%:%
-%:%1591=1652%:%
-%:%1592=1653%:%
-%:%1593=1654%:%
-%:%1594=1655%:%
-%:%1595=1656%:%
-%:%1596=1657%:%
-%:%1597=1658%:%
-%:%1598=1659%:%
-%:%1599=1660%:%
-%:%1600=1661%:%
-%:%1601=1662%:%
-%:%1601=1778%:%
-%:%1602=1779%:%
-%:%1603=1780%:%
-%:%1604=1781%:%
-%:%1605=1782%:%
-%:%1605=1943%:%
-%:%1606=1944%:%
-%:%1607=1945%:%
-%:%1608=1946%:%
-%:%1609=1947%:%
-%:%1610=1948%:%
-%:%1611=1949%:%
-%:%1612=1950%:%
-%:%1613=1951%:%
-%:%1614=1952%:%
-%:%1615=1953%:%
-%:%1616=1954%:%
-%:%1617=1955%:%
-%:%1618=1956%:%
-%:%1619=1957%:%
-%:%1620=1958%:%
-%:%1621=1959%:%
-%:%1622=1960%:%
-%:%1623=1961%:%
-%:%1624=1962%:%
-%:%1624=1963%:%
-%:%1625=1964%:%
-%:%1625=1965%:%
-%:%1625=1966%:%
-%:%1626=1967%:%
-%:%1626=1968%:%
-%:%1627=1969%:%
-%:%1628=1970%:%
-%:%1629=1971%:%
-%:%1630=1972%:%
-%:%1631=1973%:%
-%:%1632=1974%:%
-%:%1633=1975%:%
-%:%1634=1976%:%
-%:%1635=1977%:%
-%:%1636=1978%:%
-%:%1637=1979%:%
-%:%1638=1980%:%
-%:%1639=1981%:%
-%:%1640=1982%:%
-%:%1641=1983%:%
-%:%1642=1984%:%
-%:%1643=1985%:%
-%:%1643=1986%:%
-%:%1644=1987%:%
-%:%1645=1988%:%
-%:%1646=1989%:%
-%:%1646=1990%:%
-%:%1646=1991%:%
-%:%1647=1992%:%
-%:%1648=1993%:%
-%:%1649=1994%:%
-%:%1649=1995%:%
-%:%1650=1996%:%
-%:%1650=1997%:%
-%:%1651=1998%:%
-%:%1652=1999%:%
-%:%1653=2000%:%
-%:%1654=2001%:%
-%:%1655=2002%:%
-%:%1656=2003%:%
-%:%1657=2004%:%
-%:%1658=2005%:%
-%:%1659=2006%:%
-%:%1660=2007%:%
-%:%1661=2008%:%
-%:%1662=2009%:%
-%:%1663=2010%:%
-%:%1664=2011%:%
-%:%1665=2012%:%
-%:%1666=2013%:%
-%:%1667=2014%:%
-%:%1676=2019%:%
-%:%1688=2023%:%
-%:%1689=2024%:%
-%:%1690=2025%:%
-%:%1691=2026%:%
-%:%1692=2027%:%
-%:%1693=2028%:%
-%:%1694=2029%:%
-%:%1695=2030%:%
-%:%1696=2031%:%
-%:%1697=2032%:%
-%:%1698=2033%:%
-%:%1699=2034%:%
-%:%1700=2035%:%
-%:%1701=2036%:%
-%:%1702=2037%:%
-%:%1703=2038%:%
-%:%1704=2039%:%
-%:%1705=2040%:%
-%:%1706=2041%:%
-%:%1707=2042%:%
-%:%1708=2043%:%
-%:%1709=2044%:%
-%:%1710=2045%:%
-%:%1711=2046%:%
-%:%1712=2047%:%
-%:%1713=2048%:%
-%:%1714=2049%:%
-%:%1715=2050%:%
-%:%1716=2051%:%
-%:%1717=2052%:%
-%:%1718=2053%:%
-%:%1719=2054%:%
-%:%1720=2055%:%
-%:%1721=2056%:%
-%:%1722=2057%:%
-%:%1723=2058%:%
-%:%1724=2059%:%
-%:%1725=2060%:%
-%:%1726=2061%:%
-%:%1727=2062%:%
-%:%1728=2063%:%
-%:%1729=2064%:%
-%:%1730=2065%:%
-%:%1731=2066%:%
-%:%1732=2067%:%
-%:%1733=2068%:%
-%:%1734=2069%:%
-%:%1735=2070%:%
-%:%1736=2071%:%
-%:%1737=2072%:%
-%:%1738=2073%:%
-%:%1739=2074%:%
-%:%1740=2075%:%
-%:%1753=2081%:%
-%:%1756=2082%:%
-%:%1757=2083%:%
\ No newline at end of file
--- a/thys2/Journal/Paper.thy~ Fri Jan 07 22:25:26 2022 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,2277 +0,0 @@
-(*<*)
-theory Paper
-imports
- "../Lexer"
- "../Simplifying"
- "../Positions"
-
- "../SizeBound"
- "HOL-Library.LaTeXsugar"
-begin
-
-lemma Suc_0_fold:
- "Suc 0 = 1"
-by simp
-
-
-
-declare [[show_question_marks = false]]
-
-syntax (latex output)
- "_Collect" :: "pttrn => bool => 'a set" ("(1{_ \<^latex>\<open>\\mbox{\\boldmath$\\mid$}\<close> _})")
- "_CollectIn" :: "pttrn => 'a set => bool => 'a set" ("(1{_ \<in> _ |e _})")
-
-syntax
- "_Not_Ex" :: "idts \<Rightarrow> bool \<Rightarrow> bool" ("(3\<nexists>_.a./ _)" [0, 10] 10)
- "_Not_Ex1" :: "pttrn \<Rightarrow> bool \<Rightarrow> bool" ("(3\<nexists>!_.a./ _)" [0, 10] 10)
-
-
-abbreviation
- "der_syn r c \<equiv> der c r"
-
-abbreviation
- "ders_syn r s \<equiv> ders s r"
-
- abbreviation
- "bder_syn r c \<equiv> bder c r"
-
-abbreviation
- "bders_syn r s \<equiv> bders r s"
-
-
-abbreviation
- "nprec v1 v2 \<equiv> \<not>(v1 :\<sqsubset>val v2)"
-
-
-
-
-notation (latex output)
- 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
- Cons ("_\<^latex>\<open>\\mbox{$\\,$}\<close>::\<^latex>\<open>\\mbox{$\\,$}\<close>_" [75,73] 73) and
-
- ZERO ("\<^bold>0" 81) and
- ONE ("\<^bold>1" 81) and
- CH ("_" [1000] 80) and
- ALT ("_ + _" [77,77] 78) and
- SEQ ("_ \<cdot> _" [77,77] 78) and
- STAR ("_\<^sup>\<star>" [79] 78) and
-
- val.Void ("Empty" 78) and
- val.Char ("Char _" [1000] 78) and
- val.Left ("Left _" [79] 78) and
- val.Right ("Right _" [1000] 78) and
- val.Seq ("Seq _ _" [79,79] 78) and
- val.Stars ("Stars _" [79] 78) and
-
- L ("L'(_')" [10] 78) and
- LV ("LV _ _" [80,73] 78) and
- der_syn ("_\\_" [79, 1000] 76) and
- ders_syn ("_\\_" [79, 1000] 76) and
- flat ("|_|" [75] 74) and
- flats ("|_|" [72] 74) and
- Sequ ("_ @ _" [78,77] 63) and
- injval ("inj _ _ _" [79,77,79] 76) and
- mkeps ("mkeps _" [79] 76) and
- length ("len _" [73] 73) and
- intlen ("len _" [73] 73) and
- set ("_" [73] 73) and
-
- Prf ("_ : _" [75,75] 75) and
- Posix ("'(_, _') \<rightarrow> _" [63,75,75] 75) and
-
- lexer ("lexer _ _" [78,78] 77) and
- F_RIGHT ("F\<^bsub>Right\<^esub> _") and
- F_LEFT ("F\<^bsub>Left\<^esub> _") and
- F_ALT ("F\<^bsub>Alt\<^esub> _ _") and
- F_SEQ1 ("F\<^bsub>Seq1\<^esub> _ _") and
- F_SEQ2 ("F\<^bsub>Seq2\<^esub> _ _") and
- F_SEQ ("F\<^bsub>Seq\<^esub> _ _") and
- simp_SEQ ("simp\<^bsub>Seq\<^esub> _ _" [1000, 1000] 1) and
- simp_ALT ("simp\<^bsub>Alt\<^esub> _ _" [1000, 1000] 1) and
- slexer ("lexer\<^sup>+" 1000) and
-
- at ("_\<^latex>\<open>\\mbox{$\\downharpoonleft$}\<close>\<^bsub>_\<^esub>") and
- lex_list ("_ \<prec>\<^bsub>lex\<^esub> _") and
- PosOrd ("_ \<prec>\<^bsub>_\<^esub> _" [77,77,77] 77) and
- PosOrd_ex ("_ \<prec> _" [77,77] 77) and
- PosOrd_ex_eq ("_ \<^latex>\<open>\\mbox{$\\preccurlyeq$}\<close> _" [77,77] 77) and
- pflat_len ("\<parallel>_\<parallel>\<^bsub>_\<^esub>") and
- nprec ("_ \<^latex>\<open>\\mbox{$\\not\\prec$}\<close> _" [77,77] 77) and
-
- bder_syn ("_\<^latex>\<open>\\mbox{$\\bbslash$}\<close>_" [79, 1000] 76) and
- bders_syn ("_\<^latex>\<open>\\mbox{$\\bbslash$}\<close>_" [79, 1000] 76) and
- intern ("_\<^latex>\<open>\\mbox{$^\\uparrow$}\<close>" [900] 80) and
- erase ("_\<^latex>\<open>\\mbox{$^\\downarrow$}\<close>" [1000] 74) and
- bnullable ("nullable\<^latex>\<open>\\mbox{$_b$}\<close> _" [1000] 80) and
- bmkeps ("mkeps\<^latex>\<open>\\mbox{$_b$}\<close> _" [1000] 80) and
- blexer ("lexer\<^latex>\<open>\\mbox{$_b$}\<close> _ _" [77, 77] 80) and
- code ("code _" [79] 74) and
-
- DUMMY ("\<^latex>\<open>\\underline{\\hspace{2mm}}\<close>")
-
-
-definition
- "match r s \<equiv> nullable (ders s r)"
-
-
-lemma LV_STAR_ONE_empty:
- shows "LV (STAR ONE) [] = {Stars []}"
-by(auto simp add: LV_def elim: Prf.cases intro: Prf.intros)
-
-
-
-(*
-comments not implemented
-
-p9. The condition "not exists s3 s4..." appears often enough (in particular in
-the proof of Lemma 3) to warrant a definition.
-
-*)
-
-
-(*>*)
-
-section\<open>Core of the proof\<close>
-text \<open>
-This paper builds on previous work by Ausaf and Urban using
-regular expression'd bit-coded derivatives to do lexing that
-is both fast and satisfies the POSIX specification.
-In their work, a bit-coded algorithm introduced by Sulzmann and Lu
-was formally verified in Isabelle, by a very clever use of
-flex function and retrieve to carefully mimic the way a value is
-built up by the injection funciton.
-
-In the previous work, Ausaf and Urban established the below equality:
-\begin{lemma}
-@{thm [mode=IfThen] MAIN_decode}
-\end{lemma}
-
-This lemma establishes a link with the lexer without bit-codes.
-
-With it we get the correctness of bit-coded algorithm.
-\begin{lemma}
-@{thm [mode=IfThen] blexer_correctness}
-\end{lemma}
-
-However what is not certain is whether we can add simplification
-to the bit-coded algorithm, without breaking the correct lexing output.
-
-
-The reason that we do need to add a simplification phase
-after each derivative step of $\textit{blexer}$ is
-because it produces intermediate
-regular expressions that can grow exponentially.
-For example, the regular expression $(a+aa)^*$ after taking
-derivative against just 10 $a$s will have size 8192.
-
-%TODO: add figure for this?
-
-
-Therefore, we insert a simplification phase
-after each derivation step, as defined below:
-\begin{lemma}
-@{thm blexer_simp_def}
-\end{lemma}
-
-The simplification function is given as follows:
-
-\begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) bsimp.simps(1)[of "bs" "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bsimp.simps(1)[of "bs" "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) bsimp.simps(2)} & $\dn$ & @{thm (rhs) bsimp.simps(2)}\\
- @{thm (lhs) bsimp.simps(3)} & $\dn$ & @{thm (rhs) bsimp.simps(3)}\\
-
-\end{tabular}
-\end{center}
-
-And the two helper functions are:
-\begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) bsimp_AALTs.simps(2)[of "bs\<^sub>1" "r" ]} & $\dn$ & @{thm (rhs) bsimp.simps(1)[of "bs\<^sub>1" "r" ]}\\
- @{thm (lhs) bsimp_AALTs.simps(2)} & $\dn$ & @{thm (rhs) bsimp.simps(2)}\\
- @{thm (lhs) bsimp_AALTs.simps(3)} & $\dn$ & @{thm (rhs) bsimp.simps(3)}\\
-
-\end{tabular}
-\end{center}
-
-
-This might sound trivial in the case of producing a YES/NO answer,
-but once we require a lexing output to be produced (which is required
-in applications like compiler front-end, malicious attack domain extraction,
-etc.), it is not straightforward if we still extract what is needed according
-to the POSIX standard.
-
-
-
-
-
-By simplification, we mean specifically the following rules:
-
-\begin{center}
- \begin{tabular}{lcl}
- @{thm[mode=Axiom] rrewrite.intros(1)[of "bs" "r\<^sub>2"]}\\
- @{thm[mode=Axiom] rrewrite.intros(2)[of "bs" "r\<^sub>1"]}\\
- @{thm[mode=Axiom] rrewrite.intros(3)[of "bs" "bs\<^sub>1" "r\<^sub>1"]}\\
- @{thm[mode=Rule] rrewrite.intros(4)[of "r\<^sub>1" "r\<^sub>2" "bs" "r\<^sub>3"]}\\
- @{thm[mode=Rule] rrewrite.intros(5)[of "r\<^sub>3" "r\<^sub>4" "bs" "r\<^sub>1"]}\\
- @{thm[mode=Rule] rrewrite.intros(6)[of "r" "r'" "bs" "rs\<^sub>1" "rs\<^sub>2"]}\\
- @{thm[mode=Axiom] rrewrite.intros(7)[of "bs" "rs\<^sub>a" "rs\<^sub>b"]}\\
- @{thm[mode=Axiom] rrewrite.intros(8)[of "bs" "rs\<^sub>a" "bs\<^sub>1" "rs\<^sub>1" "rs\<^sub>b"]}\\
- @{thm[mode=Axiom] rrewrite.intros(9)[of "bs" "bs\<^sub>1" "rs"]}\\
- @{thm[mode=Axiom] rrewrite.intros(10)[of "bs" ]}\\
- @{thm[mode=Axiom] rrewrite.intros(11)[of "bs" "r\<^sub>1"]}\\
- @{thm[mode=Rule] rrewrite.intros(12)[of "a\<^sub>1" "a\<^sub>2" "bs" "rs\<^sub>a" "rs\<^sub>b" "rs\<^sub>c"]}\\
-
- \end{tabular}
-\end{center}
-
-
-And these can be made compact by the following simplification function:
-
-where the function $\mathit{bsimp_AALTs}$
-
-The core idea of the proof is that two regular expressions,
-if "isomorphic" up to a finite number of rewrite steps, will
-remain "isomorphic" when we take the same sequence of
-derivatives on both of them.
-This can be expressed by the following rewrite relation lemma:
-\begin{lemma}
-@{thm [mode=IfThen] central}
-\end{lemma}
-
-This isomorphic relation implies a property that leads to the
-correctness result:
-if two (nullable) regular expressions are "rewritable" in many steps
-from one another,
-then a call to function $\textit{bmkeps}$ gives the same
-bit-sequence :
-\begin{lemma}
-@{thm [mode=IfThen] rewrites_bmkeps}
-\end{lemma}
-
-Given the same bit-sequence, the decode function
-will give out the same value, which is the output
-of both lexers:
-\begin{lemma}
-@{thm blexer_def}
-\end{lemma}
-
-\begin{lemma}
-@{thm blexer_simp_def}
-\end{lemma}
-
-And that yields the correctness result:
-\begin{lemma}
-@{thm blexersimp_correctness}
-\end{lemma}
-
-The nice thing about the aove
-\<close>
-
-
-
-section \<open>Introduction\<close>
-
-
-text \<open>
-
-
-Brzozowski \cite{Brzozowski1964} introduced the notion of the {\em
-derivative} @{term "der c r"} of a regular expression \<open>r\<close> w.r.t.\
-a character~\<open>c\<close>, and showed that it gave a simple solution to the
-problem of matching a string @{term s} with a regular expression @{term
-r}: if the derivative of @{term r} w.r.t.\ (in succession) all the
-characters of the string matches the empty string, then @{term r}
-matches @{term s} (and {\em vice versa}). The derivative has the
-property (which may almost be regarded as its specification) that, for
-every string @{term s} and regular expression @{term r} and character
-@{term c}, one has @{term "cs \<in> L(r)"} if and only if \mbox{@{term "s \<in> L(der c r)"}}.
-The beauty of Brzozowski's derivatives is that
-they are neatly expressible in any functional language, and easily
-definable and reasoned about in theorem provers---the definitions just
-consist of inductive datatypes and simple recursive functions. A
-mechanised correctness proof of Brzozowski's matcher in for example HOL4
-has been mentioned by Owens and Slind~\cite{Owens2008}. Another one in
-Isabelle/HOL is part of the work by Krauss and Nipkow \cite{Krauss2011}.
-And another one in Coq is given by Coquand and Siles \cite{Coquand2012}.
-
-If a regular expression matches a string, then in general there is more
-than one way of how the string is matched. There are two commonly used
-disambiguation strategies to generate a unique answer: one is called
-GREEDY matching \cite{Frisch2004} and the other is POSIX
-matching~\cite{POSIX,Kuklewicz,OkuiSuzuki2010,Sulzmann2014,Vansummeren2006}.
-For example consider the string @{term xy} and the regular expression
-\mbox{@{term "STAR (ALT (ALT x y) xy)"}}. Either the string can be
-matched in two `iterations' by the single letter-regular expressions
-@{term x} and @{term y}, or directly in one iteration by @{term xy}. The
-first case corresponds to GREEDY matching, which first matches with the
-left-most symbol and only matches the next symbol in case of a mismatch
-(this is greedy in the sense of preferring instant gratification to
-delayed repletion). The second case is POSIX matching, which prefers the
-longest match.
-
-In the context of lexing, where an input string needs to be split up
-into a sequence of tokens, POSIX is the more natural disambiguation
-strategy for what programmers consider basic syntactic building blocks
-in their programs. These building blocks are often specified by some
-regular expressions, say \<open>r\<^bsub>key\<^esub>\<close> and \<open>r\<^bsub>id\<^esub>\<close> for recognising keywords and identifiers,
-respectively. There are a few underlying (informal) rules behind
-tokenising a string in a POSIX \cite{POSIX} fashion:
-
-\begin{itemize}
-\item[$\bullet$] \emph{The Longest Match Rule} (or \emph{``{M}aximal {M}unch {R}ule''}):
-The longest initial substring matched by any regular expression is taken as
-next token.\smallskip
-
-\item[$\bullet$] \emph{Priority Rule:}
-For a particular longest initial substring, the first (leftmost) regular expression
-that can match determines the token.\smallskip
-
-\item[$\bullet$] \emph{Star Rule:} A subexpression repeated by ${}^\star$ shall
-not match an empty string unless this is the only match for the repetition.\smallskip
-
-\item[$\bullet$] \emph{Empty String Rule:} An empty string shall be considered to
-be longer than no match at all.
-\end{itemize}
-
-\noindent Consider for example a regular expression \<open>r\<^bsub>key\<^esub>\<close> for recognising keywords such as \<open>if\<close>,
-\<open>then\<close> and so on; and \<open>r\<^bsub>id\<^esub>\<close>
-recognising identifiers (say, a single character followed by
-characters or numbers). Then we can form the regular expression
-\<open>(r\<^bsub>key\<^esub> + r\<^bsub>id\<^esub>)\<^sup>\<star>\<close>
-and use POSIX matching to tokenise strings, say \<open>iffoo\<close> and
-\<open>if\<close>. For \<open>iffoo\<close> we obtain by the Longest Match Rule
-a single identifier token, not a keyword followed by an
-identifier. For \<open>if\<close> we obtain by the Priority Rule a keyword
-token, not an identifier token---even if \<open>r\<^bsub>id\<^esub>\<close>
-matches also. By the Star Rule we know \<open>(r\<^bsub>key\<^esub> +
-r\<^bsub>id\<^esub>)\<^sup>\<star>\<close> matches \<open>iffoo\<close>,
-respectively \<open>if\<close>, in exactly one `iteration' of the star. The
-Empty String Rule is for cases where, for example, the regular expression
-\<open>(a\<^sup>\<star>)\<^sup>\<star>\<close> matches against the
-string \<open>bc\<close>. Then the longest initial matched substring is the
-empty string, which is matched by both the whole regular expression
-and the parenthesised subexpression.
-
-
-One limitation of Brzozowski's matcher is that it only generates a
-YES/NO answer for whether a string is being matched by a regular
-expression. Sulzmann and Lu~\cite{Sulzmann2014} extended this matcher
-to allow generation not just of a YES/NO answer but of an actual
-matching, called a [lexical] {\em value}. Assuming a regular
-expression matches a string, values encode the information of
-\emph{how} the string is matched by the regular expression---that is,
-which part of the string is matched by which part of the regular
-expression. For this consider again the string \<open>xy\<close> and
-the regular expression \mbox{\<open>(x + (y + xy))\<^sup>\<star>\<close>}
-(this time fully parenthesised). We can view this regular expression
-as tree and if the string \<open>xy\<close> is matched by two Star
-`iterations', then the \<open>x\<close> is matched by the left-most
-alternative in this tree and the \<open>y\<close> by the right-left alternative. This
-suggests to record this matching as
-
-\begin{center}
-@{term "Stars [Left(Char x), Right(Left(Char y))]"}
-\end{center}
-
-\noindent where @{const Stars}, \<open>Left\<close>, \<open>Right\<close> and \<open>Char\<close> are constructors for values. \<open>Stars\<close> records how many
-iterations were used; \<open>Left\<close>, respectively \<open>Right\<close>, which
-alternative is used. This `tree view' leads naturally to the idea that
-regular expressions act as types and values as inhabiting those types
-(see, for example, \cite{HosoyaVouillonPierce2005}). The value for
-matching \<open>xy\<close> in a single `iteration', i.e.~the POSIX value,
-would look as follows
-
-\begin{center}
-@{term "Stars [Seq (Char x) (Char y)]"}
-\end{center}
-
-\noindent where @{const Stars} has only a single-element list for the
-single iteration and @{const Seq} indicates that @{term xy} is matched
-by a sequence regular expression.
-
-%, which we will in what follows
-%write more formally as @{term "SEQ x y"}.
-
-
-Sulzmann and Lu give a simple algorithm to calculate a value that
-appears to be the value associated with POSIX matching. The challenge
-then is to specify that value, in an algorithm-independent fashion,
-and to show that Sulzmann and Lu's derivative-based algorithm does
-indeed calculate a value that is correct according to the
-specification. The answer given by Sulzmann and Lu
-\cite{Sulzmann2014} is to define a relation (called an ``order
-relation'') on the set of values of @{term r}, and to show that (once
-a string to be matched is chosen) there is a maximum element and that
-it is computed by their derivative-based algorithm. This proof idea is
-inspired by work of Frisch and Cardelli \cite{Frisch2004} on a GREEDY
-regular expression matching algorithm. However, we were not able to
-establish transitivity and totality for the ``order relation'' by
-Sulzmann and Lu. There are some inherent problems with their approach
-(of which some of the proofs are not published in
-\cite{Sulzmann2014}); perhaps more importantly, we give in this paper
-a simple inductive (and algorithm-independent) definition of what we
-call being a {\em POSIX value} for a regular expression @{term r} and
-a string @{term s}; we show that the algorithm by Sulzmann and Lu
-computes such a value and that such a value is unique. Our proofs are
-both done by hand and checked in Isabelle/HOL. The experience of
-doing our proofs has been that this mechanical checking was absolutely
-essential: this subject area has hidden snares. This was also noted by
-Kuklewicz \cite{Kuklewicz} who found that nearly all POSIX matching
-implementations are ``buggy'' \cite[Page 203]{Sulzmann2014} and by
-Grathwohl et al \cite[Page 36]{CrashCourse2014} who wrote:
-
-\begin{quote}
-\it{}``The POSIX strategy is more complicated than the greedy because of
-the dependence on information about the length of matched strings in the
-various subexpressions.''
-\end{quote}
-
-
-
-\noindent {\bf Contributions:} We have implemented in Isabelle/HOL the
-derivative-based regular expression matching algorithm of
-Sulzmann and Lu \cite{Sulzmann2014}. We have proved the correctness of this
-algorithm according to our specification of what a POSIX value is (inspired
-by work of Vansummeren \cite{Vansummeren2006}). Sulzmann
-and Lu sketch in \cite{Sulzmann2014} an informal correctness proof: but to
-us it contains unfillable gaps.\footnote{An extended version of
-\cite{Sulzmann2014} is available at the website of its first author; this
-extended version already includes remarks in the appendix that their
-informal proof contains gaps, and possible fixes are not fully worked out.}
-Our specification of a POSIX value consists of a simple inductive definition
-that given a string and a regular expression uniquely determines this value.
-We also show that our definition is equivalent to an ordering
-of values based on positions by Okui and Suzuki \cite{OkuiSuzuki2010}.
-
-%Derivatives as calculated by Brzozowski's method are usually more complex
-%regular expressions than the initial one; various optimisations are
-%possible. We prove the correctness when simplifications of @{term "ALT ZERO r"},
-%@{term "ALT r ZERO"}, @{term "SEQ ONE r"} and @{term "SEQ r ONE"} to
-%@{term r} are applied.
-
-We extend our results to ??? Bitcoded version??
-
-\<close>
-
-
-
-
-section \<open>Preliminaries\<close>
-
-text \<open>\noindent Strings in Isabelle/HOL are lists of characters with
-the empty string being represented by the empty list, written @{term
-"[]"}, and list-cons being written as @{term "DUMMY # DUMMY"}. Often
-we use the usual bracket notation for lists also for strings; for
-example a string consisting of just a single character @{term c} is
-written @{term "[c]"}. We use the usual definitions for
-\emph{prefixes} and \emph{strict prefixes} of strings. By using the
-type @{type char} for characters we have a supply of finitely many
-characters roughly corresponding to the ASCII character set. Regular
-expressions are defined as usual as the elements of the following
-inductive datatype:
-
- \begin{center}
- \<open>r :=\<close>
- @{const "ZERO"} $\mid$
- @{const "ONE"} $\mid$
- @{term "CH c"} $\mid$
- @{term "ALT r\<^sub>1 r\<^sub>2"} $\mid$
- @{term "SEQ r\<^sub>1 r\<^sub>2"} $\mid$
- @{term "STAR r"}
- \end{center}
-
- \noindent where @{const ZERO} stands for the regular expression that does
- not match any string, @{const ONE} for the regular expression that matches
- only the empty string and @{term c} for matching a character literal. The
- language of a regular expression is also defined as usual by the
- recursive function @{term L} with the six clauses:
-
- \begin{center}
- \begin{tabular}{l@ {\hspace{4mm}}rcl}
- \textit{(1)} & @{thm (lhs) L.simps(1)} & $\dn$ & @{thm (rhs) L.simps(1)}\\
- \textit{(2)} & @{thm (lhs) L.simps(2)} & $\dn$ & @{thm (rhs) L.simps(2)}\\
- \textit{(3)} & @{thm (lhs) L.simps(3)} & $\dn$ & @{thm (rhs) L.simps(3)}\\
- \textit{(4)} & @{thm (lhs) L.simps(4)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ &
- @{thm (rhs) L.simps(4)[of "r\<^sub>1" "r\<^sub>2"]}\\
- \textit{(5)} & @{thm (lhs) L.simps(5)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ &
- @{thm (rhs) L.simps(5)[of "r\<^sub>1" "r\<^sub>2"]}\\
- \textit{(6)} & @{thm (lhs) L.simps(6)} & $\dn$ & @{thm (rhs) L.simps(6)}\\
- \end{tabular}
- \end{center}
-
- \noindent In clause \textit{(4)} we use the operation @{term "DUMMY ;;
- DUMMY"} for the concatenation of two languages (it is also list-append for
- strings). We use the star-notation for regular expressions and for
- languages (in the last clause above). The star for languages is defined
- inductively by two clauses: \<open>(i)\<close> the empty string being in
- the star of a language and \<open>(ii)\<close> if @{term "s\<^sub>1"} is in a
- language and @{term "s\<^sub>2"} in the star of this language, then also @{term
- "s\<^sub>1 @ s\<^sub>2"} is in the star of this language. It will also be convenient
- to use the following notion of a \emph{semantic derivative} (or \emph{left
- quotient}) of a language defined as
- %
- \begin{center}
- @{thm Der_def}\;.
- \end{center}
-
- \noindent
- For semantic derivatives we have the following equations (for example
- mechanically proved in \cite{Krauss2011}):
- %
- \begin{equation}\label{SemDer}
- \begin{array}{lcl}
- @{thm (lhs) Der_null} & \dn & @{thm (rhs) Der_null}\\
- @{thm (lhs) Der_empty} & \dn & @{thm (rhs) Der_empty}\\
- @{thm (lhs) Der_char} & \dn & @{thm (rhs) Der_char}\\
- @{thm (lhs) Der_union} & \dn & @{thm (rhs) Der_union}\\
- @{thm (lhs) Der_Sequ} & \dn & @{thm (rhs) Der_Sequ}\\
- @{thm (lhs) Der_star} & \dn & @{thm (rhs) Der_star}
- \end{array}
- \end{equation}
-
-
- \noindent \emph{\Brz's derivatives} of regular expressions
- \cite{Brzozowski1964} can be easily defined by two recursive functions:
- the first is from regular expressions to booleans (implementing a test
- when a regular expression can match the empty string), and the second
- takes a regular expression and a character to a (derivative) regular
- expression:
-
- \begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) nullable.simps(1)} & $\dn$ & @{thm (rhs) nullable.simps(1)}\\
- @{thm (lhs) nullable.simps(2)} & $\dn$ & @{thm (rhs) nullable.simps(2)}\\
- @{thm (lhs) nullable.simps(3)} & $\dn$ & @{thm (rhs) nullable.simps(3)}\\
- @{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"]}\\
- @{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"]}\\
- @{thm (lhs) nullable.simps(6)} & $\dn$ & @{thm (rhs) nullable.simps(6)}\medskip\\
-
-% \end{tabular}
-% \end{center}
-
-% \begin{center}
-% \begin{tabular}{lcl}
-
- @{thm (lhs) der.simps(1)} & $\dn$ & @{thm (rhs) der.simps(1)}\\
- @{thm (lhs) der.simps(2)} & $\dn$ & @{thm (rhs) der.simps(2)}\\
- @{thm (lhs) der.simps(3)} & $\dn$ & @{thm (rhs) der.simps(3)}\\
- @{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"]}\\
- @{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"]}\\
- @{thm (lhs) der.simps(6)} & $\dn$ & @{thm (rhs) der.simps(6)}
- \end{tabular}
- \end{center}
-
- \noindent
- We may extend this definition to give derivatives w.r.t.~strings:
-
- \begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) ders.simps(1)} & $\dn$ & @{thm (rhs) ders.simps(1)}\\
- @{thm (lhs) ders.simps(2)} & $\dn$ & @{thm (rhs) ders.simps(2)}\\
- \end{tabular}
- \end{center}
-
- \noindent Given the equations in \eqref{SemDer}, it is a relatively easy
- exercise in mechanical reasoning to establish that
-
- \begin{proposition}\label{derprop}\mbox{}\\
- \begin{tabular}{ll}
- \textit{(1)} & @{thm (lhs) nullable_correctness} if and only if
- @{thm (rhs) nullable_correctness}, and \\
- \textit{(2)} & @{thm[mode=IfThen] der_correctness}.
- \end{tabular}
- \end{proposition}
-
- \noindent With this in place it is also very routine to prove that the
- regular expression matcher defined as
- %
- \begin{center}
- @{thm match_def}
- \end{center}
-
- \noindent gives a positive answer if and only if @{term "s \<in> L r"}.
- Consequently, this regular expression matching algorithm satisfies the
- usual specification for regular expression matching. While the matcher
- above calculates a provably correct YES/NO answer for whether a regular
- expression matches a string or not, the novel idea of Sulzmann and Lu
- \cite{Sulzmann2014} is to append another phase to this algorithm in order
- to calculate a [lexical] value. We will explain the details next.
-
-\<close>
-
-section \<open>POSIX Regular Expression Matching\label{posixsec}\<close>
-
-text \<open>
-
- There have been many previous works that use values for encoding
- \emph{how} a regular expression matches a string.
- The clever idea by Sulzmann and Lu \cite{Sulzmann2014} is to
- define a function on values that mirrors (but inverts) the
- construction of the derivative on regular expressions. \emph{Values}
- are defined as the inductive datatype
-
- \begin{center}
- \<open>v :=\<close>
- @{const "Void"} $\mid$
- @{term "val.Char c"} $\mid$
- @{term "Left v"} $\mid$
- @{term "Right v"} $\mid$
- @{term "Seq v\<^sub>1 v\<^sub>2"} $\mid$
- @{term "Stars vs"}
- \end{center}
-
- \noindent where we use @{term vs} to stand for a list of
- values. (This is similar to the approach taken by Frisch and
- Cardelli for GREEDY matching \cite{Frisch2004}, and Sulzmann and Lu
- for POSIX matching \cite{Sulzmann2014}). The string underlying a
- value can be calculated by the @{const flat} function, written
- @{term "flat DUMMY"} and defined as:
-
- \begin{center}
- \begin{tabular}[t]{lcl}
- @{thm (lhs) flat.simps(1)} & $\dn$ & @{thm (rhs) flat.simps(1)}\\
- @{thm (lhs) flat.simps(2)} & $\dn$ & @{thm (rhs) flat.simps(2)}\\
- @{thm (lhs) flat.simps(3)} & $\dn$ & @{thm (rhs) flat.simps(3)}\\
- @{thm (lhs) flat.simps(4)} & $\dn$ & @{thm (rhs) flat.simps(4)}
- \end{tabular}\hspace{14mm}
- \begin{tabular}[t]{lcl}
- @{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"]}\\
- @{thm (lhs) flat.simps(6)} & $\dn$ & @{thm (rhs) flat.simps(6)}\\
- @{thm (lhs) flat.simps(7)} & $\dn$ & @{thm (rhs) flat.simps(7)}\\
- \end{tabular}
- \end{center}
-
- \noindent We will sometimes refer to the underlying string of a
- value as \emph{flattened value}. We will also overload our notation and
- use @{term "flats vs"} for flattening a list of values and concatenating
- the resulting strings.
-
- Sulzmann and Lu define
- inductively an \emph{inhabitation relation} that associates values to
- regular expressions. We define this relation as
- follows:\footnote{Note that the rule for @{term Stars} differs from
- our earlier paper \cite{AusafDyckhoffUrban2016}. There we used the
- original definition by Sulzmann and Lu which does not require that
- the values @{term "v \<in> set vs"} flatten to a non-empty
- string. The reason for introducing the more restricted version of
- lexical values is convenience later on when reasoning about an
- ordering relation for values.}
-
- \begin{center}
- \begin{tabular}{c@ {\hspace{12mm}}c}\label{prfintros}
- \\[-8mm]
- @{thm[mode=Axiom] Prf.intros(4)} &
- @{thm[mode=Axiom] Prf.intros(5)[of "c"]}\\[4mm]
- @{thm[mode=Rule] Prf.intros(2)[of "v\<^sub>1" "r\<^sub>1" "r\<^sub>2"]} &
- @{thm[mode=Rule] Prf.intros(3)[of "v\<^sub>2" "r\<^sub>1" "r\<^sub>2"]}\\[4mm]
- @{thm[mode=Rule] Prf.intros(1)[of "v\<^sub>1" "r\<^sub>1" "v\<^sub>2" "r\<^sub>2"]} &
- @{thm[mode=Rule] Prf.intros(6)[of "vs"]}
- \end{tabular}
- \end{center}
-
- \noindent where in the clause for @{const "Stars"} we use the
- notation @{term "v \<in> set vs"} for indicating that \<open>v\<close> is a
- member in the list \<open>vs\<close>. We require in this rule that every
- value in @{term vs} flattens to a non-empty string. The idea is that
- @{term "Stars"}-values satisfy the informal Star Rule (see Introduction)
- where the $^\star$ does not match the empty string unless this is
- the only match for the repetition. Note also that no values are
- associated with the regular expression @{term ZERO}, and that the
- only value associated with the regular expression @{term ONE} is
- @{term Void}. It is routine to establish how values ``inhabiting''
- a regular expression correspond to the language of a regular
- expression, namely
-
- \begin{proposition}\label{inhabs}
- @{thm L_flat_Prf}
- \end{proposition}
-
- \noindent
- Given a regular expression \<open>r\<close> and a string \<open>s\<close>, we define the
- set of all \emph{Lexical Values} inhabited by \<open>r\<close> with the underlying string
- being \<open>s\<close>:\footnote{Okui and Suzuki refer to our lexical values
- as \emph{canonical values} in \cite{OkuiSuzuki2010}. The notion of \emph{non-problematic
- values} by Cardelli and Frisch \cite{Frisch2004} is related, but not identical
- to our lexical values.}
-
- \begin{center}
- @{thm LV_def}
- \end{center}
-
- \noindent The main property of @{term "LV r s"} is that it is alway finite.
-
- \begin{proposition}
- @{thm LV_finite}
- \end{proposition}
-
- \noindent This finiteness property does not hold in general if we
- remove the side-condition about @{term "flat v \<noteq> []"} in the
- @{term Stars}-rule above. For example using Sulzmann and Lu's
- less restrictive definition, @{term "LV (STAR ONE) []"} would contain
- infinitely many values, but according to our more restricted
- definition only a single value, namely @{thm LV_STAR_ONE_empty}.
-
- If a regular expression \<open>r\<close> matches a string \<open>s\<close>, then
- generally the set @{term "LV r s"} is not just a singleton set. In
- case of POSIX matching the problem is to calculate the unique lexical value
- that satisfies the (informal) POSIX rules from the Introduction.
- Graphically the POSIX value calculation algorithm by Sulzmann and Lu
- can be illustrated by the picture in Figure~\ref{Sulz} where the
- path from the left to the right involving @{term
- derivatives}/@{const nullable} is the first phase of the algorithm
- (calculating successive \Brz's derivatives) and @{const
- mkeps}/\<open>inj\<close>, the path from right to left, the second
- phase. This picture shows the steps required when a regular
- expression, say \<open>r\<^sub>1\<close>, matches the string @{term
- "[a,b,c]"}. We first build the three derivatives (according to
- @{term a}, @{term b} and @{term c}). We then use @{const nullable}
- to find out whether the resulting derivative regular expression
- @{term "r\<^sub>4"} can match the empty string. If yes, we call the
- function @{const mkeps} that produces a value @{term "v\<^sub>4"}
- for how @{term "r\<^sub>4"} can match the empty string (taking into
- account the POSIX constraints in case there are several ways). This
- function is defined by the clauses:
-
-\begin{figure}[t]
-\begin{center}
-\begin{tikzpicture}[scale=2,node distance=1.3cm,
- every node/.style={minimum size=6mm}]
-\node (r1) {@{term "r\<^sub>1"}};
-\node (r2) [right=of r1]{@{term "r\<^sub>2"}};
-\draw[->,line width=1mm](r1)--(r2) node[above,midway] {@{term "der a DUMMY"}};
-\node (r3) [right=of r2]{@{term "r\<^sub>3"}};
-\draw[->,line width=1mm](r2)--(r3) node[above,midway] {@{term "der b DUMMY"}};
-\node (r4) [right=of r3]{@{term "r\<^sub>4"}};
-\draw[->,line width=1mm](r3)--(r4) node[above,midway] {@{term "der c DUMMY"}};
-\draw (r4) node[anchor=west] {\;\raisebox{3mm}{@{term nullable}}};
-\node (v4) [below=of r4]{@{term "v\<^sub>4"}};
-\draw[->,line width=1mm](r4) -- (v4);
-\node (v3) [left=of v4] {@{term "v\<^sub>3"}};
-\draw[->,line width=1mm](v4)--(v3) node[below,midway] {\<open>inj r\<^sub>3 c\<close>};
-\node (v2) [left=of v3]{@{term "v\<^sub>2"}};
-\draw[->,line width=1mm](v3)--(v2) node[below,midway] {\<open>inj r\<^sub>2 b\<close>};
-\node (v1) [left=of v2] {@{term "v\<^sub>1"}};
-\draw[->,line width=1mm](v2)--(v1) node[below,midway] {\<open>inj r\<^sub>1 a\<close>};
-\draw (r4) node[anchor=north west] {\;\raisebox{-8mm}{@{term "mkeps"}}};
-\end{tikzpicture}
-\end{center}
-\mbox{}\\[-13mm]
-
-\caption{The two phases of the algorithm by Sulzmann \& Lu \cite{Sulzmann2014},
-matching the string @{term "[a,b,c]"}. The first phase (the arrows from
-left to right) is \Brz's matcher building successive derivatives. If the
-last regular expression is @{term nullable}, then the functions of the
-second phase are called (the top-down and right-to-left arrows): first
-@{term mkeps} calculates a value @{term "v\<^sub>4"} witnessing
-how the empty string has been recognised by @{term "r\<^sub>4"}. After
-that the function @{term inj} ``injects back'' the characters of the string into
-the values.
-\label{Sulz}}
-\end{figure}
-
- \begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) mkeps.simps(1)} & $\dn$ & @{thm (rhs) mkeps.simps(1)}\\
- @{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"]}\\
- @{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"]}\\
- @{thm (lhs) mkeps.simps(4)} & $\dn$ & @{thm (rhs) mkeps.simps(4)}\\
- \end{tabular}
- \end{center}
-
- \noindent Note that this function needs only to be partially defined,
- namely only for regular expressions that are nullable. In case @{const
- nullable} fails, the string @{term "[a,b,c]"} cannot be matched by @{term
- "r\<^sub>1"} and the null value @{term "None"} is returned. Note also how this function
- makes some subtle choices leading to a POSIX value: for example if an
- alternative regular expression, say @{term "ALT r\<^sub>1 r\<^sub>2"}, can
- match the empty string and furthermore @{term "r\<^sub>1"} can match the
- empty string, then we return a \<open>Left\<close>-value. The \<open>Right\<close>-value will only be returned if @{term "r\<^sub>1"} cannot match the empty
- string.
-
- The most interesting idea from Sulzmann and Lu \cite{Sulzmann2014} is
- the construction of a value for how @{term "r\<^sub>1"} can match the
- string @{term "[a,b,c]"} from the value how the last derivative, @{term
- "r\<^sub>4"} in Fig.~\ref{Sulz}, can match the empty string. Sulzmann and
- Lu achieve this by stepwise ``injecting back'' the characters into the
- values thus inverting the operation of building derivatives, but on the level
- of values. The corresponding function, called @{term inj}, takes three
- arguments, a regular expression, a character and a value. For example in
- the first (or right-most) @{term inj}-step in Fig.~\ref{Sulz} the regular
- expression @{term "r\<^sub>3"}, the character @{term c} from the last
- derivative step and @{term "v\<^sub>4"}, which is the value corresponding
- to the derivative regular expression @{term "r\<^sub>4"}. The result is
- the new value @{term "v\<^sub>3"}. The final result of the algorithm is
- the value @{term "v\<^sub>1"}. The @{term inj} function is defined by recursion on regular
- expressions and by analysing the shape of values (corresponding to
- the derivative regular expressions).
- %
- \begin{center}
- \begin{tabular}{l@ {\hspace{5mm}}lcl}
- \textit{(1)} & @{thm (lhs) injval.simps(1)} & $\dn$ & @{thm (rhs) injval.simps(1)}\\
- \textit{(2)} & @{thm (lhs) injval.simps(2)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1"]} & $\dn$ &
- @{thm (rhs) injval.simps(2)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1"]}\\
- \textit{(3)} & @{thm (lhs) injval.simps(3)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]} & $\dn$ &
- @{thm (rhs) injval.simps(3)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]}\\
- \textit{(4)} & @{thm (lhs) injval.simps(4)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]} & $\dn$
- & @{thm (rhs) injval.simps(4)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]}\\
- \textit{(5)} & @{thm (lhs) injval.simps(5)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]} & $\dn$
- & @{thm (rhs) injval.simps(5)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]}\\
- \textit{(6)} & @{thm (lhs) injval.simps(6)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]} & $\dn$
- & @{thm (rhs) injval.simps(6)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]}\\
- \textit{(7)} & @{thm (lhs) injval.simps(7)[of "r" "c" "v" "vs"]} & $\dn$
- & @{thm (rhs) injval.simps(7)[of "r" "c" "v" "vs"]}\\
- \end{tabular}
- \end{center}
-
- \noindent To better understand what is going on in this definition it
- might be instructive to look first at the three sequence cases (clauses
- \textit{(4)} -- \textit{(6)}). In each case we need to construct an ``injected value'' for
- @{term "SEQ r\<^sub>1 r\<^sub>2"}. This must be a value of the form @{term
- "Seq DUMMY DUMMY"}\,. Recall the clause of the \<open>derivative\<close>-function
- for sequence regular expressions:
-
- \begin{center}
- @{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"]}
- \end{center}
-
- \noindent Consider first the \<open>else\<close>-branch where the derivative is @{term
- "SEQ (der c r\<^sub>1) r\<^sub>2"}. The corresponding value must therefore
- be of the form @{term "Seq v\<^sub>1 v\<^sub>2"}, which matches the left-hand
- side in clause~\textit{(4)} of @{term inj}. In the \<open>if\<close>-branch the derivative is an
- alternative, namely @{term "ALT (SEQ (der c r\<^sub>1) r\<^sub>2) (der c
- r\<^sub>2)"}. This means we either have to consider a \<open>Left\<close>- or
- \<open>Right\<close>-value. In case of the \<open>Left\<close>-value we know further it
- must be a value for a sequence regular expression. Therefore the pattern
- we match in the clause \textit{(5)} is @{term "Left (Seq v\<^sub>1 v\<^sub>2)"},
- while in \textit{(6)} it is just @{term "Right v\<^sub>2"}. One more interesting
- point is in the right-hand side of clause \textit{(6)}: since in this case the
- regular expression \<open>r\<^sub>1\<close> does not ``contribute'' to
- matching the string, that means it only matches the empty string, we need to
- call @{const mkeps} in order to construct a value for how @{term "r\<^sub>1"}
- can match this empty string. A similar argument applies for why we can
- expect in the left-hand side of clause \textit{(7)} that the value is of the form
- @{term "Seq v (Stars vs)"}---the derivative of a star is @{term "SEQ (der c r)
- (STAR r)"}. Finally, the reason for why we can ignore the second argument
- in clause \textit{(1)} of @{term inj} is that it will only ever be called in cases
- where @{term "c=d"}, but the usual linearity restrictions in patterns do
- not allow us to build this constraint explicitly into our function
- definition.\footnote{Sulzmann and Lu state this clause as @{thm (lhs)
- injval.simps(1)[of "c" "c"]} $\dn$ @{thm (rhs) injval.simps(1)[of "c"]},
- but our deviation is harmless.}
-
- The idea of the @{term inj}-function to ``inject'' a character, say
- @{term c}, into a value can be made precise by the first part of the
- following lemma, which shows that the underlying string of an injected
- value has a prepended character @{term c}; the second part shows that
- the underlying string of an @{const mkeps}-value is always the empty
- string (given the regular expression is nullable since otherwise
- \<open>mkeps\<close> might not be defined).
-
- \begin{lemma}\mbox{}\smallskip\\\label{Prf_injval_flat}
- \begin{tabular}{ll}
- (1) & @{thm[mode=IfThen] Prf_injval_flat}\\
- (2) & @{thm[mode=IfThen] mkeps_flat}
- \end{tabular}
- \end{lemma}
-
- \begin{proof}
- Both properties are by routine inductions: the first one can, for example,
- be proved by induction over the definition of @{term derivatives}; the second by
- an induction on @{term r}. There are no interesting cases.\qed
- \end{proof}
-
- Having defined the @{const mkeps} and \<open>inj\<close> function we can extend
- \Brz's matcher so that a value is constructed (assuming the
- regular expression matches the string). The clauses of the Sulzmann and Lu lexer are
-
- \begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) lexer.simps(1)} & $\dn$ & @{thm (rhs) lexer.simps(1)}\\
- @{thm (lhs) lexer.simps(2)} & $\dn$ & \<open>case\<close> @{term "lexer (der c r) s"} \<open>of\<close>\\
- & & \phantom{$|$} @{term "None"} \<open>\<Rightarrow>\<close> @{term None}\\
- & & $|$ @{term "Some v"} \<open>\<Rightarrow>\<close> @{term "Some (injval r c v)"}
- \end{tabular}
- \end{center}
-
- \noindent If the regular expression does not match the string, @{const None} is
- returned. If the regular expression \emph{does}
- match the string, then @{const Some} value is returned. One important
- virtue of this algorithm is that it can be implemented with ease in any
- functional programming language and also in Isabelle/HOL. In the remaining
- part of this section we prove that this algorithm is correct.
-
- The well-known idea of POSIX matching is informally defined by some
- rules such as the Longest Match and Priority Rules (see
- Introduction); as correctly argued in \cite{Sulzmann2014}, this
- needs formal specification. Sulzmann and Lu define an ``ordering
- relation'' between values and argue that there is a maximum value,
- as given by the derivative-based algorithm. In contrast, we shall
- introduce a simple inductive definition that specifies directly what
- a \emph{POSIX value} is, incorporating the POSIX-specific choices
- into the side-conditions of our rules. Our definition is inspired by
- the matching relation given by Vansummeren~\cite{Vansummeren2006}.
- The relation we define is ternary and
- written as \mbox{@{term "s \<in> r \<rightarrow> v"}}, relating
- strings, regular expressions and values; the inductive rules are given in
- Figure~\ref{POSIXrules}.
- We can prove that given a string @{term s} and regular expression @{term
- r}, the POSIX value @{term v} is uniquely determined by @{term "s \<in> r \<rightarrow> v"}.
-
- %
- \begin{figure}[t]
- \begin{center}
- \begin{tabular}{c}
- @{thm[mode=Axiom] Posix.intros(1)}\<open>P\<close>@{term "ONE"} \qquad
- @{thm[mode=Axiom] Posix.intros(2)}\<open>P\<close>@{term "c"}\medskip\\
- @{thm[mode=Rule] Posix.intros(3)[of "s" "r\<^sub>1" "v" "r\<^sub>2"]}\<open>P+L\<close>\qquad
- @{thm[mode=Rule] Posix.intros(4)[of "s" "r\<^sub>2" "v" "r\<^sub>1"]}\<open>P+R\<close>\medskip\\
- $\mprset{flushleft}
- \inferrule
- {@{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
- @{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"]} \\\\
- @{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"]}}
- {@{thm (concl) Posix.intros(5)[of "s\<^sub>1" "r\<^sub>1" "v\<^sub>1" "s\<^sub>2" "r\<^sub>2" "v\<^sub>2"]}}$\<open>PS\<close>\\
- @{thm[mode=Axiom] Posix.intros(7)}\<open>P[]\<close>\medskip\\
- $\mprset{flushleft}
- \inferrule
- {@{thm (prem 1) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]} \qquad
- @{thm (prem 2) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]} \qquad
- @{thm (prem 3) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]} \\\\
- @{thm (prem 4) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]}}
- {@{thm (concl) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]}}$\<open>P\<star>\<close>
- \end{tabular}
- \end{center}
- \caption{Our inductive definition of POSIX values.}\label{POSIXrules}
- \end{figure}
-
-
-
- \begin{theorem}\mbox{}\smallskip\\\label{posixdeterm}
- \begin{tabular}{ll}
- (1) & If @{thm (prem 1) Posix1(1)} then @{thm (concl)
- Posix1(1)} and @{thm (concl) Posix1(2)}.\\
- (2) & @{thm[mode=IfThen] Posix_determ(1)[of _ _ "v" "v'"]}
- \end{tabular}
- \end{theorem}
-
- \begin{proof} Both by induction on the definition of @{term "s \<in> r \<rightarrow> v"}.
- The second parts follows by a case analysis of @{term "s \<in> r \<rightarrow> v'"} and
- the first part.\qed
- \end{proof}
-
- \noindent
- We claim that our @{term "s \<in> r \<rightarrow> v"} relation captures the idea behind the four
- informal POSIX rules shown in the Introduction: Consider for example the
- rules \<open>P+L\<close> and \<open>P+R\<close> where the POSIX value for a string
- and an alternative regular expression, that is @{term "(s, ALT r\<^sub>1 r\<^sub>2)"},
- is specified---it is always a \<open>Left\<close>-value, \emph{except} when the
- string to be matched is not in the language of @{term "r\<^sub>1"}; only then it
- is a \<open>Right\<close>-value (see the side-condition in \<open>P+R\<close>).
- Interesting is also the rule for sequence regular expressions (\<open>PS\<close>). The first two premises state that @{term "v\<^sub>1"} and @{term "v\<^sub>2"}
- are the POSIX values for @{term "(s\<^sub>1, r\<^sub>1)"} and @{term "(s\<^sub>2, r\<^sub>2)"}
- respectively. Consider now the third premise and note that the POSIX value
- of this rule should match the string \mbox{@{term "s\<^sub>1 @ s\<^sub>2"}}. According to the
- Longest Match Rule, we want that the @{term "s\<^sub>1"} is the longest initial
- split of \mbox{@{term "s\<^sub>1 @ s\<^sub>2"}} such that @{term "s\<^sub>2"} is still recognised
- by @{term "r\<^sub>2"}. Let us assume, contrary to the third premise, that there
- \emph{exist} an @{term "s\<^sub>3"} and @{term "s\<^sub>4"} such that @{term "s\<^sub>2"}
- can be split up into a non-empty string @{term "s\<^sub>3"} and a possibly empty
- string @{term "s\<^sub>4"}. Moreover the longer string @{term "s\<^sub>1 @ s\<^sub>3"} can be
- matched by \<open>r\<^sub>1\<close> and the shorter @{term "s\<^sub>4"} can still be
- matched by @{term "r\<^sub>2"}. In this case @{term "s\<^sub>1"} would \emph{not} be the
- longest initial split of \mbox{@{term "s\<^sub>1 @ s\<^sub>2"}} and therefore @{term "Seq v\<^sub>1
- v\<^sub>2"} cannot be a POSIX value for @{term "(s\<^sub>1 @ s\<^sub>2, SEQ r\<^sub>1 r\<^sub>2)"}.
- The main point is that our side-condition ensures the Longest
- Match Rule is satisfied.
-
- A similar condition is imposed on the POSIX value in the \<open>P\<star>\<close>-rule. Also there we want that @{term "s\<^sub>1"} is the longest initial
- split of @{term "s\<^sub>1 @ s\<^sub>2"} and furthermore the corresponding value
- @{term v} cannot be flattened to the empty string. In effect, we require
- that in each ``iteration'' of the star, some non-empty substring needs to
- be ``chipped'' away; only in case of the empty string we accept @{term
- "Stars []"} as the POSIX value. Indeed we can show that our POSIX values
- are lexical values which exclude those \<open>Stars\<close> that contain subvalues
- that flatten to the empty string.
-
- \begin{lemma}\label{LVposix}
- @{thm [mode=IfThen] Posix_LV}
- \end{lemma}
-
- \begin{proof}
- By routine induction on @{thm (prem 1) Posix_LV}.\qed
- \end{proof}
-
- \noindent
- Next is the lemma that shows the function @{term "mkeps"} calculates
- the POSIX value for the empty string and a nullable regular expression.
-
- \begin{lemma}\label{lemmkeps}
- @{thm[mode=IfThen] Posix_mkeps}
- \end{lemma}
-
- \begin{proof}
- By routine induction on @{term r}.\qed
- \end{proof}
-
- \noindent
- The central lemma for our POSIX relation is that the \<open>inj\<close>-function
- preserves POSIX values.
-
- \begin{lemma}\label{Posix2}
- @{thm[mode=IfThen] Posix_injval}
- \end{lemma}
-
- \begin{proof}
- By induction on \<open>r\<close>. We explain two cases.
-
- \begin{itemize}
- \item[$\bullet$] Case @{term "r = ALT r\<^sub>1 r\<^sub>2"}. There are
- two subcases, namely \<open>(a)\<close> \mbox{@{term "v = Left v'"}} and @{term
- "s \<in> der c r\<^sub>1 \<rightarrow> v'"}; and \<open>(b)\<close> @{term "v = Right v'"}, @{term
- "s \<notin> L (der c r\<^sub>1)"} and @{term "s \<in> der c r\<^sub>2 \<rightarrow> v'"}. In \<open>(a)\<close> we
- know @{term "s \<in> der c r\<^sub>1 \<rightarrow> v'"}, from which we can infer @{term "(c # s)
- \<in> r\<^sub>1 \<rightarrow> injval r\<^sub>1 c v'"} by induction hypothesis and hence @{term "(c #
- s) \<in> ALT r\<^sub>1 r\<^sub>2 \<rightarrow> injval (ALT r\<^sub>1 r\<^sub>2) c (Left v')"} as needed. Similarly
- in subcase \<open>(b)\<close> where, however, in addition we have to use
- Proposition~\ref{derprop}(2) in order to infer @{term "c # s \<notin> L r\<^sub>1"} from @{term
- "s \<notin> L (der c r\<^sub>1)"}.\smallskip
-
- \item[$\bullet$] Case @{term "r = SEQ r\<^sub>1 r\<^sub>2"}. There are three subcases:
-
- \begin{quote}
- \begin{description}
- \item[\<open>(a)\<close>] @{term "v = Left (Seq v\<^sub>1 v\<^sub>2)"} and @{term "nullable r\<^sub>1"}
- \item[\<open>(b)\<close>] @{term "v = Right v\<^sub>1"} and @{term "nullable r\<^sub>1"}
- \item[\<open>(c)\<close>] @{term "v = Seq v\<^sub>1 v\<^sub>2"} and @{term "\<not> nullable r\<^sub>1"}
- \end{description}
- \end{quote}
-
- \noindent For \<open>(a)\<close> we know @{term "s\<^sub>1 \<in> der c r\<^sub>1 \<rightarrow> v\<^sub>1"} and
- @{term "s\<^sub>2 \<in> r\<^sub>2 \<rightarrow> v\<^sub>2"} as well as
- %
- \[@{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)"}\]
-
- \noindent From the latter we can infer by Proposition~\ref{derprop}(2):
- %
- \[@{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)"}\]
-
- \noindent We can use the induction hypothesis for \<open>r\<^sub>1\<close> to obtain
- @{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
- @{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 \<open>(c)\<close>
- is similar.
-
- For \<open>(b)\<close> we know @{term "s \<in> der c r\<^sub>2 \<rightarrow> v\<^sub>1"} and
- @{term "s\<^sub>1 @ s\<^sub>2 \<notin> L (SEQ (der c r\<^sub>1) r\<^sub>2)"}. From the former
- we have @{term "(c # s) \<in> r\<^sub>2 \<rightarrow> (injval r\<^sub>2 c v\<^sub>1)"} by induction hypothesis
- for @{term "r\<^sub>2"}. From the latter we can infer
- %
- \[@{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)"}\]
-
- \noindent By Lemma~\ref{lemmkeps} we know @{term "[] \<in> r\<^sub>1 \<rightarrow> (mkeps r\<^sub>1)"}
- holds. Putting this all together, we can conclude with @{term "(c #
- 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.
-
- Finally suppose @{term "r = STAR r\<^sub>1"}. This case is very similar to the
- sequence case, except that we need to also ensure that @{term "flat (injval r\<^sub>1
- c v\<^sub>1) \<noteq> []"}. This follows from @{term "(c # s\<^sub>1)
- \<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
- r\<^sub>1 \<rightarrow> v\<^sub>1"} and the induction hypothesis).\qed
- \end{itemize}
- \end{proof}
-
- \noindent
- With Lemma~\ref{Posix2} in place, it is completely routine to establish
- that the Sulzmann and Lu lexer satisfies our specification (returning
- the null value @{term "None"} iff the string is not in the language of the regular expression,
- and returning a unique POSIX value iff the string \emph{is} in the language):
-
- \begin{theorem}\mbox{}\smallskip\\\label{lexercorrect}
- \begin{tabular}{ll}
- (1) & @{thm (lhs) lexer_correct_None} if and only if @{thm (rhs) lexer_correct_None}\\
- (2) & @{thm (lhs) lexer_correct_Some} if and only if @{thm (rhs) lexer_correct_Some}\\
- \end{tabular}
- \end{theorem}
-
- \begin{proof}
- By induction on @{term s} using Lemma~\ref{lemmkeps} and \ref{Posix2}.\qed
- \end{proof}
-
- \noindent In \textit{(2)} we further know by Theorem~\ref{posixdeterm} that the
- value returned by the lexer must be unique. A simple corollary
- of our two theorems is:
-
- \begin{corollary}\mbox{}\smallskip\\\label{lexercorrectcor}
- \begin{tabular}{ll}
- (1) & @{thm (lhs) lexer_correctness(2)} if and only if @{thm (rhs) lexer_correctness(2)}\\
- (2) & @{thm (lhs) lexer_correctness(1)} if and only if @{thm (rhs) lexer_correctness(1)}\\
- \end{tabular}
- \end{corollary}
-
- \noindent This concludes our correctness proof. Note that we have
- not changed the algorithm of Sulzmann and Lu,\footnote{All
- deviations we introduced are harmless.} but introduced our own
- specification for what a correct result---a POSIX value---should be.
- In the next section we show that our specification coincides with
- another one given by Okui and Suzuki using a different technique.
-
-\<close>
-
-section \<open>Ordering of Values according to Okui and Suzuki\<close>
-
-text \<open>
-
- While in the previous section we have defined POSIX values directly
- in terms of a ternary relation (see inference rules in Figure~\ref{POSIXrules}),
- Sulzmann and Lu took a different approach in \cite{Sulzmann2014}:
- they introduced an ordering for values and identified POSIX values
- as the maximal elements. An extended version of \cite{Sulzmann2014}
- is available at the website of its first author; this includes more
- details of their proofs, but which are evidently not in final form
- yet. Unfortunately, we were not able to verify claims that their
- ordering has properties such as being transitive or having maximal
- elements.
-
- Okui and Suzuki \cite{OkuiSuzuki2010,OkuiSuzukiTech} described
- another ordering of values, which they use to establish the
- correctness of their automata-based algorithm for POSIX matching.
- Their ordering resembles some aspects of the one given by Sulzmann
- and Lu, but overall is quite different. To begin with, Okui and
- Suzuki identify POSIX values as minimal, rather than maximal,
- elements in their ordering. A more substantial difference is that
- the ordering by Okui and Suzuki uses \emph{positions} in order to
- identify and compare subvalues. Positions are lists of natural
- numbers. This allows them to quite naturally formalise the Longest
- Match and Priority rules of the informal POSIX standard. Consider
- for example the value @{term v}
-
- \begin{center}
- @{term "v == Stars [Seq (Char x) (Char y), Char z]"}
- \end{center}
-
- \noindent
- At position \<open>[0,1]\<close> of this value is the
- subvalue \<open>Char y\<close> and at position \<open>[1]\<close> the
- subvalue @{term "Char z"}. At the `root' position, or empty list
- @{term "[]"}, is the whole value @{term v}. Positions such as \<open>[0,1,0]\<close> or \<open>[2]\<close> are outside of \<open>v\<close>. If it exists, the subvalue of @{term v} at a position \<open>p\<close>, written @{term "at v p"}, can be recursively defined by
-
- \begin{center}
- \begin{tabular}{r@ {\hspace{0mm}}lcl}
- @{term v} & \<open>\<downharpoonleft>\<^bsub>[]\<^esub>\<close> & \<open>\<equiv>\<close>& @{thm (rhs) at.simps(1)}\\
- @{term "Left v"} & \<open>\<downharpoonleft>\<^bsub>0::ps\<^esub>\<close> & \<open>\<equiv>\<close>& @{thm (rhs) at.simps(2)}\\
- @{term "Right v"} & \<open>\<downharpoonleft>\<^bsub>1::ps\<^esub>\<close> & \<open>\<equiv>\<close> &
- @{thm (rhs) at.simps(3)[simplified Suc_0_fold]}\\
- @{term "Seq v\<^sub>1 v\<^sub>2"} & \<open>\<downharpoonleft>\<^bsub>0::ps\<^esub>\<close> & \<open>\<equiv>\<close> &
- @{thm (rhs) at.simps(4)[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]} \\
- @{term "Seq v\<^sub>1 v\<^sub>2"} & \<open>\<downharpoonleft>\<^bsub>1::ps\<^esub>\<close>
- & \<open>\<equiv>\<close> &
- @{thm (rhs) at.simps(5)[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2", simplified Suc_0_fold]} \\
- @{term "Stars vs"} & \<open>\<downharpoonleft>\<^bsub>n::ps\<^esub>\<close> & \<open>\<equiv>\<close>& @{thm (rhs) at.simps(6)}\\
- \end{tabular}
- \end{center}
-
- \noindent In the last clause we use Isabelle's notation @{term "vs ! n"} for the
- \<open>n\<close>th element in a list. The set of positions inside a value \<open>v\<close>,
- written @{term "Pos v"}, is given by
-
- \begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) Pos.simps(1)} & \<open>\<equiv>\<close> & @{thm (rhs) Pos.simps(1)}\\
- @{thm (lhs) Pos.simps(2)} & \<open>\<equiv>\<close> & @{thm (rhs) Pos.simps(2)}\\
- @{thm (lhs) Pos.simps(3)} & \<open>\<equiv>\<close> & @{thm (rhs) Pos.simps(3)}\\
- @{thm (lhs) Pos.simps(4)} & \<open>\<equiv>\<close> & @{thm (rhs) Pos.simps(4)}\\
- @{thm (lhs) Pos.simps(5)[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}
- & \<open>\<equiv>\<close>
- & @{thm (rhs) Pos.simps(5)[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}\\
- @{thm (lhs) Pos_stars} & \<open>\<equiv>\<close> & @{thm (rhs) Pos_stars}\\
- \end{tabular}
- \end{center}
-
- \noindent
- whereby \<open>len\<close> in the last clause stands for the length of a list. Clearly
- for every position inside a value there exists a subvalue at that position.
-
-
- To help understanding the ordering of Okui and Suzuki, consider again
- the earlier value
- \<open>v\<close> and compare it with the following \<open>w\<close>:
-
- \begin{center}
- \begin{tabular}{l}
- @{term "v == Stars [Seq (Char x) (Char y), Char z]"}\\
- @{term "w == Stars [Char x, Char y, Char z]"}
- \end{tabular}
- \end{center}
-
- \noindent Both values match the string \<open>xyz\<close>, that means if
- we flatten these values at their respective root position, we obtain
- \<open>xyz\<close>. However, at position \<open>[0]\<close>, \<open>v\<close> matches
- \<open>xy\<close> whereas \<open>w\<close> matches only the shorter \<open>x\<close>. So
- according to the Longest Match Rule, we should prefer \<open>v\<close>,
- rather than \<open>w\<close> as POSIX value for string \<open>xyz\<close> (and
- corresponding regular expression). In order to
- formalise this idea, Okui and Suzuki introduce a measure for
- subvalues at position \<open>p\<close>, called the \emph{norm} of \<open>v\<close>
- at position \<open>p\<close>. We can define this measure in Isabelle as an
- integer as follows
-
- \begin{center}
- @{thm pflat_len_def}
- \end{center}
-
- \noindent where we take the length of the flattened value at
- position \<open>p\<close>, provided the position is inside \<open>v\<close>; if
- not, then the norm is \<open>-1\<close>. The default for outside
- positions is crucial for the POSIX requirement of preferring a
- \<open>Left\<close>-value over a \<open>Right\<close>-value (if they can match the
- same string---see the Priority Rule from the Introduction). For this
- consider
-
- \begin{center}
- @{term "v == Left (Char x)"} \qquad and \qquad @{term "w == Right (Char x)"}
- \end{center}
-
- \noindent Both values match \<open>x\<close>. At position \<open>[0]\<close>
- the norm of @{term v} is \<open>1\<close> (the subvalue matches \<open>x\<close>),
- but the norm of \<open>w\<close> is \<open>-1\<close> (the position is outside
- \<open>w\<close> according to how we defined the `inside' positions of
- \<open>Left\<close>- and \<open>Right\<close>-values). Of course at position
- \<open>[1]\<close>, the norms @{term "pflat_len v [1]"} and @{term
- "pflat_len w [1]"} are reversed, but the point is that subvalues
- will be analysed according to lexicographically ordered
- positions. According to this ordering, the position \<open>[0]\<close>
- takes precedence over \<open>[1]\<close> and thus also \<open>v\<close> will be
- preferred over \<open>w\<close>. The lexicographic ordering of positions, written
- @{term "DUMMY \<sqsubset>lex DUMMY"}, can be conveniently formalised
- by three inference rules
-
- \begin{center}
- \begin{tabular}{ccc}
- @{thm [mode=Axiom] lex_list.intros(1)}\hspace{1cm} &
- @{thm [mode=Rule] lex_list.intros(3)[where ?p1.0="p\<^sub>1" and ?p2.0="p\<^sub>2" and
- ?ps1.0="ps\<^sub>1" and ?ps2.0="ps\<^sub>2"]}\hspace{1cm} &
- @{thm [mode=Rule] lex_list.intros(2)[where ?ps1.0="ps\<^sub>1" and ?ps2.0="ps\<^sub>2"]}
- \end{tabular}
- \end{center}
-
- With the norm and lexicographic order in place,
- we can state the key definition of Okui and Suzuki
- \cite{OkuiSuzuki2010}: a value @{term "v\<^sub>1"} is \emph{smaller at position \<open>p\<close>} than
- @{term "v\<^sub>2"}, written @{term "v\<^sub>1 \<sqsubset>val p v\<^sub>2"},
- if and only if $(i)$ the norm at position \<open>p\<close> is
- greater in @{term "v\<^sub>1"} (that is the string @{term "flat (at v\<^sub>1 p)"} is longer
- than @{term "flat (at v\<^sub>2 p)"}) and $(ii)$ all subvalues at
- positions that are inside @{term "v\<^sub>1"} or @{term "v\<^sub>2"} and that are
- lexicographically smaller than \<open>p\<close>, we have the same norm, namely
-
- \begin{center}
- \begin{tabular}{c}
- @{thm (lhs) PosOrd_def[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}
- \<open>\<equiv>\<close>
- $\begin{cases}
- (i) & @{term "pflat_len v\<^sub>1 p > pflat_len v\<^sub>2 p"} \quad\text{and}\smallskip \\
- (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)"}
- \end{cases}$
- \end{tabular}
- \end{center}
-
- \noindent The position \<open>p\<close> in this definition acts as the
- \emph{first distinct position} of \<open>v\<^sub>1\<close> and \<open>v\<^sub>2\<close>, where both values match strings of different length
- \cite{OkuiSuzuki2010}. Since at \<open>p\<close> the values \<open>v\<^sub>1\<close> and \<open>v\<^sub>2\<close> match different strings, the
- ordering is irreflexive. Derived from the definition above
- are the following two orderings:
-
- \begin{center}
- \begin{tabular}{l}
- @{thm PosOrd_ex_def[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}\\
- @{thm PosOrd_ex_eq_def[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}
- \end{tabular}
- \end{center}
-
- While we encountered a number of obstacles for establishing properties like
- transitivity for the ordering of Sulzmann and Lu (and which we failed
- to overcome), it is relatively straightforward to establish this
- property for the orderings
- @{term "DUMMY :\<sqsubset>val DUMMY"} and @{term "DUMMY :\<sqsubseteq>val DUMMY"}
- by Okui and Suzuki.
-
- \begin{lemma}[Transitivity]\label{transitivity}
- @{thm [mode=IfThen] PosOrd_trans[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2" and ?v3.0="v\<^sub>3"]}
- \end{lemma}
-
- \begin{proof} From the assumption we obtain two positions \<open>p\<close>
- and \<open>q\<close>, where the values \<open>v\<^sub>1\<close> and \<open>v\<^sub>2\<close> (respectively \<open>v\<^sub>2\<close> and \<open>v\<^sub>3\<close>) are `distinct'. Since \<open>\<prec>\<^bsub>lex\<^esub>\<close> is trichotomous, we need to consider
- three cases, namely @{term "p = q"}, @{term "p \<sqsubset>lex q"} and
- @{term "q \<sqsubset>lex p"}. Let us look at the first case. Clearly
- @{term "pflat_len v\<^sub>2 p < pflat_len v\<^sub>1 p"} and @{term
- "pflat_len v\<^sub>3 p < pflat_len v\<^sub>2 p"} imply @{term
- "pflat_len v\<^sub>3 p < pflat_len v\<^sub>1 p"}. It remains to show
- that for a @{term "p' \<in> Pos v\<^sub>1 \<union> Pos v\<^sub>3"}
- with @{term "p' \<sqsubset>lex p"} that @{term "pflat_len v\<^sub>1
- p' = pflat_len v\<^sub>3 p'"} holds. Suppose @{term "p' \<in> Pos
- v\<^sub>1"}, then we can infer from the first assumption that @{term
- "pflat_len v\<^sub>1 p' = pflat_len v\<^sub>2 p'"}. But this means
- that @{term "p'"} must be in @{term "Pos v\<^sub>2"} too (the norm
- cannot be \<open>-1\<close> given @{term "p' \<in> Pos v\<^sub>1"}).
- Hence we can use the second assumption and
- infer @{term "pflat_len v\<^sub>2 p' = pflat_len v\<^sub>3 p'"},
- which concludes this case with @{term "v\<^sub>1 :\<sqsubset>val
- v\<^sub>3"}. The reasoning in the other cases is similar.\qed
- \end{proof}
-
- \noindent
- The proof for $\preccurlyeq$ is similar and omitted.
- It is also straightforward to show that \<open>\<prec>\<close> and
- $\preccurlyeq$ are partial orders. Okui and Suzuki furthermore show that they
- are linear orderings for lexical values \cite{OkuiSuzuki2010} of a given
- regular expression and given string, but we have not formalised this in Isabelle. It is
- not essential for our results. What we are going to show below is
- that for a given \<open>r\<close> and \<open>s\<close>, the orderings have a unique
- minimal element on the set @{term "LV r s"}, which is the POSIX value
- we defined in the previous section. We start with two properties that
- show how the length of a flattened value relates to the \<open>\<prec>\<close>-ordering.
-
- \begin{proposition}\mbox{}\smallskip\\\label{ordlen}
- \begin{tabular}{@ {}ll}
- (1) &
- @{thm [mode=IfThen] PosOrd_shorterE[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}\\
- (2) &
- @{thm [mode=IfThen] PosOrd_shorterI[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}
- \end{tabular}
- \end{proposition}
-
- \noindent Both properties follow from the definition of the ordering. Note that
- \textit{(2)} entails that a value, say @{term "v\<^sub>2"}, whose underlying
- string is a strict prefix of another flattened value, say @{term "v\<^sub>1"}, then
- @{term "v\<^sub>1"} must be smaller than @{term "v\<^sub>2"}. For our proofs it
- will be useful to have the following properties---in each case the underlying strings
- of the compared values are the same:
-
- \begin{proposition}\mbox{}\smallskip\\\label{ordintros}
- \begin{tabular}{ll}
- \textit{(1)} &
- @{thm [mode=IfThen] PosOrd_Left_Right[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}\\
- \textit{(2)} & If
- @{thm (prem 1) PosOrd_Left_eq[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]} \;then\;
- @{thm (lhs) PosOrd_Left_eq[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]} \;iff\;
- @{thm (rhs) PosOrd_Left_eq[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}\\
- \textit{(3)} & If
- @{thm (prem 1) PosOrd_Right_eq[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]} \;then\;
- @{thm (lhs) PosOrd_Right_eq[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]} \;iff\;
- @{thm (rhs) PosOrd_Right_eq[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}\\
- \textit{(4)} & If
- @{thm (prem 1) PosOrd_Seq_eq[where ?v2.0="v\<^sub>2" and ?w2.0="w\<^sub>2"]} \;then\;
- @{thm (lhs) PosOrd_Seq_eq[where ?v2.0="v\<^sub>2" and ?w2.0="w\<^sub>2"]} \;iff\;
- @{thm (rhs) PosOrd_Seq_eq[where ?v2.0="v\<^sub>2" and ?w2.0="w\<^sub>2"]}\\
- \textit{(5)} & If
- @{thm (prem 2) PosOrd_SeqI1[simplified, where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2" and
- ?w1.0="w\<^sub>1" and ?w2.0="w\<^sub>2"]} \;and\;
- @{thm (prem 1) PosOrd_SeqI1[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2" and
- ?w1.0="w\<^sub>1" and ?w2.0="w\<^sub>2"]} \;then\;
- @{thm (concl) PosOrd_SeqI1[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2" and
- ?w1.0="w\<^sub>1" and ?w2.0="w\<^sub>2"]}\\
- \textit{(6)} & If
- @{thm (prem 1) PosOrd_Stars_append_eq[where ?vs1.0="vs\<^sub>1" and ?vs2.0="vs\<^sub>2"]} \;then\;
- @{thm (lhs) PosOrd_Stars_append_eq[where ?vs1.0="vs\<^sub>1" and ?vs2.0="vs\<^sub>2"]} \;iff\;
- @{thm (rhs) PosOrd_Stars_append_eq[where ?vs1.0="vs\<^sub>1" and ?vs2.0="vs\<^sub>2"]}\\
-
- \textit{(7)} & If
- @{thm (prem 2) PosOrd_StarsI[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2" and
- ?vs1.0="vs\<^sub>1" and ?vs2.0="vs\<^sub>2"]} \;and\;
- @{thm (prem 1) PosOrd_StarsI[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2" and
- ?vs1.0="vs\<^sub>1" and ?vs2.0="vs\<^sub>2"]} \;then\;
- @{thm (concl) PosOrd_StarsI[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2" and
- ?vs1.0="vs\<^sub>1" and ?vs2.0="vs\<^sub>2"]}\\
- \end{tabular}
- \end{proposition}
-
- \noindent One might prefer that statements \textit{(4)} and \textit{(5)}
- (respectively \textit{(6)} and \textit{(7)})
- are combined into a single \textit{iff}-statement (like the ones for \<open>Left\<close> and \<open>Right\<close>). Unfortunately this cannot be done easily: such
- a single statement would require an additional assumption about the
- two values @{term "Seq v\<^sub>1 v\<^sub>2"} and @{term "Seq w\<^sub>1 w\<^sub>2"}
- being inhabited by the same regular expression. The
- complexity of the proofs involved seems to not justify such a
- `cleaner' single statement. The statements given are just the properties that
- allow us to establish our theorems without any difficulty. The proofs
- for Proposition~\ref{ordintros} are routine.
-
-
- Next we establish how Okui and Suzuki's orderings relate to our
- definition of POSIX values. Given a \<open>POSIX\<close> value \<open>v\<^sub>1\<close>
- for \<open>r\<close> and \<open>s\<close>, then any other lexical value \<open>v\<^sub>2\<close> in @{term "LV r s"} is greater or equal than \<open>v\<^sub>1\<close>, namely:
-
-
- \begin{theorem}\label{orderone}
- @{thm [mode=IfThen] Posix_PosOrd[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}
- \end{theorem}
-
- \begin{proof} By induction on our POSIX rules. By
- Theorem~\ref{posixdeterm} and the definition of @{const LV}, it is clear
- that \<open>v\<^sub>1\<close> and \<open>v\<^sub>2\<close> have the same
- underlying string @{term s}. The three base cases are
- straightforward: for example for @{term "v\<^sub>1 = Void"}, we have
- that @{term "v\<^sub>2 \<in> LV ONE []"} must also be of the form
- \mbox{@{term "v\<^sub>2 = Void"}}. Therefore we have @{term
- "v\<^sub>1 :\<sqsubseteq>val v\<^sub>2"}. The inductive cases for
- \<open>r\<close> being of the form @{term "ALT r\<^sub>1 r\<^sub>2"} and
- @{term "SEQ r\<^sub>1 r\<^sub>2"} are as follows:
-
-
- \begin{itemize}
-
- \item[$\bullet$] Case \<open>P+L\<close> with @{term "s \<in> (ALT r\<^sub>1 r\<^sub>2)
- \<rightarrow> (Left w\<^sub>1)"}: In this case the value
- @{term "v\<^sub>2"} is either of the
- form @{term "Left w\<^sub>2"} or @{term "Right w\<^sub>2"}. In the
- latter case we can immediately conclude with \mbox{@{term "v\<^sub>1
- :\<sqsubseteq>val v\<^sub>2"}} since a \<open>Left\<close>-value with the
- same underlying string \<open>s\<close> is always smaller than a
- \<open>Right\<close>-value by Proposition~\ref{ordintros}\textit{(1)}.
- In the former case we have @{term "w\<^sub>2
- \<in> LV r\<^sub>1 s"} and can use the induction hypothesis to infer
- @{term "w\<^sub>1 :\<sqsubseteq>val w\<^sub>2"}. Because @{term
- "w\<^sub>1"} and @{term "w\<^sub>2"} have the same underlying string
- \<open>s\<close>, we can conclude with @{term "Left w\<^sub>1
- :\<sqsubseteq>val Left w\<^sub>2"} using
- Proposition~\ref{ordintros}\textit{(2)}.\smallskip
-
- \item[$\bullet$] Case \<open>P+R\<close> with @{term "s \<in> (ALT r\<^sub>1 r\<^sub>2)
- \<rightarrow> (Right w\<^sub>1)"}: This case similar to the previous
- case, except that we additionally know @{term "s \<notin> L
- r\<^sub>1"}. This is needed when @{term "v\<^sub>2"} is of the form
- \mbox{@{term "Left w\<^sub>2"}}. Since \mbox{@{term "flat v\<^sub>2 = flat
- w\<^sub>2"} \<open>= s\<close>} and @{term "\<Turnstile> w\<^sub>2 :
- r\<^sub>1"}, we can derive a contradiction for \mbox{@{term "s \<notin> L
- r\<^sub>1"}} using
- Proposition~\ref{inhabs}. So also in this case \mbox{@{term "v\<^sub>1
- :\<sqsubseteq>val v\<^sub>2"}}.\smallskip
-
- \item[$\bullet$] Case \<open>PS\<close> with @{term "(s\<^sub>1 @
- s\<^sub>2) \<in> (SEQ r\<^sub>1 r\<^sub>2) \<rightarrow> (Seq
- w\<^sub>1 w\<^sub>2)"}: We can assume @{term "v\<^sub>2 = Seq
- (u\<^sub>1) (u\<^sub>2)"} with @{term "\<Turnstile> u\<^sub>1 :
- r\<^sub>1"} and \mbox{@{term "\<Turnstile> u\<^sub>2 :
- r\<^sub>2"}}. We have @{term "s\<^sub>1 @ s\<^sub>2 = (flat
- u\<^sub>1) @ (flat u\<^sub>2)"}. By the side-condition of the
- \<open>PS\<close>-rule we know that either @{term "s\<^sub>1 = flat
- u\<^sub>1"} or that @{term "flat u\<^sub>1"} is a strict prefix of
- @{term "s\<^sub>1"}. In the latter case we can infer @{term
- "w\<^sub>1 :\<sqsubset>val u\<^sub>1"} by
- Proposition~\ref{ordlen}\textit{(2)} and from this @{term "v\<^sub>1
- :\<sqsubseteq>val v\<^sub>2"} by Proposition~\ref{ordintros}\textit{(5)}
- (as noted above @{term "v\<^sub>1"} and @{term "v\<^sub>2"} must have the
- same underlying string).
- In the former case we know
- @{term "u\<^sub>1 \<in> LV r\<^sub>1 s\<^sub>1"} and @{term
- "u\<^sub>2 \<in> LV r\<^sub>2 s\<^sub>2"}. With this we can use the
- induction hypotheses to infer @{term "w\<^sub>1 :\<sqsubseteq>val
- u\<^sub>1"} and @{term "w\<^sub>2 :\<sqsubseteq>val u\<^sub>2"}. By
- Proposition~\ref{ordintros}\textit{(4,5)} we can again infer
- @{term "v\<^sub>1 :\<sqsubseteq>val
- v\<^sub>2"}.
-
- \end{itemize}
-
- \noindent The case for \<open>P\<star>\<close> is similar to the \<open>PS\<close>-case and omitted.\qed
- \end{proof}
-
- \noindent This theorem shows that our \<open>POSIX\<close> value for a
- regular expression \<open>r\<close> and string @{term s} is in fact a
- minimal element of the values in \<open>LV r s\<close>. By
- Proposition~\ref{ordlen}\textit{(2)} we also know that any value in
- \<open>LV r s'\<close>, with @{term "s'"} being a strict prefix, cannot be
- smaller than \<open>v\<^sub>1\<close>. The next theorem shows the
- opposite---namely any minimal element in @{term "LV r s"} must be a
- \<open>POSIX\<close> value. This can be established by induction on \<open>r\<close>, but the proof can be drastically simplified by using the fact
- from the previous section about the existence of a \<open>POSIX\<close> value
- whenever a string @{term "s \<in> L r"}.
-
-
- \begin{theorem}
- @{thm [mode=IfThen] PosOrd_Posix[where ?v1.0="v\<^sub>1"]}
- \end{theorem}
-
- \begin{proof}
- If @{thm (prem 1) PosOrd_Posix[where ?v1.0="v\<^sub>1"]} then
- @{term "s \<in> L r"} by Proposition~\ref{inhabs}. Hence by Theorem~\ref{lexercorrect}(2)
- there exists a
- \<open>POSIX\<close> value @{term "v\<^sub>P"} with @{term "s \<in> r \<rightarrow> v\<^sub>P"}
- and by Lemma~\ref{LVposix} we also have \mbox{@{term "v\<^sub>P \<in> LV r s"}}.
- By Theorem~\ref{orderone} we therefore have
- @{term "v\<^sub>P :\<sqsubseteq>val v\<^sub>1"}. If @{term "v\<^sub>P = v\<^sub>1"} then
- we are done. Otherwise we have @{term "v\<^sub>P :\<sqsubset>val v\<^sub>1"}, which
- however contradicts the second assumption about @{term "v\<^sub>1"} being the smallest
- element in @{term "LV r s"}. So we are done in this case too.\qed
- \end{proof}
-
- \noindent
- From this we can also show
- that if @{term "LV r s"} is non-empty (or equivalently @{term "s \<in> L r"}) then
- it has a unique minimal element:
-
- \begin{corollary}
- @{thm [mode=IfThen] Least_existence1}
- \end{corollary}
-
-
-
- \noindent To sum up, we have shown that the (unique) minimal elements
- of the ordering by Okui and Suzuki are exactly the \<open>POSIX\<close>
- values we defined inductively in Section~\ref{posixsec}. This provides
- an independent confirmation that our ternary relation formalises the
- informal POSIX rules.
-
-\<close>
-
-section \<open>Bitcoded Lexing\<close>
-
-
-
-
-text \<open>
-
-Incremental calculation of the value. To simplify the proof we first define the function
-@{const flex} which calculates the ``iterated'' injection function. With this we can
-rewrite the lexer as
-
-\begin{center}
-@{thm lexer_flex}
-\end{center}
-
-
-\<close>
-
-section \<open>Optimisations\<close>
-
-text \<open>
-
- Derivatives as calculated by \Brz's method are usually more complex
- regular expressions than the initial one; the result is that the
- derivative-based matching and lexing algorithms are often abysmally slow.
- However, various optimisations are possible, such as the simplifications
- of @{term "ALT ZERO r"}, @{term "ALT r ZERO"}, @{term "SEQ ONE r"} and
- @{term "SEQ r ONE"} to @{term r}. These simplifications can speed up the
- algorithms considerably, as noted in \cite{Sulzmann2014}. One of the
- advantages of having a simple specification and correctness proof is that
- the latter can be refined to prove the correctness of such simplification
- steps. While the simplification of regular expressions according to
- rules like
-
- \begin{equation}\label{Simpl}
- \begin{array}{lcllcllcllcl}
- @{term "ALT ZERO r"} & \<open>\<Rightarrow>\<close> & @{term r} \hspace{8mm}%\\
- @{term "ALT r ZERO"} & \<open>\<Rightarrow>\<close> & @{term r} \hspace{8mm}%\\
- @{term "SEQ ONE r"} & \<open>\<Rightarrow>\<close> & @{term r} \hspace{8mm}%\\
- @{term "SEQ r ONE"} & \<open>\<Rightarrow>\<close> & @{term r}
- \end{array}
- \end{equation}
-
- \noindent is well understood, there is an obstacle with the POSIX value
- calculation algorithm by Sulzmann and Lu: if we build a derivative regular
- expression and then simplify it, we will calculate a POSIX value for this
- simplified derivative regular expression, \emph{not} for the original (unsimplified)
- derivative regular expression. Sulzmann and Lu \cite{Sulzmann2014} overcome this obstacle by
- not just calculating a simplified regular expression, but also calculating
- a \emph{rectification function} that ``repairs'' the incorrect value.
-
- The rectification functions can be (slightly clumsily) implemented in
- Isabelle/HOL as follows using some auxiliary functions:
-
- \begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) F_RIGHT.simps(1)} & $\dn$ & \<open>Right (f v)\<close>\\
- @{thm (lhs) F_LEFT.simps(1)} & $\dn$ & \<open>Left (f v)\<close>\\
-
- @{thm (lhs) F_ALT.simps(1)} & $\dn$ & \<open>Right (f\<^sub>2 v)\<close>\\
- @{thm (lhs) F_ALT.simps(2)} & $\dn$ & \<open>Left (f\<^sub>1 v)\<close>\\
-
- @{thm (lhs) F_SEQ1.simps(1)} & $\dn$ & \<open>Seq (f\<^sub>1 ()) (f\<^sub>2 v)\<close>\\
- @{thm (lhs) F_SEQ2.simps(1)} & $\dn$ & \<open>Seq (f\<^sub>1 v) (f\<^sub>2 ())\<close>\\
- @{thm (lhs) F_SEQ.simps(1)} & $\dn$ & \<open>Seq (f\<^sub>1 v\<^sub>1) (f\<^sub>2 v\<^sub>2)\<close>\medskip\\
- %\end{tabular}
- %
- %\begin{tabular}{lcl}
- @{term "simp_ALT (ZERO, DUMMY) (r\<^sub>2, f\<^sub>2)"} & $\dn$ & @{term "(r\<^sub>2, F_RIGHT f\<^sub>2)"}\\
- @{term "simp_ALT (r\<^sub>1, f\<^sub>1) (ZERO, DUMMY)"} & $\dn$ & @{term "(r\<^sub>1, F_LEFT f\<^sub>1)"}\\
- @{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)"}\\
- @{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)"}\\
- @{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)"}\\
- @{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)"}\\
- \end{tabular}
- \end{center}
-
- \noindent
- The functions \<open>simp\<^bsub>Alt\<^esub>\<close> and \<open>simp\<^bsub>Seq\<^esub>\<close> encode the simplification rules
- in \eqref{Simpl} and compose the rectification functions (simplifications can occur
- deep inside the regular expression). The main simplification function is then
-
- \begin{center}
- \begin{tabular}{lcl}
- @{term "simp (ALT r\<^sub>1 r\<^sub>2)"} & $\dn$ & @{term "simp_ALT (simp r\<^sub>1) (simp r\<^sub>2)"}\\
- @{term "simp (SEQ r\<^sub>1 r\<^sub>2)"} & $\dn$ & @{term "simp_SEQ (simp r\<^sub>1) (simp r\<^sub>2)"}\\
- @{term "simp r"} & $\dn$ & @{term "(r, id)"}\\
- \end{tabular}
- \end{center}
-
- \noindent where @{term "id"} stands for the identity function. The
- function @{const simp} returns a simplified regular expression and a corresponding
- rectification function. Note that we do not simplify under stars: this
- seems to slow down the algorithm, rather than speed it up. The optimised
- lexer is then given by the clauses:
-
- \begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) slexer.simps(1)} & $\dn$ & @{thm (rhs) slexer.simps(1)}\\
- @{thm (lhs) slexer.simps(2)} & $\dn$ &
- \<open>let (r\<^sub>s, f\<^sub>r) = simp (r \<close>$\backslash$\<open> c) in\<close>\\
- & & \<open>case\<close> @{term "slexer r\<^sub>s s"} \<open>of\<close>\\
- & & \phantom{$|$} @{term "None"} \<open>\<Rightarrow>\<close> @{term None}\\
- & & $|$ @{term "Some v"} \<open>\<Rightarrow>\<close> \<open>Some (inj r c (f\<^sub>r v))\<close>
- \end{tabular}
- \end{center}
-
- \noindent
- In the second clause we first calculate the derivative @{term "der c r"}
- and then simpli
-
-text \<open>
-
-Incremental calculation of the value. To simplify the proof we first define the function
-@{const flex} which calculates the ``iterated'' injection function. With this we can
-rewrite the lexer as
-
-\begin{center}
-@{thm lexer_flex}
-\end{center}
-
-\begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) code.simps(1)} & $\dn$ & @{thm (rhs) code.simps(1)}\\
- @{thm (lhs) code.simps(2)} & $\dn$ & @{thm (rhs) code.simps(2)}\\
- @{thm (lhs) code.simps(3)} & $\dn$ & @{thm (rhs) code.simps(3)}\\
- @{thm (lhs) code.simps(4)} & $\dn$ & @{thm (rhs) code.simps(4)}\\
- @{thm (lhs) code.simps(5)[of "v\<^sub>1" "v\<^sub>2"]} & $\dn$ & @{thm (rhs) code.simps(5)[of "v\<^sub>1" "v\<^sub>2"]}\\
- @{thm (lhs) code.simps(6)} & $\dn$ & @{thm (rhs) code.simps(6)}\\
- @{thm (lhs) code.simps(7)} & $\dn$ & @{thm (rhs) code.simps(7)}
-\end{tabular}
-\end{center}
-
-\begin{center}
-\begin{tabular}{lcl}
- @{term areg} & $::=$ & @{term "AZERO"}\\
- & $\mid$ & @{term "AONE bs"}\\
- & $\mid$ & @{term "ACHAR bs c"}\\
- & $\mid$ & @{term "AALT bs r1 r2"}\\
- & $\mid$ & @{term "ASEQ bs r\<^sub>1 r\<^sub>2"}\\
- & $\mid$ & @{term "ASTAR bs r"}
-\end{tabular}
-\end{center}
-
-
-\begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) intern.simps(1)} & $\dn$ & @{thm (rhs) intern.simps(1)}\\
- @{thm (lhs) intern.simps(2)} & $\dn$ & @{thm (rhs) intern.simps(2)}\\
- @{thm (lhs) intern.simps(3)} & $\dn$ & @{thm (rhs) intern.simps(3)}\\
- @{thm (lhs) intern.simps(4)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) intern.simps(4)[of "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) intern.simps(5)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) intern.simps(5)[of "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) intern.simps(6)} & $\dn$ & @{thm (rhs) intern.simps(6)}\\
-\end{tabular}
-\end{center}
-
-\begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) erase.simps(1)} & $\dn$ & @{thm (rhs) erase.simps(1)}\\
- @{thm (lhs) erase.simps(2)[of bs]} & $\dn$ & @{thm (rhs) erase.simps(2)[of bs]}\\
- @{thm (lhs) erase.simps(3)[of bs]} & $\dn$ & @{thm (rhs) erase.simps(3)[of bs]}\\
- @{thm (lhs) erase.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) erase.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) erase.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) erase.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) erase.simps(6)[of bs]} & $\dn$ & @{thm (rhs) erase.simps(6)[of bs]}\\
-\end{tabular}
-\end{center}
-
-Some simple facts about erase
-
-\begin{lemma}\mbox{}\\
-@{thm erase_bder}\\
-@{thm erase_intern}
-\end{lemma}
-
-\begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) bnullable.simps(1)} & $\dn$ & @{thm (rhs) bnullable.simps(1)}\\
- @{thm (lhs) bnullable.simps(2)} & $\dn$ & @{thm (rhs) bnullable.simps(2)}\\
- @{thm (lhs) bnullable.simps(3)} & $\dn$ & @{thm (rhs) bnullable.simps(3)}\\
- @{thm (lhs) bnullable.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bnullable.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) bnullable.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bnullable.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) bnullable.simps(6)} & $\dn$ & @{thm (rhs) bnullable.simps(6)}\medskip\\
-
-% \end{tabular}
-% \end{center}
-
-% \begin{center}
-% \begin{tabular}{lcl}
-
- @{thm (lhs) bder.simps(1)} & $\dn$ & @{thm (rhs) bder.simps(1)}\\
- @{thm (lhs) bder.simps(2)} & $\dn$ & @{thm (rhs) bder.simps(2)}\\
- @{thm (lhs) bder.simps(3)} & $\dn$ & @{thm (rhs) bder.simps(3)}\\
- @{thm (lhs) bder.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bder.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) bder.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bder.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) bder.simps(6)} & $\dn$ & @{thm (rhs) bder.simps(6)}
- \end{tabular}
- \end{center}
-
-
-\begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) bmkeps.simps(1)} & $\dn$ & @{thm (rhs) bmkeps.simps(1)}\\
- @{thm (lhs) bmkeps.simps(2)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bmkeps.simps(2)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) bmkeps.simps(3)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bmkeps.simps(3)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) bmkeps.simps(4)} & $\dn$ & @{thm (rhs) bmkeps.simps(4)}\medskip\\
-\end{tabular}
-\end{center}
-
-
-@{thm [mode=IfThen] bder_retrieve}
-
-By induction on \<open>r\<close>
-
-\begin{theorem}[Main Lemma]\mbox{}\\
-@{thm [mode=IfThen] MAIN_decode}
-\end{theorem}
-
-\noindent
-Definition of the bitcoded lexer
-
-@{thm blexer_def}
-
-
-\begin{theorem}
-@{thm blexer_correctness}
-\end{theorem}
-
-\<close>
-
-section \<open>Optimisations\<close>
-
-text \<open>
-
- Derivatives as calculated by \Brz's method are usually more complex
- regular expressions than the initial one; the result is that the
- derivative-based matching and lexing algorithms are often abysmally slow.
- However, various optimisations are possible, such as the simplifications
- of @{term "ALT ZERO r"}, @{term "ALT r ZERO"}, @{term "SEQ ONE r"} and
- @{term "SEQ r ONE"} to @{term r}. These simplifications can speed up the
- algorithms considerably, as noted in \cite{Sulzmann2014}. One of the
- advantages of having a simple specification and correctness proof is that
- the latter can be refined to prove the correctness of such simplification
- steps. While the simplification of regular expressions according to
- rules like
-
- \begin{equation}\label{Simpl}
- \begin{array}{lcllcllcllcl}
- @{term "ALT ZERO r"} & \<open>\<Rightarrow>\<close> & @{term r} \hspace{8mm}%\\
- @{term "ALT r ZERO"} & \<open>\<Rightarrow>\<close> & @{term r} \hspace{8mm}%\\
- @{term "SEQ ONE r"} & \<open>\<Rightarrow>\<close> & @{term r} \hspace{8mm}%\\
- @{term "SEQ r ONE"} & \<open>\<Rightarrow>\<close> & @{term r}
- \end{array}
- \end{equation}
-
- \noindent is well understood, there is an obstacle with the POSIX value
- calculation algorithm by Sulzmann and Lu: if we build a derivative regular
- expression and then simplify it, we will calculate a POSIX value for this
- simplified derivative regular expression, \emph{not} for the original (unsimplified)
- derivative regular expression. Sulzmann and Lu \cite{Sulzmann2014} overcome this obstacle by
- not just calculating a simplified regular expression, but also calculating
- a \emph{rectification function} that ``repairs'' the incorrect value.
-
- The rectification functions can be (slightly clumsily) implemented in
- Isabelle/HOL as follows using some auxiliary functions:
-
- \begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) F_RIGHT.simps(1)} & $\dn$ & \<open>Right (f v)\<close>\\
- @{thm (lhs) F_LEFT.simps(1)} & $\dn$ & \<open>Left (f v)\<close>\\
-
- @{thm (lhs) F_ALT.simps(1)} & $\dn$ & \<open>Right (f\<^sub>2 v)\<close>\\
- @{thm (lhs) F_ALT.simps(2)} & $\dn$ & \<open>Left (f\<^sub>1 v)\<close>\\
-
- @{thm (lhs) F_SEQ1.simps(1)} & $\dn$ & \<open>Seq (f\<^sub>1 ()) (f\<^sub>2 v)\<close>\\
- @{thm (lhs) F_SEQ2.simps(1)} & $\dn$ & \<open>Seq (f\<^sub>1 v) (f\<^sub>2 ())\<close>\\
- @{thm (lhs) F_SEQ.simps(1)} & $\dn$ & \<open>Seq (f\<^sub>1 v\<^sub>1) (f\<^sub>2 v\<^sub>2)\<close>\medskip\\
- %\end{tabular}
- %
- %\begin{tabular}{lcl}
- @{term "simp_ALT (ZERO, DUMMY) (r\<^sub>2, f\<^sub>2)"} & $\dn$ & @{term "(r\<^sub>2, F_RIGHT f\<^sub>2)"}\\
- @{term "simp_ALT (r\<^sub>1, f\<^sub>1) (ZERO, DUMMY)"} & $\dn$ & @{term "(r\<^sub>1, F_LEFT f\<^sub>1)"}\\
- @{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)"}\\
- @{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)"}\\
- @{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)"}\\
- @{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)"}\\
- \end{tabular}
- \end{center}
-
- \noindent
- The functions \<open>simp\<^bsub>Alt\<^esub>\<close> and \<open>simp\<^bsub>Seq\<^esub>\<close> encode the simplification rules
- in \eqref{Simpl} and compose the rectification functions (simplifications can occur
- deep inside the regular expression). The main simplification function is then
-
- \begin{center}
- \begin{tabular}{lcl}
- @{term "simp (ALT r\<^sub>1 r\<^sub>2)"} & $\dn$ & @{term "simp_ALT (simp r\<^sub>1) (simp r\<^sub>2)"}\\
- @{term "simp (SEQ r\<^sub>1 r\<^sub>2)"} & $\dn$ & @{term "simp_SEQ (simp r\<^sub>1) (simp r\<^sub>2)"}\\
- @{term "simp r"} & $\dn$ & @{term "(r, id)"}\\
- \end{tabular}
- \end{center}
-
- \noindent where @{term "id"} stands for the identity function. The
- function @{const simp} returns a simplified regular expression and a corresponding
- rectification function. Note that we do not simplify under stars: this
- seems to slow down the algorithm, rather than speed it up. The optimised
- lexer is then given by the clauses:
-
- \begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) slexer.simps(1)} & $\dn$ & @{thm (rhs) slexer.simps(1)}\\
- @{thm (lhs) slexer.simps(2)} & $\dn$ &
- \<open>let (r\<^sub>s, f\<^sub>r) = simp (r \<close>$\backslash$\<open> c) in\<close>\\
- & & \<open>case\<close> @{term "slexer r\<^sub>s s"} \<open>of\<close>\\
- & & \phantom{$|$} @{term "None"} \<open>\<Rightarrow>\<close> @{term None}\\
- & & $|$ @{term "Some v"} \<open>\<Rightarrow>\<close> \<open>Some (inj r c (f\<^sub>r v))\<close>
- \end{tabular}
- \end{center}
-
- \noindent
- In the second clause we first calculate the derivative @{term "der c r"}
- and then simplify the result. This gives us a simplified derivative
- \<open>r\<^sub>s\<close> and a rectification function \<open>f\<^sub>r\<close>. The lexer
- is then recursively called with the simplified derivative, but before
- we inject the character @{term c} into the value @{term v}, we need to rectify
- @{term v} (that is construct @{term "f\<^sub>r v"}). Before we can establish the correctness
- of @{term "slexer"}, we need to show that simplification preserves the language
- and simplification preserves our POSIX relation once the value is rectified
- (recall @{const "simp"} generates a (regular expression, rectification function) pair):
-
- \begin{lemma}\mbox{}\smallskip\\\label{slexeraux}
- \begin{tabular}{ll}
- (1) & @{thm L_fst_simp[symmetric]}\\
- (2) & @{thm[mode=IfThen] Posix_simp}
- \end{tabular}
- \end{lemma}
-
- \begin{proof} Both are by induction on \<open>r\<close>. There is no
- interesting case for the first statement. For the second statement,
- of interest are the @{term "r = ALT r\<^sub>1 r\<^sub>2"} and @{term "r = SEQ r\<^sub>1
- r\<^sub>2"} cases. In each case we have to analyse four subcases whether
- @{term "fst (simp r\<^sub>1)"} and @{term "fst (simp r\<^sub>2)"} equals @{const
- ZERO} (respectively @{const ONE}). For example for @{term "r = ALT
- r\<^sub>1 r\<^sub>2"}, consider the subcase @{term "fst (simp r\<^sub>1) = ZERO"} and
- @{term "fst (simp r\<^sub>2) \<noteq> ZERO"}. By assumption we know @{term "s \<in>
- 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"}
- 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"}
- we know @{term "L (fst (simp r\<^sub>1)) = {}"}. By the first statement
- @{term "L r\<^sub>1"} is the empty set, meaning (**) @{term "s \<notin> L r\<^sub>1"}.
- Taking (*) and (**) together gives by the \mbox{\<open>P+R\<close>}-rule
- @{term "s \<in> ALT r\<^sub>1 r\<^sub>2 \<rightarrow> Right (snd (simp r\<^sub>2) v)"}. In turn this
- 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.
- The other cases are similar.\qed
- \end{proof}
-
- \noindent We can now prove relatively straightforwardly that the
- optimised lexer produces the expected result:
-
- \begin{theorem}
- @{thm slexer_correctness}
- \end{theorem}
-
- \begin{proof} By induction on @{term s} generalising over @{term
- r}. The case @{term "[]"} is trivial. For the cons-case suppose the
- string is of the form @{term "c # s"}. By induction hypothesis we
- know @{term "slexer r s = lexer r s"} holds for all @{term r} (in
- particular for @{term "r"} being the derivative @{term "der c
- r"}). Let @{term "r\<^sub>s"} be the simplified derivative regular expression, that is @{term
- "fst (simp (der c r))"}, and @{term "f\<^sub>r"} be the rectification
- function, that is @{term "snd (simp (der c r))"}. We distinguish the cases
- whether (*) @{term "s \<in> L (der c r)"} or not. In the first case we
- have by Theorem~\ref{lexercorrect}(2) a value @{term "v"} so that @{term
- "lexer (der c r) s = Some v"} and @{term "s \<in> der c r \<rightarrow> v"} hold.
- By Lemma~\ref{slexeraux}(1) we can also infer from~(*) that @{term "s
- \<in> L r\<^sub>s"} holds. Hence we know by Theorem~\ref{lexercorrect}(2) that
- there exists a @{term "v'"} with @{term "lexer r\<^sub>s s = Some v'"} and
- @{term "s \<in> r\<^sub>s \<rightarrow> v'"}. From the latter we know by
- Lemma~\ref{slexeraux}(2) that @{term "s \<in> der c r \<rightarrow> (f\<^sub>r v')"} holds.
- By the uniqueness of the POSIX relation (Theorem~\ref{posixdeterm}) we
- can infer that @{term v} is equal to @{term "f\<^sub>r v'"}---that is the
- rectification function applied to @{term "v'"}
- produces the original @{term "v"}. Now the case follows by the
- definitions of @{const lexer} and @{const slexer}.
-
- In the second case where @{term "s \<notin> L (der c r)"} we have that
- @{term "lexer (der c r) s = None"} by Theorem~\ref{lexercorrect}(1). We
- also know by Lemma~\ref{slexeraux}(1) that @{term "s \<notin> L r\<^sub>s"}. Hence
- @{term "lexer r\<^sub>s s = None"} by Theorem~\ref{lexercorrect}(1) and
- by IH then also @{term "slexer r\<^sub>s s = None"}. With this we can
- conclude in this case too.\qed
-
- \end{proof}
-
-\<close>
-fy the result. This gives us a simplified derivative
- \<open>r\<^sub>s\<close> and a rectification function \<open>f\<^sub>r\<close>. The lexer
- is then recursively called with the simplified derivative, but before
- we inject the character @{term c} into the value @{term v}, we need to rectify
- @{term v} (that is construct @{term "f\<^sub>r v"}). Before we can establish the correctness
- of @{term "slexer"}, we need to show that simplification preserves the language
- and simplification preserves our POSIX relation once the value is rectified
- (recall @{const "simp"} generates a (regular expression, rectification function) pair):
-
- \begin{lemma}\mbox{}\smallskip\\\label{slexeraux}
- \begin{tabular}{ll}
- (1) & @{thm L_fst_simp[symmetric]}\\
- (2) & @{thm[mode=IfThen] Posix_simp}
- \end{tabular}
- \end{lemma}
-
- \begin{proof} Both are by induction on \<open>r\<close>. There is no
- interesting case for the first statement. For the second statement,
- of interest are the @{term "r = ALT r\<^sub>1 r\<^sub>2"} and @{term "r = SEQ r\<^sub>1
- r\<^sub>2"} cases. In each case we have to analyse four subcases whether
- @{term "fst (simp r\<^sub>1)"} and @{term "fst (simp r\<^sub>2)"} equals @{const
- ZERO} (respectively @{const ONE}). For example for @{term "r = ALT
- r\<^sub>1 r\<^sub>2"}, consider the subcase @{term "fst (simp r\<^sub>1) = ZERO"} and
- @{term "fst (simp r\<^sub>2) \<noteq> ZERO"}. By assumption we know @{term "s \<in>
- 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"}
- 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"}
- we know @{term "L (fst (simp r\<^sub>1)) = {}"}. By the first statement
- @{term "L r\<^sub>1"} is the empty set, meaning (**) @{term "s \<notin> L r\<^sub>1"}.
- Taking (*) and (**) together gives by the \mbox{\<open>P+R\<close>}-rule
- @{term "s \<in> ALT r\<^sub>1 r\<^sub>2 \<rightarrow> Right (snd (simp r\<^sub>2) v)"}. In turn this
- 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.
- The other cases are similar.\qed
- \end{proof}
-
- \noindent We can now prove relatively straightforwardly that the
- optimised lexer produces the expected result:
-
- \begin{theorem}
- @{thm slexer_correctness}
- \end{theorem}
-
- \begin{proof} By induction on @{term s} generalising over @{term
- r}. The case @{term "[]"} is trivial. For the cons-case suppose the
- string is of the form @{term "c # s"}. By induction hypothesis we
- know @{term "slexer r s = lexer r s"} holds for all @{term r} (in
- particular for @{term "r"} being the derivative @{term "der c
- r"}). Let @{term "r\<^sub>s"} be the simplified derivative regular expression, that is @{term
- "fst (simp (der c r))"}, and @{term "f\<^sub>r"} be the rectification
- function, that is @{term "snd (simp (der c r))"}. We distinguish the cases
- whether (*) @{term "s \<in> L (der c r)"} or not. In the first case we
- have by Theorem~\ref{lexercorrect}(2) a value @{term "v"} so that @{term
- "lexer (der c r) s = Some v"} and @{term "s \<in> der c r \<rightarrow> v"} hold.
- By Lemma~\ref{slexeraux}(1) we can also infer from~(*) that @{term "s
- \<in> L r\<^sub>s"} holds. Hence we know by Theorem~\ref{lexercorrect}(2) that
- there exists a @{term "v'"} with @{term "lexer r\<^sub>s s = Some v'"} and
- @{term "s \<in> r\<^sub>s \<rightarrow> v'"}. From the latter we know by
- Lemma~\ref{slexeraux}(2) that @{term "s \<in> der c r \<rightarrow> (f\<^sub>r v')"} holds.
- By the uniqueness of the POSIX relation (Theorem~\ref{posixdeterm}) we
- can infer that @{term v} is equal to @{term "f\<^sub>r v'"}---that is the
- rectification function applied to @{term "v'"}
- produces the original @{term "v"}. Now the case follows by the
- definitions of @{const lexer} and @{const slexer}.
-
- In the second case where @{term "s \<notin> L (der c r)"} we have that
- @{term "lexer (der c r) s = None"} by Theorem~\ref{lexercorrect}(1). We
- also know by Lemma~\ref{slexeraux}(1) that @{term "s \<notin> L r\<^sub>s"}. Hence
- @{term "lexer r\<^sub>s s = None"} by Theorem~\ref{lexercorrect}(1) and
- by IH then also @{term "slexer r\<^sub>s s = None"}. With this we can
- conclude in this case too.\qed
-
- \end{proof}
-
-\<close>
-
-
-section \<open>HERE\<close>
-
-text \<open>
-
- \begin{lemma}
- @{thm [mode=IfThen] bder_retrieve}
- \end{lemma}
-
- \begin{proof}
- By induction on the definition of @{term "erase r"}. The cases for rule 1) and 2) are
- straightforward as @{term "der c ZERO"} and @{term "der c ONE"} are both equal to
- @{term ZERO}. This means @{term "\<Turnstile> v : ZERO"} cannot hold. Similarly in case of rule 3)
- where @{term r} is of the form @{term "ACHAR d"} with @{term "c = d"}. Then by assumption
- we know @{term "\<Turnstile> v : ONE"}, which implies @{term "v = Void"}. The equation follows by
- simplification of left- and right-hand side. In case @{term "c \<noteq> d"} we have again
- @{term "\<Turnstile> v : ZERO"}, which cannot hold.
-
- For rule 4a) we have again @{term "\<Turnstile> v : ZERO"}. The property holds by IH for rule 4b).
- The induction hypothesis is
- \[
- @{term "retrieve (bder c r) v = retrieve r (injval (erase r) c v)"}
- \]
- which is what left- and right-hand side simplify to. The slightly more interesting case
- is for 4c). By assumption we have
- @{term "\<Turnstile> v : ALT (der c (erase r\<^sub>1)) (der c (erase (AALTs bs (r\<^sub>2 # rs))))"}. This means we
- have either (*) @{term "\<Turnstile> v1 : der c (erase r\<^sub>1)"} with @{term "v = Left v1"} or
- (**) @{term "\<Turnstile> v2 : der c (erase (AALTs bs (r\<^sub>2 # rs)))"} with @{term "v = Right v2"}.
- The former case is straightforward by simplification. The second case is \ldots TBD.
-
- Rule 5) TBD.
-
- Finally for rule 6) the reasoning is as follows: By assumption we have
- @{term "\<Turnstile> v : SEQ (der c (erase r)) (STAR (erase r))"}. This means we also have
- @{term "v = Seq v1 v2"}, @{term "\<Turnstile> v1 : der c (erase r)"} and @{term "v2 = Stars vs"}.
- We want to prove
- \begin{align}
- & @{term "retrieve (ASEQ bs (fuse [Z] (bder c r)) (ASTAR [] r)) v"}\\
- &= @{term "retrieve (ASTAR bs r) (injval (STAR (erase r)) c v)"}
- \end{align}
- The right-hand side @{term inj}-expression is equal to
- @{term "Stars (injval (erase r) c v1 # vs)"}, which means the @{term retrieve}-expression
- simplifies to
- \[
- @{term "bs @ [Z] @ retrieve r (injval (erase r) c v1) @ retrieve (ASTAR [] r) (Stars vs)"}
- \]
- The left-hand side (3) above simplifies to
- \[
- @{term "bs @ retrieve (fuse [Z] (bder c r)) v1 @ retrieve (ASTAR [] r) (Stars vs)"}
- \]
- We can move out the @{term "fuse [Z]"} and then use the IH to show that left-hand side
- and right-hand side are equal. This completes the proof.
- \end{proof}
-
-
-
- \bibliographystyle{plain}
- \bibliography{root}
-
-\<close>
-(*<*)
-end
-(*>*)
-
-(*
-
-\begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) code.simps(1)} & $\dn$ & @{thm (rhs) code.simps(1)}\\
- @{thm (lhs) code.simps(2)} & $\dn$ & @{thm (rhs) code.simps(2)}\\
- @{thm (lhs) code.simps(3)} & $\dn$ & @{thm (rhs) code.simps(3)}\\
- @{thm (lhs) code.simps(4)} & $\dn$ & @{thm (rhs) code.simps(4)}\\
- @{thm (lhs) code.simps(5)[of "v\<^sub>1" "v\<^sub>2"]} & $\dn$ & @{thm (rhs) code.simps(5)[of "v\<^sub>1" "v\<^sub>2"]}\\
- @{thm (lhs) code.simps(6)} & $\dn$ & @{thm (rhs) code.simps(6)}\\
- @{thm (lhs) code.simps(7)} & $\dn$ & @{thm (rhs) code.simps(7)}
-\end{tabular}
-\end{center}
-
-\begin{center}
-\begin{tabular}{lcl}
- @{term areg} & $::=$ & @{term "AZERO"}\\
- & $\mid$ & @{term "AONE bs"}\\
- & $\mid$ & @{term "ACHAR bs c"}\\
- & $\mid$ & @{term "AALT bs r1 r2"}\\
- & $\mid$ & @{term "ASEQ bs r\<^sub>1 r\<^sub>2"}\\
- & $\mid$ & @{term "ASTAR bs r"}
-\end{tabular}
-\end{center}
-
-
-\begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) intern.simps(1)} & $\dn$ & @{thm (rhs) intern.simps(1)}\\
- @{thm (lhs) intern.simps(2)} & $\dn$ & @{thm (rhs) intern.simps(2)}\\
- @{thm (lhs) intern.simps(3)} & $\dn$ & @{thm (rhs) intern.simps(3)}\\
- @{thm (lhs) intern.simps(4)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) intern.simps(4)[of "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) intern.simps(5)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) intern.simps(5)[of "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) intern.simps(6)} & $\dn$ & @{thm (rhs) intern.simps(6)}\\
-\end{tabular}
-\end{center}
-
-\begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) erase.simps(1)} & $\dn$ & @{thm (rhs) erase.simps(1)}\\
- @{thm (lhs) erase.simps(2)[of bs]} & $\dn$ & @{thm (rhs) erase.simps(2)[of bs]}\\
- @{thm (lhs) erase.simps(3)[of bs]} & $\dn$ & @{thm (rhs) erase.simps(3)[of bs]}\\
- @{thm (lhs) erase.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) erase.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) erase.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) erase.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) erase.simps(6)[of bs]} & $\dn$ & @{thm (rhs) erase.simps(6)[of bs]}\\
-\end{tabular}
-\end{center}
-
-Some simple facts about erase
-
-\begin{lemma}\mbox{}\\
-@{thm erase_bder}\\
-@{thm erase_intern}
-\end{lemma}
-
-\begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) bnullable.simps(1)} & $\dn$ & @{thm (rhs) bnullable.simps(1)}\\
- @{thm (lhs) bnullable.simps(2)} & $\dn$ & @{thm (rhs) bnullable.simps(2)}\\
- @{thm (lhs) bnullable.simps(3)} & $\dn$ & @{thm (rhs) bnullable.simps(3)}\\
- @{thm (lhs) bnullable.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bnullable.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) bnullable.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bnullable.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) bnullable.simps(6)} & $\dn$ & @{thm (rhs) bnullable.simps(6)}\medskip\\
-
-% \end{tabular}
-% \end{center}
-
-% \begin{center}
-% \begin{tabular}{lcl}
-
- @{thm (lhs) bder.simps(1)} & $\dn$ & @{thm (rhs) bder.simps(1)}\\
- @{thm (lhs) bder.simps(2)} & $\dn$ & @{thm (rhs) bder.simps(2)}\\
- @{thm (lhs) bder.simps(3)} & $\dn$ & @{thm (rhs) bder.simps(3)}\\
- @{thm (lhs) bder.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bder.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) bder.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bder.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) bder.simps(6)} & $\dn$ & @{thm (rhs) bder.simps(6)}
- \end{tabular}
- \end{center}
-
-
-\begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) bmkeps.simps(1)} & $\dn$ & @{thm (rhs) bmkeps.simps(1)}\\
- @{thm (lhs) bmkeps.simps(2)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bmkeps.simps(2)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) bmkeps.simps(3)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bmkeps.simps(3)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) bmkeps.simps(4)} & $\dn$ & @{thm (rhs) bmkeps.simps(4)}\medskip\\
-\end{tabular}
-\end{center}
-
-
-@{thm [mode=IfThen] bder_retrieve}
-
-By induction on \<open>r\<close>
-
-\begin{theorem}[Main Lemma]\mbox{}\\
-@{thm [mode=IfThen] MAIN_decode}
-\end{theorem}
-
-\noindent
-Definition of the bitcoded lexer
-
-@{thm blexer_def}
-
-
-\begin{theorem}
-@{thm blexer_correctness}
-\end{theorem}
-
-
-
-
-
-\begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) code.simps(1)} & $\dn$ & @{thm (rhs) code.simps(1)}\\
- @{thm (lhs) code.simps(2)} & $\dn$ & @{thm (rhs) code.simps(2)}\\
- @{thm (lhs) code.simps(3)} & $\dn$ & @{thm (rhs) code.simps(3)}\\
- @{thm (lhs) code.simps(4)} & $\dn$ & @{thm (rhs) code.simps(4)}\\
- @{thm (lhs) code.simps(5)[of "v\<^sub>1" "v\<^sub>2"]} & $\dn$ & @{thm (rhs) code.simps(5)[of "v\<^sub>1" "v\<^sub>2"]}\\
- @{thm (lhs) code.simps(6)} & $\dn$ & @{thm (rhs) code.simps(6)}\\
- @{thm (lhs) code.simps(7)} & $\dn$ & @{thm (rhs) code.simps(7)}
-\end{tabular}
-\end{center}
-
-\begin{center}
-\begin{tabular}{lcl}
- @{term areg} & $::=$ & @{term "AZERO"}\\
- & $\mid$ & @{term "AONE bs"}\\
- & $\mid$ & @{term "ACHAR bs c"}\\
- & $\mid$ & @{term "AALT bs r1 r2"}\\
- & $\mid$ & @{term "ASEQ bs r\<^sub>1 r\<^sub>2"}\\
- & $\mid$ & @{term "ASTAR bs r"}
-\end{tabular}
-\end{center}
-
-
-\begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) intern.simps(1)} & $\dn$ & @{thm (rhs) intern.simps(1)}\\
- @{thm (lhs) intern.simps(2)} & $\dn$ & @{thm (rhs) intern.simps(2)}\\
- @{thm (lhs) intern.simps(3)} & $\dn$ & @{thm (rhs) intern.simps(3)}\\
- @{thm (lhs) intern.simps(4)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) intern.simps(4)[of "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) intern.simps(5)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) intern.simps(5)[of "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) intern.simps(6)} & $\dn$ & @{thm (rhs) intern.simps(6)}\\
-\end{tabular}
-\end{center}
-
-\begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) erase.simps(1)} & $\dn$ & @{thm (rhs) erase.simps(1)}\\
- @{thm (lhs) erase.simps(2)[of bs]} & $\dn$ & @{thm (rhs) erase.simps(2)[of bs]}\\
- @{thm (lhs) erase.simps(3)[of bs]} & $\dn$ & @{thm (rhs) erase.simps(3)[of bs]}\\
- @{thm (lhs) erase.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) erase.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) erase.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) erase.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) erase.simps(6)[of bs]} & $\dn$ & @{thm (rhs) erase.simps(6)[of bs]}\\
-\end{tabular}
-\end{center}
-
-Some simple facts about erase
-
-\begin{lemma}\mbox{}\\
-@{thm erase_bder}\\
-@{thm erase_intern}
-\end{lemma}
-
-\begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) bnullable.simps(1)} & $\dn$ & @{thm (rhs) bnullable.simps(1)}\\
- @{thm (lhs) bnullable.simps(2)} & $\dn$ & @{thm (rhs) bnullable.simps(2)}\\
- @{thm (lhs) bnullable.simps(3)} & $\dn$ & @{thm (rhs) bnullable.simps(3)}\\
- @{thm (lhs) bnullable.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bnullable.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) bnullable.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bnullable.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) bnullable.simps(6)} & $\dn$ & @{thm (rhs) bnullable.simps(6)}\medskip\\
-
-% \end{tabular}
-% \end{center}
-
-% \begin{center}
-% \begin{tabular}{lcl}
-
- @{thm (lhs) bder.simps(1)} & $\dn$ & @{thm (rhs) bder.simps(1)}\\
- @{thm (lhs) bder.simps(2)} & $\dn$ & @{thm (rhs) bder.simps(2)}\\
- @{thm (lhs) bder.simps(3)} & $\dn$ & @{thm (rhs) bder.simps(3)}\\
- @{thm (lhs) bder.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bder.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) bder.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bder.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) bder.simps(6)} & $\dn$ & @{thm (rhs) bder.simps(6)}
- \end{tabular}
- \end{center}
-
-
-\begin{center}
- \begin{tabular}{lcl}
- @{thm (lhs) bmkeps.simps(1)} & $\dn$ & @{thm (rhs) bmkeps.simps(1)}\\
- @{thm (lhs) bmkeps.simps(2)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bmkeps.simps(2)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) bmkeps.simps(3)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bmkeps.simps(3)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
- @{thm (lhs) bmkeps.simps(4)} & $\dn$ & @{thm (rhs) bmkeps.simps(4)}\medskip\\
-\end{tabular}
-\end{center}
-
-
-@{thm [mode=IfThen] bder_retrieve}
-
-By induction on \<open>r\<close>
-
-\begin{theorem}[Main Lemma]\mbox{}\\
-@{thm [mode=IfThen] MAIN_decode}
-\end{theorem}
-
-\noindent
-Definition of the bitcoded lexer
-
-@{thm blexer_def}
-
-
-\begin{theorem}
-@{thm blexer_correctness}
-\end{theorem}
-
-\<close>
-\<close>*)
\ No newline at end of file
--- a/thys2/Journal/isabelle.sty Fri Jan 07 22:25:26 2022 +0000
+++ b/thys2/Journal/isabelle.sty Fri Jan 07 22:28:23 2022 +0000
@@ -39,6 +39,36 @@
\DeclareRobustCommand{\isactrlesup}{\egroup\egroup\endmath\egroup}
\newcommand{\isactrlbold}[1]{{\bfseries\upshape\boldmath#1}}
+%blackboard-bold (requires font txmia from pxfonts)
+\DeclareSymbolFont{bbbfont}{U}{txmia}{m}{it}
+\SetSymbolFont{bbbfont}{bold}{U}{txmia}{bx}{it}
+\DeclareMathSymbol{\bbbA}{\mathord}{bbbfont}{129}
+\DeclareMathSymbol{\bbbB}{\mathord}{bbbfont}{130}
+\DeclareMathSymbol{\bbbC}{\mathord}{bbbfont}{131}
+\DeclareMathSymbol{\bbbD}{\mathord}{bbbfont}{132}
+\DeclareMathSymbol{\bbbE}{\mathord}{bbbfont}{133}
+\DeclareMathSymbol{\bbbF}{\mathord}{bbbfont}{134}
+\DeclareMathSymbol{\bbbG}{\mathord}{bbbfont}{135}
+\DeclareMathSymbol{\bbbH}{\mathord}{bbbfont}{136}
+\DeclareMathSymbol{\bbbI}{\mathord}{bbbfont}{137}
+\DeclareMathSymbol{\bbbJ}{\mathord}{bbbfont}{138}
+\DeclareMathSymbol{\bbbK}{\mathord}{bbbfont}{139}
+\DeclareMathSymbol{\bbbL}{\mathord}{bbbfont}{140}
+\DeclareMathSymbol{\bbbM}{\mathord}{bbbfont}{141}
+\DeclareMathSymbol{\bbbN}{\mathord}{bbbfont}{142}
+\DeclareMathSymbol{\bbbO}{\mathord}{bbbfont}{143}
+\DeclareMathSymbol{\bbbP}{\mathord}{bbbfont}{144}
+\DeclareMathSymbol{\bbbQ}{\mathord}{bbbfont}{145}
+\DeclareMathSymbol{\bbbR}{\mathord}{bbbfont}{146}
+\DeclareMathSymbol{\bbbS}{\mathord}{bbbfont}{147}
+\DeclareMathSymbol{\bbbT}{\mathord}{bbbfont}{148}
+\DeclareMathSymbol{\bbbU}{\mathord}{bbbfont}{149}
+\DeclareMathSymbol{\bbbV}{\mathord}{bbbfont}{150}
+\DeclareMathSymbol{\bbbW}{\mathord}{bbbfont}{151}
+\DeclareMathSymbol{\bbbX}{\mathord}{bbbfont}{152}
+\DeclareMathSymbol{\bbbY}{\mathord}{bbbfont}{153}
+\DeclareMathSymbol{\bbbZ}{\mathord}{bbbfont}{154}
+
\newenvironment{isaantiq}{{\isacharat\isacharbraceleft}}{{\isacharbraceright}}
\newdimen\isa@parindent\newdimen\isa@parskip
@@ -112,8 +142,8 @@
\chardef\isachartilde=`\~%
\def\isacharverbatimopen{\isacharbraceleft\isacharasterisk}%
\def\isacharverbatimclose{\isacharasterisk\isacharbraceright}%
-\def\isacartoucheopen{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\langle$}}}%
-\def\isacartoucheclose{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\rangle$}}}%
+\def\isacartoucheopen{\isatext{\guilsinglleft}}%
+\def\isacartoucheclose{\isatext{\guilsinglright}}%
}
@@ -143,6 +173,12 @@
\newcommand{\isamarkupcmt}[1]{{\isastylecmt--- #1}}
+% index entries
+
+\newcommand{\isaindexdef}[1]{\textbf{#1}}
+\newcommand{\isaindexref}[1]{#1}
+
+
% styles
\def\isabellestyle#1{\csname isabellestyle#1\endcsname}
@@ -205,8 +241,8 @@
\def\isacharbar{\isamath{\mid}}%
\def\isacharbraceright{\isamath{\}}}%
\def\isachartilde{\isamath{{}\sp{\sim}}}%
-\def\isacharbackquoteopen{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\langle$}}}%
-\def\isacharbackquoteclose{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\rangle$}}}%
+\def\isacharbackquoteopen{\isatext{\guilsinglleft}}%
+\def\isacharbackquoteclose{\isatext{\guilsinglright}}%
}
\newcommand{\isabellestyleliteral}{%
@@ -239,29 +275,8 @@
\newcommand{\isamarkupcancel}[1]{\isa{\xout{#1}}}
-% tagged regions
-
-%plain TeX version of comment package -- much faster!
-\let\isafmtname\fmtname\def\fmtname{plain}
-\usepackage{comment}
-\let\fmtname\isafmtname
+% tags
\newcommand{\isafold}[1]{\emph{$\langle\mathord{\mathit{#1}}\rangle$}}
-\newcommand{\isakeeptag}[1]%
-{\includecomment{isadelim#1}\includecomment{isatag#1}\csarg\def{isafold#1}{}}
-\newcommand{\isadroptag}[1]%
-{\excludecomment{isadelim#1}\excludecomment{isatag#1}\csarg\def{isafold#1}{}}
-\newcommand{\isafoldtag}[1]%
-{\includecomment{isadelim#1}\excludecomment{isatag#1}\csarg\def{isafold#1}{\isafold{#1}}}
-
-\isakeeptag{document}
-\isakeeptag{theory}
-\isakeeptag{proof}
-\isakeeptag{ML}
-\isakeeptag{visible}
-\isadroptag{invisible}
-\isakeeptag{important}
-\isakeeptag{unimportant}
-
\IfFileExists{isabelletags.sty}{\usepackage{isabelletags}}{}
--- a/thys2/Journal/isabellesym.sty Fri Jan 07 22:25:26 2022 +0000
+++ b/thys2/Journal/isabellesym.sty Fri Jan 07 22:28:23 2022 +0000
@@ -150,12 +150,32 @@
\newcommand{\isasymPhi}{\isamath{\Phi}}
\newcommand{\isasymPsi}{\isamath{\Psi}}
\newcommand{\isasymOmega}{\isamath{\Omega}}
-\newcommand{\isasymbool}{\isamath{\mathrm{I}\mkern-3.8mu\mathrm{B}}}
-\newcommand{\isasymcomplex}{\isamath{\mathrm{C}\mkern-15mu{\phantom{\mathrm{t}}\vrule}\mkern9mu}}
-\newcommand{\isasymnat}{\isamath{\mathrm{I}\mkern-3.8mu\mathrm{N}}}
-\newcommand{\isasymrat}{\isamath{\mathrm{Q}\mkern-16mu{\phantom{\mathrm{t}}\vrule}\mkern10mu}}
-\newcommand{\isasymreal}{\isamath{\mathrm{I}\mkern-3.8mu\mathrm{R}}}
-\newcommand{\isasymint}{\isamath{\mathsf{Z}\mkern-7.5mu\mathsf{Z}}}
+\newcommand{\isasymbbbA}{\isamath{\bbbA}} %requires font txmia from txfonts
+\newcommand{\isasymbool}{\isamath{\bbbB}} %requires font txmia from txfonts
+\newcommand{\isasymcomplex}{\isamath{\bbbC}} %requires font txmia from txfonts
+\newcommand{\isasymbbbD}{\isamath{\bbbD}} %requires font txmia from txfonts
+\newcommand{\isasymbbbE}{\isamath{\bbbE}} %requires font txmia from txfonts
+\newcommand{\isasymbbbF}{\isamath{\bbbF}} %requires font txmia from txfonts
+\newcommand{\isasymbbbG}{\isamath{\bbbG}} %requires font txmia from txfonts
+\newcommand{\isasymbbbH}{\isamath{\bbbH}} %requires font txmia from txfonts
+\newcommand{\isasymbbbI}{\isamath{\bbbI}} %requires font txmia from txfonts
+\newcommand{\isasymbbbJ}{\isamath{\bbbJ}} %requires font txmia from txfonts
+\newcommand{\isasymbbbK}{\isamath{\bbbK}} %requires font txmia from txfonts
+\newcommand{\isasymbbbL}{\isamath{\bbbL}} %requires font txmia from txfonts
+\newcommand{\isasymbbbM}{\isamath{\bbbM}} %requires font txmia from txfonts
+\newcommand{\isasymnat}{\isamath{\bbbN}} %requires font txmia from txfonts
+\newcommand{\isasymbbbO}{\isamath{\bbbO}} %requires font txmia from txfonts
+\newcommand{\isasymbbbP}{\isamath{\bbbP}} %requires font txmia from txfonts
+\newcommand{\isasymrat}{\isamath{\bbbQ}} %requires font txmia from txfonts
+\newcommand{\isasymreal}{\isamath{\bbbR}} %requires font txmia from txfonts
+\newcommand{\isasymbbbS}{\isamath{\bbbS}} %requires font txmia from txfonts
+\newcommand{\isasymbbbT}{\isamath{\bbbT}} %requires font txmia from txfonts
+\newcommand{\isasymbbbU}{\isamath{\bbbU}} %requires font txmia from txfonts
+\newcommand{\isasymbbbV}{\isamath{\bbbV}} %requires font txmia from txfonts
+\newcommand{\isasymbbbW}{\isamath{\bbbW}} %requires font txmia from txfonts
+\newcommand{\isasymbbbX}{\isamath{\bbbX}} %requires font txmia from txfonts
+\newcommand{\isasymbbbY}{\isamath{\bbbY}} %requires font txmia from txfonts
+\newcommand{\isasymint}{\isamath{\bbbZ}} %requires font txmia from txfonts
\newcommand{\isasymleftarrow}{\isamath{\leftarrow}}
\newcommand{\isasymrightarrow}{\isamath{\rightarrow}}
\newcommand{\isasymlongleftarrow}{\isamath{\longleftarrow}}
@@ -206,14 +226,16 @@
\newcommand{\isasymrceil}{\isamath{\rceil}}
\newcommand{\isasymlfloor}{\isamath{\lfloor}}
\newcommand{\isasymrfloor}{\isamath{\rfloor}}
-\newcommand{\isasymlparr}{\isamath{\mathopen{(\mkern-3mu\mid}}}
-\newcommand{\isasymrparr}{\isamath{\mathclose{\mid\mkern-3mu)}}}
+\newcommand{\isasymlparr}{\isamath{\mathopen{(\mkern-3.3mu\mid}}}
+\newcommand{\isasymrparr}{\isamath{\mathclose{\mid\mkern-3.3mu)}}}
\newcommand{\isasymlbrakk}{\isamath{\mathopen{\lbrack\mkern-3mu\lbrack}}}
\newcommand{\isasymrbrakk}{\isamath{\mathclose{\rbrack\mkern-3mu\rbrack}}}
-\newcommand{\isasymlbrace}{\isamath{\mathopen{\lbrace\mkern-4.5mu\mid}}}
-\newcommand{\isasymrbrace}{\isamath{\mathclose{\mid\mkern-4.5mu\rbrace}}}
-\newcommand{\isasymguillemotleft}{\isatext{\flqq}} %requires babel
-\newcommand{\isasymguillemotright}{\isatext{\frqq}} %requires babel
+\newcommand{\isasymlbrace}{\isamath{\mathopen{\lbrace\mkern-4.3mu\mid}}}
+\newcommand{\isasymrbrace}{\isamath{\mathclose{\mid\mkern-4.3mu\rbrace}}}
+\newcommand{\isasymlblot}{\isamath{{\langle}\mkern -3.5mu{|}}}
+\newcommand{\isasymrblot}{\isamath{{|}\mkern -3.5mu{\rangle}}}
+\newcommand{\isasymguillemotleft}{\isatext{\guillemotleft}}
+\newcommand{\isasymguillemotright}{\isatext{\guillemotright}}
\newcommand{\isasymbottom}{\isamath{\bot}}
\newcommand{\isasymtop}{\isamath{\top}}
\newcommand{\isasymand}{\isamath{\wedge}}
@@ -281,6 +303,9 @@
\newcommand{\isasympreceq}{\isamath{\preceq}}
\newcommand{\isasymsucceq}{\isamath{\succeq}}
\newcommand{\isasymparallel}{\isamath{\parallel}}
+\newcommand{\isasymParallel}{\isamath{\bigparallel}} %requires stmaryrd
+\newcommand{\isasyminterleace}{\isamath{\interleave}} %requires stmaryrd
+\newcommand{\isasymsslash}{\isamath{\sslash}} %requires stmaryrd
\newcommand{\isasymbar}{\isamath{\mid}}
\newcommand{\isasymbbar}{\isamath{[\mskip-1.5mu]}}
\newcommand{\isasymplusminus}{\isamath{\pm}}
@@ -364,14 +389,44 @@
\newcommand{\isasymsome}{\isamath{\epsilon\,}}
\newcommand{\isasymbind}{\isamath{\mathbin{>\!\!\!>\mkern-6.7mu=}}}
\newcommand{\isasymthen}{\isamath{\mathbin{>\!\!\!>}}}
-\newcommand{\isasymopen}{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\langle$}}}
-\newcommand{\isasymclose}{\isatext{\raise.3ex\hbox{$\scriptscriptstyle\rangle$}}}
+
+%Z notation
+\newcommand{\isaZhbar}[1]{\rlap{\raise.0001ex\hbox{\isamath{-}}}#1}
+\newcommand{\isaZpvbar}[1]{\ooalign{\hfil\isamath{\mapstochar\mkern 5mu}\hfil\cr#1}}
+\newcommand{\isaZfvbar}[1]{\ooalign{\hfil\isamath{\mapstochar\mkern 3mu\mapstochar\mkern 5mu}\hfil\cr#1}}
+\newcommand{\isaZdarrow}[3]{\ooalign{\isamath{#1}\hfil\cr\isamath{\mkern#3mu\isamath{#2}}}}
+\newcommand{\isasymZcomp}{\isamath{\fatsemi}} %requires stmaryrd
+\newcommand{\isasymZinj}{\isamath{\rightarrowtail}} %requires amssymb
+\newcommand{\isasymZpinj}{\isaZpvbar{\isamath{\rightarrowtail}}} %requires amssymb
+\newcommand{\isasymZfinj}{\isaZfvbar{\isasymZinj}} %requires amssymb
+\newcommand{\isasymZsurj}{\isaZdarrow{\rightarrow}{\rightarrow}{4}} %requires amssymb
+\newcommand{\isasymZpsurj}{\isaZpvbar{\isasymZsurj}} %requires amssymb
+\newcommand{\isasymZbij}{\isaZdarrow{\rightarrowtail}{\rightarrow}{5}} %requires amssymb
+\newcommand{\isasymZpfun}{\isaZpvbar{\isamath{\rightarrow}}}
+\newcommand{\isasymZffun}{\isaZfvbar{\isamath{\rightarrow}}}
+\newcommand{\isasymZdres}{\isamath{\lhd}} %requires amssymb
+\newcommand{\isasymZndres}{\isaZhbar{\isamath{\lhd}}} %requires amssymb
+\newcommand{\isasymZrres}{\isamath{\rhd}} %requires amssymb
+\newcommand{\isasymZnrres}{\isaZhbar{\isamath{\rhd}}} %requires amssymb
+\newcommand{\isasymZspot}{\isamath{\bullet}}
+\newcommand{\isasymZproject}{\isamath{\upharpoonright}} %requires amssymb
+\newcommand{\isasymZsemi}{\isatext{\raise 0.66ex\hbox{\oalign{\hfil\isamath{\scriptscriptstyle\mathrm{o}}\hfil\cr\hfil\isamath{\scriptscriptstyle\mathrm{9}}\hfil}}}}
+\newcommand{\isasymZtypecolon}{\isatext{\raise 0.6ex\hbox{\oalign{\hfil\isamath{\scriptscriptstyle\mathrm{o}}\hfil\cr\hfil\isamath{\scriptscriptstyle\mathrm{o}}\hfil}}}}
+\newcommand{\isasymZhide}{\isamath{\backslash}}
+\newcommand{\isasymZcat}{\isatext{\raise 0.8ex\hbox{\isamath{\mathchar\frown}}}}
+\newcommand{\isasymZinbag}{\isatext{\ooalign{\isamath{\sqsubset\mkern-1mu}\cr\isamath{-\mkern-1mu}\cr}}}
+
\newcommand{\isasymhole}{\isatext{\normalfont\rmfamily\wasylozenge}} %requires wasysym
\newcommand{\isasymnewline}{\isatext{\fbox{$\hookleftarrow$}}}
\newcommand{\isasymcomment}{\isatext{\isastylecmt---}}
\newcommand{\isasymproof}{\isamath{\,\langle\mathit{proof}\rangle}}
-
+\newcommand{\isasymopen}{\isatext{\guilsinglleft}}
+\newcommand{\isasymclose}{\isatext{\guilsinglright}}
+\newcommand{\isasymcheckmark}{\isatext{\ding{51}}} %requires pifont
+\newcommand{\isasymcrossmark}{\isatext{\ding{55}}} %requires pifont
\newcommand{\isactrlmarker}{\isatext{\ding{48}}} %requires pifont
+\newcommand{\isactrltry}{\isakeywordcontrol{try}}
+\newcommand{\isactrlcan}{\isakeywordcontrol{can}}
\newcommand{\isactrlassert}{\isakeywordcontrol{assert}}
\newcommand{\isactrlcancel}{\isakeywordcontrol{cancel}}
\newcommand{\isactrlbinding}{\isakeywordcontrol{binding}}
@@ -389,9 +444,12 @@
\newcommand{\isactrldir}{\isakeywordcontrol{dir}}
\newcommand{\isactrlfile}{\isakeywordcontrol{file}}
\newcommand{\isactrlhere}{\isakeywordcontrol{here}}
+\newcommand{\isactrlinstantiate}{\isakeywordcontrol{instantiate}}
\newcommand{\isactrlkeyword}{\isakeywordcontrol{keyword}}
\newcommand{\isactrllatex}{\isakeywordcontrol{latex}}
\newcommand{\isactrllocale}{\isakeywordcontrol{locale}}
+\newcommand{\isactrlmakeUNDERSCOREjudgment}{\isakeywordcontrol{make{\isacharunderscore}judgment}}
+\newcommand{\isactrldestUNDERSCOREjudgment}{\isakeywordcontrol{dest{\isacharunderscore}judgment}}
\newcommand{\isactrlmakeUNDERSCOREstring}{\isakeywordcontrol{make{\isacharunderscore}string}}
\newcommand{\isactrlmasterUNDERSCOREdir}{\isakeywordcontrol{master{\isacharunderscore}dir}}
\newcommand{\isactrlmethod}{\isakeywordcontrol{method}}
@@ -420,8 +478,19 @@
\newcommand{\isactrltypeUNDERSCOREname}{\isakeywordcontrol{type{\isacharunderscore}name}}
\newcommand{\isactrltypeUNDERSCOREsyntax}{\isakeywordcontrol{type{\isacharunderscore}syntax}}
\newcommand{\isactrlundefined}{\isakeywordcontrol{undefined}}
+\newcommand{\isactrltvar}{\isakeywordcontrol{tvar}}
+\newcommand{\isactrlvar}{\isakeywordcontrol{var}}
+\newcommand{\isactrlConst}{\isakeywordcontrol{Const}}
+\newcommand{\isactrlConstUNDERSCORE}{\isakeywordcontrol{Const{\isacharunderscore}}}
+\newcommand{\isactrlConstUNDERSCOREfn}{\isakeywordcontrol{Const{\isacharunderscore}fn}}
+\newcommand{\isactrlType}{\isakeywordcontrol{Type}}
+\newcommand{\isactrlTypeUNDERSCOREfn}{\isakeywordcontrol{Type{\isacharunderscore}fn}}
\newcommand{\isactrlcode}{\isakeywordcontrol{code}}
\newcommand{\isactrlcomputation}{\isakeywordcontrol{computation}}
\newcommand{\isactrlcomputationUNDERSCOREconv}{\isakeywordcontrol{computation{\isacharunderscore}conv}}
\newcommand{\isactrlcomputationUNDERSCOREcheck}{\isakeywordcontrol{computation{\isacharunderscore}check}}
+\newcommand{\isactrlifUNDERSCORElinux}{\isakeywordcontrol{if{\isacharunderscore}linux}}
+\newcommand{\isactrlifUNDERSCOREmacos}{\isakeywordcontrol{if{\isacharunderscore}macos}}
+\newcommand{\isactrlifUNDERSCOREwindows}{\isakeywordcontrol{if{\isacharunderscore}windows}}
+\newcommand{\isactrlifUNDERSCOREunix}{\isakeywordcontrol{if{\isacharunderscore}unix}}
--- a/thys2/Journal/isabelletags.sty Fri Jan 07 22:25:26 2022 +0000
+++ b/thys2/Journal/isabelletags.sty Fri Jan 07 22:28:23 2022 +0000
@@ -1,1 +1,20 @@
+%plain TeX version of comment package -- much faster!
+\let\isafmtname\fmtname\def\fmtname{plain}
+\usepackage{comment}
+\let\fmtname\isafmtname
+\newcommand{\isakeeptag}[1]%
+{\includecomment{isadelim#1}\includecomment{isatag#1}\csarg\def{isafold#1}{}}
+\newcommand{\isadroptag}[1]%
+{\excludecomment{isadelim#1}\excludecomment{isatag#1}\csarg\def{isafold#1}{}}
+\newcommand{\isafoldtag}[1]%
+{\includecomment{isadelim#1}\excludecomment{isatag#1}\csarg\def{isafold#1}{\isafold{#1}}}
+
+\isakeeptag{ML}
+\isakeeptag{document}
+\isakeeptag{important}
+\isadroptag{invisible}
+\isakeeptag{proof}
+\isakeeptag{theory}
+\isakeeptag{unimportant}
+\isakeeptag{visible}
--- a/thys2/Journal/root.aux Fri Jan 07 22:25:26 2022 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,96 +0,0 @@
-\relax
-\providecommand\hyper@newdestlabel[2]{}
-\providecommand\HyperFirstAtBeginDocument{\AtBeginDocument}
-\HyperFirstAtBeginDocument{\ifx\hyper@anchor\@undefined
-\global\let\oldcontentsline\contentsline
-\gdef\contentsline#1#2#3#4{\oldcontentsline{#1}{#2}{#3}}
-\global\let\oldnewlabel\newlabel
-\gdef\newlabel#1#2{\newlabelxx{#1}#2}
-\gdef\newlabelxx#1#2#3#4#5#6{\oldnewlabel{#1}{{#2}{#3}}}
-\AtEndDocument{\ifx\hyper@anchor\@undefined
-\let\contentsline\oldcontentsline
-\let\newlabel\oldnewlabel
-\fi}
-\fi}
-\global\let\hyper@last\relax
-\gdef\HyperFirstAtBeginDocument#1{#1}
-\providecommand\HyField@AuxAddToFields[1]{}
-\providecommand\HyField@AuxAddToCoFields[2]{}
-\citation{AusafDyckhoffUrban2016}
-\citation{OkuiSuzuki2010,OkuiSuzukiTech}
-\citation{Sulzmann2014}
-\@writefile{toc}{\contentsline {section}{\numberline {1}Bit-Encodings}{2}{section.1.1}}
-\@writefile{toc}{\contentsline {section}{\numberline {2}Annotated Regular Expressions}{3}{section.1.2}}
-\citation{Brzozowski1964}
-\citation{Owens2008}
-\citation{Krauss2011}
-\citation{Coquand2012}
-\@writefile{toc}{\contentsline {section}{\numberline {3}Introduction}{41}{section.1.3}}
-\citation{Frisch2004}
-\citation{POSIX,Kuklewicz,OkuiSuzuki2010,Sulzmann2014,Vansummeren2006}
-\citation{POSIX}
-\citation{Sulzmann2014}
-\citation{HosoyaVouillonPierce2005}
-\citation{Sulzmann2014}
-\citation{Frisch2004}
-\citation{Sulzmann2014}
-\citation{Kuklewicz}
-\citation{Sulzmann2014}
-\citation{CrashCourse2014}
-\citation{Sulzmann2014}
-\citation{Vansummeren2006}
-\citation{Sulzmann2014}
-\citation{Sulzmann2014}
-\citation{OkuiSuzuki2010}
-\@writefile{toc}{\contentsline {section}{\numberline {4}Preliminaries}{44}{section.1.4}}
-\citation{Krauss2011}
-\citation{Brzozowski1964}
-\newlabel{SemDer}{{1}{45}{Preliminaries}{equation.1.4.1}{}}
-\citation{Sulzmann2014}
-\citation{Sulzmann2014}
-\citation{Frisch2004}
-\citation{Sulzmann2014}
-\citation{AusafDyckhoffUrban2016}
-\newlabel{derprop}{{1}{46}{Preliminaries}{proposition.1.1}{}}
-\newlabel{posixsec}{{5}{46}{POSIX Regular Expression Matching}{section.1.5}{}}
-\@writefile{toc}{\contentsline {section}{\numberline {5}POSIX Regular Expression Matching}{46}{section.1.5}}
-\citation{OkuiSuzuki2010}
-\citation{Frisch2004}
-\citation{Sulzmann2014}
-\citation{Sulzmann2014}
-\newlabel{prfintros}{{5}{47}{POSIX Regular Expression Matching}{section.1.5}{}}
-\newlabel{inhabs}{{2}{47}{POSIX Regular Expression Matching}{proposition.1.2}{}}
-\citation{Sulzmann2014}
-\@writefile{lof}{\contentsline {figure}{\numberline {1}{\ignorespaces The two phases of the algorithm by Sulzmann \& Lu \cite {Sulzmann2014}, matching the string \emph {\sfcode 42 1000 \sfcode 63 1000 \sfcode 33 1000 \sfcode 58 1000 \sfcode 59 1000 \sfcode 44 1000 \it {\emph {$[$}}{\kern 0pt}a{\emph {$\mathord ,$}}{\kern 0pt}\ b{\emph {$\mathord ,$}}{\kern 0pt}\ c{\emph {$]$}}{\kern 0pt}}. The first phase (the arrows from left to right) is Brzozowski's matcher building successive derivatives. If the last regular expression is \emph {\sfcode 42 1000 \sfcode 63 1000 \sfcode 33 1000 \sfcode 58 1000 \sfcode 59 1000 \sfcode 44 1000 \it nullable}, then the functions of the second phase are called (the top-down and right-to-left arrows): first \emph {\sfcode 42 1000 \sfcode 63 1000 \sfcode 33 1000 \sfcode 58 1000 \sfcode 59 1000 \sfcode 44 1000 \it mkeps} calculates a value \emph {\sfcode 42 1000 \sfcode 63 1000 \sfcode 33 1000 \sfcode 58 1000 \sfcode 59 1000 \sfcode 44 1000 \it v\emph {\isascriptstyle ${}\sb {4}$}} witnessing how the empty string has been recognised by \emph {\sfcode 42 1000 \sfcode 63 1000 \sfcode 33 1000 \sfcode 58 1000 \sfcode 59 1000 \sfcode 44 1000 \it r\emph {\isascriptstyle ${}\sb {4}$}}. After that the function \emph {\sfcode 42 1000 \sfcode 63 1000 \sfcode 33 1000 \sfcode 58 1000 \sfcode 59 1000 \sfcode 44 1000 \it inj} ``injects back'' the characters of the string into the values. }}{48}{figure.1.1}}
-\newlabel{Sulz}{{1}{48}{The two phases of the algorithm by Sulzmann \& Lu \cite {Sulzmann2014}, matching the string \isa {{\isacharbrackleft }{\kern 0pt}a{\isacharcomma }{\kern 0pt}\ b{\isacharcomma }{\kern 0pt}\ c{\isacharbrackright }{\kern 0pt}}. The first phase (the arrows from left to right) is \Brz 's matcher building successive derivatives. If the last regular expression is \isa {nullable}, then the functions of the second phase are called (the top-down and right-to-left arrows): first \isa {mkeps} calculates a value \isa {v\isactrlsub {\isadigit {4}}} witnessing how the empty string has been recognised by \isa {r\isactrlsub {\isadigit {4}}}. After that the function \isa {inj} ``injects back'' the characters of the string into the values}{figure.1.1}{}}
-\newlabel{Prf_injval_flat}{{2}{49}{POSIX Regular Expression Matching}{lemma.1.2}{}}
-\citation{Sulzmann2014}
-\citation{Vansummeren2006}
-\newlabel{posixdeterm}{{1}{50}{POSIX Regular Expression Matching}{theorem.1.5.1}{}}
-\@writefile{lof}{\contentsline {figure}{\numberline {2}{\ignorespaces Our inductive definition of POSIX values.}}{51}{figure.1.2}}
-\newlabel{POSIXrules}{{2}{51}{Our inductive definition of POSIX values}{figure.1.2}{}}
-\newlabel{LVposix}{{3}{51}{POSIX Regular Expression Matching}{lemma.1.3}{}}
-\newlabel{lemmkeps}{{4}{51}{POSIX Regular Expression Matching}{lemma.1.4}{}}
-\newlabel{Posix2}{{5}{52}{POSIX Regular Expression Matching}{lemma.1.5}{}}
-\newlabel{lexercorrect}{{2}{52}{POSIX Regular Expression Matching}{theorem.1.5.2}{}}
-\citation{Sulzmann2014}
-\citation{Sulzmann2014}
-\citation{OkuiSuzuki2010,OkuiSuzukiTech}
-\newlabel{lexercorrectcor}{{1}{53}{POSIX Regular Expression Matching}{corollary.1.1}{}}
-\@writefile{toc}{\contentsline {section}{\numberline {6}Ordering of Values according to Okui and Suzuki}{53}{section.1.6}}
-\citation{OkuiSuzuki2010}
-\citation{OkuiSuzuki2010}
-\citation{OkuiSuzuki2010}
-\newlabel{transitivity}{{6}{55}{Ordering of Values according to Okui and Suzuki}{lemma.1.6}{}}
-\newlabel{ordlen}{{4}{56}{Ordering of Values according to Okui and Suzuki}{proposition.1.4}{}}
-\newlabel{ordintros}{{5}{56}{Ordering of Values according to Okui and Suzuki}{proposition.1.5}{}}
-\newlabel{orderone}{{3}{56}{Ordering of Values according to Okui and Suzuki}{theorem.1.6.3}{}}
-\citation{Sulzmann2014}
-\citation{Sulzmann2014}
-\@writefile{toc}{\contentsline {section}{\numberline {7}Bitcoded Lexing}{58}{section.1.7}}
-\@writefile{toc}{\contentsline {section}{\numberline {8}Optimisations}{58}{section.1.8}}
-\newlabel{Simpl}{{2}{58}{Optimisations}{equation.1.8.2}{}}
-\newlabel{slexeraux}{{7}{63}{Optimisations}{lemma.1.7}{}}
-\@writefile{toc}{\contentsline {section}{\numberline {9}HERE}{64}{section.1.9}}
-\bibstyle{plain}
-\bibdata{root}
--- a/thys2/Journal/root.log Fri Jan 07 22:25:26 2022 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1976 +0,0 @@
-This is LuaTeX, Version 1.07.0 (TeX Live 2018) (format=lualatex 2019.2.7) 1 NOV 2021 11:19
- restricted system commands enabled.
-**root.tex
-(./root.tex
-LaTeX2e <2018-04-01> patch level 2
-Lua module: luaotfload-main 2017/01/29 2.80001 OpenType layout system.
-Lua module: lualibs 2017-02-01 2.5 ConTeXt Lua standard libraries.
-Lua module: lualibs-extended 2017-02-01 2.5 ConTeXt Lua libraries -- extended co
-llection.(using write cache: /Users/cstan/Library/texlive/2018/texmf-var/luatex-
-cache/generic)(using read cache: /usr/local/texlive/2018/texmf-var/luatex-cache/
-generic /Users/cstan/Library/texlive/2018/texmf-var/luatex-cache/generic)
-luaotfload | conf : Root cache directory is /Users/cstan/Library/texlive/2018/te
-xmf-var/luatex-cache/generic/names.
-luaotfload | init : Loading fontloader “fontloader-2017-02-11.lua” from kpse
--resolved path “/usr/local/texlive/2018/texmf-dist/tex/luatex/luaotfload/fontl
-oader-2017-02-11.lua”.
-Lua-only attribute luaotfload@state = 1
-Lua-only attribute luaotfload@noligature = 2
-Lua-only attribute luaotfload@syllabe = 3
-luaotfload | init : Context OpenType loader version “3.027”
-Inserting `luaotfload.node_processor' at position 1 in `pre_linebreak_filter'.
-Inserting `luaotfload.node_processor' at position 1 in `hpack_filter'.
-Inserting `luaotfload.define_font' at position 1 in `define_font'.
-Lua-only attribute luaotfload_color_attribute = 4
-luaotfload | conf : Root cache directory is /Users/cstan/Library/texlive/2018/te
-xmf-var/luatex-cache/generic/names.
-Inserting `luaotfload.aux.set_sscale_dimens' at position 1 in `luaotfload.patch_
-font'.
-Inserting `luaotfload.aux.patch_cambria_domh' at position 2 in `luaotfload.patch
-_font'.
-Inserting `luaotfload.aux.fixup_fontdata' at position 1 in `luaotfload.patch_fon
-t_unsafe'.
-Inserting `luaotfload.aux.set_capheight' at position 3 in `luaotfload.patch_font
-'.
-Inserting `luaotfload.rewrite_fontname' at position 4 in `luaotfload.patch_font'
-.
-luaotfload | main : initialization completed in 0.120 seconds
-Babel <3.18> and hyphenation patterns for 1 language(s) loaded.
-(./llncs.cls
-Document Class: llncs 2002/01/28 v2.13
- LaTeX document class for Lecture Notes in Computer Science
-(/usr/local/texlive/2018/texmf-dist/tex/latex/base/article.cls
-Document Class: article 2014/09/29 v1.4h Standard LaTeX document class
-(/usr/local/texlive/2018/texmf-dist/tex/latex/base/size10.clo
-File: size10.clo 2014/09/29 v1.4h Standard LaTeX file (size option)
-luaotfload | db : Font names database loaded from /Users/cstan/Library/texlive/2
-018/texmf-var/luatex-cache/generic/names/luaotfload-names.luc(load luc: /Users/c
-stan/Library/texlive/2018/texmf-var/luatex-cache/generic/fonts/otl/lmroman10-reg
-ular.luc))
-\c@part=\count80
-\c@section=\count81
-\c@subsection=\count82
-\c@subsubsection=\count83
-\c@paragraph=\count84
-\c@subparagraph=\count85
-\c@figure=\count86
-\c@table=\count87
-\abovecaptionskip=\skip41
-\belowcaptionskip=\skip42
-\bibindent=\dimen102
-) (/usr/local/texlive/2018/texmf-dist/tex/latex/tools/multicol.sty
-Package: multicol 2018/04/01 v1.8r multicolumn formatting (FMi)
-\c@tracingmulticols=\count88
-\mult@box=\box26
-\multicol@leftmargin=\dimen103
-\c@unbalance=\count89
-\c@collectmore=\count90
-\doublecol@number=\count91
-\multicoltolerance=\count92
-\multicolpretolerance=\count93
-\full@width=\dimen104
-\page@free=\dimen105
-\premulticols=\dimen106
-\postmulticols=\dimen107
-\multicolsep=\skip43
-\multicolbaselineskip=\skip44
-\partial@page=\box27
-\last@line=\box28
-\maxbalancingoverflow=\dimen108
-\mult@rightbox=\box29
-\mult@grightbox=\box30
-\mult@gfirstbox=\box31
-\mult@firstbox=\box32
-\@tempa=\box33
-\@tempa=\box34
-\@tempa=\box35
-\@tempa=\box36
-\@tempa=\box37
-\@tempa=\box38
-\@tempa=\box39
-\@tempa=\box40
-\@tempa=\box41
-\@tempa=\box42
-\@tempa=\box43
-\@tempa=\box44
-\@tempa=\box45
-\@tempa=\box46
-\@tempa=\box47
-\@tempa=\box48
-\@tempa=\box49
-\c@columnbadness=\count94
-\c@finalcolumnbadness=\count95
-\last@try=\dimen109
-\multicolovershoot=\dimen110
-\multicolundershoot=\dimen111
-\mult@nat@firstbox=\box50
-\colbreak@box=\box51
-\mc@col@check@num=\count96
-)
-\c@chapter=\count97
-LaTeX Font Info: Redeclaring math symbol \Gamma on input line 360.
-LaTeX Font Info: Redeclaring math symbol \Delta on input line 361.
-LaTeX Font Info: Redeclaring math symbol \Theta on input line 362.
-LaTeX Font Info: Redeclaring math symbol \Lambda on input line 363.
-LaTeX Font Info: Redeclaring math symbol \Xi on input line 364.
-LaTeX Font Info: Redeclaring math symbol \Pi on input line 365.
-LaTeX Font Info: Redeclaring math symbol \Sigma on input line 366.
-LaTeX Font Info: Redeclaring math symbol \Upsilon on input line 367.
-LaTeX Font Info: Redeclaring math symbol \Phi on input line 368.
-LaTeX Font Info: Redeclaring math symbol \Psi on input line 369.
-LaTeX Font Info: Redeclaring math symbol \Omega on input line 370.
-\tocchpnum=\dimen112
-\tocsecnum=\dimen113
-\tocsectotal=\dimen114
-\tocsubsecnum=\dimen115
-\tocsubsectotal=\dimen116
-\tocsubsubsecnum=\dimen117
-\tocsubsubsectotal=\dimen118
-\tocparanum=\dimen119
-\tocparatotal=\dimen120
-\tocsubparanum=\dimen121
-\@tempcntc=\count98
-\fnindent=\dimen122
-\c@@inst=\count99
-\c@@auth=\count100
-\c@auco=\count101
-\instindent=\dimen123
-\authrun=\box52
-\authorrunning=\toks14
-\tocauthor=\toks15
-\titrun=\box53
-\titlerunning=\toks16
-\toctitle=\toks17
-\c@theorem=\count102
-\c@case=\count103
-\c@conjecture=\count104
-\c@corollary=\count105
-\c@definition=\count106
-\c@example=\count107
-\c@exercise=\count108
-\c@lemma=\count109
-\c@note=\count110
-\c@problem=\count111
-\c@property=\count112
-\c@proposition=\count113
-\c@question=\count114
-\c@solution=\count115
-\c@remark=\count116
-\headlineindent=\dimen124
-) (/usr/local/texlive/2018/texmf-dist/tex/latex/psnfss/times.sty
-Package: times 2005/04/12 PSNFSS-v9.2a (SPQR)
-)
-(./isabelle.sty
-\isa@parindent=\dimen125
-\isa@parskip=\dimen126
- (/usr/local/texlive/2018/texmf-dist/tex/generic/ulem/ulem.sty
-\UL@box=\box54
-\UL@hyphenbox=\box55
-\UL@skip=\skip45
-\UL@hook=\toks18
-\UL@height=\dimen127
-\UL@pe=\count117
-\UL@pixel=\dimen128
-\ULC@box=\box56
-Package: ulem 2012/05/18
-\ULdepth=\dimen129
-)
-(./comment.sty
-\CommentStream=\write3
- Excluding comment 'comment')
-Including comment 'isadelimdocument' Including comment 'isatagdocument'
-Including comment 'isadelimtheory' Including comment 'isatagtheory'
-Including comment 'isadelimproof' Including comment 'isatagproof'
-Including comment 'isadelimML' Including comment 'isatagML'
-Including comment 'isadelimvisible' Including comment 'isatagvisible'
-Excluding comment 'isadeliminvisible' Excluding comment 'isataginvisible'
-Including comment 'isadelimimportant' Including comment 'isatagimportant'
-Including comment 'isadelimunimportant' Including comment 'isatagunimportant')
-(./isabelletags.sty) (./isabellesym.sty)
-(/usr/local/texlive/2018/texmf-dist/tex/latex/amsmath/amsmath.sty
-Package: amsmath 2017/09/02 v2.17a AMS math features
-\@mathmargin=\skip46
-
-For additional information on amsmath, use the `?' option.
-(/usr/local/texlive/2018/texmf-dist/tex/latex/amsmath/amstext.sty
-Package: amstext 2000/06/29 v2.01 AMS text
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/amsmath/amsgen.sty
-File: amsgen.sty 1999/11/30 v2.0 generic functions
-\@emptytoks=\toks19
-\ex@=\dimen130
-))
-(/usr/local/texlive/2018/texmf-dist/tex/latex/amsmath/amsbsy.sty
-Package: amsbsy 1999/11/29 v1.2d Bold Symbols
-\pmbraise@=\dimen131
-)
-(/usr/local/texlive/2018/texmf-dist/tex/latex/amsmath/amsopn.sty
-Package: amsopn 2016/03/08 v2.02 operator names
-)
-\inf@bad=\count118
-LaTeX Info: Redefining \frac on input line 213.
-\uproot@=\count119
-\leftroot@=\count120
-LaTeX Info: Redefining \overline on input line 375.
-\classnum@=\count121
-\DOTSCASE@=\count122
-LaTeX Info: Redefining \ldots on input line 472.
-LaTeX Info: Redefining \dots on input line 475.
-LaTeX Info: Redefining \cdots on input line 596.
-\Mathstrutbox@=\box57
-\strutbox@=\box58
-\big@size=\dimen132
-LaTeX Font Info: Redeclaring font encoding OML on input line 712.
-LaTeX Font Info: Redeclaring font encoding OMS on input line 713.
-
-
-Package amsmath Warning: Unable to redefine math accent \vec.
-
-\macc@depth=\count123
-\c@MaxMatrixCols=\count124
-\dotsspace@=\muskip10
-\c@parentequation=\count125
-\dspbrk@lvl=\count126
-\tag@help=\toks20
-\row@=\count127
-\column@=\count128
-\maxfields@=\count129
-\andhelp@=\toks21
-\eqnshift@=\dimen133
-\alignsep@=\dimen134
-\tagshift@=\dimen135
-\tagwidth@=\dimen136
-\totwidth@=\dimen137
-\lineht@=\dimen138
-\@envbody=\toks22
-\multlinegap=\skip47
-\multlinetaggap=\skip48
-\mathdisplay@stack=\toks23
-LaTeX Info: Redefining \[ on input line 2817.
-LaTeX Info: Redefining \] on input line 2818.
-) (/usr/local/texlive/2018/texmf-dist/tex/latex/amsfonts/amssymb.sty
-Package: amssymb 2013/01/14 v3.01 AMS font symbols
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/amsfonts/amsfonts.sty
-Package: amsfonts 2013/01/14 v3.01 Basic AMSFonts support
-\symAMSa=\mathgroup4
-\symAMSb=\mathgroup5
-LaTeX Font Info: Overwriting math alphabet `\mathfrak' in version `bold'
-(Font) U/euf/m/n --> U/euf/b/n on input line 106.
-))
-(/usr/local/texlive/2018/texmf-dist/tex/latex/mathpartir/mathpartir.sty
-Package: mathpartir 2016/02/24 version 1.3.2 Math Paragraph for Typesetting Infe
-rence Rules
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/graphics/keyval.sty
-Package: keyval 2014/10/28 v1.15 key=value parser (DPC)
-\KV@toks@=\toks24
-)
-\mpr@andskip=\skip49
-\mpr@lista=\toks25
-\mpr@listb=\toks26
-\mpr@hlist=\box59
-\mpr@vlist=\box60
-\mpr@right=\box61
-)
-(/usr/local/texlive/2018/texmf-dist/tex/latex/pgf/frontendlayer/tikz.sty
-(/usr/local/texlive/2018/texmf-dist/tex/latex/pgf/basiclayer/pgf.sty
-(/usr/local/texlive/2018/texmf-dist/tex/latex/pgf/utilities/pgfrcs.sty
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/utilities/pgfutil-common.te
-x
-\pgfutil@everybye=\toks27
-\pgfutil@tempdima=\dimen139
-\pgfutil@tempdimb=\dimen140
-
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/utilities/pgfutil-common-li
-sts.tex))
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/utilities/pgfutil-latex.def
-\pgfutil@abb=\box62
-(/usr/local/texlive/2018/texmf-dist/tex/latex/ms/everyshi.sty
-Package: everyshi 2001/05/15 v3.00 EveryShipout Package (MS)
-))
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/utilities/pgfrcs.code.tex
-Package: pgfrcs 2015/08/07 v3.0.1a (rcs-revision 1.31)
-))
-Package: pgf 2015/08/07 v3.0.1a (rcs-revision 1.15)
-(/usr/local/texlive/2018/texmf-dist/tex/latex/pgf/basiclayer/pgfcore.sty
-(/usr/local/texlive/2018/texmf-dist/tex/latex/graphics/graphicx.sty
-Package: graphicx 2017/06/01 v1.1a Enhanced LaTeX Graphics (DPC,SPQR)
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/graphics/graphics.sty
-Package: graphics 2017/06/25 v1.2c Standard LaTeX Graphics (DPC,SPQR)
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/graphics/trig.sty
-Package: trig 2016/01/03 v1.10 sin cos tan (DPC)
-)
-(/usr/local/texlive/2018/texmf-dist/tex/latex/graphics-cfg/graphics.cfg
-File: graphics.cfg 2016/06/04 v1.11 sample graphics configuration
-)
-Package graphics Info: Driver file: luatex.def on input line 99.
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/graphics-def/luatex.def
-File: luatex.def 2018/01/08 v1.0l Graphics/color driver for luatex
-))
-\Gin@req@height=\dimen141
-\Gin@req@width=\dimen142
-)
-(/usr/local/texlive/2018/texmf-dist/tex/latex/pgf/systemlayer/pgfsys.sty
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/systemlayer/pgfsys.code.tex
-Package: pgfsys 2014/07/09 v3.0.1a (rcs-revision 1.48)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/utilities/pgfkeys.code.tex
-\pgfkeys@pathtoks=\toks28
-\pgfkeys@temptoks=\toks29
-
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/utilities/pgfkeysfiltered.c
-ode.tex
-\pgfkeys@tmptoks=\toks30
-))
-\pgf@x=\dimen143
-\pgf@y=\dimen144
-\pgf@xa=\dimen145
-\pgf@ya=\dimen146
-\pgf@xb=\dimen147
-\pgf@yb=\dimen148
-\pgf@xc=\dimen149
-\pgf@yc=\dimen150
-\w@pgf@writea=\write4
-\r@pgf@reada=\read1
-\c@pgf@counta=\count130
-\c@pgf@countb=\count131
-\c@pgf@countc=\count132
-\c@pgf@countd=\count133
-\t@pgf@toka=\toks31
-\t@pgf@tokb=\toks32
-\t@pgf@tokc=\toks33
-
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/systemlayer/pgf.cfg
-File: pgf.cfg 2008/05/14 (rcs-revision 1.7)
-)
-Driver file for pgf: pgfsys-luatex.def
-
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/systemlayer/pgfsys-luatex.d
-ef
-File: pgfsys-luatex.def 2014/10/11 (rcs-revision 1.35)
-
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/systemlayer/pgfsys-common-p
-df.def
-File: pgfsys-common-pdf.def 2013/10/10 (rcs-revision 1.13)
-)))
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/systemlayer/pgfsyssoftpath.
-code.tex
-File: pgfsyssoftpath.code.tex 2013/09/09 (rcs-revision 1.9)
-\pgfsyssoftpath@smallbuffer@items=\count134
-\pgfsyssoftpath@bigbuffer@items=\count135
-)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/systemlayer/pgfsysprotocol.
-code.tex
-File: pgfsysprotocol.code.tex 2006/10/16 (rcs-revision 1.4)
-)) (/usr/local/texlive/2018/texmf-dist/tex/latex/xcolor/xcolor.sty
-Package: xcolor 2016/05/11 v2.12 LaTeX color extensions (UK)
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/graphics-cfg/color.cfg
-File: color.cfg 2016/01/02 v1.6 sample color configuration
-)
-Package xcolor Info: Driver file: luatex.def on input line 225.
-Package xcolor Info: Model `cmy' substituted by `cmy0' on input line 1348.
-Package xcolor Info: Model `hsb' substituted by `rgb' on input line 1352.
-Package xcolor Info: Model `RGB' extended on input line 1364.
-Package xcolor Info: Model `HTML' substituted by `rgb' on input line 1366.
-Package xcolor Info: Model `Hsb' substituted by `hsb' on input line 1367.
-Package xcolor Info: Model `tHsb' substituted by `hsb' on input line 1368.
-Package xcolor Info: Model `HSB' substituted by `hsb' on input line 1369.
-Package xcolor Info: Model `Gray' substituted by `gray' on input line 1370.
-Package xcolor Info: Model `wave' substituted by `hsb' on input line 1371.
-)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/basiclayer/pgfcore.code.tex
-Package: pgfcore 2010/04/11 v3.0.1a (rcs-revision 1.7)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/math/pgfmath.code.tex
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/math/pgfmathcalc.code.tex
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/math/pgfmathutil.code.tex)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/math/pgfmathparser.code.tex
-\pgfmath@dimen=\dimen151
-\pgfmath@count=\count136
-\pgfmath@box=\box63
-\pgfmath@toks=\toks34
-\pgfmath@stack@operand=\toks35
-\pgfmath@stack@operation=\toks36
-)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/math/pgfmathfunctions.code.
-tex
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/math/pgfmathfunctions.basic
-.code.tex)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/math/pgfmathfunctions.trigo
-nometric.code.tex)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/math/pgfmathfunctions.rando
-m.code.tex)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/math/pgfmathfunctions.compa
-rison.code.tex)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/math/pgfmathfunctions.base.
-code.tex)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/math/pgfmathfunctions.round
-.code.tex)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/math/pgfmathfunctions.misc.
-code.tex)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/math/pgfmathfunctions.integ
-erarithmetics.code.tex)))
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/math/pgfmathfloat.code.tex
-\c@pgfmathroundto@lastzeros=\count137
-))
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/basiclayer/pgfcorepoints.co
-de.tex
-File: pgfcorepoints.code.tex 2013/10/07 (rcs-revision 1.27)
-\pgf@picminx=\dimen152
-\pgf@picmaxx=\dimen153
-\pgf@picminy=\dimen154
-\pgf@picmaxy=\dimen155
-\pgf@pathminx=\dimen156
-\pgf@pathmaxx=\dimen157
-\pgf@pathminy=\dimen158
-\pgf@pathmaxy=\dimen159
-\pgf@xx=\dimen160
-\pgf@xy=\dimen161
-\pgf@yx=\dimen162
-\pgf@yy=\dimen163
-\pgf@zx=\dimen164
-\pgf@zy=\dimen165
-)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/basiclayer/pgfcorepathconst
-ruct.code.tex
-File: pgfcorepathconstruct.code.tex 2013/10/07 (rcs-revision 1.29)
-\pgf@path@lastx=\dimen166
-\pgf@path@lasty=\dimen167
-)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/basiclayer/pgfcorepathusage
-.code.tex
-File: pgfcorepathusage.code.tex 2014/11/02 (rcs-revision 1.24)
-\pgf@shorten@end@additional=\dimen168
-\pgf@shorten@start@additional=\dimen169
-)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/basiclayer/pgfcorescopes.co
-de.tex
-File: pgfcorescopes.code.tex 2015/05/08 (rcs-revision 1.46)
-\pgfpic=\box64
-\pgf@hbox=\box65
-\pgf@layerbox@main=\box66
-\pgf@picture@serial@count=\count138
-)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/basiclayer/pgfcoregraphicst
-ate.code.tex
-File: pgfcoregraphicstate.code.tex 2014/11/02 (rcs-revision 1.12)
-\pgflinewidth=\dimen170
-)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/basiclayer/pgfcoretransform
-ations.code.tex
-File: pgfcoretransformations.code.tex 2015/08/07 (rcs-revision 1.20)
-\pgf@pt@x=\dimen171
-\pgf@pt@y=\dimen172
-\pgf@pt@temp=\dimen173
-)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/basiclayer/pgfcorequick.cod
-e.tex
-File: pgfcorequick.code.tex 2008/10/09 (rcs-revision 1.3)
-)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/basiclayer/pgfcoreobjects.c
-ode.tex
-File: pgfcoreobjects.code.tex 2006/10/11 (rcs-revision 1.2)
-)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/basiclayer/pgfcorepathproce
-ssing.code.tex
-File: pgfcorepathprocessing.code.tex 2013/09/09 (rcs-revision 1.9)
-)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/basiclayer/pgfcorearrows.co
-de.tex
-File: pgfcorearrows.code.tex 2015/05/14 (rcs-revision 1.43)
-\pgfarrowsep=\dimen174
-)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/basiclayer/pgfcoreshade.cod
-e.tex
-File: pgfcoreshade.code.tex 2013/07/15 (rcs-revision 1.15)
-\pgf@max=\dimen175
-\pgf@sys@shading@range@num=\count139
-)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/basiclayer/pgfcoreimage.cod
-e.tex
-File: pgfcoreimage.code.tex 2013/07/15 (rcs-revision 1.18)
-
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/basiclayer/pgfcoreexternal.
-code.tex
-File: pgfcoreexternal.code.tex 2014/07/09 (rcs-revision 1.21)
-\pgfexternal@startupbox=\box67
-))
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/basiclayer/pgfcorelayers.co
-de.tex
-File: pgfcorelayers.code.tex 2013/07/18 (rcs-revision 1.7)
-)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/basiclayer/pgfcoretranspare
-ncy.code.tex
-File: pgfcoretransparency.code.tex 2013/09/30 (rcs-revision 1.5)
-)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/basiclayer/pgfcorepatterns.
-code.tex
-File: pgfcorepatterns.code.tex 2013/11/07 (rcs-revision 1.5)
-)))
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/modules/pgfmoduleshapes.cod
-e.tex
-File: pgfmoduleshapes.code.tex 2014/03/21 (rcs-revision 1.35)
-\pgfnodeparttextbox=\box68
-)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/modules/pgfmoduleplot.code.
-tex
-File: pgfmoduleplot.code.tex 2015/08/03 (rcs-revision 1.13)
-)
-(/usr/local/texlive/2018/texmf-dist/tex/latex/pgf/compatibility/pgfcomp-version
--0-65.sty
-Package: pgfcomp-version-0-65 2007/07/03 v3.0.1a (rcs-revision 1.7)
-\pgf@nodesepstart=\dimen176
-\pgf@nodesepend=\dimen177
-)
-(/usr/local/texlive/2018/texmf-dist/tex/latex/pgf/compatibility/pgfcomp-version
--1-18.sty
-Package: pgfcomp-version-1-18 2007/07/23 v3.0.1a (rcs-revision 1.1)
-))
-(/usr/local/texlive/2018/texmf-dist/tex/latex/pgf/utilities/pgffor.sty
-(/usr/local/texlive/2018/texmf-dist/tex/latex/pgf/utilities/pgfkeys.sty
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/utilities/pgfkeys.code.tex))
- (/usr/local/texlive/2018/texmf-dist/tex/latex/pgf/math/pgfmath.sty
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/math/pgfmath.code.tex))
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/utilities/pgffor.code.tex
-Package: pgffor 2013/12/13 v3.0.1a (rcs-revision 1.25)
-
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/math/pgfmath.code.tex)
-\pgffor@iter=\dimen178
-\pgffor@skip=\dimen179
-\pgffor@stack=\toks37
-\pgffor@toks=\toks38
-))
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/frontendlayer/tikz/tikz.cod
-e.tex
-Package: tikz 2015/08/07 v3.0.1a (rcs-revision 1.151)
-
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/libraries/pgflibraryplothan
-dlers.code.tex
-File: pgflibraryplothandlers.code.tex 2013/08/31 v3.0.1a (rcs-revision 1.20)
-\pgf@plot@mark@count=\count140
-\pgfplotmarksize=\dimen180
-)
-\tikz@lastx=\dimen181
-\tikz@lasty=\dimen182
-\tikz@lastxsaved=\dimen183
-\tikz@lastysaved=\dimen184
-\tikzleveldistance=\dimen185
-\tikzsiblingdistance=\dimen186
-\tikz@figbox=\box69
-\tikz@figbox@bg=\box70
-\tikz@tempbox=\box71
-\tikz@tempbox@bg=\box72
-\tikztreelevel=\count141
-\tikznumberofchildren=\count142
-\tikznumberofcurrentchild=\count143
-\tikz@fig@count=\count144
-
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/modules/pgfmodulematrix.cod
-e.tex
-File: pgfmodulematrix.code.tex 2013/09/17 (rcs-revision 1.8)
-\pgfmatrixcurrentrow=\count145
-\pgfmatrixcurrentcolumn=\count146
-\pgf@matrix@numberofcolumns=\count147
-)
-\tikz@expandcount=\count148
-
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/frontendlayer/tikz/librarie
-s/tikzlibrarytopaths.code.tex
-File: tikzlibrarytopaths.code.tex 2008/06/17 v3.0.1a (rcs-revision 1.2)
-)))
-(/usr/local/texlive/2018/texmf-dist/tex/generic/pgf/frontendlayer/tikz/librarie
-s/tikzlibrarypositioning.code.tex
-File: tikzlibrarypositioning.code.tex 2008/10/06 v3.0.1a (rcs-revision 1.7)
-) (./pdfsetup.sty)
-(/usr/local/texlive/2018/texmf-dist/tex/latex/hyperref/hyperref.sty
-Package: hyperref 2018/02/06 v6.86b Hypertext links for LaTeX
-
-(/usr/local/texlive/2018/texmf-dist/tex/generic/oberdiek/hobsub-hyperref.sty
-Package: hobsub-hyperref 2016/05/16 v1.14 Bundle oberdiek, subset hyperref (HO)
-
-(/usr/local/texlive/2018/texmf-dist/tex/generic/oberdiek/hobsub-generic.sty
-Package: hobsub-generic 2016/05/16 v1.14 Bundle oberdiek, subset generic (HO)
-Package: hobsub 2016/05/16 v1.14 Construct package bundles (HO)
-Package: infwarerr 2016/05/16 v1.4 Providing info/warning/error messages (HO)
-Package: ltxcmds 2016/05/16 v1.23 LaTeX kernel commands for general use (HO)
-Package: ifluatex 2016/05/16 v1.4 Provides the ifluatex switch (HO)
-Package ifluatex Info: LuaTeX detected.
-Package: ifvtex 2016/05/16 v1.6 Detect VTeX and its facilities (HO)
-Package ifvtex Info: VTeX not detected.
-Package: intcalc 2016/05/16 v1.2 Expandable calculations with integers (HO)
-Package: ifpdf 2017/03/15 v3.2 Provides the ifpdf switch
-Package: etexcmds 2016/05/16 v1.6 Avoid name clashes with e-TeX commands (HO)
-Package: kvsetkeys 2016/05/16 v1.17 Key value parser (HO)
-Package: kvdefinekeys 2016/05/16 v1.4 Define keys (HO)
-Package: luatex-loader 2016/05/16 v0.6 Lua module loader (HO)
-
-(/usr/local/texlive/2018/texmf-dist/scripts/oberdiek/oberdiek.luatex.lua)
-Package: pdftexcmds 2018/01/30 v0.27 Utility functions of pdfTeX for LuaTeX (HO)
-
-Package pdftexcmds Info: \pdf@primitive is available.
-Package pdftexcmds Info: \pdf@ifprimitive is available.
-Package pdftexcmds Info: \pdfdraftmode found.
-\pdftexcmds@toks=\toks39
-Package: pdfescape 2016/05/16 v1.14 Implements pdfTeX's escape features (HO)
-Package: bigintcalc 2016/05/16 v1.4 Expandable calculations on big integers (HO)
-
-Package: bitset 2016/05/16 v1.2 Handle bit-vector datatype (HO)
-Package: uniquecounter 2016/05/16 v1.3 Provide unlimited unique counter (HO)
-)
-Package hobsub Info: Skipping package `hobsub' (already loaded).
-Package: letltxmacro 2016/05/16 v1.5 Let assignment for LaTeX macros (HO)
-Package: hopatch 2016/05/16 v1.3 Wrapper for package hooks (HO)
-Package: xcolor-patch 2016/05/16 xcolor patch
-Package: atveryend 2016/05/16 v1.9 Hooks at the very end of document (HO)
-Package atveryend Info: \enddocument detected (standard20110627).
-Package: atbegshi 2016/06/09 v1.18 At begin shipout hook (HO)
-Package: refcount 2016/05/16 v3.5 Data extraction from label references (HO)
-Package: hycolor 2016/05/16 v1.8 Color options for hyperref/bookmark (HO)
-)
-(/usr/local/texlive/2018/texmf-dist/tex/generic/ifxetex/ifxetex.sty
-Package: ifxetex 2010/09/12 v0.6 Provides ifxetex conditional
-)
-(/usr/local/texlive/2018/texmf-dist/tex/latex/oberdiek/auxhook.sty
-Package: auxhook 2016/05/16 v1.4 Hooks for auxiliary files (HO)
-)
-(/usr/local/texlive/2018/texmf-dist/tex/latex/oberdiek/kvoptions.sty
-Package: kvoptions 2016/05/16 v3.12 Key value format for package options (HO)
-)
-\@linkdim=\dimen187
-\Hy@linkcounter=\count149
-\Hy@pagecounter=\count150
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/hyperref/pd1enc.def
-File: pd1enc.def 2018/02/06 v6.86b Hyperref: PDFDocEncoding definition (HO)
-)
-\Hy@SavedSpaceFactor=\count151
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/latexconfig/hyperref.cfg
-File: hyperref.cfg 2002/06/06 v1.2 hyperref configuration of TeXLive
-)
-Package hyperref Info: Option `colorlinks' set `true' on input line 4383.
-Package hyperref Info: Option `pdfpagelabels' set `true' on input line 4383.
-Package hyperref Info: Hyper figures OFF on input line 4509.
-Package hyperref Info: Link nesting OFF on input line 4514.
-Package hyperref Info: Hyper index ON on input line 4517.
-Package hyperref Info: Plain pages OFF on input line 4524.
-Package hyperref Info: Backreferencing OFF on input line 4529.
-Package hyperref Info: Implicit mode ON; LaTeX internals redefined.
-Package hyperref Info: Bookmarks ON on input line 4762.
-\c@Hy@tempcnt=\count152
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/url/url.sty
-\Urlmuskip=\muskip11
-Package: url 2013/09/16 ver 3.4 Verb mode for urls, etc.
-)
-LaTeX Info: Redefining \url on input line 5115.
-\XeTeXLinkMargin=\dimen188
-\Fld@menulength=\count153
-\Field@Width=\dimen189
-\Fld@charsize=\dimen190
-Package hyperref Info: Hyper figures OFF on input line 6369.
-Package hyperref Info: Link nesting OFF on input line 6374.
-Package hyperref Info: Hyper index ON on input line 6377.
-Package hyperref Info: backreferencing OFF on input line 6384.
-Package hyperref Info: Link coloring ON on input line 6387.
-Package hyperref Info: Link coloring with OCG OFF on input line 6394.
-Package hyperref Info: PDF/A mode OFF on input line 6399.
-LaTeX Info: Redefining \ref on input line 6439.
-LaTeX Info: Redefining \pageref on input line 6443.
-\Hy@abspage=\count154
-\c@Item=\count155
-\c@Hfootnote=\count156
-)
-Package hyperref Info: Driver (autodetected): hluatex.
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/hyperref/hluatex.def
-File: hluatex.def 2018/02/06 v6.86b Hyperref driver for luaTeX
-\Fld@listcount=\count157
-\c@bookmark@seq@number=\count158
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/oberdiek/rerunfilecheck.sty
-Package: rerunfilecheck 2016/05/16 v1.8 Rerun checks for auxiliary files (HO)
-Package uniquecounter Info: New unique counter `rerunfilecheck' on input line 28
-2.
-)
-\Hy@SectionHShift=\skip50
-)
-(/usr/local/texlive/2018/texmf-dist/tex/latex/stmaryrd/stmaryrd.sty
-Package: stmaryrd 1994/03/03 St Mary's Road symbol package
-\symstmry=\mathgroup6
-LaTeX Font Info: Overwriting symbol font `stmry' in version `bold'
-(Font) U/stmry/m/n --> U/stmry/b/n on input line 89.
-)
-(/usr/local/texlive/2018/texmf-dist/tex/latex/tipa/tipa.sty
-Package: tipa 2002/08/08 TIPA version 1.1
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/base/fontenc.sty
-Package: fontenc 2017/04/05 v2.0i Standard LaTeX package
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/tipa/t3enc.def
-File: t3enc.def 2001/12/31 T3 encoding
-(load luc: /Users/cstan/Library/texlive/2018/texmf-var/luatex-cache/generic/font
-s/otl/lmromanslant10-regular.luc)(load luc: /Users/cstan/Library/texlive/2018/te
-xmf-var/luatex-cache/generic/fonts/otl/lmroman10-italic.luc)(load luc: /Users/cs
-tan/Library/texlive/2018/texmf-var/luatex-cache/generic/fonts/otl/lmroman10-bold
-.luc)
-LaTeX Font Info: Try loading font information for TU+phv on input line 357.
-LaTeX Font Info: No file TUphv.fd. on input line 357.
-
-
-LaTeX Font Warning: Font shape `TU/phv/m/n' undefined
-(Font) using `TU/lmr/m/n' instead on input line 357.
-
-) (/usr/local/texlive/2018/texmf-dist/tex/latex/base/tuenc.def
-File: tuenc.def 2017/04/05 v2.0i Standard LaTeX file
-LaTeX Font Info: Redeclaring font encoding TU on input line 82.
-)))
-(/usr/local/texlive/2018/texmf-dist/tex/latex/psnfss/mathpazo.sty
-Package: mathpazo 2005/04/12 PSNFSS-v9.2a Palatino w/ Pazo Math (D.Puga, WaS)
-\symupright=\mathgroup7
-)
-(/usr/local/texlive/2018/texmf-dist/tex/latex/fontspec/fontspec.sty
-(/usr/local/texlive/2018/texmf-dist/tex/latex/l3packages/xparse/xparse.sty
-(/usr/local/texlive/2018/texmf-dist/tex/latex/l3kernel/expl3.sty
-Package: expl3 2018/03/05 L3 programming layer (loader)
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/l3kernel/expl3-code.tex
-Package: expl3 2018/03/05 L3 programming layer (code)
-\ucharcat@table=\catcodetable5
-\c_max_int=\count159
-\l_tmpa_int=\count160
-\l_tmpb_int=\count161
-\g_tmpa_int=\count162
-\g_tmpb_int=\count163
-\g__intarray_font_int=\count164
-\g__prg_map_int=\count165
-\c_log_iow=\count166
-\l_iow_line_count_int=\count167
-\l__iow_line_target_int=\count168
-\l__iow_one_indent_int=\count169
-\l__iow_indent_int=\count170
-\c_zero_dim=\dimen191
-\c_max_dim=\dimen192
-\l_tmpa_dim=\dimen193
-\l_tmpb_dim=\dimen194
-\g_tmpa_dim=\dimen195
-\g_tmpb_dim=\dimen196
-\c_zero_skip=\skip51
-\c_max_skip=\skip52
-\l_tmpa_skip=\skip53
-\l_tmpb_skip=\skip54
-\g_tmpa_skip=\skip55
-\g_tmpb_skip=\skip56
-\c_zero_muskip=\muskip12
-\c_max_muskip=\muskip13
-\l_tmpa_muskip=\muskip14
-\l_tmpb_muskip=\muskip15
-\g_tmpa_muskip=\muskip16
-\g_tmpb_muskip=\muskip17
-\l_keys_choice_int=\count171
-\c__fp_leading_shift_int=\count172
-\c__fp_middle_shift_int=\count173
-\c__fp_trailing_shift_int=\count174
-\c__fp_big_leading_shift_int=\count175
-\c__fp_big_middle_shift_int=\count176
-\c__fp_big_trailing_shift_int=\count177
-\c__fp_Bigg_leading_shift_int=\count178
-\c__fp_Bigg_middle_shift_int=\count179
-\c__fp_Bigg_trailing_shift_int=\count180
-\c__fp_rand_size_int=\count181
-\c__fp_rand_four_int=\count182
-\c__fp_rand_eight_int=\count183
-\l__sort_length_int=\count184
-\l__sort_min_int=\count185
-\l__sort_top_int=\count186
-\l__sort_max_int=\count187
-\l__sort_true_max_int=\count188
-\l__sort_block_int=\count189
-\l__sort_begin_int=\count190
-\l__sort_end_int=\count191
-\l__sort_A_int=\count192
-\l__sort_B_int=\count193
-\l__sort_C_int=\count194
-\l__tl_build_start_index_int=\count195
-\l__tl_build_index_int=\count196
-\l__tl_analysis_normal_int=\count197
-\l__tl_analysis_index_int=\count198
-\l__tl_analysis_nesting_int=\count199
-\l__tl_analysis_type_int=\count266
-\l__regex_internal_a_int=\count267
-\l__regex_internal_b_int=\count268
-\l__regex_internal_c_int=\count269
-\l__regex_balance_int=\count270
-\l__regex_group_level_int=\count271
-\l__regex_mode_int=\count272
-\c__regex_cs_in_class_mode_int=\count273
-\c__regex_cs_mode_int=\count274
-\l__regex_catcodes_int=\count275
-\l__regex_default_catcodes_int=\count276
-\c__regex_catcode_L_int=\count277
-\c__regex_catcode_O_int=\count278
-\c__regex_catcode_A_int=\count279
-\c__regex_all_catcodes_int=\count280
-\l__regex_show_lines_int=\count281
-\l__regex_min_state_int=\count282
-\l__regex_max_state_int=\count283
-\l__regex_left_state_int=\count284
-\l__regex_right_state_int=\count285
-\l__regex_capturing_group_int=\count286
-\l__regex_min_pos_int=\count287
-\l__regex_max_pos_int=\count288
-\l__regex_curr_pos_int=\count289
-\l__regex_start_pos_int=\count290
-\l__regex_success_pos_int=\count291
-\l__regex_curr_char_int=\count292
-\l__regex_curr_catcode_int=\count293
-\l__regex_last_char_int=\count294
-\l__regex_case_changed_char_int=\count295
-\l__regex_curr_state_int=\count296
-\l__regex_step_int=\count297
-\l__regex_min_active_int=\count298
-\l__regex_max_active_int=\count299
-\l__regex_replacement_csnames_int=\count300
-\l__regex_match_count_int=\count301
-\l__regex_min_submatch_int=\count302
-\l__regex_submatch_int=\count303
-\l__regex_zeroth_submatch_int=\count304
-\g__regex_trace_regex_int=\count305
-\c_empty_box=\box73
-\l_tmpa_box=\box74
-\l_tmpb_box=\box75
-\g_tmpa_box=\box76
-\g_tmpb_box=\box77
-\l__box_top_dim=\dimen197
-\l__box_bottom_dim=\dimen198
-\l__box_left_dim=\dimen199
-\l__box_right_dim=\dimen256
-\l__box_top_new_dim=\dimen257
-\l__box_bottom_new_dim=\dimen258
-\l__box_left_new_dim=\dimen259
-\l__box_right_new_dim=\dimen260
-\l__box_internal_box=\box78
-\l__coffin_internal_box=\box79
-\l__coffin_internal_dim=\dimen261
-\l__coffin_offset_x_dim=\dimen262
-\l__coffin_offset_y_dim=\dimen263
-\l__coffin_x_dim=\dimen264
-\l__coffin_y_dim=\dimen265
-\l__coffin_x_prime_dim=\dimen266
-\l__coffin_y_prime_dim=\dimen267
-\c_empty_coffin=\box80
-\l__coffin_aligned_coffin=\box81
-\l__coffin_aligned_internal_coffin=\box82
-\l_tmpa_coffin=\box83
-\l_tmpb_coffin=\box84
-\l__coffin_display_coffin=\box85
-\l__coffin_display_coord_coffin=\box86
-\l__coffin_display_pole_coffin=\box87
-\l__coffin_display_offset_dim=\dimen268
-\l__coffin_display_x_dim=\dimen269
-\l__coffin_display_y_dim=\dimen270
-\l__coffin_bounding_shift_dim=\dimen271
-\l__coffin_left_corner_dim=\dimen272
-\l__coffin_right_corner_dim=\dimen273
-\l__coffin_bottom_corner_dim=\dimen274
-\l__coffin_top_corner_dim=\dimen275
-\l__coffin_scaled_total_height_dim=\dimen276
-\l__coffin_scaled_width_dim=\dimen277
-)
-(/usr/local/texlive/2018/texmf-dist/tex/latex/l3kernel/l3pdfmode.def
-File: l3pdfmode.def 2017/03/18 v L3 Experimental driver: PDF mode
-\l__driver_color_stack_int=\count306
-))
-Package: xparse 2018/02/21 L3 Experimental document command parser
-\l__xparse_current_arg_int=\count307
-\g__xparse_grabber_int=\count308
-\l__xparse_m_args_int=\count309
-\l__xparse_mandatory_args_int=\count310
-\l__xparse_v_nesting_int=\count311
-)
-Package: fontspec 2017/11/09 v2.6g Font selection for XeLaTeX and LuaLaTeX
-Lua module: fontspec 2017/11/09 2.6g Font selection for XeLaTeX and LuaLaTeX
-(/usr/local/texlive/2018/texmf-dist/tex/latex/fontspec/fontspec-luatex.sty
-Package: fontspec-luatex 2017/11/09 v2.6g Font selection for XeLaTeX and LuaLaTe
-X
-\l__fontspec_script_int=\count312
-\l__fontspec_language_int=\count313
-\l__fontspec_strnum_int=\count314
-\l__fontspec_tmp_int=\count315
-\l__fontspec_em_int=\count316
-\l__fontspec_emdef_int=\count317
-\l__fontspec_strong_int=\count318
-\l__fontspec_strongdef_int=\count319
-\l__fontspec_tmpa_dim=\dimen278
-\l__fontspec_tmpb_dim=\dimen279
-\l__fontspec_tmpc_dim=\dimen280
-\g__file_internal_ior=\read2
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/base/fontenc.sty
-Package: fontenc 2017/04/05 v2.0i Standard LaTeX package
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/base/tuenc.def
-File: tuenc.def 2017/04/05 v2.0i Standard LaTeX file
-LaTeX Font Info: Redeclaring font encoding TU on input line 82.
-))
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \fontspec with sig. 'O{}mO{}' on line 545.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \setmainfont with sig. 'O{}mO{}' on line 549.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \setsansfont with sig. 'O{}mO{}' on line 553.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \setmonofont with sig. 'O{}mO{}' on line 557.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \setmathrm with sig. 'O{}mO{}' on line 561.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \setboldmathrm with sig. 'O{}mO{}' on line 565.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \setmathsf with sig. 'O{}mO{}' on line 569.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \setmathtt with sig. 'O{}mO{}' on line 573.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \setromanfont with sig. 'O{}mO{}' on line 577.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \newfontfamily with sig. 'mO{}mO{}' on line 581.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \newfontface with sig. 'mO{}mO{}' on line 585.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \defaultfontfeatures with sig. 't+om' on line 589.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \addfontfeatures with sig. 'm' on line 593.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \addfontfeature with sig. 'm' on line 597.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \newfontfeature with sig. 'mm' on line 601.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \newAATfeature with sig. 'mmmm' on line 605.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \newopentypefeature with sig. 'mmm' on line 609.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \newICUfeature with sig. 'mmm' on line 613.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \aliasfontfeature with sig. 'mm' on line 617.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \aliasfontfeatureoption with sig. 'mmm' on line 621.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \newfontscript with sig. 'mm' on line 625.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \newfontlanguage with sig. 'mm' on line 629.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \DeclareFontsExtensions with sig. 'm' on line 633.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \IfFontFeatureActiveTF with sig. 'mmm' on line 637.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \EncodingCommand with sig. 'mO{}m' on line 3366.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \EncodingAccent with sig. 'mm' on line 3372.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \EncodingSymbol with sig. 'mm' on line 3378.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \EncodingComposite with sig. 'mmm' on line 3384.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \EncodingCompositeCommand with sig. 'mmm' on line 3390.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \DeclareUnicodeEncoding with sig. 'mm' on line 3415.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \UndeclareSymbol with sig. 'm' on line 3421.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \UndeclareAccent with sig. 'm' on line 3427.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \UndeclareCommand with sig. 'm' on line 3433.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \UndeclareComposite with sig. 'mm' on line 3440.
-.................................................
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/fontspec/fontspec.cfg)
-LaTeX Info: Redefining \itshape on input line 3625.
-LaTeX Info: Redefining \slshape on input line 3630.
-LaTeX Info: Redefining \scshape on input line 3635.
-LaTeX Info: Redefining \upshape on input line 3640.
-LaTeX Info: Redefining \em on input line 3670.
-LaTeX Info: Redefining \emph on input line 3695.
-LaTeX Info: Redefining \- on input line 3749.
-.................................................
-. LaTeX info: "xparse/redefine-command"
-.
-. Redefining command \oldstylenums with sig. 'm' on line 3844.
-.................................................
-.................................................
-. LaTeX info: "xparse/define-command"
-.
-. Defining command \liningnums with sig. 'm' on line 3848.
-.................................................
-))
-\c@falsehood=\count320
-\c@conject=\count321
-
-(./root.aux)
-\openout1 = root.aux
-
-LaTeX Font Info: Checking defaults for OML/cmm/m/it on input line 42.
-LaTeX Font Info: ... okay on input line 42.
-LaTeX Font Info: Checking defaults for T1/cmr/m/n on input line 42.
-LaTeX Font Info: ... okay on input line 42.
-LaTeX Font Info: Checking defaults for OT1/cmr/m/n on input line 42.
-LaTeX Font Info: ... okay on input line 42.
-LaTeX Font Info: Checking defaults for OMS/cmsy/m/n on input line 42.
-LaTeX Font Info: ... okay on input line 42.
-LaTeX Font Info: Checking defaults for TU/lmr/m/n on input line 42.
-LaTeX Font Info: ... okay on input line 42.
-LaTeX Font Info: Checking defaults for OMX/cmex/m/n on input line 42.
-LaTeX Font Info: ... okay on input line 42.
-LaTeX Font Info: Checking defaults for U/cmr/m/n on input line 42.
-LaTeX Font Info: ... okay on input line 42.
-LaTeX Font Info: Checking defaults for PD1/pdf/m/n on input line 42.
-LaTeX Font Info: ... okay on input line 42.
-LaTeX Font Info: Checking defaults for T3/cmr/m/n on input line 42.
-LaTeX Font Info: Try loading font information for T3+cmr on input line 42.
- (/usr/local/texlive/2018/texmf-dist/tex/latex/tipa/t3cmr.fd
-File: t3cmr.fd 2001/12/31 TIPA font definitions
-)
-LaTeX Font Info: ... okay on input line 42.
-
-ABD: EveryShipout initializing macros
-(/usr/local/texlive/2018/texmf-dist/tex/context/base/mkii/supp-pdf.mkii
-[Loading MPS to PDF converter (version 2006.09.02).]
-\scratchcounter=\count322
-\scratchdimen=\dimen281
-\scratchbox=\box88
-\nofMPsegments=\count323
-\nofMParguments=\count324
-\everyMPshowfont=\toks40
-\MPscratchCnt=\count325
-\MPscratchDim=\dimen282
-\MPnumerator=\count326
-\makeMPintoPDFobject=\count327
-\everyMPtoPDFconversion=\toks41
-) (/usr/local/texlive/2018/texmf-dist/tex/latex/oberdiek/epstopdf-base.sty
-Package: epstopdf-base 2016/05/15 v2.6 Base part for package epstopdf
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/oberdiek/grfext.sty
-Package: grfext 2016/05/16 v1.2 Manage graphics extensions (HO)
-)
-Package epstopdf-base Info: Redefining graphics rule for `.eps' on input line 43
-8.
-Package grfext Info: Graphics extension search list:
-(grfext) [.pdf,.png,.jpg,.mps,.jpeg,.jbig2,.jb2,.PDF,.PNG,.JPG,.JPEG
-,.JBIG2,.JB2,.eps]
-(grfext) \AppendGraphicsExtensions on input line 456.
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/latexconfig/epstopdf-sys.cfg
-File: epstopdf-sys.cfg 2010/07/13 v1.3 Configuration of (r)epstopdf for TeX Live
-
-))
-\AtBeginShipoutBox=\box89
-Package hyperref Info: Link coloring ON on input line 42.
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/hyperref/nameref.sty
-Package: nameref 2016/05/21 v2.44 Cross-referencing by name of section
-
-(/usr/local/texlive/2018/texmf-dist/tex/generic/oberdiek/gettitlestring.sty
-Package: gettitlestring 2016/05/16 v1.5 Cleanup title references (HO)
-)
-\c@section@level=\count328
-)
-LaTeX Info: Redefining \ref on input line 42.
-LaTeX Info: Redefining \pageref on input line 42.
-LaTeX Info: Redefining \nameref on input line 42.
-
-(./root.out) (./root.out)
-\@outlinefile=\write5
-
-\openout5 = root.out
-(load luc: /Users/cstan/Library/texlive/2018/texmf-var/luatex-cache/generic/font
-s/otl/lmroman9-regular.luc)
-LaTeX Font Info: Try loading font information for OT1+pplx on input line 57.
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/psnfss/ot1pplx.fd
-File: ot1pplx.fd 2004/09/06 font definitions for OT1/pplx.
-)
-LaTeX Font Info: Try loading font information for OML+zplm on input line 57.
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/psnfss/omlzplm.fd
-File: omlzplm.fd 2002/09/08 Fontinst v1.914 font definitions for OML/zplm.
-)
-LaTeX Font Info: Try loading font information for OMS+zplm on input line 57.
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/psnfss/omszplm.fd
-File: omszplm.fd 2002/09/08 Fontinst v1.914 font definitions for OMS/zplm.
-)
-LaTeX Font Info: Try loading font information for OMX+zplm on input line 57.
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/psnfss/omxzplm.fd
-File: omxzplm.fd 2002/09/08 Fontinst v1.914 font definitions for OMX/zplm.
-)
-LaTeX Font Info: Font shape `U/msa/m/n' will be
-(Font) scaled to size 9.37807pt on input line 57.
-LaTeX Font Info: Font shape `U/msa/m/n' will be
-(Font) scaled to size 7.29405pt on input line 57.
-LaTeX Font Info: Font shape `U/msa/m/n' will be
-(Font) scaled to size 5.21004pt on input line 57.
-LaTeX Font Info: Font shape `U/msb/m/n' will be
-(Font) scaled to size 9.37807pt on input line 57.
-LaTeX Font Info: Font shape `U/msb/m/n' will be
-(Font) scaled to size 7.29405pt on input line 57.
-LaTeX Font Info: Font shape `U/msb/m/n' will be
-(Font) scaled to size 5.21004pt on input line 57.
-LaTeX Font Info: Try loading font information for U+stmry on input line 57.
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/stmaryrd/Ustmry.fd)
-LaTeX Font Info: Try loading font information for OT1+zplm on input line 57.
-
-(/usr/local/texlive/2018/texmf-dist/tex/latex/psnfss/ot1zplm.fd
-File: ot1zplm.fd 2002/09/08 Fontinst v1.914 font definitions for OT1/zplm.
-)(load luc: /Users/cstan/Library/texlive/2018/texmf-var/luatex-cache/generic/fon
-ts/otl/lmroman7-regular.luc)(load luc: /Users/cstan/Library/texlive/2018/texmf-v
-ar/luatex-cache/generic/fonts/otl/lmroman9-bold.luc)
-
-LaTeX Warning: Citation `AusafDyckhoffUrban2016' on page 1 undefined on input li
-ne 57.
-
-
-LaTeX Warning: Citation `OkuiSuzuki2010' on page 1 undefined on input line 57.
-
-
-LaTeX Warning: Citation `OkuiSuzukiTech' on page 1 undefined on input line 57.
-
-
-LaTeX Warning: Citation `Sulzmann2014' on page 1 undefined on input line 57.
-
-(load luc: /Users/cstan/Library/texlive/2018/texmf-var/luatex-cache/generic/font
-s/otl/lmroman12-regular.luc)(load luc: /Users/cstan/Library/texlive/2018/texmf-v
-ar/luatex-cache/generic/fonts/otl/lmroman12-bold.luc)
-LaTeX Font Info: Font shape `U/msa/m/n' will be
-(Font) scaled to size 10.42007pt on input line 69.
-LaTeX Font Info: Font shape `U/msa/m/n' will be
-(Font) scaled to size 7.91925pt on input line 69.
-LaTeX Font Info: Font shape `U/msa/m/n' will be
-(Font) scaled to size 6.25204pt on input line 69.
-LaTeX Font Info: Font shape `U/msb/m/n' will be
-(Font) scaled to size 10.42007pt on input line 69.
-LaTeX Font Info: Font shape `U/msb/m/n' will be
-(Font) scaled to size 7.91925pt on input line 69.
-LaTeX Font Info: Font shape `U/msb/m/n' will be
-(Font) scaled to size 6.25204pt on input line 69.
-(load luc: /Users/cstan/Library/texlive/2018/texmf-var/luatex-cache/generic/font
-s/otl/lmmono9-regular.luc)(load luc: /Users/cstan/Library/texlive/2018/texmf-var
-/luatex-cache/generic/fonts/otl/lmroman9-italic.luc) (./session.tex (./RegLangs
-.tex) (./Spec.tex)
-(./Lexer.tex) (./Simplifying.tex) (./Sulzmann.tex) (./Positions.tex)
-(./SizeBound.tex [1
-
-{/usr/local/texlive/2018/texmf-var/fonts/map/pdftex/updmap/pdftex.map}](load luc
-: /Users/cstan/Library/texlive/2018/texmf-var/luatex-cache/generic/fonts/otl/lmr
-oman8-italic.luc)(load luc: /Users/cstan/Library/texlive/2018/texmf-var/luatex-c
-ache/generic/fonts/otl/lmroman7-italic.luc) [2] [3]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [4]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [5]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [6]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [7]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [8]
-Overfull \hbox (10.05669pt too wide) in paragraph at lines 776--778
-[][] \TU/lmr/bx/n/10 ap-ply$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 auto simp add
-$\OT1/zplm/m/n/10 :$ \TU/lmr/m/it/10 bnul-lable[]correctness mkeps[]nullable re
--trieve[]fuse2$\OT1/zplm/m/n/10 )$[]
- []
-
-
-Overfull \hbox (24.20656pt too wide) in paragraph at lines 780--782
-[][] \TU/lmr/bx/n/10 ap-ply $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 smt bnul-lable
-$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 4$
-\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/10 bnul-lable[]correctness erase$\OT1/pplx/m/n
-/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 5$\OT1/zplm/m/n/
-10 )$ \TU/lmr/m/it/10 erase$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/
-m/n/10 ($\TU/lmr/m/it/10 6$\OT1/zplm/m/n/10 )$
- []
-
-
-Overfull \hbox (21.52992pt too wide) in paragraph at lines 786--788
-[][] \TU/lmr/bx/n/10 ap-ply $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 smt ap-pend[]N
-il2 bnul-lable$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10 ($\TU/
-lmr/m/it/10 4$\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/10 bnul-lable[]correctness erase
-$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 5$
-\OT1/zplm/m/n/10 )$
- []
-
-
-Overfull \hbox (1.98334pt too wide) in paragraph at lines 814--816
-[][] \TU/lmr/bx/n/10 ap-ply $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 metis ap-pend[
-]Nil2 bnul-lable[]correctness erase$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 simps$\O
-T1/zplm/m/n/10 ($\TU/lmr/m/it/10 6$\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/10 erase[]f
-use
- []
-
-
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [9]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [10]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [11]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [12]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [13]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [14]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [15]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [16]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [17]
-Overfull \hbox (45.00008pt too wide) in paragraph at lines 1700--1713
-[][] \TU/lmr/bx/n/10 by $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 metis L[]bders[]si
-mp bnul-lable[]correctness lexer$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 simps$\OT1/
-zplm/m/n/10 ($\TU/lmr/m/it/10 1$\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/10 lexer[]corr
-ect[]None
- []
-
-
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [18]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [19]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [20]
-Overfull \hbox (72.57005pt too wide) in paragraph at lines 2035--2037
-[][] \TU/lmr/bx/n/10 ap-ply $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 metis ap-pend
-$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 assoc b2 bnul-lable[]correctness erase[]fus
-e bnul-lable[]Hdbmkeps[]Hd$\OT1/zplm/m/n/10 )$[]
- []
-
-
-Overfull \hbox (10.72014pt too wide) in paragraph at lines 2120--2122
-[][] \TU/lmr/bx/n/10 ap-ply $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 metis L[]erase
-[]AALTs L[]erase[]flts L[]flat[]Prf1 L[]flat[]Prf2
- []
-
-
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [21]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [22]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [23]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [24]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [25]
-Overfull \hbox (6.40976pt too wide) in paragraph at lines 2633--2641
-[][] \TU/lmr/bx/n/10 ap-ply$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 induction ra$ [
-] $r1 rb$ [] $AZERO ar-bi-trary$\OT1/zplm/m/n/10 :$ \TU/lmr/m/it/10 bs1 r1 r2 r
-ule$\OT1/zplm/m/n/10 :$ \TU/lmr/m/it/10 rrewrites$\OT1/pplx/m/n/10 .$\TU/lmr/m/
-it/10 induct$\OT1/zplm/m/n/10 )$[]
- []
-
-
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [26]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [27]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [28]
-Overfull \hbox (9.2633pt too wide) in paragraph at lines 2921--2923
-[][] \TU/lmr/bx/n/10 ap-ply $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 metis ap-pend[
-]Nil flts$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10 ($\TU/lmr/m
-/it/10 2$\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/10 many[]steps[]later rrewrite$\OT1/p
-plx/m/n/10 .$\TU/lmr/m/it/10 intros$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 7$\OT1/z
-plm/m/n/10 )$$)$[]
- []
-
-
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [29]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [30]
-Overfull \hbox (47.09676pt too wide) in paragraph at lines 3142--3144
-[][] \TU/lmr/bx/n/10 ap-ply $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 smt $\OT1/zplm
-/m/n/10 ($\TU/lmr/m/it/10 verit$\OT1/pplx/m/n/10 ,$ \TU/lmr/m/it/10 ccfv[]thres
-hold$\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/10 ap-pend[]Cons ap-pend[]assoc ap-pend[]
-self[]conv2
- []
-
-
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [31]
-Overfull \hbox (6.43988pt too wide) in paragraph at lines 3238--3240
-[][] \TU/lmr/bx/n/10 ap-ply$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 subgoal[]tac AA
-LTs bs1 $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 flts $\OT1/zplm/m/n/10 ($\TU/lmr/m/
-it/10 map bsimp rs$\OT1/zplm/m/n/10 )$$)$ $\U/msa/m/n/10 $$\OMS/zplm/m/n/10 $
- \TU/lmr/m/it/10 AALTs bs1 $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 distinctBy
- []
-
-
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [32]
-Overfull \hbox (27.40672pt too wide) in paragraph at lines 3304--3311
-[][] \TU/lmr/bx/n/10 by $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 smt $\OT1/zplm/m/n
-/10 ($\TU/lmr/m/it/10 z3$\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/10 Un[]iff bnul-lable
-[]correctness in-sert[]iff list$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 set$\OT1/zpl
-m/m/n/10 ($\TU/lmr/m/it/10 2$\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/10 qq3 set[]appen
-d$\OT1/zplm/m/n/10 )$[]
- []
-
-
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [33]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [34]
-Overfull \hbox (11.15656pt too wide) in paragraph at lines 3559--3561
-[][]\TU/lmr/bx/n/10 lemma \TU/lmr/m/it/10 third[]segment[]bnullable$\OT1/zplm/m
-/n/10 :$ $[]$\TU/lmr/m/it/10 bnul-lable $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 AAL
-Ts bs $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 rs1$\OT1/pplx/m/n/10 @$\TU/lmr/m/it/1
-0 rs2$\OT1/pplx/m/n/10 @$\TU/lmr/m/it/10 rs3$\OT1/zplm/m/n/10 )$$)$$\OT1/pplx/m
-/n/10 ;$ $\OMS/zplm/m/n/10 :$\TU/lmr/m/it/10 bnul-
- []
-
-
-Overfull \hbox (30.67331pt too wide) in paragraph at lines 3565--3576
-[][] \TU/lmr/bx/n/10 by $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 metis ap-pend$\OT1
-/pplx/m/n/10 .$\TU/lmr/m/it/10 left[]neutral ap-pend[]Cons bnul-lable$\OT1/pplx
-/m/n/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 1$\OT1/zplm/
-m/n/10 )$ \TU/lmr/m/it/10 bnul-lable[]segment
- []
-
-
-Overfull \hbox (6.39323pt too wide) in paragraph at lines 3580--3582
-[][]\TU/lmr/bx/n/10 lemma \TU/lmr/m/it/10 third[]segment[]bmkeps$\OT1/zplm/m/n/
-10 :$ $[]$\TU/lmr/m/it/10 bnul-lable $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 AALTs
- bs $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 rs1$\OT1/pplx/m/n/10 @$\TU/lmr/m/it/10
-rs2$\OT1/pplx/m/n/10 @$\TU/lmr/m/it/10 rs3$\OT1/zplm/m/n/10 )$$)$$\OT1/pplx/m/n
-/10 ;$ $\OMS/zplm/m/n/10 :$\TU/lmr/m/it/10 bnul-
- []
-
-
-Overfull \hbox (25.44008pt too wide) in paragraph at lines 3601--3603
-[][] \TU/lmr/bx/n/10 ap-ply $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 metis ap-pend$
-\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 assoc in[]set[]conv[]decomp r2 third[]segmen
-t[]bnullable$\OT1/zplm/m/n/10 )$[]
- []
-
-
-Overfull \hbox (31.5301pt too wide) in paragraph at lines 3643--3645
-[][] \TU/lmr/bx/n/10 ap-ply $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 smt $\OT1/zplm
-/m/n/10 ($\TU/lmr/m/it/10 verit$\OT1/pplx/m/n/10 ,$ \TU/lmr/m/it/10 ccfv[]thres
-hold$\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/10 ap-pend[]eq[]append[]conv2 list$\OT1/p
-plx/m/n/10 .$\TU/lmr/m/it/10 set[]intros$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 1$\
-OT1/zplm/m/n/10 )$
- []
-
-
-Overfull \hbox (1.18669pt too wide) in paragraph at lines 3643--3645
-\TU/lmr/m/it/10 qq2 qq3 rewritenul-lable rrewrite$\OT1/pplx/m/n/10 .$\TU/lmr/m/
-it/10 intros$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 8$\OT1/zplm/m/n/10 )$ \TU/lmr/m
-/it/10 self[]append[]conv2 spillbmkep-slistr$\OT1/zplm/m/n/10 )$[]
- []
-
-
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [35]
-Overfull \hbox (12.19328pt too wide) in paragraph at lines 3654--3656
-[][] \TU/lmr/bx/n/10 ap-ply $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 metis ap-pend[
-]Cons ap-pend[]Nil bnul-lable$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 simps$\OT1/zpl
-m/m/n/10 ($\TU/lmr/m/it/10 1$\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/10 bnul-lable[]se
-gment
- []
-
-
-Overfull \hbox (94.37671pt too wide) in paragraph at lines 3657--3664
-[][] \TU/lmr/bx/n/10 by $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 metis bnul-lable$\
-OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 4$\O
-T1/zplm/m/n/10 )$ \TU/lmr/m/it/10 rewrite[]non[]nullable rrewrite$\OT1/pplx/m/n
-/10 .$\TU/lmr/m/it/10 intros$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 10$\OT1/zplm/m/
-n/10 )$ \TU/lmr/m/it/10 third[]segment[]bmkeps$\OT1/zplm/m/n/10 )$[]
- []
-
-
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [36]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [37]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [38]
-Overfull \hbox (17.27003pt too wide) in paragraph at lines 4067--4069
-[][] \TU/lmr/bx/n/10 ap-ply $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 meson con-tex-
-trewrites1 r[]in[]rstar rrewrite$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 intros$\OT1
-/zplm/m/n/10 ($\TU/lmr/m/it/10 11$\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/10 rrewrite$
-\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 intros$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 2$
-\OT1/zplm/m/n/10 )$
- []
-
-
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [39]
-Overfull \hbox (52.87671pt too wide) in paragraph at lines 4155--4157
-[][] \TU/lmr/bx/n/10 ap-ply $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 metis bder$\OT
-1/pplx/m/n/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 4$\OT1
-/zplm/m/n/10 )$ \TU/lmr/m/it/10 bder[]fuse[]list map[]map r[]in[]rstar rrewrite
-$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 intros$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 9
-$\OT1/zplm/m/n/10 )$$)$[]
- []
-
-
-Overfull \hbox (30.75671pt too wide) in paragraph at lines 4158--4160
-[][] \TU/lmr/bx/n/10 ap-ply $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 metis List$\OT
-1/pplx/m/n/10 .$\TU/lmr/m/it/10 map$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 composit
-ionality bder$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10 ($\TU/l
-mr/m/it/10 4$\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/10 bder[]fuse[]list r[]in[]rstar
- []
-
-
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [40]
-Overfull \hbox (31.29991pt too wide) in paragraph at lines 4223--4230
-[][] \TU/lmr/bx/n/10 by $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 metis bders$\OT1/p
-plx/m/n/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 1$\OT1/zp
-lm/m/n/10 )$ \TU/lmr/m/it/10 bders$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 simps$\OT
-1/zplm/m/n/10 ($\TU/lmr/m/it/10 2$\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/10 bders[]si
-mp$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10
-1$\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/10 bders[]simp$\OT1/pplx/m/n/10 .$\TU/lmr/m/
-it/10 simps$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 2$\OT1/zplm/m/n/10 )$
- []
-
-) (./Paper.tex
-
-LaTeX Warning: Citation `Brzozowski1964' on page 41 undefined on input line 78.
-
-
-LaTeX Warning: Citation `Owens2008' on page 41 undefined on input line 92.
-
-
-LaTeX Warning: Citation `Krauss2011' on page 41 undefined on input line 93.
-
-
-LaTeX Warning: Citation `Coquand2012' on page 41 undefined on input line 94.
-
-[41]
-
-LaTeX Warning: Citation `Frisch2004' on page 42 undefined on input line 99.
-
-
-LaTeX Warning: Citation `POSIX' on page 42 undefined on input line 100.
-
-
-LaTeX Warning: Citation `Kuklewicz' on page 42 undefined on input line 100.
-
-
-LaTeX Warning: Citation `OkuiSuzuki2010' on page 42 undefined on input line 100.
-
-
-
-LaTeX Warning: Citation `Sulzmann2014' on page 42 undefined on input line 100.
-
-
-LaTeX Warning: Citation `Vansummeren2006' on page 42 undefined on input line 100
-.
-
-
-LaTeX Warning: Citation `POSIX' on page 42 undefined on input line 117.
-
-[42]
-
-LaTeX Warning: Citation `Sulzmann2014' on page 43 undefined on input line 156.
-
-
-LaTeX Warning: Citation `HosoyaVouillonPierce2005' on page 43 undefined on input
- line 178.
-
-
-LaTeX Warning: Citation `Sulzmann2014' on page 43 undefined on input line 200.
-
-
-LaTeX Warning: Citation `Frisch2004' on page 43 undefined on input line 204.
-
-
-LaTeX Warning: Citation `Sulzmann2014' on page 43 undefined on input line 209.
-
-
-LaTeX Warning: Citation `Kuklewicz' on page 43 undefined on input line 217.
-
-
-LaTeX Warning: Citation `Sulzmann2014' on page 43 undefined on input line 218.
-
-
-LaTeX Warning: Citation `CrashCourse2014' on page 43 undefined on input line 219
-.
-
-[43]
-
-LaTeX Warning: Citation `Sulzmann2014' on page 44 undefined on input line 231.
-
-
-LaTeX Warning: Citation `Vansummeren2006' on page 44 undefined on input line 233
-.
-
-
-LaTeX Warning: Citation `Sulzmann2014' on page 44 undefined on input line 234.
-
-(load luc: /Users/cstan/Library/texlive/2018/texmf-var/luatex-cache/generic/font
-s/otl/lmroman8-regular.luc)
-
-LaTeX Warning: Citation `Sulzmann2014' on page 44 undefined on input line 238.
-
-
-LaTeX Warning: Citation `OkuiSuzuki2010' on page 44 undefined on input line 242.
-
-
-(load luc: /Users/cstan/Library/texlive/2018/texmf-var/luatex-cache/generic/font
-s/otl/lmroman10-bolditalic.luc) [44]
-
-LaTeX Font Warning: Font shape `U/stmry/b/n' undefined
-(Font) using `U/stmry/m/n' instead on input line 321.
-
-
-LaTeX Warning: Citation `Krauss2011' on page 45 undefined on input line 326.
-
-(load luc: /Users/cstan/Library/texlive/2018/texmf-var/luatex-cache/generic/font
-s/otl/lmroman6-regular.luc)
-
-LaTeX Warning: Citation `Brzozowski1964' on page 45 undefined on input line 341.
-
-
-
-Overfull \hbox (1.4002pt too wide) in paragraph at lines 348--369
- []
- []
-
-[45]
-
-LaTeX Warning: Citation `Sulzmann2014' on page 46 undefined on input line 404.
-
-
-LaTeX Warning: Citation `Sulzmann2014' on page 46 undefined on input line 428.
-
-
-LaTeX Warning: Citation `Frisch2004' on page 46 undefined on input line 445.
-
-
-LaTeX Warning: Citation `Sulzmann2014' on page 46 undefined on input line 446.
-
-
-LaTeX Warning: Citation `AusafDyckhoffUrban2016' on page 46 undefined on input l
-ine 478.
-
-[46]
-
-LaTeX Warning: Citation `OkuiSuzuki2010' on page 47 undefined on input line 515.
-
-
-
-LaTeX Warning: Citation `Frisch2004' on page 47 undefined on input line 515.
-
-
-Underfull \hbox (badness 10000) in paragraph at lines 575--576
-
- []
-
-
-LaTeX Warning: Citation `Sulzmann2014' on page 47 undefined on input line 586.
-
-
-LaTeX Warning: Citation `Sulzmann2014' on page 47 undefined on input line 586.
-
-[47]
-
-LaTeX Warning: Citation `Sulzmann2014' on page 48 undefined on input line 606.
-
-
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [48]
-[49]
-
-LaTeX Warning: Citation `Sulzmann2014' on page 50 undefined on input line 714.
-
-
-LaTeX Warning: Citation `Vansummeren2006' on page 50 undefined on input line 721
-.
-
-(load luc: /Users/cstan/Library/texlive/2018/texmf-var/luatex-cache/generic/font
-s/otl/lmroman5-regular.luc) [50] [51] [52]
-
-LaTeX Warning: Citation `Sulzmann2014' on page 53 undefined on input line 934.
-
-
-LaTeX Warning: Citation `Sulzmann2014' on page 53 undefined on input line 936.
-
-
-LaTeX Warning: Citation `OkuiSuzuki2010' on page 53 undefined on input line 943.
-
-
-
-LaTeX Warning: Citation `OkuiSuzukiTech' on page 53 undefined on input line 943.
-
-
-[53] [54]
-
-LaTeX Warning: Citation `OkuiSuzuki2010' on page 55 undefined on input line 1065
-.
-
-
-LaTeX Warning: Citation `OkuiSuzuki2010' on page 55 undefined on input line 1086
-.
-
-
-LaTeX Warning: Citation `OkuiSuzuki2010' on page 55 undefined on input line 1126
-.
-
-[55]
-Overfull \hbox (47.18065pt too wide) in paragraph at lines 1150--1180
-[] []
- []
-
-[56] [57]
-
-LaTeX Warning: Citation `Sulzmann2014' on page 58 undefined on input line 1347.
-
-(load luc: /Users/cstan/Library/texlive/2018/texmf-var/luatex-cache/generic/font
-s/otl/lmroman8-bold.luc)(load luc: /Users/cstan/Library/texlive/2018/texmf-var/l
-uatex-cache/generic/fonts/otl/lmroman5-bold.luc)
-
-LaTeX Warning: Citation `Sulzmann2014' on page 58 undefined on input line 1366.
-
-
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [58]
-Overfull \hbox (4.61574pt too wide) in paragraph at lines 1430--1431
-\OT1/zplm/m/n/10 $$$\OMS/zplm/m/n/10 n$\TU/lmr/m/it/10 dn$\OT1/zplm/m/n/10 $$ $
-\OT1/pplx/m/n/10 &$ $@$$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/10 thm $\OT1/zplm/m/n/1
-0 ($\TU/lmr/m/it/10 rhs$\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/10 in-tern$\OT1/pplx/m
-/n/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 5$\OT1/zplm/m/
-n/10 )$$[$\TU/lmr/m/it/10 of r$[][]$ r$[][]$$\OT1/zplm/m/n/10 ]$$\OMS/zplm/m/n/
-10 g$$n$$n$ $\OT1/pplx/m/n/10 @$$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/10 thm $\OT1/z
-plm/m/n/10 ($\TU/lmr/m/it/10 lhs$\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/10 in-tern$\O
-T1/pplx/m/n/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 6$\OT
-1/zplm/m/n/10 )$$\OMS/zplm/m/n/10 g$
- []
-
-
-Overfull \hbox (43.28516pt too wide) in paragraph at lines 1430--1431
-\OT1/pplx/m/n/10 &$ $\OT1/zplm/m/n/10 $$$\OMS/zplm/m/n/10 n$\TU/lmr/m/it/10 dn$
-\OT1/zplm/m/n/10 $$ $\OT1/pplx/m/n/10 &$ $@$$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/10
- thm $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 rhs$\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/1
-0 in-tern$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10 ($\TU/lmr/m
-/it/10 6$\OT1/zplm/m/n/10 )$$\OMS/zplm/m/n/10 g$$n$$n$ $n$\TU/lmr/m/it/10 end$\
-OMS/zplm/m/n/10 f$\TU/lmr/m/it/10 tabular$\OMS/zplm/m/n/10 g$ $n$\TU/lmr/m/it/1
-0 end$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/10 center$\OMS/zplm/m/n/10 g$ $n$\TU/lmr
-/m/it/10 begin$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/10 center$\OMS/zplm/m/n/10 g$
- []
-
-
-Overfull \hbox (31.37523pt too wide) in paragraph at lines 1430--1431
-\OMS/zplm/m/n/10 n$\TU/lmr/m/it/10 end$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/10 tabul
-ar$\OMS/zplm/m/n/10 g$ $n$\TU/lmr/m/it/10 end$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/1
-0 center$\OMS/zplm/m/n/10 g$ \TU/lmr/m/it/10 Some sim-ple facts about erase $
-\OMS/zplm/m/n/10 n$\TU/lmr/m/it/10 begin$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/10 lem
-ma$\OMS/zplm/m/n/10 g$$n$\TU/lmr/m/it/10 mbox$\OMS/zplm/m/n/10 f$$g$$n$$n$
- []
-
-
-Overfull \hbox (4.57928pt too wide) in paragraph at lines 1430--1431
-\OT1/pplx/m/n/10 @$$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/10 thm $\OT1/zplm/m/n/10 ($
-\TU/lmr/m/it/10 rhs$\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/10 bnul-lable$\OT1/pplx/m/
-n/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 4$\OT1/zplm/m/n
-/10 )$$[$\TU/lmr/m/it/10 of bs r$[][]$ r$[][]$$\OT1/zplm/m/n/10 ]$$\OMS/zplm/m/
-n/10 g$$n$$n$ $\OT1/pplx/m/n/10 @$$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/10 thm $\OT1
-/zplm/m/n/10 ($\TU/lmr/m/it/10 lhs$\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/10 bnul-lab
-le$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10
-5$\OT1/zplm/m/n/10 )$$[$\TU/lmr/m/it/10 of
- []
-
-
-Overfull \hbox (7.7924pt too wide) in paragraph at lines 1430--1431
-\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 lhs$\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/10 bnul
--lable$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10 ($\TU/lmr/m/it
-/10 6$\OT1/zplm/m/n/10 )$$\OMS/zplm/m/n/10 g$ $\OT1/pplx/m/n/10 &$ $\OT1/zplm/m
-/n/10 $$$\OMS/zplm/m/n/10 n$\TU/lmr/m/it/10 dn$\OT1/zplm/m/n/10 $$ $\OT1/pplx/m
-/n/10 &$ $@$$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/10 thm $\OT1/zplm/m/n/10 ($\TU/lmr
-/m/it/10 rhs$\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/10 bnul-lable$\OT1/pplx/m/n/10 .$
-\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 6$\OT1/zplm/m/n/10 )$$
-\OMS/zplm/m/n/10 g$$n$\TU/lmr/m/it/10 medskip$\OMS/zplm/m/n/10 n$$n$
- []
-
-
-Overfull \hbox (34.28575pt too wide) in paragraph at lines 1430--1431
-\TU/lmr/m/it/10 bder$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10
-($\TU/lmr/m/it/10 3$\OT1/zplm/m/n/10 )$$\OMS/zplm/m/n/10 g$ $\OT1/pplx/m/n/10 &
-$ $\OT1/zplm/m/n/10 $$$\OMS/zplm/m/n/10 n$\TU/lmr/m/it/10 dn$\OT1/zplm/m/n/10 $
-$ $\OT1/pplx/m/n/10 &$ $@$$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/10 thm $\OT1/zplm/m/
-n/10 ($\TU/lmr/m/it/10 rhs$\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/10 bder$\OT1/pplx/m
-/n/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 3$\OT1/zplm/m/
-n/10 )$$\OMS/zplm/m/n/10 g$$n$$n$ $\OT1/pplx/m/n/10 @$$\OMS/zplm/m/n/10 f$\TU/l
-mr/m/it/10 thm $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 lhs$\OT1/zplm/m/n/10 )$ \TU/
-lmr/m/it/10 bder$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10 ($\T
-U/lmr/m/it/10 4$\OT1/zplm/m/n/10 )$$[$\TU/lmr/m/it/10 of
- []
-
-
-Overfull \hbox (3.47198pt too wide) in paragraph at lines 1430--1431
-\OT1/pplx/m/n/10 @$$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/10 thm $\OT1/zplm/m/n/10 ($
-\TU/lmr/m/it/10 rhs$\OT1/zplm/m/n/10 )$ \TU/lmr/m/it/10 bmkeps$\OT1/pplx/m/n/10
- .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 4$\OT1/zplm/m/n/10
-)$$\OMS/zplm/m/n/10 g$$n$\TU/lmr/m/it/10 medskip$\OMS/zplm/m/n/10 n$$n$ $n$\TU/
-lmr/m/it/10 end$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/10 tabular$\OMS/zplm/m/n/10 g$
-$n$\TU/lmr/m/it/10 end$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/10 center$\OMS/zplm/m/n/
-10 g$ $\OT1/pplx/m/n/10 @$$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/10 thm
- []
-
-
-Overfull \hbox (2.7528pt too wide) in paragraph at lines 1430--1431
-\OMS/zplm/m/n/10 n$\TU/lmr/m/it/10 noindent Def-i-ni-tion of the bit-coded lexe
-r $\OT1/pplx/m/n/10 @$$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/10 thm blexer[]def$\OMS
-/zplm/m/n/10 g$ $n$\TU/lmr/m/it/10 begin$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/10 t
-heorem$\OMS/zplm/m/n/10 g$
- []
-
-[59]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [60]
-Overfull \hbox (2.67558pt too wide) in paragraph at lines 1434--1443
-\TU/lmr/m/it/10 ac-cord-ing to rules like $\OMS/zplm/m/n/10 n$\TU/lmr/m/it/10
-begin$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/10 equation$\OMS/zplm/m/n/10 g$$n$\TU/lmr
-/m/it/10 label$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/10 Simpl$\OMS/zplm/m/n/10 g$ $n$
-\TU/lmr/m/it/10 begin$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/10 array$\OMS/zplm/m/n/10
- g$$f$\TU/lmr/m/it/10 lcllcllcllcl$\OMS/zplm/m/n/10 g$
- []
-
-
-Overfull \hbox (35.68385pt too wide) in paragraph at lines 1434--1443
-\TU/lmr/m/it/10 F[]SEQ1$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/
-10 ($\TU/lmr/m/it/10 1$\OT1/zplm/m/n/10 )$$\OMS/zplm/m/n/10 g$ $\OT1/pplx/m/n/1
-0 &$ $\OT1/zplm/m/n/10 $$$\OMS/zplm/m/n/10 n$\TU/lmr/m/it/10 dn$\OT1/zplm/m/n/1
-0 $$ $\OT1/pplx/m/n/10 &$ []\TU/lmr/m/it/10 Seq $\OT1/zplm/m/n/10 ($\TU/lmr/m/i
-t/10 f$[][]$ $\OT1/zplm/m/n/10 ($$)$$)$ $($\TU/lmr/m/it/10 f$[][]$ v$\OT1/zplm/
-m/n/10 )$[]$\OMS/zplm/m/n/10 n$$n$ $\OT1/pplx/m/n/10 @$$\OMS/zplm/m/n/10 f$\TU/
-lmr/m/it/10 thm $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 lhs$\OT1/zplm/m/n/10 )$ \TU
-/lmr/m/it/10 F[]SEQ2$\OT1/pplx/m/n/10 .$\TU/lmr/m/it/10 simps$\OT1/zplm/m/n/10
-($\TU/lmr/m/it/10 1$\OT1/zplm/m/n/10 )$$\OMS/zplm/m/n/10 g$
- []
-
-
-Overfull \hbox (18.66588pt too wide) in paragraph at lines 1434--1443
-\OT1/pplx/m/n/10 &$ $@$$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/10 thm L[]fst[]simp$\OT
-1/zplm/m/n/10 [$\TU/lmr/m/it/10 symmetric$\OT1/zplm/m/n/10 ]$$\OMS/zplm/m/n/10
-g$$n$$n$ $\OT1/zplm/m/n/10 ($\TU/lmr/m/it/10 2$\OT1/zplm/m/n/10 )$ $\OT1/pplx/m
-/n/10 &$ $@$$\OMS/zplm/m/n/10 f$\TU/lmr/m/it/10 thm$\OT1/zplm/m/n/10 [$\TU/lmr/
-m/it/10 mode$\OT1/zplm/m/n/10 =$\TU/lmr/m/it/10 IfThen$\OT1/zplm/m/n/10 ]$ \TU/
-lmr/m/it/10 Posix[]simp$\OMS/zplm/m/n/10 g$
- []
-
-
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [61]
-Underfull \vbox (badness 10000) has occurred while \output is active []
-
- [62]
-[63] [64]
-No file root.bbl.
-)) [65]
-Package atveryend Info: Empty hook `BeforeClearDocument' on input line 96.
-Package atveryend Info: Empty hook `AfterLastShipout' on input line 96.
- (./root.aux)
-Package atveryend Info: Executing hook `AtVeryEndDocument' on input line 96.
-Package atveryend Info: Executing hook `AtEndAfterFileList' on input line 96.
-Package rerunfilecheck Info: File `root.out' has not changed.
-(rerunfilecheck) Checksum: F68486A1B88421815EBCC3CB75200893;512.
-
-
-LaTeX Font Warning: Size substitutions with differences
-(Font) up to 2.26395pt have occurred.
-
-
-LaTeX Font Warning: Some font shapes were not available, defaults substituted.
-
-Package atveryend Info: Empty hook `AtVeryVeryEnd' on input line 96.
-)
-
-Here is how much of LuaTeX's memory you used:
- 29101 strings out of 494413
- 156265,794899 words of node,token memory allocated
- 1590 words of node memory still in use:
- 9 hlist, 2 vlist, 2 rule, 11 disc, 21 glue, 7 kern, 98 glyph, 33 attribute, 5
-7 glue_spec, 33 attribute_list, 1 write nodes
- avail lists: 2:6992,3:291,4:52,5:2358,6:47,7:12582,8:52,9:593,10:23,11:398
- 31954 multiletter control sequences out of 65536+600000
- 282 fonts using 23215399 bytes
- 55i,23n,68p,45885b,993s stack positions out of 5000i,500n,10000p,200000b,100000s
-</usr/local/texlive/2018/texmf-dist/fonts/opentype/public/lm/lmroman7-regular.ot
-f></usr/local/texlive/2018/texmf-dist/fonts/opentype/public/lm/lmroman8-regular.
-otf></usr/local/texlive/2018/texmf-dist/fonts/opentype/public/lm/lmroman7-italic
-.otf></usr/local/texlive/2018/texmf-dist/fonts/opentype/public/lm/lmroman10-ital
-ic.otf></usr/local/texlive/2018/texmf-dist/fonts/opentype/public/lm/lmroman10-bo
-ld.otf>{/usr/local/texlive/2018/texmf-dist/fonts/enc/dvips/base/8r.enc}</usr/loc
-al/texlive/2018/texmf-dist/fonts/opentype/public/lm/lmroman9-italic.otf></usr/lo
-cal/texlive/2018/texmf-dist/fonts/opentype/public/lm/lmroman9-bold.otf></usr/loc
-al/texlive/2018/texmf-dist/fonts/opentype/public/lm/lmmono9-regular.otf></usr/lo
-cal/texlive/2018/texmf-dist/fonts/opentype/public/lm/lmroman9-regular.otf></usr/
-local/texlive/2018/texmf-dist/fonts/opentype/public/lm/lmroman10-regular.otf></u
-sr/local/texlive/2018/texmf-dist/fonts/opentype/public/lm/lmroman12-bold.otf></u
-sr/local/texlive/2018/texmf-dist/fonts/type1/public/amsfonts/cm/cmbsy10.pfb></us
-r/local/texlive/2018/texmf-dist/fonts/type1/public/amsfonts/cm/cmex10.pfb></usr/
-local/texlive/2018/texmf-dist/fonts/type1/public/amsfonts/cm/cmmi10.pfb></usr/lo
-cal/texlive/2018/texmf-dist/fonts/type1/public/amsfonts/cm/cmr10.pfb></usr/local
-/texlive/2018/texmf-dist/fonts/type1/public/amsfonts/cm/cmsy10.pfb></usr/local/t
-exlive/2018/texmf-dist/fonts/type1/public/amsfonts/symbols/msam10.pfb></usr/loca
-l/texlive/2018/texmf-dist/fonts/type1/public/amsfonts/symbols/msbm10.pfb></usr/l
-ocal/texlive/2018/texmf-dist/fonts/type1/public/stmaryrd/stmary10.pfb></usr/loca
-l/texlive/2018/texmf-dist/fonts/type1/urw/palatino/uplr8a.pfb></usr/local/texliv
-e/2018/texmf-dist/fonts/type1/urw/palatino/uplri8a.pfb>
-Output written on root.pdf (65 pages, 375799 bytes).
-
-PDF statistics: 527 PDF objects out of 1000 (max. 8388607)
- 412 compressed objects within 5 object streams
- 107 named destinations out of 1000 (max. 131072)
- 88 words of extra memory for PDF output out of 10000 (max. 100000000)
-
--- a/thys2/Journal/root.out Fri Jan 07 22:25:26 2022 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,9 +0,0 @@
-\BOOKMARK [1][-]{section.1.1}{Bit-Encodings}{}% 1
-\BOOKMARK [1][-]{section.1.2}{Annotated Regular Expressions}{}% 2
-\BOOKMARK [1][-]{section.1.3}{Introduction}{}% 3
-\BOOKMARK [1][-]{section.1.4}{Preliminaries}{}% 4
-\BOOKMARK [1][-]{section.1.5}{POSIX Regular Expression Matching}{}% 5
-\BOOKMARK [1][-]{section.1.6}{Ordering of Values according to Okui and Suzuki}{}% 6
-\BOOKMARK [1][-]{section.1.7}{Bitcoded Lexing}{}% 7
-\BOOKMARK [1][-]{section.1.8}{Optimisations}{}% 8
-\BOOKMARK [1][-]{section.1.9}{HERE}{}% 9
Binary file thys2/Journal/root.pdf has changed
--- a/thys2/Journal/session.tex Fri Jan 07 22:25:26 2022 +0000
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,8 +0,0 @@
-\input{RegLangs.tex}
-\input{Spec.tex}
-\input{Lexer.tex}
-\input{Simplifying.tex}
-\input{Sulzmann.tex}
-\input{Positions.tex}
-\input{SizeBound.tex}
-\input{Paper.tex}
Binary file thys2/Journal/session_graph.pdf has changed
Binary file thys2/journal.pdf has changed