added all files in Journal folder
authorChengsong
Mon, 01 Nov 2021 10:40:21 +0000
changeset 369 e00950ba4514
parent 368 56781ad291cf
child 370 5499ba68188c
added all files in Journal folder
thys2/Journal/Lexer.tex
thys2/Journal/Paper.tex
thys2/Journal/Paper.thy~
thys2/Journal/PaperExt.thy
thys2/Journal/Positions.tex
thys2/Journal/RegLangs.tex
thys2/Journal/Simplifying.tex
thys2/Journal/SizeBound.tex
thys2/Journal/Spec.tex
thys2/Journal/Sulzmann.tex
thys2/Journal/comment.sty
thys2/Journal/document/llncs.cls
thys2/Journal/document/root.bib
thys2/Journal/document/root.tex
thys2/Journal/isabelle.sty
thys2/Journal/isabellesym.sty
thys2/Journal/isabelletags.sty
thys2/Journal/llncs.cls
thys2/Journal/pdfsetup.sty
thys2/Journal/railsetup.sty
thys2/Journal/root.aux
thys2/Journal/root.bib
thys2/Journal/root.log
thys2/Journal/root.out
thys2/Journal/root.pdf
thys2/Journal/session.tex
thys2/Journal/session_graph.pdf
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys2/Journal/Paper.tex	Mon Nov 01 10:40:21 2021 +0000
@@ -0,0 +1,3061 @@
+%
+\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{Introduction%
+}
+\isamarkuptrue%
+%
+\endisatagdocument
+{\isafolddocument}%
+%
+\isadelimdocument
+%
+\endisadelimdocument
+%
+\begin{isamarkuptext}%
+This works builds on previous work by Ausaf and Urban using 
+regular expression'd bit-coded derivatives to do lexing that 
+is both fast and satisfied 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}\ {\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}
+
+This lemma links the derivative of a bit-coded regular expression with
+the regular expression itself before the derivative. 
+
+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
+%
+\end{isabellebody}%
+\endinput
+%:%file=~/Dropbox/Workspace/journalpaper/lexing/thys2/Journal/Paper.thy%:%
+%:%50=136%:%
+%:%62=141%:%
+%:%63=142%:%
+%:%64=143%:%
+%:%65=144%:%
+%:%66=145%:%
+%:%67=146%:%
+%:%68=147%:%
+%:%69=148%:%
+%:%70=149%:%
+%:%71=150%:%
+%:%72=151%:%
+%:%73=152%:%
+%:%74=153%:%
+%:%75=154%:%
+%:%76=155%:%
+%:%77=156%:%
+%:%78=157%:%
+%:%79=158%:%
+%:%80=159%:%
+%:%81=160%:%
+%:%81=161%:%
+%:%82=162%:%
+%:%83=163%:%
+%:%84=164%:%
+%:%85=165%:%
+%:%86=166%:%
+%:%87=167%:%
+%:%88=168%:%
+%:%89=169%:%
+%:%90=170%:%
+%:%91=171%:%
+%:%92=172%:%
+%:%93=173%:%
+%:%94=174%:%
+%:%95=175%:%
+%:%96=176%:%
+%:%97=177%:%
+%:%98=178%:%
+%:%99=179%:%
+%:%100=180%:%
+%:%101=181%:%
+%:%102=182%:%
+%:%103=183%:%
+%:%104=184%:%
+%:%105=185%:%
+%:%106=186%:%
+%:%107=187%:%
+%:%108=188%:%
+%:%109=189%:%
+%:%110=190%:%
+%:%111=191%:%
+%:%112=192%:%
+%:%113=193%:%
+%:%114=194%:%
+%:%115=195%:%
+%:%116=196%:%
+%:%117=197%:%
+%:%118=198%:%
+%:%119=199%:%
+%:%120=200%:%
+%:%121=201%:%
+%:%122=202%:%
+%:%123=203%:%
+%:%124=204%:%
+%:%125=205%:%
+%:%126=206%:%
+%:%127=207%:%
+%:%128=208%:%
+%:%129=209%:%
+%:%130=210%:%
+%:%131=211%:%
+%:%132=212%:%
+%:%133=213%:%
+%:%134=214%:%
+%:%135=215%:%
+%:%136=216%:%
+%:%137=217%:%
+%:%138=218%:%
+%:%139=219%:%
+%:%140=220%:%
+%:%141=221%:%
+%:%142=222%:%
+%:%143=223%:%
+%:%144=224%:%
+%:%145=225%:%
+%:%145=226%:%
+%:%146=227%:%
+%:%147=228%:%
+%:%148=229%:%
+%:%149=230%:%
+%:%150=231%:%
+%:%151=232%:%
+%:%152=233%:%
+%:%153=234%:%
+%:%154=235%:%
+%:%155=236%:%
+%:%156=237%:%
+%:%157=238%:%
+%:%158=239%:%
+%:%159=240%:%
+%:%160=241%:%
+%:%161=242%:%
+%:%162=243%:%
+%:%163=244%:%
+%:%164=245%:%
+%:%165=246%:%
+%:%166=247%:%
+%:%167=248%:%
+%:%168=249%:%
+%:%169=250%:%
+%:%170=251%:%
+%:%171=252%:%
+%:%172=253%:%
+%:%173=254%:%
+%:%174=255%:%
+%:%175=256%:%
+%:%176=257%:%
+%:%177=258%:%
+%:%178=259%:%
+%:%179=260%:%
+%:%180=261%:%
+%:%181=262%:%
+%:%182=263%:%
+%:%183=264%:%
+%:%184=265%:%
+%:%185=266%:%
+%:%186=267%:%
+%:%187=268%:%
+%:%188=269%:%
+%:%189=270%:%
+%:%190=271%:%
+%:%191=272%:%
+%:%192=273%:%
+%:%193=274%:%
+%:%194=275%:%
+%:%195=276%:%
+%:%196=277%:%
+%:%197=278%:%
+%:%198=279%:%
+%:%199=280%:%
+%:%200=281%:%
+%:%201=282%:%
+%:%202=283%:%
+%:%203=284%:%
+%:%204=285%:%
+%:%205=286%:%
+%:%206=287%:%
+%:%207=288%:%
+%:%208=289%:%
+%:%209=290%:%
+%:%210=291%:%
+%:%211=292%:%
+%:%212=293%:%
+%:%213=294%:%
+%:%214=295%:%
+%:%215=296%:%
+%:%216=297%:%
+%:%217=298%:%
+%:%218=299%:%
+%:%219=300%:%
+%:%220=301%:%
+%:%221=302%:%
+%:%222=303%:%
+%:%223=304%:%
+%:%224=305%:%
+%:%225=306%:%
+%:%226=307%:%
+%:%227=308%:%
+%:%228=309%:%
+%:%229=310%:%
+%:%230=311%:%
+%:%231=312%:%
+%:%232=313%:%
+%:%233=314%:%
+%:%234=315%:%
+%:%235=316%:%
+%:%236=317%:%
+%:%237=318%:%
+%:%238=319%:%
+%:%239=320%:%
+%:%240=321%:%
+%:%241=322%:%
+%:%242=323%:%
+%:%243=324%:%
+%:%244=325%:%
+%:%245=326%:%
+%:%246=327%:%
+%:%247=328%:%
+%:%248=329%:%
+%:%249=330%:%
+%:%250=331%:%
+%:%259=335%:%
+%:%271=337%:%
+%:%272=338%:%
+%:%272=339%:%
+%:%273=340%:%
+%:%274=341%:%
+%:%275=342%:%
+%:%276=343%:%
+%:%277=344%:%
+%:%278=345%:%
+%:%279=346%:%
+%:%280=347%:%
+%:%281=348%:%
+%:%282=349%:%
+%:%283=350%:%
+%:%284=351%:%
+%:%285=352%:%
+%:%286=353%:%
+%:%287=354%:%
+%:%288=355%:%
+%:%289=356%:%
+%:%290=357%:%
+%:%291=358%:%
+%:%292=359%:%
+%:%293=360%:%
+%:%294=361%:%
+%:%295=362%:%
+%:%296=363%:%
+%:%297=364%:%
+%:%298=365%:%
+%:%299=366%:%
+%:%300=367%:%
+%:%301=368%:%
+%:%302=369%:%
+%:%303=370%:%
+%:%304=371%:%
+%:%305=372%:%
+%:%306=373%:%
+%:%307=374%:%
+%:%308=375%:%
+%:%309=376%:%
+%:%310=377%:%
+%:%311=378%:%
+%:%311=379%:%
+%:%312=380%:%
+%:%313=381%:%
+%:%314=382%:%
+%:%315=383%:%
+%:%316=384%:%
+%:%316=385%:%
+%:%317=386%:%
+%:%318=387%:%
+%:%319=388%:%
+%:%320=389%:%
+%:%321=390%:%
+%:%322=391%:%
+%:%323=392%:%
+%:%324=393%:%
+%:%325=394%:%
+%:%326=395%:%
+%:%327=396%:%
+%:%328=397%:%
+%:%329=398%:%
+%:%330=399%:%
+%:%331=400%:%
+%:%332=401%:%
+%:%333=402%:%
+%:%334=403%:%
+%:%335=404%:%
+%:%336=405%:%
+%:%337=406%:%
+%:%338=407%:%
+%:%339=408%:%
+%:%340=409%:%
+%:%341=410%:%
+%:%342=411%:%
+%:%343=412%:%
+%:%344=413%:%
+%:%345=414%:%
+%:%346=415%:%
+%:%347=416%:%
+%:%348=417%:%
+%:%349=418%:%
+%:%350=419%:%
+%:%351=420%:%
+%:%352=421%:%
+%:%353=422%:%
+%:%354=423%:%
+%:%355=424%:%
+%:%356=425%:%
+%:%357=426%:%
+%:%358=427%:%
+%:%359=428%:%
+%:%360=429%:%
+%:%361=430%:%
+%:%362=431%:%
+%:%363=432%:%
+%:%364=433%:%
+%:%365=434%:%
+%:%366=435%:%
+%:%367=436%:%
+%:%368=437%:%
+%:%369=438%:%
+%:%370=439%:%
+%:%371=440%:%
+%:%372=441%:%
+%:%373=442%:%
+%:%374=443%:%
+%:%375=444%:%
+%:%376=445%:%
+%:%377=446%:%
+%:%378=447%:%
+%:%379=448%:%
+%:%380=449%:%
+%:%381=450%:%
+%:%382=451%:%
+%:%383=452%:%
+%:%384=453%:%
+%:%385=454%:%
+%:%386=455%:%
+%:%387=456%:%
+%:%388=457%:%
+%:%389=458%:%
+%:%390=459%:%
+%:%391=460%:%
+%:%392=461%:%
+%:%393=462%:%
+%:%394=463%:%
+%:%395=464%:%
+%:%396=465%:%
+%:%397=466%:%
+%:%398=467%:%
+%:%399=468%:%
+%:%400=469%:%
+%:%401=470%:%
+%:%402=471%:%
+%:%403=472%:%
+%:%404=473%:%
+%:%405=474%:%
+%:%414=478%:%
+%:%426=482%:%
+%:%427=483%:%
+%:%428=484%:%
+%:%429=485%:%
+%:%430=486%:%
+%:%431=487%:%
+%:%432=488%:%
+%:%433=489%:%
+%:%434=490%:%
+%:%435=491%:%
+%:%436=492%:%
+%:%437=493%:%
+%:%438=494%:%
+%:%439=495%:%
+%:%440=496%:%
+%:%441=497%:%
+%:%442=498%:%
+%:%443=499%:%
+%:%444=500%:%
+%:%445=501%:%
+%:%446=502%:%
+%:%447=503%:%
+%:%448=504%:%
+%:%449=505%:%
+%:%450=506%:%
+%:%451=507%:%
+%:%452=508%:%
+%:%453=509%:%
+%:%454=510%:%
+%:%455=511%:%
+%:%456=512%:%
+%:%457=513%:%
+%:%458=514%:%
+%:%459=515%:%
+%:%460=516%:%
+%:%461=517%:%
+%:%462=518%:%
+%:%463=519%:%
+%:%464=520%:%
+%:%465=521%:%
+%:%466=522%:%
+%:%467=523%:%
+%:%468=524%:%
+%:%469=525%:%
+%:%470=526%:%
+%:%471=527%:%
+%:%472=528%:%
+%:%473=529%:%
+%:%474=530%:%
+%:%475=531%:%
+%:%476=532%:%
+%:%477=533%:%
+%:%478=534%:%
+%:%479=535%:%
+%:%480=536%:%
+%:%481=537%:%
+%:%482=538%:%
+%:%483=539%:%
+%:%484=540%:%
+%:%485=541%:%
+%:%486=542%:%
+%:%487=543%:%
+%:%488=544%:%
+%:%489=545%:%
+%:%490=546%:%
+%:%491=547%:%
+%:%492=548%:%
+%:%493=549%:%
+%:%494=550%:%
+%:%495=551%:%
+%:%496=552%:%
+%:%497=553%:%
+%:%498=554%:%
+%:%499=555%:%
+%:%500=556%:%
+%:%501=557%:%
+%:%502=558%:%
+%:%503=559%:%
+%:%504=560%:%
+%:%505=561%:%
+%:%506=562%:%
+%:%507=563%:%
+%:%508=564%:%
+%:%509=565%:%
+%:%510=566%:%
+%:%511=567%:%
+%:%512=568%:%
+%:%513=569%:%
+%:%514=570%:%
+%:%515=571%:%
+%:%516=572%:%
+%:%517=573%:%
+%:%518=574%:%
+%:%519=575%:%
+%:%520=576%:%
+%:%521=577%:%
+%:%522=578%:%
+%:%523=579%:%
+%:%524=580%:%
+%:%525=581%:%
+%:%526=582%:%
+%:%527=583%:%
+%:%528=584%:%
+%:%529=585%:%
+%:%530=586%:%
+%:%531=587%:%
+%:%532=588%:%
+%:%533=589%:%
+%:%534=590%:%
+%:%535=591%:%
+%:%536=592%:%
+%:%537=593%:%
+%:%538=594%:%
+%:%539=595%:%
+%:%540=596%:%
+%:%540=597%:%
+%:%541=598%:%
+%:%541=599%:%
+%:%542=600%:%
+%:%543=601%:%
+%:%543=602%:%
+%:%544=603%:%
+%:%545=604%:%
+%:%546=605%:%
+%:%547=606%:%
+%:%548=607%:%
+%:%549=608%:%
+%:%550=609%:%
+%:%551=610%:%
+%:%552=611%:%
+%:%553=612%:%
+%:%554=613%:%
+%:%555=614%:%
+%:%556=615%:%
+%:%557=616%:%
+%:%558=617%:%
+%:%559=618%:%
+%:%560=619%:%
+%:%561=620%:%
+%:%562=621%:%
+%:%563=622%:%
+%:%564=623%:%
+%:%565=624%:%
+%:%566=625%:%
+%:%567=626%:%
+%:%568=627%:%
+%:%569=628%:%
+%:%570=629%:%
+%:%571=630%:%
+%:%572=631%:%
+%:%573=632%:%
+%:%574=633%:%
+%:%575=634%:%
+%:%576=635%:%
+%:%577=636%:%
+%:%578=637%:%
+%:%579=638%:%
+%:%580=639%:%
+%:%581=640%:%
+%:%582=641%:%
+%:%583=642%:%
+%:%584=643%:%
+%:%585=644%:%
+%:%586=645%:%
+%:%587=646%:%
+%:%588=647%:%
+%:%589=648%:%
+%:%590=649%:%
+%:%591=650%:%
+%:%592=651%:%
+%:%593=652%:%
+%:%594=653%:%
+%:%595=654%:%
+%:%596=655%:%
+%:%597=656%:%
+%:%598=657%:%
+%:%599=658%:%
+%:%599=659%:%
+%:%599=660%:%
+%:%600=661%:%
+%:%601=662%:%
+%:%602=663%:%
+%:%603=664%:%
+%:%604=665%:%
+%:%605=666%:%
+%:%606=667%:%
+%:%607=668%:%
+%:%608=669%:%
+%:%608=670%:%
+%:%609=671%:%
+%:%610=672%:%
+%:%611=673%:%
+%:%612=674%:%
+%:%613=675%:%
+%:%614=676%:%
+%:%615=677%:%
+%:%616=678%:%
+%:%617=679%:%
+%:%618=680%:%
+%:%619=681%:%
+%:%620=682%:%
+%:%621=683%:%
+%:%622=684%:%
+%:%623=685%:%
+%:%624=686%:%
+%:%625=687%:%
+%:%626=688%:%
+%:%627=689%:%
+%:%628=690%:%
+%:%629=691%:%
+%:%630=692%:%
+%:%631=693%:%
+%:%632=694%:%
+%:%633=695%:%
+%:%634=696%:%
+%:%635=697%:%
+%:%636=698%:%
+%:%637=699%:%
+%:%638=700%:%
+%:%639=701%:%
+%:%640=702%:%
+%:%641=703%:%
+%:%642=704%:%
+%:%643=705%:%
+%:%643=706%:%
+%:%644=707%:%
+%:%645=708%:%
+%:%646=709%:%
+%:%647=710%:%
+%:%648=711%:%
+%:%649=712%:%
+%:%650=713%:%
+%:%650=714%:%
+%:%651=715%:%
+%:%652=716%:%
+%:%653=717%:%
+%:%653=718%:%
+%:%654=719%:%
+%:%655=720%:%
+%:%656=721%:%
+%:%657=722%:%
+%:%658=723%:%
+%:%659=724%:%
+%:%660=725%:%
+%:%661=726%:%
+%:%662=727%:%
+%:%663=728%:%
+%:%664=729%:%
+%:%664=730%:%
+%:%665=731%:%
+%:%666=732%:%
+%:%667=733%:%
+%:%668=734%:%
+%:%668=735%:%
+%:%669=736%:%
+%:%670=737%:%
+%:%671=738%:%
+%:%672=739%:%
+%:%673=740%:%
+%:%674=741%:%
+%:%675=742%:%
+%:%676=743%:%
+%:%677=744%:%
+%:%678=745%:%
+%:%679=746%:%
+%:%680=747%:%
+%:%681=748%:%
+%:%682=749%:%
+%:%683=750%:%
+%:%684=751%:%
+%:%685=752%:%
+%:%686=753%:%
+%:%687=754%:%
+%:%688=755%:%
+%:%689=756%:%
+%:%690=757%:%
+%:%691=758%:%
+%:%692=759%:%
+%:%693=760%:%
+%:%694=761%:%
+%:%695=762%:%
+%:%696=763%:%
+%:%697=764%:%
+%:%698=765%:%
+%:%699=766%:%
+%:%700=767%:%
+%:%701=768%:%
+%:%702=769%:%
+%:%703=770%:%
+%:%704=771%:%
+%:%705=772%:%
+%:%706=773%:%
+%:%707=774%:%
+%:%708=775%:%
+%:%709=776%:%
+%:%710=777%:%
+%:%711=778%:%
+%:%712=779%:%
+%:%713=780%:%
+%:%714=781%:%
+%:%715=782%:%
+%:%716=783%:%
+%:%717=784%:%
+%:%718=785%:%
+%:%719=786%:%
+%:%720=787%:%
+%:%721=788%:%
+%:%722=789%:%
+%:%723=790%:%
+%:%724=791%:%
+%:%725=792%:%
+%:%726=793%:%
+%:%726=794%:%
+%:%727=795%:%
+%:%728=796%:%
+%:%729=797%:%
+%:%730=798%:%
+%:%731=799%:%
+%:%732=800%:%
+%:%733=801%:%
+%:%734=802%:%
+%:%735=803%:%
+%:%736=804%:%
+%:%737=805%:%
+%:%738=806%:%
+%:%739=807%:%
+%:%740=808%:%
+%:%741=809%:%
+%:%742=810%:%
+%:%743=811%:%
+%:%744=812%:%
+%:%745=813%:%
+%:%746=814%:%
+%:%747=815%:%
+%:%748=816%:%
+%:%749=817%:%
+%:%750=818%:%
+%:%751=819%:%
+%:%752=820%:%
+%:%753=821%:%
+%:%754=822%:%
+%:%755=823%:%
+%:%756=824%:%
+%:%757=825%:%
+%:%758=826%:%
+%:%759=827%:%
+%:%759=828%:%
+%:%760=829%:%
+%:%761=830%:%
+%:%762=831%:%
+%:%763=832%:%
+%:%764=833%:%
+%:%765=834%:%
+%:%766=835%:%
+%:%767=836%:%
+%:%768=837%:%
+%:%769=838%:%
+%:%770=839%:%
+%:%771=840%:%
+%:%772=841%:%
+%:%773=842%:%
+%:%774=843%:%
+%:%775=844%:%
+%:%776=845%:%
+%:%777=846%:%
+%:%778=847%:%
+%:%779=848%:%
+%:%780=849%:%
+%:%781=850%:%
+%:%782=851%:%
+%:%783=852%:%
+%:%784=853%:%
+%:%785=854%:%
+%:%786=855%:%
+%:%787=856%:%
+%:%788=857%:%
+%:%789=858%:%
+%:%789=859%:%
+%:%790=860%:%
+%:%791=861%:%
+%:%792=862%:%
+%:%793=863%:%
+%:%794=864%:%
+%:%795=865%:%
+%:%796=866%:%
+%:%797=867%:%
+%:%797=868%:%
+%:%798=869%:%
+%:%799=870%:%
+%:%800=871%:%
+%:%801=872%:%
+%:%802=873%:%
+%:%803=874%:%
+%:%804=875%:%
+%:%805=876%:%
+%:%806=877%:%
+%:%807=878%:%
+%:%808=879%:%
+%:%809=880%:%
+%:%810=881%:%
+%:%811=882%:%
+%:%812=883%:%
+%:%813=884%:%
+%:%814=885%:%
+%:%815=886%:%
+%:%816=887%:%
+%:%817=888%:%
+%:%818=889%:%
+%:%819=890%:%
+%:%820=891%:%
+%:%821=892%:%
+%:%822=893%:%
+%:%823=894%:%
+%:%824=895%:%
+%:%825=896%:%
+%:%826=897%:%
+%:%827=898%:%
+%:%828=899%:%
+%:%829=900%:%
+%:%830=901%:%
+%:%831=902%:%
+%:%832=903%:%
+%:%833=904%:%
+%:%834=905%:%
+%:%834=906%:%
+%:%834=907%:%
+%:%835=908%:%
+%:%835=909%:%
+%:%835=910%:%
+%:%836=911%:%
+%:%837=912%:%
+%:%837=913%:%
+%:%838=914%:%
+%:%839=915%:%
+%:%840=916%:%
+%:%841=917%:%
+%:%842=918%:%
+%:%843=919%:%
+%:%844=920%:%
+%:%845=921%:%
+%:%846=922%:%
+%:%847=923%:%
+%:%848=924%:%
+%:%849=925%:%
+%:%850=926%:%
+%:%851=927%:%
+%:%852=928%:%
+%:%853=929%:%
+%:%854=930%:%
+%:%855=931%:%
+%:%856=932%:%
+%:%857=933%:%
+%:%858=934%:%
+%:%859=935%:%
+%:%860=936%:%
+%:%861=937%:%
+%:%862=938%:%
+%:%863=939%:%
+%:%864=940%:%
+%:%865=941%:%
+%:%866=942%:%
+%:%867=943%:%
+%:%868=944%:%
+%:%869=945%:%
+%:%870=946%:%
+%:%871=947%:%
+%:%871=948%:%
+%:%872=949%:%
+%:%873=950%:%
+%:%874=951%:%
+%:%874=952%:%
+%:%874=953%:%
+%:%874=954%:%
+%:%875=955%:%
+%:%876=956%:%
+%:%877=957%:%
+%:%878=958%:%
+%:%879=959%:%
+%:%880=960%:%
+%:%881=961%:%
+%:%882=962%:%
+%:%883=963%:%
+%:%884=964%:%
+%:%885=965%:%
+%:%886=966%:%
+%:%887=967%:%
+%:%888=968%:%
+%:%889=969%:%
+%:%890=970%:%
+%:%891=971%:%
+%:%892=972%:%
+%:%893=973%:%
+%:%894=974%:%
+%:%895=975%:%
+%:%896=976%:%
+%:%897=977%:%
+%:%898=978%:%
+%:%899=979%:%
+%:%900=980%:%
+%:%901=981%:%
+%:%902=982%:%
+%:%903=983%:%
+%:%904=984%:%
+%:%905=985%:%
+%:%906=986%:%
+%:%907=987%:%
+%:%908=988%:%
+%:%909=989%:%
+%:%910=990%:%
+%:%911=991%:%
+%:%920=995%:%
+%:%932=999%:%
+%:%933=1000%:%
+%:%934=1001%:%
+%:%935=1002%:%
+%:%936=1003%:%
+%:%937=1004%:%
+%:%938=1005%:%
+%:%939=1006%:%
+%:%940=1007%:%
+%:%941=1008%:%
+%:%942=1009%:%
+%:%943=1010%:%
+%:%944=1011%:%
+%:%945=1012%:%
+%:%946=1013%:%
+%:%947=1014%:%
+%:%948=1015%:%
+%:%949=1016%:%
+%:%950=1017%:%
+%:%951=1018%:%
+%:%952=1019%:%
+%:%953=1020%:%
+%:%954=1021%:%
+%:%955=1022%:%
+%:%956=1023%:%
+%:%957=1024%:%
+%:%958=1025%:%
+%:%959=1026%:%
+%:%960=1027%:%
+%:%961=1028%:%
+%:%962=1029%:%
+%:%963=1030%:%
+%:%964=1031%:%
+%:%965=1032%:%
+%:%966=1033%:%
+%:%967=1034%:%
+%:%968=1035%:%
+%:%969=1036%:%
+%:%970=1037%:%
+%:%971=1038%:%
+%:%972=1039%:%
+%:%973=1040%:%
+%:%974=1041%:%
+%:%975=1042%:%
+%:%976=1043%:%
+%:%977=1044%:%
+%:%978=1045%:%
+%:%979=1046%:%
+%:%980=1047%:%
+%:%981=1048%:%
+%:%982=1049%:%
+%:%983=1050%:%
+%:%984=1051%:%
+%:%985=1052%:%
+%:%986=1053%:%
+%:%987=1054%:%
+%:%988=1055%:%
+%:%989=1056%:%
+%:%990=1057%:%
+%:%991=1058%:%
+%:%992=1059%:%
+%:%993=1060%:%
+%:%994=1061%:%
+%:%995=1062%:%
+%:%996=1063%:%
+%:%997=1064%:%
+%:%998=1065%:%
+%:%999=1066%:%
+%:%1000=1067%:%
+%:%1001=1068%:%
+%:%1002=1069%:%
+%:%1003=1070%:%
+%:%1004=1071%:%
+%:%1005=1072%:%
+%:%1006=1073%:%
+%:%1007=1074%:%
+%:%1008=1075%:%
+%:%1009=1076%:%
+%:%1010=1077%:%
+%:%1011=1078%:%
+%:%1012=1079%:%
+%:%1013=1080%:%
+%:%1014=1081%:%
+%:%1015=1082%:%
+%:%1016=1083%:%
+%:%1017=1084%:%
+%:%1018=1085%:%
+%:%1019=1086%:%
+%:%1020=1087%:%
+%:%1021=1088%:%
+%:%1022=1089%:%
+%:%1023=1090%:%
+%:%1024=1091%:%
+%:%1025=1092%:%
+%:%1026=1093%:%
+%:%1027=1094%:%
+%:%1028=1095%:%
+%:%1029=1096%:%
+%:%1030=1097%:%
+%:%1031=1098%:%
+%:%1032=1099%:%
+%:%1033=1100%:%
+%:%1034=1101%:%
+%:%1035=1102%:%
+%:%1036=1103%:%
+%:%1037=1104%:%
+%:%1038=1105%:%
+%:%1039=1106%:%
+%:%1040=1107%:%
+%:%1041=1108%:%
+%:%1042=1109%:%
+%:%1043=1110%:%
+%:%1044=1111%:%
+%:%1045=1112%:%
+%:%1046=1113%:%
+%:%1047=1114%:%
+%:%1047=1115%:%
+%:%1048=1116%:%
+%:%1049=1117%:%
+%:%1050=1118%:%
+%:%1051=1119%:%
+%:%1052=1120%:%
+%:%1053=1121%:%
+%:%1054=1122%:%
+%:%1055=1123%:%
+%:%1056=1124%:%
+%:%1057=1125%:%
+%:%1058=1126%:%
+%:%1058=1127%:%
+%:%1059=1128%:%
+%:%1060=1129%:%
+%:%1061=1130%:%
+%:%1062=1131%:%
+%:%1063=1132%:%
+%:%1064=1133%:%
+%:%1065=1134%:%
+%:%1066=1135%:%
+%:%1067=1136%:%
+%:%1068=1137%:%
+%:%1069=1138%:%
+%:%1070=1139%:%
+%:%1071=1140%:%
+%:%1072=1141%:%
+%:%1073=1142%:%
+%:%1074=1143%:%
+%:%1075=1144%:%
+%:%1076=1145%:%
+%:%1077=1146%:%
+%:%1078=1147%:%
+%:%1079=1148%:%
+%:%1080=1149%:%
+%:%1081=1150%:%
+%:%1082=1151%:%
+%:%1083=1152%:%
+%:%1084=1153%:%
+%:%1085=1154%:%
+%:%1086=1155%:%
+%:%1087=1156%:%
+%:%1088=1157%:%
+%:%1089=1158%:%
+%:%1090=1159%:%
+%:%1091=1160%:%
+%:%1092=1161%:%
+%:%1093=1162%:%
+%:%1094=1163%:%
+%:%1095=1164%:%
+%:%1096=1165%:%
+%:%1097=1166%:%
+%:%1098=1167%:%
+%:%1099=1168%:%
+%:%1100=1169%:%
+%:%1101=1170%:%
+%:%1102=1171%:%
+%:%1103=1172%:%
+%:%1104=1173%:%
+%:%1105=1174%:%
+%:%1106=1175%:%
+%:%1107=1176%:%
+%:%1108=1177%:%
+%:%1109=1178%:%
+%:%1110=1179%:%
+%:%1111=1180%:%
+%:%1112=1181%:%
+%:%1112=1182%:%
+%:%1112=1183%:%
+%:%1113=1184%:%
+%:%1114=1185%:%
+%:%1114=1186%:%
+%:%1114=1187%:%
+%:%1114=1188%:%
+%:%1115=1189%:%
+%:%1116=1190%:%
+%:%1117=1191%:%
+%:%1118=1192%:%
+%:%1119=1193%:%
+%:%1119=1194%:%
+%:%1120=1195%:%
+%:%1121=1196%:%
+%:%1122=1197%:%
+%:%1123=1198%:%
+%:%1124=1199%:%
+%:%1125=1200%:%
+%:%1126=1201%:%
+%:%1127=1202%:%
+%:%1128=1203%:%
+%:%1129=1204%:%
+%:%1130=1205%:%
+%:%1131=1206%:%
+%:%1132=1207%:%
+%:%1133=1208%:%
+%:%1134=1209%:%
+%:%1135=1210%:%
+%:%1136=1211%:%
+%:%1137=1212%:%
+%:%1138=1213%:%
+%:%1139=1214%:%
+%:%1140=1215%:%
+%:%1141=1216%:%
+%:%1142=1217%:%
+%:%1143=1218%:%
+%:%1144=1219%:%
+%:%1145=1220%:%
+%:%1146=1221%:%
+%:%1147=1222%:%
+%:%1148=1223%:%
+%:%1149=1224%:%
+%:%1150=1225%:%
+%:%1151=1226%:%
+%:%1152=1227%:%
+%:%1153=1228%:%
+%:%1154=1229%:%
+%:%1155=1230%:%
+%:%1156=1231%:%
+%:%1157=1232%:%
+%:%1158=1233%:%
+%:%1159=1234%:%
+%:%1160=1235%:%
+%:%1161=1236%:%
+%:%1162=1237%:%
+%:%1163=1238%:%
+%:%1164=1239%:%
+%:%1165=1240%:%
+%:%1166=1241%:%
+%:%1167=1242%:%
+%:%1167=1243%:%
+%:%1168=1244%:%
+%:%1168=1245%:%
+%:%1169=1246%:%
+%:%1169=1247%:%
+%:%1170=1248%:%
+%:%1171=1249%:%
+%:%1172=1250%:%
+%:%1173=1251%:%
+%:%1174=1252%:%
+%:%1175=1253%:%
+%:%1176=1254%:%
+%:%1176=1255%:%
+%:%1177=1256%:%
+%:%1177=1257%:%
+%:%1178=1258%:%
+%:%1178=1259%:%
+%:%1179=1260%:%
+%:%1180=1261%:%
+%:%1181=1262%:%
+%:%1182=1263%:%
+%:%1183=1264%:%
+%:%1184=1265%:%
+%:%1185=1266%:%
+%:%1186=1267%:%
+%:%1187=1268%:%
+%:%1188=1269%:%
+%:%1189=1270%:%
+%:%1190=1271%:%
+%:%1191=1272%:%
+%:%1192=1273%:%
+%:%1193=1274%:%
+%:%1194=1275%:%
+%:%1195=1276%:%
+%:%1196=1277%:%
+%:%1197=1278%:%
+%:%1198=1279%:%
+%:%1199=1280%:%
+%:%1200=1281%:%
+%:%1201=1282%:%
+%:%1202=1283%:%
+%:%1203=1284%:%
+%:%1204=1285%:%
+%:%1205=1286%:%
+%:%1206=1287%:%
+%:%1207=1288%:%
+%:%1208=1289%:%
+%:%1209=1290%:%
+%:%1209=1291%:%
+%:%1210=1292%:%
+%:%1211=1293%:%
+%:%1212=1294%:%
+%:%1213=1295%:%
+%:%1214=1296%:%
+%:%1215=1297%:%
+%:%1216=1298%:%
+%:%1216=1299%:%
+%:%1217=1300%:%
+%:%1218=1301%:%
+%:%1219=1302%:%
+%:%1219=1303%:%
+%:%1220=1304%:%
+%:%1221=1305%:%
+%:%1222=1306%:%
+%:%1222=1307%:%
+%:%1223=1308%:%
+%:%1223=1309%:%
+%:%1224=1310%:%
+%:%1224=1311%:%
+%:%1225=1312%:%
+%:%1226=1313%:%
+%:%1227=1314%:%
+%:%1227=1315%:%
+%:%1228=1316%:%
+%:%1228=1317%:%
+%:%1229=1318%:%
+%:%1229=1319%:%
+%:%1229=1320%:%
+%:%1229=1321%:%
+%:%1230=1322%:%
+%:%1230=1323%:%
+%:%1231=1324%:%
+%:%1232=1325%:%
+%:%1232=1327%:%
+%:%1232=1328%:%
+%:%1232=1329%:%
+%:%1232=1330%:%
+%:%1232=1331%:%
+%:%1233=1332%:%
+%:%1233=1333%:%
+%:%1234=1334%:%
+%:%1234=1335%:%
+%:%1235=1336%:%
+%:%1235=1337%:%
+%:%1236=1338%:%
+%:%1237=1339%:%
+%:%1238=1340%:%
+%:%1239=1341%:%
+%:%1239=1342%:%
+%:%1240=1343%:%
+%:%1240=1344%:%
+%:%1241=1345%:%
+%:%1242=1346%:%
+%:%1242=1347%:%
+%:%1243=1348%:%
+%:%1244=1349%:%
+%:%1245=1350%:%
+%:%1246=1351%:%
+%:%1247=1352%:%
+%:%1248=1353%:%
+%:%1249=1354%:%
+%:%1250=1355%:%
+%:%1251=1356%:%
+%:%1252=1357%:%
+%:%1253=1358%:%
+%:%1254=1359%:%
+%:%1255=1360%:%
+%:%1256=1361%:%
+%:%1257=1362%:%
+%:%1258=1363%:%
+%:%1259=1364%:%
+%:%1260=1365%:%
+%:%1261=1366%:%
+%:%1262=1367%:%
+%:%1263=1368%:%
+%:%1264=1369%:%
+%:%1265=1370%:%
+%:%1266=1371%:%
+%:%1267=1372%:%
+%:%1268=1373%:%
+%:%1269=1374%:%
+%:%1270=1375%:%
+%:%1271=1376%:%
+%:%1272=1377%:%
+%:%1273=1378%:%
+%:%1274=1379%:%
+%:%1275=1380%:%
+%:%1276=1381%:%
+%:%1277=1382%:%
+%:%1278=1383%:%
+%:%1279=1384%:%
+%:%1280=1385%:%
+%:%1281=1386%:%
+%:%1282=1387%:%
+%:%1283=1388%:%
+%:%1284=1389%:%
+%:%1285=1390%:%
+%:%1286=1391%:%
+%:%1287=1392%:%
+%:%1288=1393%:%
+%:%1289=1394%:%
+%:%1290=1395%:%
+%:%1291=1396%:%
+%:%1292=1397%:%
+%:%1293=1398%:%
+%:%1302=1402%:%
+%:%1314=1409%:%
+%:%1315=1410%:%
+%:%1316=1411%:%
+%:%1317=1412%:%
+%:%1318=1413%:%
+%:%1319=1414%:%
+%:%1320=1415%:%
+%:%1329=1420%:%
+%:%1341=1424%:%
+%:%1342=1425%:%
+%:%1343=1426%:%
+%:%1344=1427%:%
+%:%1345=1428%:%
+%:%1346=1429%:%
+%:%1347=1430%:%
+%:%1348=1431%:%
+%:%1349=1432%:%
+%:%1350=1433%:%
+%:%1351=1434%:%
+%:%1352=1435%:%
+%:%1353=1436%:%
+%:%1354=1437%:%
+%:%1355=1438%:%
+%:%1356=1439%:%
+%:%1357=1440%:%
+%:%1358=1441%:%
+%:%1359=1442%:%
+%:%1360=1443%:%
+%:%1361=1444%:%
+%:%1362=1445%:%
+%:%1363=1446%:%
+%:%1364=1447%:%
+%:%1365=1448%:%
+%:%1366=1449%:%
+%:%1367=1450%:%
+%:%1368=1451%:%
+%:%1369=1452%:%
+%:%1370=1453%:%
+%:%1371=1454%:%
+%:%1372=1455%:%
+%:%1373=1456%:%
+%:%1374=1457%:%
+%:%1375=1458%:%
+%:%1376=1459%:%
+%:%1377=1460%:%
+%:%1378=1461%:%
+%:%1379=1462%:%
+%:%1380=1463%:%
+%:%1381=1464%:%
+%:%1382=1465%:%
+%:%1383=1466%:%
+%:%1384=1467%:%
+%:%1385=1468%:%
+%:%1386=1469%:%
+%:%1387=1470%:%
+%:%1388=1471%:%
+%:%1389=1472%:%
+%:%1390=1473%:%
+%:%1391=1474%:%
+%:%1392=1475%:%
+%:%1393=1476%:%
+%:%1394=1477%:%
+%:%1395=1478%:%
+%:%1396=1479%:%
+%:%1397=1480%:%
+%:%1398=1481%:%
+%:%1399=1482%:%
+%:%1400=1483%:%
+%:%1401=1484%:%
+%:%1402=1485%:%
+%:%1403=1486%:%
+%:%1404=1487%:%
+%:%1405=1488%:%
+%:%1406=1489%:%
+%:%1407=1490%:%
+%:%1408=1491%:%
+%:%1409=1492%:%
+%:%1410=1493%:%
+%:%1411=1494%:%
+%:%1412=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%:%
+%:%1430=1629%:%
+%:%1431=1630%:%
+%:%1432=1631%:%
+%:%1433=1632%:%
+%:%1434=1633%:%
+%:%1434=1794%:%
+%:%1435=1795%:%
+%:%1436=1796%:%
+%:%1437=1797%:%
+%:%1438=1798%:%
+%:%1439=1799%:%
+%:%1440=1800%:%
+%:%1441=1801%:%
+%:%1442=1802%:%
+%:%1443=1803%:%
+%:%1444=1804%:%
+%:%1445=1805%:%
+%:%1446=1806%:%
+%:%1447=1807%:%
+%:%1448=1808%:%
+%:%1449=1809%:%
+%:%1450=1810%:%
+%:%1451=1811%:%
+%:%1452=1812%:%
+%:%1453=1813%:%
+%:%1453=1814%:%
+%:%1454=1815%:%
+%:%1454=1816%:%
+%:%1454=1817%:%
+%:%1455=1818%:%
+%:%1455=1819%:%
+%:%1456=1820%:%
+%:%1457=1821%:%
+%:%1458=1822%:%
+%:%1459=1823%:%
+%:%1460=1824%:%
+%:%1461=1825%:%
+%:%1462=1826%:%
+%:%1463=1827%:%
+%:%1464=1828%:%
+%:%1465=1829%:%
+%:%1466=1830%:%
+%:%1467=1831%:%
+%:%1468=1832%:%
+%:%1469=1833%:%
+%:%1470=1834%:%
+%:%1471=1835%:%
+%:%1472=1836%:%
+%:%1472=1837%:%
+%:%1473=1838%:%
+%:%1474=1839%:%
+%:%1475=1840%:%
+%:%1475=1841%:%
+%:%1475=1842%:%
+%:%1476=1843%:%
+%:%1477=1844%:%
+%:%1478=1845%:%
+%:%1478=1846%:%
+%:%1479=1847%:%
+%:%1479=1848%:%
+%:%1480=1849%:%
+%:%1481=1850%:%
+%:%1482=1851%:%
+%:%1483=1852%:%
+%:%1484=1853%:%
+%:%1485=1854%:%
+%:%1486=1855%:%
+%:%1487=1856%:%
+%:%1488=1857%:%
+%:%1489=1858%:%
+%:%1490=1859%:%
+%:%1491=1860%:%
+%:%1492=1861%:%
+%:%1493=1862%:%
+%:%1494=1863%:%
+%:%1495=1864%:%
+%:%1496=1865%:%
+%:%1505=1870%:%
+%:%1517=1874%:%
+%:%1518=1875%:%
+%:%1519=1876%:%
+%:%1520=1877%:%
+%:%1521=1878%:%
+%:%1522=1879%:%
+%:%1523=1880%:%
+%:%1524=1881%:%
+%:%1525=1882%:%
+%:%1526=1883%:%
+%:%1527=1884%:%
+%:%1528=1885%:%
+%:%1529=1886%:%
+%:%1530=1887%:%
+%:%1531=1888%:%
+%:%1532=1889%:%
+%:%1533=1890%:%
+%:%1534=1891%:%
+%:%1535=1892%:%
+%:%1536=1893%:%
+%:%1537=1894%:%
+%:%1538=1895%:%
+%:%1539=1896%:%
+%:%1540=1897%:%
+%:%1541=1898%:%
+%:%1542=1899%:%
+%:%1543=1900%:%
+%:%1544=1901%:%
+%:%1545=1902%:%
+%:%1546=1903%:%
+%:%1547=1904%:%
+%:%1548=1905%:%
+%:%1549=1906%:%
+%:%1550=1907%:%
+%:%1551=1908%:%
+%:%1552=1909%:%
+%:%1553=1910%:%
+%:%1554=1911%:%
+%:%1555=1912%:%
+%:%1556=1913%:%
+%:%1557=1914%:%
+%:%1558=1915%:%
+%:%1559=1916%:%
+%:%1560=1917%:%
+%:%1561=1918%:%
+%:%1562=1919%:%
+%:%1563=1920%:%
+%:%1564=1921%:%
+%:%1565=1922%:%
+%:%1566=1923%:%
+%:%1567=1924%:%
+%:%1568=1925%:%
+%:%1569=1926%:%
+%:%1582=1932%:%
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys2/Journal/Paper.thy~	Mon Nov 01 10:40:21 2021 +0000
@@ -0,0 +1,2051 @@
+(*<*)
+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>Introduction\<close>
+
+
+text \<open>
+Trying out after lualatex.
+
+
+uhuhuhu
+This works builds on previous work by Ausaf and Urban using 
+regular expression'd bit-coded derivatives to do lexing that 
+is both fast and satisfied 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] bder_retrieve}
+\end{lemma}
+
+This lemma links the regular expression
+
+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}
+
+\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 "ACH 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 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 "ACH 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{center}
+  \begin{tabular}{llcl}
+  1) & @{thm (lhs) erase.simps(1)} & $\dn$ & @{thm (rhs) erase.simps(1)}\\
+  2) & @{thm (lhs) erase.simps(2)[of bs]} & $\dn$ & @{thm (rhs) erase.simps(2)[of bs]}\\
+  3) & @{thm (lhs) erase.simps(3)[of bs]} & $\dn$ & @{thm (rhs) erase.simps(3)[of bs]}\\
+  4a) & @{term "erase (AALTs bs [])"} & $\dn$ & @{term ZERO}\\
+  4b) & @{term "erase (AALTs bs [r])"} & $\dn$ & @{term "erase r"}\\
+  4c) & @{term "erase (AALTs bs (r\<^sub>1#r\<^sub>2#rs))"} & $\dn$ & 
+        @{term "ALT (erase r\<^sub>1) (erase (AALTs bs (r\<^sub>2#rs)))"}\\
+  5) & @{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"]}\\
+  6) & @{thm (lhs) erase.simps(6)[of bs]} & $\dn$ & @{thm (rhs) erase.simps(6)[of bs]}\\
+  \end{tabular}
+  \end{center}
+
+  \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 "ACH 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
+(*>*)
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys2/Journal/PaperExt.thy	Mon Nov 01 10:40:21 2021 +0000
@@ -0,0 +1,323 @@
+(*<*)
+theory PaperExt
+imports 
+   (*"../LexerExt"*)
+   (*"../PositionsExt"*)
+   "~~/src/HOL/Library/LaTeXsugar"
+begin
+(*>*)
+
+(*
+declare [[show_question_marks = false]]
+declare [[eta_contract = false]]
+
+abbreviation 
+  "der_syn r c \<equiv> der c r"
+
+abbreviation 
+  "ders_syn r s \<equiv> ders s r"
+
+
+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 
+  CHAR ("_" [1000] 80) and
+  ALT ("_ + _" [77,77] 78) and
+  SEQ ("_ \<cdot> _" [77,77] 78) and
+  STAR ("_\<^sup>\<star>" [78] 78) and
+  NTIMES ("_\<^bsup>'{_'}\<^esup>" [78, 50] 80) and
+  FROMNTIMES ("_\<^bsup>'{_..'}\<^esup>" [78, 50] 80) and
+  UPNTIMES ("_\<^bsup>'{.._'}\<^esup>" [78, 50] 80) and
+  NMTIMES ("_\<^bsup>'{_.._'}\<^esup>" [78, 50,50] 80) 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 _" [1000] 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 _ _ _" [81,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
+  DUMMY ("\<^latex>\<open>\\underline{\\hspace{2mm}}\<close>")
+*)  
+(*>*)
+
+(*
+section {* Extensions*}
+
+text {*
+
+  A strong point in favour of
+  Sulzmann and Lu's algorithm is that it can be extended in various
+  ways.  If we are interested in tokenising a string, then we need to not just
+  split up the string into tokens, but also ``classify'' the tokens (for
+  example whether it is a keyword or an identifier). This can be done with
+  only minor modifications to the algorithm by introducing \emph{record
+  regular expressions} and \emph{record values} (for example
+  \cite{Sulzmann2014b}):
+
+  \begin{center}  
+  @{text "r :="}
+  @{text "..."} $\mid$
+  @{text "(l : r)"} \qquad\qquad
+  @{text "v :="}
+  @{text "..."} $\mid$
+  @{text "(l : v)"}
+  \end{center}
+  
+  \noindent where @{text l} is a label, say a string, @{text r} a regular
+  expression and @{text v} a value. All functions can be smoothly extended
+  to these regular expressions and values. For example \mbox{@{text "(l :
+  r)"}} is nullable iff @{term r} is, and so on. The purpose of the record
+  regular expression is to mark certain parts of a regular expression and
+  then record in the calculated value which parts of the string were matched
+  by this part. The label can then serve as classification for the tokens.
+  For this recall the regular expression @{text "(r\<^bsub>key\<^esub> + r\<^bsub>id\<^esub>)\<^sup>\<star>"} for
+  keywords and identifiers from the Introduction. With the record regular
+  expression we can form \mbox{@{text "((key : r\<^bsub>key\<^esub>) + (id : r\<^bsub>id\<^esub>))\<^sup>\<star>"}}
+  and then traverse the calculated value and only collect the underlying
+  strings in record values. With this we obtain finite sequences of pairs of
+  labels and strings, for example
+
+  \[@{text "(l\<^sub>1 : s\<^sub>1), ..., (l\<^sub>n : s\<^sub>n)"}\]
+  
+  \noindent from which tokens with classifications (keyword-token,
+  identifier-token and so on) can be extracted.
+
+  In the context of POSIX matching, it is also interesting to study additional
+  constructors about bounded-repetitions of regular expressions. For this let
+  us extend the results from the previous section to the following four
+  additional regular expression constructors:
+
+  \begin{center}
+  \begin{tabular}{lcrl@ {\hspace{12mm}}l}
+  @{text r} & @{text ":="} & $\ldots\mid$ & @{term "NTIMES r n"} & exactly-@{text n}-times\\
+            &              & $\mid$   & @{term "UPNTIMES r n"} & upto-@{text n}-times\\
+	    &              & $\mid$   & @{term "FROMNTIMES r n"} & from-@{text n}-times\\
+	    &              & $\mid$   & @{term "NMTIMES r n m"} & between-@{text nm}-times\\
+  \end{tabular}
+  \end{center}
+
+  \noindent
+  We will call them \emph{bounded regular expressions}.
+  With the help of the power operator (definition ommited) for sets of strings, the languages
+  recognised by these regular expression can be defined in Isabelle as follows:
+
+  \begin{center}
+  \begin{tabular}{lcl} 
+  @{thm (lhs) L.simps(8)} & $\dn$ & @{thm (rhs) L.simps(8)}\\
+  @{thm (lhs) L.simps(7)} & $\dn$ & @{thm (rhs) L.simps(7)}\\
+  @{thm (lhs) L.simps(9)} & $\dn$ & @{thm (rhs) L.simps(9)}\\
+  @{thm (lhs) L.simps(10)} & $\dn$ & @{thm (rhs) L.simps(10)}\\
+  \end{tabular}
+  \end{center}
+
+  \noindent This definition implies that in the last clause @{term
+  "NMTIMES r n m"} matches no string in case @{term "m < n"}, because
+  then the interval @{term "{n..m}"} is empty.  While the language
+  recognised by these regular expressions is straightforward, some
+  care is needed for how to define the corresponding lexical
+  values. First, with a slight abuse of language, we will (re)use
+  values of the form @{term "Stars vs"} for values inhabited in
+  bounded regular expressions. Second, we need to introduce inductive
+  rules for extending our inhabitation relation shown on
+  Page~\ref{prfintros}, from which we then derived our notion of
+  lexical values. Given the rule for @{term "STAR r"}, the rule for
+  @{term "UPNTIMES r n"} just requires additionally that the length of
+  the list of values must be smaller or equal to @{term n}:
+
+  \begin{center}
+  @{thm[mode=Rule] Prf.intros(7)[of "vs"]}
+  \end{center}
+
+  \noindent Like in the @{term "STAR r"}-rule, we require with the
+  left-premise that some non-empty part of the string is `chipped'
+  away by \emph{every} value in @{text vs}, that means the corresponding
+  values do not flatten to the empty string. In the rule for @{term
+  "NTIMES r n"} (that is exactly-@{term n}-times @{text r}) we clearly need to require
+  that the length of the list of values equals to @{text n}. But enforcing
+  that every of these @{term n} values `chipps' away some part of a string
+  would be too strong. 
+
+
+  \begin{center}
+  @{thm[mode=Rule] Prf.intros(8)[of "vs\<^sub>1" r "vs\<^sub>2"]}
+  \end{center}
+
+  \begin{center}
+  \begin{tabular}{lcl} 
+  @{thm (lhs) der.simps(8)} & $\dn$ & @{thm (rhs) der.simps(8)}\\
+  @{thm (lhs) der.simps(7)} & $\dn$ & @{thm (rhs) der.simps(7)}\\
+  @{thm (lhs) der.simps(9)} & $\dn$ & @{thm (rhs) der.simps(9)}\\
+  @{thm (lhs) der.simps(10)} & $\dn$ & @{thm (rhs) der.simps(10)}\\
+  \end{tabular}
+  \end{center} 
+
+
+  \begin{center}
+  \begin{tabular}{lcl}
+  @{thm (lhs) mkeps.simps(5)} & $\dn$ & @{thm (rhs) mkeps.simps(5)}\\
+  @{thm (lhs) mkeps.simps(6)} & $\dn$ & @{thm (rhs) mkeps.simps(6)}\\
+  @{thm (lhs) mkeps.simps(7)} & $\dn$ & @{thm (rhs) mkeps.simps(7)}\\
+  @{thm (lhs) mkeps.simps(8)} & $\dn$ & @{thm (rhs) mkeps.simps(8)}\\
+  \end{tabular}
+  \end{center}
+
+  \begin{center}
+  \begin{tabular}{lcl}
+  @{thm (lhs) injval.simps(8)} & $\dn$ & @{thm (rhs) injval.simps(8)}\\
+  @{thm (lhs) injval.simps(9)} & $\dn$ & @{thm (rhs) injval.simps(9)}\\
+  @{thm (lhs) injval.simps(10)} & $\dn$ & @{thm (rhs) injval.simps(10)}\\
+  @{thm (lhs) injval.simps(11)} & $\dn$ & @{thm (rhs) injval.simps(11)}\\
+  \end{tabular}
+  \end{center}
+  
+
+  @{thm [mode=Rule] Posix_NTIMES1[of "s\<^sub>1" r v "s\<^sub>2"]}
+
+  @{thm [mode=Rule] Posix_NTIMES2}
+
+  @{thm [mode=Rule] Posix_UPNTIMES1[of "s\<^sub>1" r v "s\<^sub>2"]}
+
+  @{thm [mode=Rule] Posix_UPNTIMES2}
+
+  @{thm [mode=Rule] Posix_FROMNTIMES1[of "s\<^sub>1" r v "s\<^sub>2"]}
+
+  @{thm [mode=Rule] Posix_FROMNTIMES3[of "s\<^sub>1" r v "s\<^sub>2"]}
+
+  @{thm [mode=Rule] Posix_FROMNTIMES2}
+
+  @{thm [mode=Rule] Posix_NMTIMES1[of "s\<^sub>1" r v "s\<^sub>2"]}
+
+  @{thm [mode=Rule] Posix_NMTIMES3[of "s\<^sub>1" r v "s\<^sub>2"]}
+
+  @{thm [mode=Rule] Posix_NMTIMES2}
+
+
+
+   @{term "\<Sum> i \<in> {m..n} . P i"}  @{term "\<Sum> i \<in> {..n} . P i"}
+  @{term "\<Union> i \<in> {m..n} . P i"}  @{term "\<Union> i \<in> {..n} . P i"}
+   @{term "\<Union> i \<in> {0::nat..n} . P i"}
+
+
+
+*}
+
+
+
+section {* The Correctness Argument by Sulzmann and Lu\label{argu} *}
+
+text {*
+%  \newcommand{\greedy}{\succcurlyeq_{gr}}
+ \newcommand{\posix}{>}
+
+  An extended version of \cite{Sulzmann2014} is available at the website of
+  its first author; this includes some ``proofs'', claimed in
+  \cite{Sulzmann2014} to be ``rigorous''. Since these are evidently not in
+  final form, we make no comment thereon, preferring to give general reasons
+  for our belief that the approach of \cite{Sulzmann2014} is problematic.
+  Their central definition is an ``ordering relation'' defined by the
+  rules (slightly adapted to fit our notation):
+
+  ??
+
+  \noindent The idea behind the rules (A1) and (A2), for example, is that a
+  @{text Left}-value is bigger than a @{text Right}-value, if the underlying
+  string of the @{text Left}-value is longer or of equal length to the
+  underlying string of the @{text Right}-value. The order is reversed,
+  however, if the @{text Right}-value can match a longer string than a
+  @{text Left}-value. In this way the POSIX value is supposed to be the
+  biggest value for a given string and regular expression.
+
+  Sulzmann and Lu explicitly refer to the paper \cite{Frisch2004} by Frisch
+  and Cardelli from where they have taken the idea for their correctness
+  proof. Frisch and Cardelli introduced a similar ordering for GREEDY
+  matching and they showed that their GREEDY matching algorithm always
+  produces a maximal element according to this ordering (from all possible
+  solutions). The only difference between their GREEDY ordering and the
+  ``ordering'' by Sulzmann and Lu is that GREEDY always prefers a @{text
+  Left}-value over a @{text Right}-value, no matter what the underlying
+  string is. This seems to be only a very minor difference, but it has
+  drastic consequences in terms of what properties both orderings enjoy.
+  What is interesting for our purposes is that the properties reflexivity,
+  totality and transitivity for this GREEDY ordering can be proved
+  relatively easily by induction.
+*}
+*)
+
+section {* Conclusion *}
+
+text {*
+
+  We have implemented the POSIX value calculation algorithm introduced by
+  Sulzmann and Lu
+  \cite{Sulzmann2014}. Our implementation is nearly identical to the
+  original and all modifications we introduced are harmless (like our char-clause for
+  @{text inj}). We have proved this algorithm to be correct, but correct
+  according to our own specification of what POSIX values are. Our
+  specification (inspired from work by Vansummeren \cite{Vansummeren2006}) appears to be
+  much simpler than in \cite{Sulzmann2014} and our proofs are nearly always
+  straightforward. We have attempted to formalise the original proof
+  by Sulzmann and Lu \cite{Sulzmann2014}, but we believe it contains
+  unfillable gaps. In the online version of \cite{Sulzmann2014}, the authors
+  already acknowledge some small problems, but our experience suggests
+  that there are more serious problems. 
+  
+  Having proved the correctness of the POSIX lexing algorithm in
+  \cite{Sulzmann2014}, which lessons have we learned? Well, this is a
+  perfect example for the importance of the \emph{right} definitions. We
+  have (on and off) explored mechanisations as soon as first versions
+  of \cite{Sulzmann2014} appeared, but have made little progress with
+  turning the relatively detailed proof sketch in \cite{Sulzmann2014} into a
+  formalisable proof. Having seen \cite{Vansummeren2006} and adapted the
+  POSIX definition given there for the algorithm by Sulzmann and Lu made all
+  the difference: the proofs, as said, are nearly straightforward. The
+  question remains whether the original proof idea of \cite{Sulzmann2014},
+  potentially using our result as a stepping stone, can be made to work?
+  Alas, we really do not know despite considerable effort.
+
+
+  Closely related to our work is an automata-based lexer formalised by
+  Nipkow \cite{Nipkow98}. This lexer also splits up strings into longest
+  initial substrings, but Nipkow's algorithm is not completely
+  computational. The algorithm by Sulzmann and Lu, in contrast, can be
+  implemented with ease in any functional language. A bespoke lexer for the
+  Imp-language is formalised in Coq as part of the Software Foundations book
+  by Pierce et al \cite{Pierce2015}. The disadvantage of such bespoke lexers is that they
+  do not generalise easily to more advanced features.
+  Our formalisation is available from the Archive of Formal Proofs \cite{aduAFP16}
+  under \url{http://www.isa-afp.org/entries/Posix-Lexing.shtml}.\medskip
+
+ \noindent
+  {\bf Acknowledgements:}
+  We are very grateful to Martin Sulzmann for his comments on our work and 
+  moreover for patiently explaining to us the details in \cite{Sulzmann2014}. We
+  also received very helpful comments from James Cheney and anonymous referees.
+  %  \small
+  \bibliographystyle{plain}
+  \bibliography{root}
+
+*}
+
+(*<*)
+end
+(*>*)
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys2/Journal/SizeBound.tex	Mon Nov 01 10:40:21 2021 +0000
@@ -0,0 +1,7263 @@
+%
+\begin{isabellebody}%
+\setisabellecontext{SizeBound}%
+%
+\isadelimtheory
+\isanewline
+%
+\endisadelimtheory
+%
+\isatagtheory
+\isacommand{theory}\isamarkupfalse%
+\ SizeBound\isanewline
+\ \ \isakeyword{imports}\ {\isachardoublequoteopen}Lexer{\isachardoublequoteclose}\ \isanewline
+\isakeyword{begin}%
+\endisatagtheory
+{\isafoldtheory}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\isadelimdocument
+%
+\endisadelimdocument
+%
+\isatagdocument
+%
+\isamarkupsection{Bit-Encodings%
+}
+\isamarkuptrue%
+%
+\endisatagdocument
+{\isafolddocument}%
+%
+\isadelimdocument
+%
+\endisadelimdocument
+\isacommand{datatype}\isamarkupfalse%
+\ bit\ {\isacharequal}{\kern0pt}\ Z\ {\isacharbar}{\kern0pt}\ S\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ code\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}val\ {\isasymRightarrow}\ bit\ list{\isachardoublequoteclose}\isanewline
+\isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}code\ Void\ {\isacharequal}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}code\ {\isacharparenleft}{\kern0pt}Char\ c{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}code\ {\isacharparenleft}{\kern0pt}Left\ v{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ Z\ {\isacharhash}{\kern0pt}\ {\isacharparenleft}{\kern0pt}code\ v{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}code\ {\isacharparenleft}{\kern0pt}Right\ v{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ S\ {\isacharhash}{\kern0pt}\ {\isacharparenleft}{\kern0pt}code\ v{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}code\ {\isacharparenleft}{\kern0pt}Seq\ v{\isadigit{1}}\ v{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}code\ v{\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharat}{\kern0pt}\ {\isacharparenleft}{\kern0pt}code\ v{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}code\ {\isacharparenleft}{\kern0pt}Stars\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}S{\isacharbrackright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}code\ {\isacharparenleft}{\kern0pt}Stars\ {\isacharparenleft}{\kern0pt}v\ {\isacharhash}{\kern0pt}\ vs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ \ {\isacharparenleft}{\kern0pt}Z\ {\isacharhash}{\kern0pt}\ code\ v{\isacharparenright}{\kern0pt}\ {\isacharat}{\kern0pt}\ code\ {\isacharparenleft}{\kern0pt}Stars\ vs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ \isanewline
+\ \ Stars{\isacharunderscore}{\kern0pt}add\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}val\ {\isasymRightarrow}\ val\ {\isasymRightarrow}\ val{\isachardoublequoteclose}\isanewline
+\isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}Stars{\isacharunderscore}{\kern0pt}add\ v\ {\isacharparenleft}{\kern0pt}Stars\ vs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ Stars\ {\isacharparenleft}{\kern0pt}v\ {\isacharhash}{\kern0pt}\ vs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isacommand{function}\isamarkupfalse%
+\isanewline
+\ \ decode{\isacharprime}{\kern0pt}\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}bit\ list\ {\isasymRightarrow}\ rexp\ {\isasymRightarrow}\ {\isacharparenleft}{\kern0pt}val\ {\isacharasterisk}{\kern0pt}\ bit\ list{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}decode{\isacharprime}{\kern0pt}\ ds\ ZERO\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}Void{\isacharcomma}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}decode{\isacharprime}{\kern0pt}\ ds\ ONE\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}Void{\isacharcomma}{\kern0pt}\ ds{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}decode{\isacharprime}{\kern0pt}\ ds\ {\isacharparenleft}{\kern0pt}CH\ d{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}Char\ d{\isacharcomma}{\kern0pt}\ ds{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}decode{\isacharprime}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}ALT\ r{\isadigit{1}}\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}Void{\isacharcomma}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}decode{\isacharprime}{\kern0pt}\ {\isacharparenleft}{\kern0pt}Z\ {\isacharhash}{\kern0pt}\ ds{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}ALT\ r{\isadigit{1}}\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}let\ {\isacharparenleft}{\kern0pt}v{\isacharcomma}{\kern0pt}\ ds{\isacharprime}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ decode{\isacharprime}{\kern0pt}\ ds\ r{\isadigit{1}}\ in\ {\isacharparenleft}{\kern0pt}Left\ v{\isacharcomma}{\kern0pt}\ ds{\isacharprime}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}decode{\isacharprime}{\kern0pt}\ {\isacharparenleft}{\kern0pt}S\ {\isacharhash}{\kern0pt}\ ds{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}ALT\ r{\isadigit{1}}\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}let\ {\isacharparenleft}{\kern0pt}v{\isacharcomma}{\kern0pt}\ ds{\isacharprime}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ decode{\isacharprime}{\kern0pt}\ ds\ r{\isadigit{2}}\ in\ {\isacharparenleft}{\kern0pt}Right\ v{\isacharcomma}{\kern0pt}\ ds{\isacharprime}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}decode{\isacharprime}{\kern0pt}\ ds\ {\isacharparenleft}{\kern0pt}SEQ\ r{\isadigit{1}}\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}let\ {\isacharparenleft}{\kern0pt}v{\isadigit{1}}{\isacharcomma}{\kern0pt}\ ds{\isacharprime}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ decode{\isacharprime}{\kern0pt}\ ds\ r{\isadigit{1}}\ in\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ let\ {\isacharparenleft}{\kern0pt}v{\isadigit{2}}{\isacharcomma}{\kern0pt}\ ds{\isacharprime}{\kern0pt}{\isacharprime}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ decode{\isacharprime}{\kern0pt}\ ds{\isacharprime}{\kern0pt}\ r{\isadigit{2}}\ in\ {\isacharparenleft}{\kern0pt}Seq\ v{\isadigit{1}}\ v{\isadigit{2}}{\isacharcomma}{\kern0pt}\ ds{\isacharprime}{\kern0pt}{\isacharprime}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}decode{\isacharprime}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}STAR\ r{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}Void{\isacharcomma}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}decode{\isacharprime}{\kern0pt}\ {\isacharparenleft}{\kern0pt}S\ {\isacharhash}{\kern0pt}\ ds{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}STAR\ r{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}Stars\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharcomma}{\kern0pt}\ ds{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}decode{\isacharprime}{\kern0pt}\ {\isacharparenleft}{\kern0pt}Z\ {\isacharhash}{\kern0pt}\ ds{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}STAR\ r{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}let\ {\isacharparenleft}{\kern0pt}v{\isacharcomma}{\kern0pt}\ ds{\isacharprime}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ decode{\isacharprime}{\kern0pt}\ ds\ r\ in\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ let\ {\isacharparenleft}{\kern0pt}vs{\isacharcomma}{\kern0pt}\ ds{\isacharprime}{\kern0pt}{\isacharprime}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ decode{\isacharprime}{\kern0pt}\ ds{\isacharprime}{\kern0pt}\ {\isacharparenleft}{\kern0pt}STAR\ r{\isacharparenright}{\kern0pt}\ \isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ in\ {\isacharparenleft}{\kern0pt}Stars{\isacharunderscore}{\kern0pt}add\ v\ vs{\isacharcomma}{\kern0pt}\ ds{\isacharprime}{\kern0pt}{\isacharprime}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{by}\isamarkupfalse%
+\ pat{\isacharunderscore}{\kern0pt}completeness\ auto%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ decode{\isacharprime}{\kern0pt}{\isacharunderscore}{\kern0pt}smaller{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}decode{\isacharprime}{\kern0pt}{\isacharunderscore}{\kern0pt}dom\ {\isacharparenleft}{\kern0pt}ds{\isacharcomma}{\kern0pt}\ r{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}length\ {\isacharparenleft}{\kern0pt}snd\ {\isacharparenleft}{\kern0pt}decode{\isacharprime}{\kern0pt}\ ds\ r{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isasymle}\ length\ ds{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ ds\ r{\isacharparenright}{\kern0pt}\isanewline
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto\ simp\ add{\isacharcolon}{\kern0pt}\ decode{\isacharprime}{\kern0pt}{\isachardot}{\kern0pt}psimps\ split{\isacharcolon}{\kern0pt}\ prod{\isachardot}{\kern0pt}split{\isacharparenright}{\kern0pt}\isanewline
+\isacommand{using}\isamarkupfalse%
+\ dual{\isacharunderscore}{\kern0pt}order{\isachardot}{\kern0pt}trans\ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}meson\ dual{\isacharunderscore}{\kern0pt}order{\isachardot}{\kern0pt}trans\ le{\isacharunderscore}{\kern0pt}SucI{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{termination}\isamarkupfalse%
+\ {\isachardoublequoteopen}decode{\isacharprime}{\kern0pt}{\isachardoublequoteclose}\ \ \isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}relation\ {\isachardoublequoteopen}inv{\isacharunderscore}{\kern0pt}image\ {\isacharparenleft}{\kern0pt}measure{\isacharparenleft}{\kern0pt}{\isacharpercent}{\kern0pt}cs{\isachardot}{\kern0pt}\ size\ cs{\isacharparenright}{\kern0pt}\ {\isacharless}{\kern0pt}{\isacharasterisk}{\kern0pt}lex{\isacharasterisk}{\kern0pt}{\isachargreater}{\kern0pt}\ measure{\isacharparenleft}{\kern0pt}{\isacharpercent}{\kern0pt}s{\isachardot}{\kern0pt}\ size\ s{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}{\isacharpercent}{\kern0pt}{\isacharparenleft}{\kern0pt}ds{\isacharcomma}{\kern0pt}r{\isacharparenright}{\kern0pt}{\isachardot}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r{\isacharcomma}{\kern0pt}ds{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\ \isanewline
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto\ dest{\isacharbang}{\kern0pt}{\isacharcolon}{\kern0pt}\ decode{\isacharprime}{\kern0pt}{\isacharunderscore}{\kern0pt}smaller{\isacharparenright}{\kern0pt}\isanewline
+\isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ less{\isacharunderscore}{\kern0pt}Suc{\isacharunderscore}{\kern0pt}eq{\isacharunderscore}{\kern0pt}le\ snd{\isacharunderscore}{\kern0pt}conv{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{definition}\isamarkupfalse%
+\isanewline
+\ \ decode\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}bit\ list\ {\isasymRightarrow}\ rexp\ {\isasymRightarrow}\ val\ option{\isachardoublequoteclose}\isanewline
+\isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}decode\ ds\ r\ {\isasymequiv}\ {\isacharparenleft}{\kern0pt}let\ {\isacharparenleft}{\kern0pt}v{\isacharcomma}{\kern0pt}\ ds{\isacharprime}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ decode{\isacharprime}{\kern0pt}\ ds\ r\ \isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ in\ {\isacharparenleft}{\kern0pt}if\ ds{\isacharprime}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ then\ Some\ v\ else\ None{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ decode{\isacharprime}{\kern0pt}{\isacharunderscore}{\kern0pt}code{\isacharunderscore}{\kern0pt}Stars{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}{\isasymforall}v{\isasymin}set\ vs{\isachardot}{\kern0pt}\ {\isasymTurnstile}\ v\ {\isacharcolon}{\kern0pt}\ r\ {\isasymand}\ {\isacharparenleft}{\kern0pt}{\isasymforall}x{\isachardot}{\kern0pt}\ decode{\isacharprime}{\kern0pt}\ {\isacharparenleft}{\kern0pt}code\ v\ {\isacharat}{\kern0pt}\ x{\isacharparenright}{\kern0pt}\ r\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}v{\isacharcomma}{\kern0pt}\ x{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isasymand}\ flat\ v\ {\isasymnoteq}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isachardoublequoteclose}\ \isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}decode{\isacharprime}{\kern0pt}\ {\isacharparenleft}{\kern0pt}code\ {\isacharparenleft}{\kern0pt}Stars\ vs{\isacharparenright}{\kern0pt}\ {\isacharat}{\kern0pt}\ ds{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}STAR\ r{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}Stars\ vs{\isacharcomma}{\kern0pt}\ ds{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ vs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ decode{\isacharprime}{\kern0pt}{\isacharunderscore}{\kern0pt}code{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}{\isasymTurnstile}\ v\ {\isacharcolon}{\kern0pt}\ r{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}decode{\isacharprime}{\kern0pt}\ {\isacharparenleft}{\kern0pt}{\isacharparenleft}{\kern0pt}code\ v{\isacharparenright}{\kern0pt}\ {\isacharat}{\kern0pt}\ ds{\isacharparenright}{\kern0pt}\ r\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}v{\isacharcomma}{\kern0pt}\ ds{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ v\ r\ arbitrary{\isacharcolon}{\kern0pt}\ ds{\isacharparenright}{\kern0pt}\ \isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ decode{\isacharprime}{\kern0pt}{\isacharunderscore}{\kern0pt}code{\isacharunderscore}{\kern0pt}Stars\ \isacommand{by}\isamarkupfalse%
+\ blast%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ decode{\isacharunderscore}{\kern0pt}code{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}{\isasymTurnstile}\ v\ {\isacharcolon}{\kern0pt}\ r{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}decode\ {\isacharparenleft}{\kern0pt}code\ v{\isacharparenright}{\kern0pt}\ r\ {\isacharequal}{\kern0pt}\ Some\ v{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\ \isacommand{unfolding}\isamarkupfalse%
+\ decode{\isacharunderscore}{\kern0pt}def\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}smt\ append{\isacharunderscore}{\kern0pt}Nil{\isadigit{2}}\ decode{\isacharprime}{\kern0pt}{\isacharunderscore}{\kern0pt}code\ old{\isachardot}{\kern0pt}prod{\isachardot}{\kern0pt}case{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isadelimdocument
+%
+\endisadelimdocument
+%
+\isatagdocument
+%
+\isamarkupsection{Annotated Regular Expressions%
+}
+\isamarkuptrue%
+%
+\endisatagdocument
+{\isafolddocument}%
+%
+\isadelimdocument
+%
+\endisadelimdocument
+\isacommand{datatype}\isamarkupfalse%
+\ arexp\ {\isacharequal}{\kern0pt}\ \isanewline
+\ \ AZERO\isanewline
+{\isacharbar}{\kern0pt}\ AONE\ {\isachardoublequoteopen}bit\ list{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ ACHAR\ {\isachardoublequoteopen}bit\ list{\isachardoublequoteclose}\ char\isanewline
+{\isacharbar}{\kern0pt}\ ASEQ\ {\isachardoublequoteopen}bit\ list{\isachardoublequoteclose}\ arexp\ arexp\isanewline
+{\isacharbar}{\kern0pt}\ AALTs\ {\isachardoublequoteopen}bit\ list{\isachardoublequoteclose}\ {\isachardoublequoteopen}arexp\ list{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ ASTAR\ {\isachardoublequoteopen}bit\ list{\isachardoublequoteclose}\ arexp\isanewline
+\isanewline
+\isacommand{abbreviation}\isamarkupfalse%
+\isanewline
+\ \ {\isachardoublequoteopen}AALT\ bs\ r{\isadigit{1}}\ r{\isadigit{2}}\ {\isasymequiv}\ AALTs\ bs\ {\isacharbrackleft}{\kern0pt}r{\isadigit{1}}{\isacharcomma}{\kern0pt}\ r{\isadigit{2}}{\isacharbrackright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ asize\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}arexp\ {\isasymRightarrow}\ nat{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}asize\ AZERO\ {\isacharequal}{\kern0pt}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}asize\ {\isacharparenleft}{\kern0pt}AONE\ cs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isadigit{1}}{\isachardoublequoteclose}\ \isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}asize\ {\isacharparenleft}{\kern0pt}ACHAR\ cs\ c{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isadigit{1}}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}asize\ {\isacharparenleft}{\kern0pt}AALTs\ cs\ rs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ Suc\ {\isacharparenleft}{\kern0pt}sum{\isacharunderscore}{\kern0pt}list\ {\isacharparenleft}{\kern0pt}map\ asize\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}asize\ {\isacharparenleft}{\kern0pt}ASEQ\ cs\ r{\isadigit{1}}\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ Suc\ {\isacharparenleft}{\kern0pt}asize\ r{\isadigit{1}}\ {\isacharplus}{\kern0pt}\ asize\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}asize\ {\isacharparenleft}{\kern0pt}ASTAR\ cs\ r{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ Suc\ {\isacharparenleft}{\kern0pt}asize\ r{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ \isanewline
+\ \ erase\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}arexp\ {\isasymRightarrow}\ rexp{\isachardoublequoteclose}\isanewline
+\isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}erase\ AZERO\ {\isacharequal}{\kern0pt}\ ZERO{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}erase\ {\isacharparenleft}{\kern0pt}AONE\ {\isacharunderscore}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ ONE{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}erase\ {\isacharparenleft}{\kern0pt}ACHAR\ {\isacharunderscore}{\kern0pt}\ c{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ CH\ c{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}erase\ {\isacharparenleft}{\kern0pt}AALTs\ {\isacharunderscore}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ ZERO{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}erase\ {\isacharparenleft}{\kern0pt}AALTs\ {\isacharunderscore}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}r{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}erase\ r{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}erase\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}r{\isacharhash}{\kern0pt}rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ ALT\ {\isacharparenleft}{\kern0pt}erase\ r{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}erase\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}erase\ {\isacharparenleft}{\kern0pt}ASEQ\ {\isacharunderscore}{\kern0pt}\ r{\isadigit{1}}\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ SEQ\ {\isacharparenleft}{\kern0pt}erase\ r{\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}erase\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}erase\ {\isacharparenleft}{\kern0pt}ASTAR\ {\isacharunderscore}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ STAR\ {\isacharparenleft}{\kern0pt}erase\ r{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ nonalt\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}arexp\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}nonalt\ {\isacharparenleft}{\kern0pt}AALTs\ bs{\isadigit{2}}\ rs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ False{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}nonalt\ r\ {\isacharequal}{\kern0pt}\ True{\isachardoublequoteclose}\isanewline
+\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ good\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}arexp\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}good\ AZERO\ {\isacharequal}{\kern0pt}\ False{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}good\ {\isacharparenleft}{\kern0pt}AONE\ cs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ True{\isachardoublequoteclose}\ \isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}good\ {\isacharparenleft}{\kern0pt}ACHAR\ cs\ c{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ True{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}good\ {\isacharparenleft}{\kern0pt}AALTs\ cs\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ False{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}good\ {\isacharparenleft}{\kern0pt}AALTs\ cs\ {\isacharbrackleft}{\kern0pt}r{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ False{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}good\ {\isacharparenleft}{\kern0pt}AALTs\ cs\ {\isacharparenleft}{\kern0pt}r{\isadigit{1}}{\isacharhash}{\kern0pt}r{\isadigit{2}}{\isacharhash}{\kern0pt}rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}{\isasymforall}r{\isacharprime}{\kern0pt}\ {\isasymin}\ set\ {\isacharparenleft}{\kern0pt}r{\isadigit{1}}{\isacharhash}{\kern0pt}r{\isadigit{2}}{\isacharhash}{\kern0pt}rs{\isacharparenright}{\kern0pt}{\isachardot}{\kern0pt}\ good\ r{\isacharprime}{\kern0pt}\ {\isasymand}\ nonalt\ r{\isacharprime}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}good\ {\isacharparenleft}{\kern0pt}ASEQ\ {\isacharunderscore}{\kern0pt}\ AZERO\ {\isacharunderscore}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ False{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}good\ {\isacharparenleft}{\kern0pt}ASEQ\ {\isacharunderscore}{\kern0pt}\ {\isacharparenleft}{\kern0pt}AONE\ {\isacharunderscore}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharunderscore}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ False{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}good\ {\isacharparenleft}{\kern0pt}ASEQ\ {\isacharunderscore}{\kern0pt}\ {\isacharunderscore}{\kern0pt}\ AZERO{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ False{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}good\ {\isacharparenleft}{\kern0pt}ASEQ\ cs\ r{\isadigit{1}}\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}good\ r{\isadigit{1}}\ {\isasymand}\ good\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}good\ {\isacharparenleft}{\kern0pt}ASTAR\ cs\ r{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ True{\isachardoublequoteclose}\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ fuse\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}bit\ list\ {\isasymRightarrow}\ arexp\ {\isasymRightarrow}\ arexp{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}fuse\ bs\ AZERO\ {\isacharequal}{\kern0pt}\ AZERO{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}fuse\ bs\ {\isacharparenleft}{\kern0pt}AONE\ cs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ AONE\ {\isacharparenleft}{\kern0pt}bs\ {\isacharat}{\kern0pt}\ cs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\ \isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}fuse\ bs\ {\isacharparenleft}{\kern0pt}ACHAR\ cs\ c{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ ACHAR\ {\isacharparenleft}{\kern0pt}bs\ {\isacharat}{\kern0pt}\ cs{\isacharparenright}{\kern0pt}\ c{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}fuse\ bs\ {\isacharparenleft}{\kern0pt}AALTs\ cs\ rs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ AALTs\ {\isacharparenleft}{\kern0pt}bs\ {\isacharat}{\kern0pt}\ cs{\isacharparenright}{\kern0pt}\ rs{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}fuse\ bs\ {\isacharparenleft}{\kern0pt}ASEQ\ cs\ r{\isadigit{1}}\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ ASEQ\ {\isacharparenleft}{\kern0pt}bs\ {\isacharat}{\kern0pt}\ cs{\isacharparenright}{\kern0pt}\ r{\isadigit{1}}\ r{\isadigit{2}}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}fuse\ bs\ {\isacharparenleft}{\kern0pt}ASTAR\ cs\ r{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ ASTAR\ {\isacharparenleft}{\kern0pt}bs\ {\isacharat}{\kern0pt}\ cs{\isacharparenright}{\kern0pt}\ r{\isachardoublequoteclose}\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ fuse{\isacharunderscore}{\kern0pt}append{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}fuse\ {\isacharparenleft}{\kern0pt}bs{\isadigit{1}}\ {\isacharat}{\kern0pt}\ bs{\isadigit{2}}{\isacharparenright}{\kern0pt}\ r\ {\isacharequal}{\kern0pt}\ fuse\ bs{\isadigit{1}}\ {\isacharparenleft}{\kern0pt}fuse\ bs{\isadigit{2}}\ r{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ r{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ intern\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}rexp\ {\isasymRightarrow}\ arexp{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}intern\ ZERO\ {\isacharequal}{\kern0pt}\ AZERO{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}intern\ ONE\ {\isacharequal}{\kern0pt}\ AONE\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}intern\ {\isacharparenleft}{\kern0pt}CH\ c{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ ACHAR\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ c{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}intern\ {\isacharparenleft}{\kern0pt}ALT\ r{\isadigit{1}}\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ AALT\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}fuse\ {\isacharbrackleft}{\kern0pt}Z{\isacharbrackright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}intern\ r{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ \isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}{\kern0pt}fuse\ {\isacharbrackleft}{\kern0pt}S{\isacharbrackright}{\kern0pt}\ \ {\isacharparenleft}{\kern0pt}intern\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}intern\ {\isacharparenleft}{\kern0pt}SEQ\ r{\isadigit{1}}\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ ASEQ\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}intern\ r{\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}intern\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}intern\ {\isacharparenleft}{\kern0pt}STAR\ r{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ ASTAR\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}intern\ r{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ retrieve\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}arexp\ {\isasymRightarrow}\ val\ {\isasymRightarrow}\ bit\ list{\isachardoublequoteclose}\ \isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}retrieve\ {\isacharparenleft}{\kern0pt}AONE\ bs{\isacharparenright}{\kern0pt}\ Void\ {\isacharequal}{\kern0pt}\ bs{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}retrieve\ {\isacharparenleft}{\kern0pt}ACHAR\ bs\ c{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}Char\ d{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bs{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}retrieve\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharbrackleft}{\kern0pt}r{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}\ v\ {\isacharequal}{\kern0pt}\ bs\ {\isacharat}{\kern0pt}\ retrieve\ r\ v{\isachardoublequoteclose}\ \isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}retrieve\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}r{\isacharhash}{\kern0pt}rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}Left\ v{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bs\ {\isacharat}{\kern0pt}\ retrieve\ r\ v{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}retrieve\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}r{\isacharhash}{\kern0pt}rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}Right\ v{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bs\ {\isacharat}{\kern0pt}\ retrieve\ {\isacharparenleft}{\kern0pt}AALTs\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}\ v{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}retrieve\ {\isacharparenleft}{\kern0pt}ASEQ\ bs\ r{\isadigit{1}}\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}Seq\ v{\isadigit{1}}\ v{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bs\ {\isacharat}{\kern0pt}\ retrieve\ r{\isadigit{1}}\ v{\isadigit{1}}\ {\isacharat}{\kern0pt}\ retrieve\ r{\isadigit{2}}\ v{\isadigit{2}}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}retrieve\ {\isacharparenleft}{\kern0pt}ASTAR\ bs\ r{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}Stars\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bs\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}S{\isacharbrackright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}retrieve\ {\isacharparenleft}{\kern0pt}ASTAR\ bs\ r{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}Stars\ {\isacharparenleft}{\kern0pt}v{\isacharhash}{\kern0pt}vs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ \isanewline
+\ \ \ \ \ bs\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}Z{\isacharbrackright}{\kern0pt}\ {\isacharat}{\kern0pt}\ retrieve\ r\ v\ {\isacharat}{\kern0pt}\ retrieve\ {\isacharparenleft}{\kern0pt}ASTAR\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ r{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}Stars\ vs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\isanewline
+\ bnullable\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}arexp\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
+\isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}bnullable\ {\isacharparenleft}{\kern0pt}AZERO{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ False{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bnullable\ {\isacharparenleft}{\kern0pt}AONE\ bs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ True{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bnullable\ {\isacharparenleft}{\kern0pt}ACHAR\ bs\ c{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ False{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}{\isasymexists}r\ {\isasymin}\ set\ rs{\isachardot}{\kern0pt}\ bnullable\ r{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bnullable\ {\isacharparenleft}{\kern0pt}ASEQ\ bs\ r{\isadigit{1}}\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}bnullable\ r{\isadigit{1}}\ {\isasymand}\ bnullable\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bnullable\ {\isacharparenleft}{\kern0pt}ASTAR\ bs\ r{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ True{\isachardoublequoteclose}\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ \isanewline
+\ \ bmkeps\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}arexp\ {\isasymRightarrow}\ bit\ list{\isachardoublequoteclose}\isanewline
+\isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}bmkeps{\isacharparenleft}{\kern0pt}AONE\ bs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bs{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bmkeps{\isacharparenleft}{\kern0pt}ASEQ\ bs\ r{\isadigit{1}}\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bs\ {\isacharat}{\kern0pt}\ {\isacharparenleft}{\kern0pt}bmkeps\ r{\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharat}{\kern0pt}\ {\isacharparenleft}{\kern0pt}bmkeps\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bmkeps{\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharbrackleft}{\kern0pt}r{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bs\ {\isacharat}{\kern0pt}\ {\isacharparenleft}{\kern0pt}bmkeps\ r{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bmkeps{\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}r{\isacharhash}{\kern0pt}rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}if\ bnullable{\isacharparenleft}{\kern0pt}r{\isacharparenright}{\kern0pt}\ then\ bs\ {\isacharat}{\kern0pt}\ {\isacharparenleft}{\kern0pt}bmkeps\ r{\isacharparenright}{\kern0pt}\ else\ {\isacharparenleft}{\kern0pt}bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bmkeps{\isacharparenleft}{\kern0pt}ASTAR\ bs\ r{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bs\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}S{\isacharbrackright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\isanewline
+\ bder\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}char\ {\isasymRightarrow}\ arexp\ {\isasymRightarrow}\ arexp{\isachardoublequoteclose}\isanewline
+\isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}bder\ c\ {\isacharparenleft}{\kern0pt}AZERO{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ AZERO{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bder\ c\ {\isacharparenleft}{\kern0pt}AONE\ bs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ AZERO{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bder\ c\ {\isacharparenleft}{\kern0pt}ACHAR\ bs\ d{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}if\ c\ {\isacharequal}{\kern0pt}\ d\ then\ AONE\ bs\ else\ AZERO{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bder\ c\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ AALTs\ bs\ {\isacharparenleft}{\kern0pt}map\ {\isacharparenleft}{\kern0pt}bder\ c{\isacharparenright}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bder\ c\ {\isacharparenleft}{\kern0pt}ASEQ\ bs\ r{\isadigit{1}}\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ \isanewline
+\ \ \ \ \ {\isacharparenleft}{\kern0pt}if\ bnullable\ r{\isadigit{1}}\isanewline
+\ \ \ \ \ \ then\ AALT\ bs\ {\isacharparenleft}{\kern0pt}ASEQ\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}bder\ c\ r{\isadigit{1}}{\isacharparenright}{\kern0pt}\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}fuse\ {\isacharparenleft}{\kern0pt}bmkeps\ r{\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}bder\ c\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ else\ ASEQ\ bs\ {\isacharparenleft}{\kern0pt}bder\ c\ r{\isadigit{1}}{\isacharparenright}{\kern0pt}\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bder\ c\ {\isacharparenleft}{\kern0pt}ASTAR\ bs\ r{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ ASEQ\ bs\ {\isacharparenleft}{\kern0pt}fuse\ {\isacharbrackleft}{\kern0pt}Z{\isacharbrackright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}bder\ c\ r{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}ASTAR\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ r{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ \isanewline
+\ \ bders\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}arexp\ {\isasymRightarrow}\ string\ {\isasymRightarrow}\ arexp{\isachardoublequoteclose}\isanewline
+\isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}bders\ r\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ r{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bders\ r\ {\isacharparenleft}{\kern0pt}c{\isacharhash}{\kern0pt}s{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bders\ {\isacharparenleft}{\kern0pt}bder\ c\ r{\isacharparenright}{\kern0pt}\ s{\isachardoublequoteclose}\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ bders{\isacharunderscore}{\kern0pt}append{\isacharcolon}{\kern0pt}\isanewline
+\ \ {\isachardoublequoteopen}bders\ r\ {\isacharparenleft}{\kern0pt}s{\isadigit{1}}\ {\isacharat}{\kern0pt}\ s{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bders\ {\isacharparenleft}{\kern0pt}bders\ r\ s{\isadigit{1}}{\isacharparenright}{\kern0pt}\ s{\isadigit{2}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ s{\isadigit{1}}\ arbitrary{\isacharcolon}{\kern0pt}\ r\ s{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharunderscore}{\kern0pt}all{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ bnullable{\isacharunderscore}{\kern0pt}correctness{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}nullable\ {\isacharparenleft}{\kern0pt}erase\ r{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bnullable\ r{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ r\ rule{\isacharcolon}{\kern0pt}\ erase{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharunderscore}{\kern0pt}all{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ erase{\isacharunderscore}{\kern0pt}fuse{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}erase\ {\isacharparenleft}{\kern0pt}fuse\ bs\ r{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ erase\ r{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ r\ rule{\isacharcolon}{\kern0pt}\ erase{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharunderscore}{\kern0pt}all{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{thm}\isamarkupfalse%
+\ Posix{\isachardot}{\kern0pt}induct\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ erase{\isacharunderscore}{\kern0pt}intern\ {\isacharbrackleft}{\kern0pt}simp{\isacharbrackright}{\kern0pt}{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}erase\ {\isacharparenleft}{\kern0pt}intern\ r{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ r{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ r{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharunderscore}{\kern0pt}all\ add{\isacharcolon}{\kern0pt}\ erase{\isacharunderscore}{\kern0pt}fuse{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ erase{\isacharunderscore}{\kern0pt}bder\ {\isacharbrackleft}{\kern0pt}simp{\isacharbrackright}{\kern0pt}{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}erase\ {\isacharparenleft}{\kern0pt}bder\ a\ r{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ der\ a\ {\isacharparenleft}{\kern0pt}erase\ r{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ r\ rule{\isacharcolon}{\kern0pt}\ erase{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharunderscore}{\kern0pt}all\ add{\isacharcolon}{\kern0pt}\ erase{\isacharunderscore}{\kern0pt}fuse\ bnullable{\isacharunderscore}{\kern0pt}correctness{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ erase{\isacharunderscore}{\kern0pt}bders\ {\isacharbrackleft}{\kern0pt}simp{\isacharbrackright}{\kern0pt}{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}erase\ {\isacharparenleft}{\kern0pt}bders\ r\ s{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ ders\ s\ {\isacharparenleft}{\kern0pt}erase\ r{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ s\ arbitrary{\isacharcolon}{\kern0pt}\ r\ {\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharunderscore}{\kern0pt}all{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ retrieve{\isacharunderscore}{\kern0pt}encode{\isacharunderscore}{\kern0pt}STARS{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}{\isasymforall}v{\isasymin}set\ vs{\isachardot}{\kern0pt}\ {\isasymTurnstile}\ v\ {\isacharcolon}{\kern0pt}\ r\ {\isasymand}\ code\ v\ {\isacharequal}{\kern0pt}\ retrieve\ {\isacharparenleft}{\kern0pt}intern\ r{\isacharparenright}{\kern0pt}\ v{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}code\ {\isacharparenleft}{\kern0pt}Stars\ vs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ retrieve\ {\isacharparenleft}{\kern0pt}ASTAR\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}intern\ r{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}Stars\ vs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ vs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharunderscore}{\kern0pt}all{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ retrieve{\isacharunderscore}{\kern0pt}fuse{\isadigit{2}}{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}{\isasymTurnstile}\ v\ {\isacharcolon}{\kern0pt}\ {\isacharparenleft}{\kern0pt}erase\ r{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}retrieve\ {\isacharparenleft}{\kern0pt}fuse\ bs\ r{\isacharparenright}{\kern0pt}\ v\ {\isacharequal}{\kern0pt}\ bs\ {\isacharat}{\kern0pt}\ retrieve\ r\ v{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ r\ arbitrary{\isacharcolon}{\kern0pt}\ v\ bs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto\ elim{\isacharcolon}{\kern0pt}\ Prf{\isacharunderscore}{\kern0pt}elims{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isadigit{4}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \isacommand{defer}\isamarkupfalse%
+\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ retrieve{\isacharunderscore}{\kern0pt}encode{\isacharunderscore}{\kern0pt}STARS\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto\ elim{\isacharbang}{\kern0pt}{\isacharcolon}{\kern0pt}\ Prf{\isacharunderscore}{\kern0pt}elims{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ vs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ x{\isadigit{2}}a{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto\ elim{\isacharbang}{\kern0pt}{\isacharcolon}{\kern0pt}\ Prf{\isacharunderscore}{\kern0pt}elims{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ list{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto\ elim{\isacharbang}{\kern0pt}{\isacharcolon}{\kern0pt}\ Prf{\isacharunderscore}{\kern0pt}elims{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ retrieve{\isacharunderscore}{\kern0pt}fuse{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}{\isasymTurnstile}\ v\ {\isacharcolon}{\kern0pt}\ r{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}retrieve\ {\isacharparenleft}{\kern0pt}fuse\ bs\ {\isacharparenleft}{\kern0pt}intern\ r{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ v\ {\isacharequal}{\kern0pt}\ bs\ {\isacharat}{\kern0pt}\ retrieve\ {\isacharparenleft}{\kern0pt}intern\ r{\isacharparenright}{\kern0pt}\ v{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\ \isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp{\isacharunderscore}{\kern0pt}all\ add{\isacharcolon}{\kern0pt}\ retrieve{\isacharunderscore}{\kern0pt}fuse{\isadigit{2}}{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ retrieve{\isacharunderscore}{\kern0pt}code{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}{\isasymTurnstile}\ v\ {\isacharcolon}{\kern0pt}\ r{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}code\ v\ {\isacharequal}{\kern0pt}\ retrieve\ {\isacharparenleft}{\kern0pt}intern\ r{\isacharparenright}{\kern0pt}\ v{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ v\ r\ {\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharunderscore}{\kern0pt}all\ add{\isacharcolon}{\kern0pt}\ retrieve{\isacharunderscore}{\kern0pt}fuse\ retrieve{\isacharunderscore}{\kern0pt}encode{\isacharunderscore}{\kern0pt}STARS{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ bnullable{\isacharunderscore}{\kern0pt}Hdbmkeps{\isacharunderscore}{\kern0pt}Hd{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}bnullable\ a{\isachardoublequoteclose}\ \isanewline
+\ \ \isakeyword{shows}\ \ {\isachardoublequoteopen}bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}a\ {\isacharhash}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bs\ {\isacharat}{\kern0pt}\ {\isacharparenleft}{\kern0pt}bmkeps\ a{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ bmkeps{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}\ bmkeps{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}\ list{\isachardot}{\kern0pt}exhaust{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ r{\isadigit{1}}{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}{\isasymnot}\ bnullable\ a{\isachardoublequoteclose}\ {\isachardoublequoteopen}bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ \ {\isachardoublequoteopen}bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}a\ {\isacharhash}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ rs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ r{\isadigit{2}}{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}x\ {\isasymin}\ set\ rs{\isachardoublequoteclose}\ {\isachardoublequoteopen}bnullable\ x{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ rs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ \ r{\isadigit{3}}{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}{\isasymnot}\ bnullable\ r{\isachardoublequoteclose}\ \isanewline
+\ \ \ \ \ \ \ \ \ \ {\isachardoublequoteopen}\ {\isasymexists}\ x\ {\isasymin}\ set\ rs{\isachardot}{\kern0pt}\ bnullable\ x{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}retrieve\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}mkeps\ {\isacharparenleft}{\kern0pt}erase\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ retrieve\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}r\ {\isacharhash}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}mkeps\ {\isacharparenleft}{\kern0pt}erase\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}r\ {\isacharhash}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ rs\ arbitrary{\isacharcolon}{\kern0pt}\ r\ bs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ bnullable{\isacharunderscore}{\kern0pt}correctness\ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto\ simp\ add{\isacharcolon}{\kern0pt}\ bnullable{\isacharunderscore}{\kern0pt}correctness\ mkeps{\isacharunderscore}{\kern0pt}nullable\ retrieve{\isacharunderscore}{\kern0pt}fuse{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ retrieve{\isacharunderscore}{\kern0pt}fuse{\isadigit{2}}{\isacharbrackleft}{\kern0pt}symmetric{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}smt\ bnullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{5}}{\isacharparenright}{\kern0pt}\ erase{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{6}}{\isacharparenright}{\kern0pt}\ insert{\isacharunderscore}{\kern0pt}iff\ list{\isachardot}{\kern0pt}exhaust\ list{\isachardot}{\kern0pt}set{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}\ mkeps{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}\ mkeps{\isacharunderscore}{\kern0pt}nullable{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}bnullable\ a{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}smt\ append{\isacharunderscore}{\kern0pt}Nil{\isadigit{2}}\ bnullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{5}}{\isacharparenright}{\kern0pt}\ erase{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{6}}{\isacharparenright}{\kern0pt}\ fuse{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}\ insert{\isacharunderscore}{\kern0pt}iff\ list{\isachardot}{\kern0pt}exhaust\ list{\isachardot}{\kern0pt}set{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}\ mkeps{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}\ mkeps{\isacharunderscore}{\kern0pt}nullable\ retrieve{\isacharunderscore}{\kern0pt}fuse{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}drule{\isacharunderscore}{\kern0pt}tac\ x{\isacharequal}{\kern0pt}{\isachardoublequoteopen}a{\isachardoublequoteclose}\ \isakeyword{in}\ meta{\isacharunderscore}{\kern0pt}spec{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}drule{\isacharunderscore}{\kern0pt}tac\ x{\isacharequal}{\kern0pt}{\isachardoublequoteopen}bs{\isachardoublequoteclose}\ \isakeyword{in}\ meta{\isacharunderscore}{\kern0pt}spec{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}drule\ meta{\isacharunderscore}{\kern0pt}mp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}drule\ meta{\isacharunderscore}{\kern0pt}mp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ retrieve{\isacharunderscore}{\kern0pt}fuse{\isadigit{2}}{\isacharbrackleft}{\kern0pt}symmetric{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ rs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ bnullable{\isacharunderscore}{\kern0pt}correctness{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ append{\isacharunderscore}{\kern0pt}Nil{\isadigit{2}}\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isacharunderscore}{\kern0pt}fuse\ fuse{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}\ list{\isachardot}{\kern0pt}set{\isacharunderscore}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}\ mkeps{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}\ mkeps{\isacharunderscore}{\kern0pt}nullable\ nullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ bnullable{\isacharunderscore}{\kern0pt}correctness{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ append{\isacharunderscore}{\kern0pt}Nil{\isadigit{2}}\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{6}}{\isacharparenright}{\kern0pt}\ erase{\isacharunderscore}{\kern0pt}fuse\ fuse{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}\ list{\isachardot}{\kern0pt}set{\isacharunderscore}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}\ mkeps{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}\ mkeps{\isacharunderscore}{\kern0pt}nullable\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ t{\isacharcolon}{\kern0pt}\ \isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}{\isasymforall}r\ {\isasymin}\ set\ rs{\isachardot}{\kern0pt}\ nullable\ {\isacharparenleft}{\kern0pt}erase\ r{\isacharparenright}{\kern0pt}\ {\isasymlongrightarrow}\ bmkeps\ r\ {\isacharequal}{\kern0pt}\ retrieve\ r\ {\isacharparenleft}{\kern0pt}mkeps\ {\isacharparenleft}{\kern0pt}erase\ r{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\ \isanewline
+\ \ \ \ \ \ \ \ \ \ {\isachardoublequoteopen}nullable\ {\isacharparenleft}{\kern0pt}erase\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}\ bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ retrieve\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}mkeps\ {\isacharparenleft}{\kern0pt}erase\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ rs\ arbitrary{\isacharcolon}{\kern0pt}\ bs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto\ simp\ add{\isacharcolon}{\kern0pt}\ bnullable{\isacharunderscore}{\kern0pt}correctness{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ rs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto\ simp\ add{\isacharcolon}{\kern0pt}\ bnullable{\isacharunderscore}{\kern0pt}correctness{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isadigit{2}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ r{\isadigit{1}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}assumption{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}drule{\isacharunderscore}{\kern0pt}tac\ x{\isacharequal}{\kern0pt}{\isachardoublequoteopen}bs{\isachardoublequoteclose}\ \isakeyword{in}\ meta{\isacharunderscore}{\kern0pt}spec{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}drule\ meta{\isacharunderscore}{\kern0pt}mp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \isacommand{prefer}\isamarkupfalse%
+\ {\isadigit{2}}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}bnullable\ a{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ bnullable{\isacharunderscore}{\kern0pt}Hdbmkeps{\isacharunderscore}{\kern0pt}Hd{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}nullable\ {\isacharparenleft}{\kern0pt}erase\ a{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{prefer}\isamarkupfalse%
+\ {\isadigit{2}}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ bnullable{\isacharunderscore}{\kern0pt}correctness\ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ {\isacharparenleft}{\kern0pt}no{\isacharunderscore}{\kern0pt}types{\isacharcomma}{\kern0pt}\ lifting{\isacharparenright}{\kern0pt}\ erase{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{5}}{\isacharparenright}{\kern0pt}\ erase{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{6}}{\isacharparenright}{\kern0pt}\ list{\isachardot}{\kern0pt}exhaust\ mkeps{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}\ retrieve{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}\ retrieve{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ r{\isadigit{1}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ r{\isadigit{2}}\ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}drule{\isacharunderscore}{\kern0pt}tac\ x{\isacharequal}{\kern0pt}{\isachardoublequoteopen}bs{\isachardoublequoteclose}\ \isakeyword{in}\ meta{\isacharunderscore}{\kern0pt}spec{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}drule\ meta{\isacharunderscore}{\kern0pt}mp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ r{\isadigit{3}}\ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ r{\isadigit{3}}\ \isacommand{by}\isamarkupfalse%
+\ blast%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ bmkeps{\isacharunderscore}{\kern0pt}retrieve{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}nullable\ {\isacharparenleft}{\kern0pt}erase\ r{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}bmkeps\ r\ {\isacharequal}{\kern0pt}\ retrieve\ r\ {\isacharparenleft}{\kern0pt}mkeps\ {\isacharparenleft}{\kern0pt}erase\ r{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ r{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{defer}\isamarkupfalse%
+\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ t{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ bder{\isacharunderscore}{\kern0pt}retrieve{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}{\isasymTurnstile}\ v\ {\isacharcolon}{\kern0pt}\ der\ c\ {\isacharparenleft}{\kern0pt}erase\ r{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}retrieve\ {\isacharparenleft}{\kern0pt}bder\ c\ r{\isacharparenright}{\kern0pt}\ v\ {\isacharequal}{\kern0pt}\ retrieve\ r\ {\isacharparenleft}{\kern0pt}injval\ {\isacharparenleft}{\kern0pt}erase\ r{\isacharparenright}{\kern0pt}\ c\ v{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ r\ arbitrary{\isacharcolon}{\kern0pt}\ v\ rule{\isacharcolon}{\kern0pt}\ erase{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}erule\ Prf{\isacharunderscore}{\kern0pt}elims{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}erule\ Prf{\isacharunderscore}{\kern0pt}elims{\isacharparenright}{\kern0pt}\ \isanewline
+\ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}c\ {\isacharequal}{\kern0pt}\ ca{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}erule\ Prf{\isacharunderscore}{\kern0pt}elims{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}erule\ Prf{\isacharunderscore}{\kern0pt}elims{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}erule\ Prf{\isacharunderscore}{\kern0pt}elims{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rename{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}r\isactrlsub {\isadigit{1}}{\isachardoublequoteclose}\ {\isachardoublequoteopen}r\isactrlsub {\isadigit{2}}{\isachardoublequoteclose}\ rs\ v{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}erule\ Prf{\isacharunderscore}{\kern0pt}elims{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ rs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}smt\ Prf{\isacharunderscore}{\kern0pt}elims{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}\ injval{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}\ injval{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}\ retrieve{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}\ retrieve{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{5}}{\isacharparenright}{\kern0pt}\ same{\isacharunderscore}{\kern0pt}append{\isacharunderscore}{\kern0pt}eq{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}nullable\ {\isacharparenleft}{\kern0pt}erase\ r{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}erule\ Prf{\isacharunderscore}{\kern0pt}elims{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}bnullable\ r{\isadigit{1}}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{prefer}\isamarkupfalse%
+\ {\isadigit{2}}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ bnullable{\isacharunderscore}{\kern0pt}correctness\ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}erule\ Prf{\isacharunderscore}{\kern0pt}elims{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}bnullable\ r{\isadigit{1}}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{prefer}\isamarkupfalse%
+\ {\isadigit{2}}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ bnullable{\isacharunderscore}{\kern0pt}correctness\ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ retrieve{\isacharunderscore}{\kern0pt}fuse{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ bmkeps{\isacharunderscore}{\kern0pt}retrieve{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}erule\ Prf{\isacharunderscore}{\kern0pt}elims{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ bnullable{\isacharunderscore}{\kern0pt}correctness\ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rename{\isacharunderscore}{\kern0pt}tac\ bs\ r\ v{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}erule\ Prf{\isacharunderscore}{\kern0pt}elims{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}clarify{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}erule\ Prf{\isacharunderscore}{\kern0pt}elims{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}clarify{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ injval{\isachardot}{\kern0pt}simps{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp\ del{\isacharcolon}{\kern0pt}\ retrieve{\isachardot}{\kern0pt}simps{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ retrieve{\isachardot}{\kern0pt}simps{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ retrieve{\isachardot}{\kern0pt}simps{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ retrieve{\isacharunderscore}{\kern0pt}fuse{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\ \ \isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ MAIN{\isacharunderscore}{\kern0pt}decode{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}{\isasymTurnstile}\ v\ {\isacharcolon}{\kern0pt}\ ders\ s\ r{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}Some\ {\isacharparenleft}{\kern0pt}flex\ r\ id\ s\ v{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ decode\ {\isacharparenleft}{\kern0pt}retrieve\ {\isacharparenleft}{\kern0pt}bders\ {\isacharparenleft}{\kern0pt}intern\ r{\isacharparenright}{\kern0pt}\ s{\isacharparenright}{\kern0pt}\ v{\isacharparenright}{\kern0pt}\ r{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\isacommand{proof}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}induct\ s\ arbitrary{\isacharcolon}{\kern0pt}\ v\ rule{\isacharcolon}{\kern0pt}\ rev{\isacharunderscore}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{case}\isamarkupfalse%
+\ Nil\isanewline
+\ \ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}{\isasymTurnstile}\ v\ {\isacharcolon}{\kern0pt}\ ders\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ r{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
+\ fact\isanewline
+\ \ \isacommand{then}\isamarkupfalse%
+\ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}{\isasymTurnstile}\ v\ {\isacharcolon}{\kern0pt}\ r{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{then}\isamarkupfalse%
+\ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}Some\ v\ {\isacharequal}{\kern0pt}\ decode\ {\isacharparenleft}{\kern0pt}retrieve\ {\isacharparenleft}{\kern0pt}intern\ r{\isacharparenright}{\kern0pt}\ v{\isacharparenright}{\kern0pt}\ r{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{using}\isamarkupfalse%
+\ decode{\isacharunderscore}{\kern0pt}code\ retrieve{\isacharunderscore}{\kern0pt}code\ \isacommand{by}\isamarkupfalse%
+\ auto\isanewline
+\ \ \isacommand{then}\isamarkupfalse%
+\ \isacommand{show}\isamarkupfalse%
+\ {\isachardoublequoteopen}Some\ {\isacharparenleft}{\kern0pt}flex\ r\ id\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ v{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ decode\ {\isacharparenleft}{\kern0pt}retrieve\ {\isacharparenleft}{\kern0pt}bders\ {\isacharparenleft}{\kern0pt}intern\ r{\isacharparenright}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}\ v{\isacharparenright}{\kern0pt}\ r{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{by}\isamarkupfalse%
+\ simp\isanewline
+\isacommand{next}\isamarkupfalse%
+\isanewline
+\ \ \isacommand{case}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}snoc\ c\ s\ v{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{have}\isamarkupfalse%
+\ IH{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymAnd}v{\isachardot}{\kern0pt}\ {\isasymTurnstile}\ v\ {\isacharcolon}{\kern0pt}\ ders\ s\ r\ {\isasymLongrightarrow}\ \isanewline
+\ \ \ \ \ Some\ {\isacharparenleft}{\kern0pt}flex\ r\ id\ s\ v{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ decode\ {\isacharparenleft}{\kern0pt}retrieve\ {\isacharparenleft}{\kern0pt}bders\ {\isacharparenleft}{\kern0pt}intern\ r{\isacharparenright}{\kern0pt}\ s{\isacharparenright}{\kern0pt}\ v{\isacharparenright}{\kern0pt}\ r{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
+\ fact\isanewline
+\ \ \isacommand{have}\isamarkupfalse%
+\ asm{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymTurnstile}\ v\ {\isacharcolon}{\kern0pt}\ ders\ {\isacharparenleft}{\kern0pt}s\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}c{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}\ r{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
+\ fact\isanewline
+\ \ \isacommand{then}\isamarkupfalse%
+\ \isacommand{have}\isamarkupfalse%
+\ asm{\isadigit{2}}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymTurnstile}\ injval\ {\isacharparenleft}{\kern0pt}ders\ s\ r{\isacharparenright}{\kern0pt}\ c\ v\ {\isacharcolon}{\kern0pt}\ ders\ s\ r{\isachardoublequoteclose}\ \isanewline
+\ \ \ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ Prf{\isacharunderscore}{\kern0pt}injval\ ders{\isacharunderscore}{\kern0pt}append{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}Some\ {\isacharparenleft}{\kern0pt}flex\ r\ id\ {\isacharparenleft}{\kern0pt}s\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}c{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}\ v{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ Some\ {\isacharparenleft}{\kern0pt}flex\ r\ id\ s\ {\isacharparenleft}{\kern0pt}injval\ {\isacharparenleft}{\kern0pt}ders\ s\ r{\isacharparenright}{\kern0pt}\ c\ v{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ flex{\isacharunderscore}{\kern0pt}append{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{also}\isamarkupfalse%
+\ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}{\isachardot}{\kern0pt}{\isachardot}{\kern0pt}{\isachardot}{\kern0pt}\ {\isacharequal}{\kern0pt}\ decode\ {\isacharparenleft}{\kern0pt}retrieve\ {\isacharparenleft}{\kern0pt}bders\ {\isacharparenleft}{\kern0pt}intern\ r{\isacharparenright}{\kern0pt}\ s{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}injval\ {\isacharparenleft}{\kern0pt}ders\ s\ r{\isacharparenright}{\kern0pt}\ c\ v{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ r{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{using}\isamarkupfalse%
+\ asm{\isadigit{2}}\ IH\ \isacommand{by}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{also}\isamarkupfalse%
+\ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}{\isachardot}{\kern0pt}{\isachardot}{\kern0pt}{\isachardot}{\kern0pt}\ {\isacharequal}{\kern0pt}\ decode\ {\isacharparenleft}{\kern0pt}retrieve\ {\isacharparenleft}{\kern0pt}bder\ c\ {\isacharparenleft}{\kern0pt}bders\ {\isacharparenleft}{\kern0pt}intern\ r{\isacharparenright}{\kern0pt}\ s{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ v{\isacharparenright}{\kern0pt}\ r{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{using}\isamarkupfalse%
+\ asm\ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp{\isacharunderscore}{\kern0pt}all\ add{\isacharcolon}{\kern0pt}\ bder{\isacharunderscore}{\kern0pt}retrieve\ ders{\isacharunderscore}{\kern0pt}append{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{finally}\isamarkupfalse%
+\ \isacommand{show}\isamarkupfalse%
+\ {\isachardoublequoteopen}Some\ {\isacharparenleft}{\kern0pt}flex\ r\ id\ {\isacharparenleft}{\kern0pt}s\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}c{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}\ v{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ \isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ decode\ {\isacharparenleft}{\kern0pt}retrieve\ {\isacharparenleft}{\kern0pt}bders\ {\isacharparenleft}{\kern0pt}intern\ r{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}s\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}c{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ v{\isacharparenright}{\kern0pt}\ r{\isachardoublequoteclose}\ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ bders{\isacharunderscore}{\kern0pt}append{\isacharparenright}{\kern0pt}\isanewline
+\isacommand{qed}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{definition}\isamarkupfalse%
+\ blex\ \isakeyword{where}\isanewline
+\ {\isachardoublequoteopen}blex\ a\ s\ {\isasymequiv}\ if\ bnullable\ {\isacharparenleft}{\kern0pt}bders\ a\ s{\isacharparenright}{\kern0pt}\ then\ Some\ {\isacharparenleft}{\kern0pt}bmkeps\ {\isacharparenleft}{\kern0pt}bders\ a\ s{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ else\ None{\isachardoublequoteclose}\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isacommand{definition}\isamarkupfalse%
+\ blexer\ \isakeyword{where}\isanewline
+\ {\isachardoublequoteopen}blexer\ r\ s\ {\isasymequiv}\ if\ bnullable\ {\isacharparenleft}{\kern0pt}bders\ {\isacharparenleft}{\kern0pt}intern\ r{\isacharparenright}{\kern0pt}\ s{\isacharparenright}{\kern0pt}\ then\ \isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ decode\ {\isacharparenleft}{\kern0pt}bmkeps\ {\isacharparenleft}{\kern0pt}bders\ {\isacharparenleft}{\kern0pt}intern\ r{\isacharparenright}{\kern0pt}\ s{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ r\ else\ None{\isachardoublequoteclose}\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ blexer{\isacharunderscore}{\kern0pt}correctness{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}blexer\ r\ s\ {\isacharequal}{\kern0pt}\ lexer\ r\ s{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+%
+\endisadelimproof
+%
+\isatagproof
+\isacommand{proof}\isamarkupfalse%
+\ {\isacharminus}{\kern0pt}\isanewline
+\ \ \isacommand{{\isacharbraceleft}{\kern0pt}}\isamarkupfalse%
+\ \isacommand{define}\isamarkupfalse%
+\ bds\ \isakeyword{where}\ {\isachardoublequoteopen}bds\ {\isasymequiv}\ bders\ {\isacharparenleft}{\kern0pt}intern\ r{\isacharparenright}{\kern0pt}\ s{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{define}\isamarkupfalse%
+\ ds\ \ \isakeyword{where}\ {\isachardoublequoteopen}ds\ {\isasymequiv}\ ders\ s\ r{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{assume}\isamarkupfalse%
+\ asm{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}nullable\ ds{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{have}\isamarkupfalse%
+\ era{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}erase\ bds\ {\isacharequal}{\kern0pt}\ ds{\isachardoublequoteclose}\ \isanewline
+\ \ \ \ \ \ \isacommand{unfolding}\isamarkupfalse%
+\ ds{\isacharunderscore}{\kern0pt}def\ bds{\isacharunderscore}{\kern0pt}def\ \isacommand{by}\isamarkupfalse%
+\ simp\isanewline
+\ \ \ \ \isacommand{have}\isamarkupfalse%
+\ mke{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymTurnstile}\ mkeps\ ds\ {\isacharcolon}{\kern0pt}\ ds{\isachardoublequoteclose}\isanewline
+\ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
+\ asm\ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ mkeps{\isacharunderscore}{\kern0pt}nullable{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}decode\ {\isacharparenleft}{\kern0pt}bmkeps\ bds{\isacharparenright}{\kern0pt}\ r\ {\isacharequal}{\kern0pt}\ decode\ {\isacharparenleft}{\kern0pt}retrieve\ bds\ {\isacharparenleft}{\kern0pt}mkeps\ ds{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ r{\isachardoublequoteclose}\isanewline
+\ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
+\ bmkeps{\isacharunderscore}{\kern0pt}retrieve\isanewline
+\ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
+\ asm\ era\ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ bmkeps{\isacharunderscore}{\kern0pt}retrieve{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{also}\isamarkupfalse%
+\ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}{\isachardot}{\kern0pt}{\isachardot}{\kern0pt}{\isachardot}{\kern0pt}\ {\isacharequal}{\kern0pt}\ \ Some\ {\isacharparenleft}{\kern0pt}flex\ r\ id\ s\ {\isacharparenleft}{\kern0pt}mkeps\ ds{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\ \ \ \ \ \ \isacommand{using}\isamarkupfalse%
+\ mke\ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp{\isacharunderscore}{\kern0pt}all\ add{\isacharcolon}{\kern0pt}\ MAIN{\isacharunderscore}{\kern0pt}decode\ ds{\isacharunderscore}{\kern0pt}def\ bds{\isacharunderscore}{\kern0pt}def{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{finally}\isamarkupfalse%
+\ \isacommand{have}\isamarkupfalse%
+\ {\isachardoublequoteopen}decode\ {\isacharparenleft}{\kern0pt}bmkeps\ bds{\isacharparenright}{\kern0pt}\ r\ {\isacharequal}{\kern0pt}\ Some\ {\isacharparenleft}{\kern0pt}flex\ r\ id\ s\ {\isacharparenleft}{\kern0pt}mkeps\ ds{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\ \isanewline
+\ \ \ \ \ \ \isacommand{unfolding}\isamarkupfalse%
+\ bds{\isacharunderscore}{\kern0pt}def\ ds{\isacharunderscore}{\kern0pt}def\ \isacommand{{\isachardot}{\kern0pt}}\isamarkupfalse%
+\isanewline
+\ \ \isacommand{{\isacharbraceright}{\kern0pt}}\isamarkupfalse%
+\isanewline
+\ \ \isacommand{then}\isamarkupfalse%
+\ \isacommand{show}\isamarkupfalse%
+\ {\isachardoublequoteopen}blexer\ r\ s\ {\isacharequal}{\kern0pt}\ lexer\ r\ s{\isachardoublequoteclose}\isanewline
+\ \ \ \ \isacommand{unfolding}\isamarkupfalse%
+\ blexer{\isacharunderscore}{\kern0pt}def\ lexer{\isacharunderscore}{\kern0pt}flex\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ bnullable{\isacharunderscore}{\kern0pt}correctness{\isacharbrackleft}{\kern0pt}symmetric{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{done}\isamarkupfalse%
+\isanewline
+\isacommand{qed}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ distinctBy\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isacharprime}{\kern0pt}a\ list\ {\isasymRightarrow}\ {\isacharparenleft}{\kern0pt}{\isacharprime}{\kern0pt}a\ {\isasymRightarrow}\ {\isacharprime}{\kern0pt}b{\isacharparenright}{\kern0pt}\ {\isasymRightarrow}\ {\isacharprime}{\kern0pt}b\ set\ {\isasymRightarrow}\ {\isacharprime}{\kern0pt}a\ list{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}distinctBy\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ f\ acc\ {\isacharequal}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}distinctBy\ {\isacharparenleft}{\kern0pt}x{\isacharhash}{\kern0pt}xs{\isacharparenright}{\kern0pt}\ f\ acc\ {\isacharequal}{\kern0pt}\ \isanewline
+\ \ \ \ \ {\isacharparenleft}{\kern0pt}if\ {\isacharparenleft}{\kern0pt}f\ x{\isacharparenright}{\kern0pt}\ {\isasymin}\ acc\ then\ distinctBy\ xs\ f\ acc\ \isanewline
+\ \ \ \ \ \ else\ x\ {\isacharhash}{\kern0pt}\ {\isacharparenleft}{\kern0pt}distinctBy\ xs\ f\ {\isacharparenleft}{\kern0pt}{\isacharbraceleft}{\kern0pt}f\ x{\isacharbraceright}{\kern0pt}\ {\isasymunion}\ acc{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ flts\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}arexp\ list\ {\isasymRightarrow}\ arexp\ list{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{where}\ \isanewline
+\ \ {\isachardoublequoteopen}flts\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}flts\ {\isacharparenleft}{\kern0pt}AZERO\ {\isacharhash}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ flts\ rs{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}flts\ {\isacharparenleft}{\kern0pt}{\isacharparenleft}{\kern0pt}AALTs\ bs\ \ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharhash}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}map\ {\isacharparenleft}{\kern0pt}fuse\ bs{\isacharparenright}{\kern0pt}\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharat}{\kern0pt}\ flts\ rs{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}flts\ {\isacharparenleft}{\kern0pt}r{\isadigit{1}}\ {\isacharhash}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ r{\isadigit{1}}\ {\isacharhash}{\kern0pt}\ flts\ rs{\isachardoublequoteclose}\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ li\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}bit\ list\ {\isasymRightarrow}\ arexp\ list\ {\isasymRightarrow}\ arexp{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}li\ {\isacharunderscore}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ AZERO{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}li\ bs\ {\isacharbrackleft}{\kern0pt}a{\isacharbrackright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ fuse\ bs\ a{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}li\ bs\ as\ {\isacharequal}{\kern0pt}\ AALTs\ bs\ as{\isachardoublequoteclose}\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ bsimp{\isacharunderscore}{\kern0pt}ASEQ\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}bit\ list\ {\isasymRightarrow}\ arexp\ {\isasymRightarrow}\ arexp\ {\isasymRightarrow}\ arexp{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}bsimp{\isacharunderscore}{\kern0pt}ASEQ\ {\isacharunderscore}{\kern0pt}\ AZERO\ {\isacharunderscore}{\kern0pt}\ {\isacharequal}{\kern0pt}\ AZERO{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bsimp{\isacharunderscore}{\kern0pt}ASEQ\ {\isacharunderscore}{\kern0pt}\ {\isacharunderscore}{\kern0pt}\ AZERO\ {\isacharequal}{\kern0pt}\ AZERO{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bsimp{\isacharunderscore}{\kern0pt}ASEQ\ bs{\isadigit{1}}\ {\isacharparenleft}{\kern0pt}AONE\ bs{\isadigit{2}}{\isacharparenright}{\kern0pt}\ r{\isadigit{2}}\ {\isacharequal}{\kern0pt}\ fuse\ {\isacharparenleft}{\kern0pt}bs{\isadigit{1}}\ {\isacharat}{\kern0pt}\ bs{\isadigit{2}}{\isacharparenright}{\kern0pt}\ r{\isadigit{2}}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bsimp{\isacharunderscore}{\kern0pt}ASEQ\ bs{\isadigit{1}}\ r{\isadigit{1}}\ r{\isadigit{2}}\ {\isacharequal}{\kern0pt}\ ASEQ\ \ bs{\isadigit{1}}\ r{\isadigit{1}}\ r{\isadigit{2}}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ bsimp{\isacharunderscore}{\kern0pt}AALTs\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}bit\ list\ {\isasymRightarrow}\ arexp\ list\ {\isasymRightarrow}\ arexp{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}bsimp{\isacharunderscore}{\kern0pt}AALTs\ {\isacharunderscore}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ AZERO{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bsimp{\isacharunderscore}{\kern0pt}AALTs\ bs{\isadigit{1}}\ {\isacharbrackleft}{\kern0pt}r{\isacharbrackright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ fuse\ bs{\isadigit{1}}\ r{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bsimp{\isacharunderscore}{\kern0pt}AALTs\ bs{\isadigit{1}}\ rs\ {\isacharequal}{\kern0pt}\ AALTs\ bs{\isadigit{1}}\ rs{\isachardoublequoteclose}\isanewline
+\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ bsimp\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}arexp\ {\isasymRightarrow}\ arexp{\isachardoublequoteclose}\ \isanewline
+\ \ \isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}bsimp\ {\isacharparenleft}{\kern0pt}ASEQ\ bs{\isadigit{1}}\ r{\isadigit{1}}\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bsimp{\isacharunderscore}{\kern0pt}ASEQ\ bs{\isadigit{1}}\ {\isacharparenleft}{\kern0pt}bsimp\ r{\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}bsimp\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bsimp\ {\isacharparenleft}{\kern0pt}AALTs\ bs{\isadigit{1}}\ rs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bsimp{\isacharunderscore}{\kern0pt}AALTs\ bs{\isadigit{1}}\ {\isacharparenleft}{\kern0pt}distinctBy\ \ {\isacharparenleft}{\kern0pt}flts\ {\isacharparenleft}{\kern0pt}map\ bsimp\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ erase\ {\isacharbraceleft}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharparenright}{\kern0pt}\ {\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bsimp\ r\ {\isacharequal}{\kern0pt}\ r{\isachardoublequoteclose}\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ \isanewline
+\ \ bders{\isacharunderscore}{\kern0pt}simp\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}arexp\ {\isasymRightarrow}\ string\ {\isasymRightarrow}\ arexp{\isachardoublequoteclose}\isanewline
+\isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}bders{\isacharunderscore}{\kern0pt}simp\ r\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ r{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}bders{\isacharunderscore}{\kern0pt}simp\ r\ {\isacharparenleft}{\kern0pt}c\ {\isacharhash}{\kern0pt}\ s{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bders{\isacharunderscore}{\kern0pt}simp\ {\isacharparenleft}{\kern0pt}bsimp\ {\isacharparenleft}{\kern0pt}bder\ c\ r{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ s{\isachardoublequoteclose}\isanewline
+\isanewline
+\isacommand{definition}\isamarkupfalse%
+\ blexer{\isacharunderscore}{\kern0pt}simp\ \isakeyword{where}\isanewline
+\ {\isachardoublequoteopen}blexer{\isacharunderscore}{\kern0pt}simp\ r\ s\ {\isasymequiv}\ if\ bnullable\ {\isacharparenleft}{\kern0pt}bders{\isacharunderscore}{\kern0pt}simp\ {\isacharparenleft}{\kern0pt}intern\ r{\isacharparenright}{\kern0pt}\ s{\isacharparenright}{\kern0pt}\ then\ \isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ decode\ {\isacharparenleft}{\kern0pt}bmkeps\ {\isacharparenleft}{\kern0pt}bders{\isacharunderscore}{\kern0pt}simp\ {\isacharparenleft}{\kern0pt}intern\ r{\isacharparenright}{\kern0pt}\ s{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ r\ else\ None{\isachardoublequoteclose}\isanewline
+\isanewline
+\isacommand{export{\isacharunderscore}{\kern0pt}code}\isamarkupfalse%
+\ bders{\isacharunderscore}{\kern0pt}simp\ \isakeyword{in}\ Scala\ \isakeyword{module{\isacharunderscore}{\kern0pt}name}\ Example\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ bders{\isacharunderscore}{\kern0pt}simp{\isacharunderscore}{\kern0pt}append{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}bders{\isacharunderscore}{\kern0pt}simp\ r\ {\isacharparenleft}{\kern0pt}s{\isadigit{1}}\ {\isacharat}{\kern0pt}\ s{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bders{\isacharunderscore}{\kern0pt}simp\ {\isacharparenleft}{\kern0pt}bders{\isacharunderscore}{\kern0pt}simp\ r\ s{\isadigit{1}}{\isacharparenright}{\kern0pt}\ s{\isadigit{2}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ s{\isadigit{1}}\ arbitrary{\isacharcolon}{\kern0pt}\ r\ s{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ L{\isacharunderscore}{\kern0pt}bsimp{\isacharunderscore}{\kern0pt}ASEQ{\isacharcolon}{\kern0pt}\isanewline
+\ \ {\isachardoublequoteopen}L\ {\isacharparenleft}{\kern0pt}SEQ\ {\isacharparenleft}{\kern0pt}erase\ r{\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}erase\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ L\ {\isacharparenleft}{\kern0pt}erase\ {\isacharparenleft}{\kern0pt}bsimp{\isacharunderscore}{\kern0pt}ASEQ\ bs\ r{\isadigit{1}}\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ bs\ r{\isadigit{1}}\ r{\isadigit{2}}\ rule{\isacharcolon}{\kern0pt}\ bsimp{\isacharunderscore}{\kern0pt}ASEQ{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharunderscore}{\kern0pt}all{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ erase{\isacharunderscore}{\kern0pt}fuse\ fuse{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ L{\isacharunderscore}{\kern0pt}bsimp{\isacharunderscore}{\kern0pt}AALTs{\isacharcolon}{\kern0pt}\isanewline
+\ \ {\isachardoublequoteopen}L\ {\isacharparenleft}{\kern0pt}erase\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ L\ {\isacharparenleft}{\kern0pt}erase\ {\isacharparenleft}{\kern0pt}bsimp{\isacharunderscore}{\kern0pt}AALTs\ bs\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ bs\ rs\ rule{\isacharcolon}{\kern0pt}\ bsimp{\isacharunderscore}{\kern0pt}AALTs{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharunderscore}{\kern0pt}all\ add{\isacharcolon}{\kern0pt}\ erase{\isacharunderscore}{\kern0pt}fuse{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ L{\isacharunderscore}{\kern0pt}erase{\isacharunderscore}{\kern0pt}AALTs{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}L\ {\isacharparenleft}{\kern0pt}erase\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isasymUnion}\ {\isacharparenleft}{\kern0pt}L\ {\isacharbackquote}{\kern0pt}\ erase\ {\isacharbackquote}{\kern0pt}\ {\isacharparenleft}{\kern0pt}set\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ rs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ rs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ L{\isacharunderscore}{\kern0pt}erase{\isacharunderscore}{\kern0pt}flts{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}{\isasymUnion}\ {\isacharparenleft}{\kern0pt}L\ {\isacharbackquote}{\kern0pt}\ erase\ {\isacharbackquote}{\kern0pt}\ {\isacharparenleft}{\kern0pt}set\ {\isacharparenleft}{\kern0pt}flts\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isasymUnion}\ {\isacharparenleft}{\kern0pt}L\ {\isacharbackquote}{\kern0pt}\ erase\ {\isacharbackquote}{\kern0pt}\ {\isacharparenleft}{\kern0pt}set\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ rs\ rule{\isacharcolon}{\kern0pt}\ flts{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharunderscore}{\kern0pt}all{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ L{\isacharunderscore}{\kern0pt}erase{\isacharunderscore}{\kern0pt}AALTs\ erase{\isacharunderscore}{\kern0pt}fuse\ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ L{\isacharunderscore}{\kern0pt}erase{\isacharunderscore}{\kern0pt}AALTs\ erase{\isacharunderscore}{\kern0pt}fuse{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ L{\isacharunderscore}{\kern0pt}erase{\isacharunderscore}{\kern0pt}dB{\isacharunderscore}{\kern0pt}acc{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}{\isacharparenleft}{\kern0pt}\ {\isasymUnion}{\isacharparenleft}{\kern0pt}L\ {\isacharbackquote}{\kern0pt}\ acc{\isacharparenright}{\kern0pt}\ {\isasymunion}\ {\isacharparenleft}{\kern0pt}\ {\isasymUnion}\ {\isacharparenleft}{\kern0pt}L\ {\isacharbackquote}{\kern0pt}\ erase\ {\isacharbackquote}{\kern0pt}\ {\isacharparenleft}{\kern0pt}set\ {\isacharparenleft}{\kern0pt}distinctBy\ rs\ erase\ acc{\isacharparenright}{\kern0pt}\ {\isacharparenright}{\kern0pt}\ {\isacharparenright}{\kern0pt}\ {\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isasymUnion}{\isacharparenleft}{\kern0pt}L\ {\isacharbackquote}{\kern0pt}\ acc{\isacharparenright}{\kern0pt}\ {\isasymunion}\ \ {\isasymUnion}\ {\isacharparenleft}{\kern0pt}L\ {\isacharbackquote}{\kern0pt}\ erase\ {\isacharbackquote}{\kern0pt}\ {\isacharparenleft}{\kern0pt}set\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induction\ rs\ arbitrary{\isacharcolon}{\kern0pt}\ acc{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}smt\ {\isacharparenleft}{\kern0pt}z{\isadigit{3}}{\isacharparenright}{\kern0pt}\ SUP{\isacharunderscore}{\kern0pt}absorb\ UN{\isacharunderscore}{\kern0pt}insert\ sup{\isacharunderscore}{\kern0pt}assoc\ sup{\isacharunderscore}{\kern0pt}commute{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ L{\isacharunderscore}{\kern0pt}erase{\isacharunderscore}{\kern0pt}dB{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}\ {\isacharparenleft}{\kern0pt}\ {\isasymUnion}\ {\isacharparenleft}{\kern0pt}L\ {\isacharbackquote}{\kern0pt}\ erase\ {\isacharbackquote}{\kern0pt}\ {\isacharparenleft}{\kern0pt}set\ {\isacharparenleft}{\kern0pt}distinctBy\ rs\ erase\ {\isacharbraceleft}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharparenright}{\kern0pt}\ {\isacharparenright}{\kern0pt}\ {\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isasymUnion}\ {\isacharparenleft}{\kern0pt}L\ {\isacharbackquote}{\kern0pt}\ erase\ {\isacharbackquote}{\kern0pt}\ {\isacharparenleft}{\kern0pt}set\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ L{\isacharunderscore}{\kern0pt}erase{\isacharunderscore}{\kern0pt}dB{\isacharunderscore}{\kern0pt}acc\ Un{\isacharunderscore}{\kern0pt}commute\ Union{\isacharunderscore}{\kern0pt}image{\isacharunderscore}{\kern0pt}empty{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ L{\isacharunderscore}{\kern0pt}bsimp{\isacharunderscore}{\kern0pt}erase{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}L\ {\isacharparenleft}{\kern0pt}erase\ r{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ L\ {\isacharparenleft}{\kern0pt}erase\ {\isacharparenleft}{\kern0pt}bsimp\ r{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ r{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto\ simp\ add{\isacharcolon}{\kern0pt}\ Sequ{\isacharunderscore}{\kern0pt}def{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ L{\isacharunderscore}{\kern0pt}bsimp{\isacharunderscore}{\kern0pt}ASEQ{\isacharbrackleft}{\kern0pt}symmetric{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto\ simp\ add{\isacharcolon}{\kern0pt}\ Sequ{\isacharunderscore}{\kern0pt}def{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ {\isacharparenleft}{\kern0pt}asm{\isacharparenright}{\kern0pt}\ \ L{\isacharunderscore}{\kern0pt}bsimp{\isacharunderscore}{\kern0pt}ASEQ{\isacharbrackleft}{\kern0pt}symmetric{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto\ simp\ add{\isacharcolon}{\kern0pt}\ Sequ{\isacharunderscore}{\kern0pt}def{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ L{\isacharunderscore}{\kern0pt}bsimp{\isacharunderscore}{\kern0pt}AALTs{\isacharbrackleft}{\kern0pt}symmetric{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{defer}\isamarkupfalse%
+\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ {\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}L{\isacharunderscore}{\kern0pt}erase{\isacharunderscore}{\kern0pt}AALTs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ L{\isacharunderscore}{\kern0pt}erase{\isacharunderscore}{\kern0pt}dB{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ L{\isacharunderscore}{\kern0pt}erase{\isacharunderscore}{\kern0pt}flts{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ L{\isacharunderscore}{\kern0pt}erase{\isacharunderscore}{\kern0pt}AALTs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ L{\isacharunderscore}{\kern0pt}erase{\isacharunderscore}{\kern0pt}AALTs\ \isacommand{by}\isamarkupfalse%
+\ blast%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ bsimp{\isacharunderscore}{\kern0pt}ASEQ{\isadigit{0}}{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}bsimp{\isacharunderscore}{\kern0pt}ASEQ\ bs\ r{\isadigit{1}}\ AZERO\ {\isacharequal}{\kern0pt}\ AZERO{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ r{\isadigit{1}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ bsimp{\isacharunderscore}{\kern0pt}ASEQ{\isadigit{1}}{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}r{\isadigit{1}}\ {\isasymnoteq}\ AZERO{\isachardoublequoteclose}\ {\isachardoublequoteopen}r{\isadigit{2}}\ {\isasymnoteq}\ AZERO{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isasymforall}bs{\isachardot}{\kern0pt}\ r{\isadigit{1}}\ {\isasymnoteq}\ AONE\ bs{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}bsimp{\isacharunderscore}{\kern0pt}ASEQ\ bs\ r{\isadigit{1}}\ r{\isadigit{2}}\ {\isacharequal}{\kern0pt}\ ASEQ\ bs\ r{\isadigit{1}}\ r{\isadigit{2}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ bs\ r{\isadigit{1}}\ r{\isadigit{2}}\ rule{\isacharcolon}{\kern0pt}\ bsimp{\isacharunderscore}{\kern0pt}ASEQ{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ bsimp{\isacharunderscore}{\kern0pt}ASEQ{\isadigit{2}}{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}bsimp{\isacharunderscore}{\kern0pt}ASEQ\ bs\ {\isacharparenleft}{\kern0pt}AONE\ bs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ r{\isadigit{2}}\ {\isacharequal}{\kern0pt}\ fuse\ {\isacharparenleft}{\kern0pt}bs\ {\isacharat}{\kern0pt}\ bs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ r{\isadigit{2}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ L{\isacharunderscore}{\kern0pt}bders{\isacharunderscore}{\kern0pt}simp{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}L\ {\isacharparenleft}{\kern0pt}erase\ {\isacharparenleft}{\kern0pt}bders{\isacharunderscore}{\kern0pt}simp\ r\ s{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ L\ {\isacharparenleft}{\kern0pt}erase\ {\isacharparenleft}{\kern0pt}bders\ r\ s{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ s\ arbitrary{\isacharcolon}{\kern0pt}\ r\ rule{\isacharcolon}{\kern0pt}\ rev{\isacharunderscore}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ ders{\isacharunderscore}{\kern0pt}append{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ bders{\isacharunderscore}{\kern0pt}simp{\isacharunderscore}{\kern0pt}append{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ L{\isacharunderscore}{\kern0pt}bsimp{\isacharunderscore}{\kern0pt}erase{\isacharbrackleft}{\kern0pt}symmetric{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ der{\isacharunderscore}{\kern0pt}correctness{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ b{\isadigit{2}}{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}bnullable\ r{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}bmkeps\ {\isacharparenleft}{\kern0pt}fuse\ bs\ r{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bs\ {\isacharat}{\kern0pt}\ bmkeps\ r{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ assms\ bmkeps{\isacharunderscore}{\kern0pt}retrieve\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isacharunderscore}{\kern0pt}fuse\ mkeps{\isacharunderscore}{\kern0pt}nullable\ retrieve{\isacharunderscore}{\kern0pt}fuse{\isadigit{2}}{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ b{\isadigit{4}}{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}bnullable\ {\isacharparenleft}{\kern0pt}bders{\isacharunderscore}{\kern0pt}simp\ r\ s{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bnullable\ {\isacharparenleft}{\kern0pt}bders\ r\ s{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ L{\isacharunderscore}{\kern0pt}bders{\isacharunderscore}{\kern0pt}simp\ bnullable{\isacharunderscore}{\kern0pt}correctness\ lexer{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}\ lexer{\isacharunderscore}{\kern0pt}correct{\isacharunderscore}{\kern0pt}None\ option{\isachardot}{\kern0pt}distinct{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ qq{\isadigit{1}}{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}{\isasymexists}r\ {\isasymin}\ set\ rs{\isachardot}{\kern0pt}\ bnullable\ r{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs\ {\isacharat}{\kern0pt}\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ rs\ arbitrary{\isacharcolon}{\kern0pt}\ rs{\isadigit{1}}\ bs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ Nil{\isacharunderscore}{\kern0pt}is{\isacharunderscore}{\kern0pt}append{\isacharunderscore}{\kern0pt}conv\ bmkeps{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}\ neq{\isacharunderscore}{\kern0pt}Nil{\isacharunderscore}{\kern0pt}conv\ bnullable{\isacharunderscore}{\kern0pt}Hdbmkeps{\isacharunderscore}{\kern0pt}Hd\ split{\isacharunderscore}{\kern0pt}list{\isacharunderscore}{\kern0pt}last{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ qq{\isadigit{2}}{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}{\isasymforall}r\ {\isasymin}\ set\ rs{\isachardot}{\kern0pt}\ {\isasymnot}\ bnullable\ r{\isachardoublequoteclose}\ {\isachardoublequoteopen}{\isasymexists}r\ {\isasymin}\ set\ rs{\isadigit{1}}{\isachardot}{\kern0pt}\ bnullable\ r{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs\ {\isacharat}{\kern0pt}\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ rs\ arbitrary{\isacharcolon}{\kern0pt}\ rs{\isadigit{1}}\ bs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ append{\isacharunderscore}{\kern0pt}assoc\ in{\isacharunderscore}{\kern0pt}set{\isacharunderscore}{\kern0pt}conv{\isacharunderscore}{\kern0pt}decomp\ r{\isadigit{1}}\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\ \ \isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ qq{\isadigit{3}}{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}{\isasymexists}r\ {\isasymin}\ set\ rs{\isachardot}{\kern0pt}\ bnullable\ r{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ rs\ arbitrary{\isacharcolon}{\kern0pt}\ bs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ nonnested\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}arexp\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}nonnested\ {\isacharparenleft}{\kern0pt}AALTs\ bs{\isadigit{2}}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ True{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}nonnested\ {\isacharparenleft}{\kern0pt}AALTs\ bs{\isadigit{2}}\ {\isacharparenleft}{\kern0pt}{\isacharparenleft}{\kern0pt}AALTs\ bs{\isadigit{1}}\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharhash}{\kern0pt}\ rs{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ False{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}nonnested\ {\isacharparenleft}{\kern0pt}AALTs\ bs{\isadigit{2}}\ {\isacharparenleft}{\kern0pt}r\ {\isacharhash}{\kern0pt}\ rs{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ nonnested\ {\isacharparenleft}{\kern0pt}AALTs\ bs{\isadigit{2}}\ rs{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}nonnested\ r\ {\isacharequal}{\kern0pt}\ True{\isachardoublequoteclose}\isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ \ k{\isadigit{0}}{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}flts\ {\isacharparenleft}{\kern0pt}r\ {\isacharhash}{\kern0pt}\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ flts\ {\isacharbrackleft}{\kern0pt}r{\isacharbrackright}{\kern0pt}\ {\isacharat}{\kern0pt}\ flts\ rs{\isadigit{1}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ r\ arbitrary{\isacharcolon}{\kern0pt}\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ \ k{\isadigit{0}}{\isadigit{0}}{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}flts\ {\isacharparenleft}{\kern0pt}rs{\isadigit{1}}\ {\isacharat}{\kern0pt}\ rs{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ flts\ rs{\isadigit{1}}\ {\isacharat}{\kern0pt}\ flts\ rs{\isadigit{2}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ rs{\isadigit{1}}\ arbitrary{\isacharcolon}{\kern0pt}\ rs{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ append{\isachardot}{\kern0pt}assoc\ k{\isadigit{0}}{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ \ k{\isadigit{0}}a{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}flts\ {\isacharbrackleft}{\kern0pt}AALTs\ bs\ rs{\isacharbrackright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ map\ {\isacharparenleft}{\kern0pt}fuse\ bs{\isacharparenright}{\kern0pt}\ \ rs{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ bsimp{\isacharunderscore}{\kern0pt}AALTs{\isacharunderscore}{\kern0pt}qq{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}{\isadigit{1}}\ {\isacharless}{\kern0pt}\ length\ rs{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}bsimp{\isacharunderscore}{\kern0pt}AALTs\ bs\ rs\ {\isacharequal}{\kern0pt}\ AALTs\ bs\ \ rs{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ \ assms\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ rs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ list{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharunderscore}{\kern0pt}all{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ bbbbs{\isadigit{1}}{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}nonalt\ r\ {\isasymor}\ {\isacharparenleft}{\kern0pt}{\isasymexists}bs\ rs{\isachardot}{\kern0pt}\ r\ \ {\isacharequal}{\kern0pt}\ AALTs\ bs\ rs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ nonalt{\isachardot}{\kern0pt}elims{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}\ \isacommand{by}\isamarkupfalse%
+\ auto%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\ \ \isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ flts{\isacharunderscore}{\kern0pt}append{\isacharcolon}{\kern0pt}\isanewline
+\ \ {\isachardoublequoteopen}flts\ {\isacharparenleft}{\kern0pt}xs{\isadigit{1}}\ {\isacharat}{\kern0pt}\ xs{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ flts\ xs{\isadigit{1}}\ {\isacharat}{\kern0pt}\ flts\ xs{\isadigit{2}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ xs{\isadigit{1}}\ \ arbitrary{\isacharcolon}{\kern0pt}\ xs{\isadigit{2}}\ \ rule{\isacharcolon}{\kern0pt}\ rev{\isacharunderscore}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ xs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ x{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ x{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ nonazero\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}arexp\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}nonazero\ AZERO\ {\isacharequal}{\kern0pt}\ False{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}nonazero\ r\ {\isacharequal}{\kern0pt}\ True{\isachardoublequoteclose}\isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ flts{\isacharunderscore}{\kern0pt}single{\isadigit{1}}{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}nonalt\ r{\isachardoublequoteclose}\ {\isachardoublequoteopen}nonazero\ r{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}flts\ {\isacharbrackleft}{\kern0pt}r{\isacharbrackright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}r{\isacharbrackright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ r{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ q{\isadigit{3}}a{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}{\isasymexists}r\ {\isasymin}\ set\ rs{\isachardot}{\kern0pt}\ bnullable\ r{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}map\ {\isacharparenleft}{\kern0pt}fuse\ bs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ {\isacharparenleft}{\kern0pt}bs{\isacharat}{\kern0pt}bs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ rs\ arbitrary{\isacharcolon}{\kern0pt}\ bs\ bs{\isadigit{1}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ append{\isacharunderscore}{\kern0pt}assoc\ b{\isadigit{2}}\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isacharunderscore}{\kern0pt}fuse\ bnullable{\isacharunderscore}{\kern0pt}Hdbmkeps{\isacharunderscore}{\kern0pt}Hd{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}bnullable\ a{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ append{\isachardot}{\kern0pt}assoc\ b{\isadigit{2}}\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isacharunderscore}{\kern0pt}fuse\ bnullable{\isacharunderscore}{\kern0pt}Hdbmkeps{\isacharunderscore}{\kern0pt}Hd{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ rs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isacharunderscore}{\kern0pt}fuse{\isacharparenright}{\kern0pt}{\isacharplus}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ qq{\isadigit{4}}{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}{\isasymexists}x{\isasymin}set\ list{\isachardot}{\kern0pt}\ bnullable\ x{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}{\isasymexists}x{\isasymin}set\ {\isacharparenleft}{\kern0pt}flts\ list{\isacharparenright}{\kern0pt}{\isachardot}{\kern0pt}\ bnullable\ x{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ list\ rule{\isacharcolon}{\kern0pt}\ flts{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ UnCI\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isacharunderscore}{\kern0pt}fuse\ imageI{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\ \ \isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ qs{\isadigit{3}}{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}{\isasymexists}r\ {\isasymin}\ set\ rs{\isachardot}{\kern0pt}\ bnullable\ r{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}flts\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ rs\ arbitrary{\isacharcolon}{\kern0pt}\ bs\ taking{\isacharcolon}{\kern0pt}\ size\ rule{\isacharcolon}{\kern0pt}\ measure{\isacharunderscore}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ x{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ a{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ r{\isadigit{1}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ bnullable{\isacharunderscore}{\kern0pt}Hdbmkeps{\isacharunderscore}{\kern0pt}Hd{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}flts\ list{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ L{\isacharunderscore}{\kern0pt}erase{\isacharunderscore}{\kern0pt}AALTs\ L{\isacharunderscore}{\kern0pt}erase{\isacharunderscore}{\kern0pt}flts\ L{\isacharunderscore}{\kern0pt}flat{\isacharunderscore}{\kern0pt}Prf{\isadigit{1}}\ L{\isacharunderscore}{\kern0pt}flat{\isacharunderscore}{\kern0pt}Prf{\isadigit{2}}\ Prf{\isacharunderscore}{\kern0pt}elims{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}\ mkeps{\isacharunderscore}{\kern0pt}nullable\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ r{\isadigit{1}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{prefer}\isamarkupfalse%
+\ {\isadigit{3}}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ bnullable{\isacharunderscore}{\kern0pt}Hdbmkeps{\isacharunderscore}{\kern0pt}Hd{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{prefer}\isamarkupfalse%
+\ {\isadigit{2}}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}{\isasymexists}x{\isasymin}set\ x{\isadigit{5}}{\isadigit{2}}{\isachardot}{\kern0pt}\ bnullable\ x{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}list{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ b{\isadigit{2}}\ fuse{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}\ q{\isadigit{3}}a\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}erule\ disjE{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ qq{\isadigit{1}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isacharunderscore}{\kern0pt}fuse{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ b{\isadigit{2}}\ fuse{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}\ q{\isadigit{3}}a\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ qq{\isadigit{1}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isacharunderscore}{\kern0pt}fuse\ image{\isacharunderscore}{\kern0pt}eqI\ set{\isacharunderscore}{\kern0pt}map{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ b{\isadigit{2}}\ fuse{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}\ q{\isadigit{3}}a\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ qq{\isadigit{1}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isacharunderscore}{\kern0pt}fuse\ image{\isacharunderscore}{\kern0pt}eqI\ set{\isacharunderscore}{\kern0pt}map{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ b{\isadigit{2}}\ fuse{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}\ q{\isadigit{3}}a\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ qq{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isacharunderscore}{\kern0pt}fuse\ imageE\ set{\isacharunderscore}{\kern0pt}map{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{prefer}\isamarkupfalse%
+\ {\isadigit{2}}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}list{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ qq{\isadigit{4}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ list{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ bnullable{\isacharunderscore}{\kern0pt}Hdbmkeps{\isacharunderscore}{\kern0pt}Hd{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}bnullable\ {\isacharparenleft}{\kern0pt}ASEQ\ x{\isadigit{4}}{\isadigit{1}}\ x{\isadigit{4}}{\isadigit{2}}\ x{\isadigit{4}}{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ list{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ bnullable{\isacharunderscore}{\kern0pt}Hdbmkeps{\isacharunderscore}{\kern0pt}Hd{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ qq{\isadigit{4}}\ r{\isadigit{1}}\ r{\isadigit{2}}\ \isacommand{by}\isamarkupfalse%
+\ auto%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isanewline
+\ \ \isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ bder{\isacharunderscore}{\kern0pt}fuse{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}bder\ c\ {\isacharparenleft}{\kern0pt}fuse\ bs\ a{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ fuse\ bs\ \ {\isacharparenleft}{\kern0pt}bder\ c\ a{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ a\ arbitrary{\isacharcolon}{\kern0pt}\ bs\ c{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharunderscore}{\kern0pt}all{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ flts{\isadigit{2}}\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}char\ {\isasymRightarrow}\ arexp\ list\ {\isasymRightarrow}\ arexp\ list{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{where}\ \isanewline
+\ \ {\isachardoublequoteopen}flts{\isadigit{2}}\ {\isacharunderscore}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}flts{\isadigit{2}}\ c\ {\isacharparenleft}{\kern0pt}AZERO\ {\isacharhash}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ flts{\isadigit{2}}\ c\ rs{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}flts{\isadigit{2}}\ c\ {\isacharparenleft}{\kern0pt}AONE\ {\isacharunderscore}{\kern0pt}\ {\isacharhash}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ flts{\isadigit{2}}\ c\ rs{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}flts{\isadigit{2}}\ c\ {\isacharparenleft}{\kern0pt}ACHAR\ bs\ d\ {\isacharhash}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}if\ c\ {\isacharequal}{\kern0pt}\ d\ then\ {\isacharparenleft}{\kern0pt}ACHAR\ bs\ d\ {\isacharhash}{\kern0pt}\ flts{\isadigit{2}}\ c\ rs{\isacharparenright}{\kern0pt}\ else\ flts{\isadigit{2}}\ c\ rs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}flts{\isadigit{2}}\ c\ {\isacharparenleft}{\kern0pt}{\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharhash}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}map\ {\isacharparenleft}{\kern0pt}fuse\ bs{\isacharparenright}{\kern0pt}\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharat}{\kern0pt}\ flts{\isadigit{2}}\ c\ rs{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}flts{\isadigit{2}}\ c\ {\isacharparenleft}{\kern0pt}ASEQ\ bs\ r{\isadigit{1}}\ r{\isadigit{2}}\ {\isacharhash}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}if\ {\isacharparenleft}{\kern0pt}bnullable{\isacharparenleft}{\kern0pt}r{\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isasymand}\ r{\isadigit{2}}\ {\isacharequal}{\kern0pt}\ AZERO{\isacharparenright}{\kern0pt}\ then\ \isanewline
+\ \ \ \ flts{\isadigit{2}}\ c\ rs\isanewline
+\ \ \ \ else\ ASEQ\ bs\ r{\isadigit{1}}\ r{\isadigit{2}}\ {\isacharhash}{\kern0pt}\ flts{\isadigit{2}}\ c\ rs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}flts{\isadigit{2}}\ c\ {\isacharparenleft}{\kern0pt}r{\isadigit{1}}\ {\isacharhash}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ r{\isadigit{1}}\ {\isacharhash}{\kern0pt}\ flts{\isadigit{2}}\ c\ rs{\isachardoublequoteclose}\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\ \isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ WQ{\isadigit{1}}{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}s\ {\isasymin}\ L\ {\isacharparenleft}{\kern0pt}der\ c\ r{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}s\ {\isasymin}\ der\ c\ r\ {\isasymrightarrow}\ mkeps\ {\isacharparenleft}{\kern0pt}ders\ s\ {\isacharparenleft}{\kern0pt}der\ c\ r{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\ \ \isacommand{oops}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ bder{\isacharunderscore}{\kern0pt}bsimp{\isacharunderscore}{\kern0pt}AALTs{\isacharcolon}{\kern0pt}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}bder\ c\ {\isacharparenleft}{\kern0pt}bsimp{\isacharunderscore}{\kern0pt}AALTs\ bs\ rs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bsimp{\isacharunderscore}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}map\ {\isacharparenleft}{\kern0pt}bder\ c{\isacharparenright}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ bs\ rs\ rule{\isacharcolon}{\kern0pt}\ bsimp{\isacharunderscore}{\kern0pt}AALTs{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\ \ \isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ bder{\isacharunderscore}{\kern0pt}fuse{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\isanewline
+\ \ \isakeyword{assumes}\ {\isachardoublequoteopen}asize\ {\isacharparenleft}{\kern0pt}bsimp\ a{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ asize\ a{\isachardoublequoteclose}\ \ {\isachardoublequoteopen}a\ {\isacharequal}{\kern0pt}\ AALTs\ bs\ {\isacharbrackleft}{\kern0pt}AALTs\ bs{\isadigit{2}}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharcomma}{\kern0pt}\ AZERO{\isacharcomma}{\kern0pt}\ AONE\ bs{\isadigit{3}}{\isacharbrackright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}bsimp\ a\ {\isacharequal}{\kern0pt}\ a{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ assms\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{oops}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isacommand{inductive}\isamarkupfalse%
+\ rrewrite{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}arexp\ {\isasymRightarrow}\ arexp\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ {\isacharparenleft}{\kern0pt}{\isachardoublequoteopen}{\isacharunderscore}{\kern0pt}\ {\isasymleadsto}\ {\isacharunderscore}{\kern0pt}{\isachardoublequoteclose}\ {\isacharbrackleft}{\kern0pt}{\isadigit{9}}{\isadigit{9}}{\isacharcomma}{\kern0pt}\ {\isadigit{9}}{\isadigit{9}}{\isacharbrackright}{\kern0pt}\ {\isadigit{9}}{\isadigit{9}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}ASEQ\ bs\ AZERO\ r{\isadigit{2}}\ {\isasymleadsto}\ AZERO{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}ASEQ\ bs\ r{\isadigit{1}}\ AZERO\ {\isasymleadsto}\ AZERO{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}ASEQ\ bs\ {\isacharparenleft}{\kern0pt}AONE\ bs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ r\ {\isasymleadsto}\ fuse\ {\isacharparenleft}{\kern0pt}bs{\isacharat}{\kern0pt}bs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ r{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}r{\isadigit{1}}\ {\isasymleadsto}\ r{\isadigit{2}}\ {\isasymLongrightarrow}\ ASEQ\ bs\ r{\isadigit{1}}\ r{\isadigit{3}}\ {\isasymleadsto}\ ASEQ\ bs\ r{\isadigit{2}}\ r{\isadigit{3}}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}r{\isadigit{3}}\ {\isasymleadsto}\ r{\isadigit{4}}\ {\isasymLongrightarrow}\ ASEQ\ bs\ r{\isadigit{1}}\ r{\isadigit{3}}\ {\isasymleadsto}\ ASEQ\ bs\ r{\isadigit{1}}\ r{\isadigit{4}}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}r\ {\isasymleadsto}\ r{\isacharprime}{\kern0pt}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs{\isadigit{1}}\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}r{\isacharbrackright}{\kern0pt}\ {\isacharat}{\kern0pt}\ rs{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isasymleadsto}\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs{\isadigit{1}}\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}r{\isacharprime}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isacharat}{\kern0pt}\ rs{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rsa{\isacharat}{\kern0pt}AZERO\ {\isacharhash}{\kern0pt}\ rsb{\isacharparenright}{\kern0pt}\ {\isasymleadsto}\ AALTs\ bs\ {\isacharparenleft}{\kern0pt}rsa{\isacharat}{\kern0pt}rsb{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rsa{\isacharat}{\kern0pt}{\isacharparenleft}{\kern0pt}AALTs\ bs{\isadigit{1}}\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharhash}{\kern0pt}\ rsb{\isacharparenright}{\kern0pt}\ {\isasymleadsto}\ AALTs\ bs\ {\isacharparenleft}{\kern0pt}rsa{\isacharat}{\kern0pt}{\isacharparenleft}{\kern0pt}map\ {\isacharparenleft}{\kern0pt}fuse\ bs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharat}{\kern0pt}rsb{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}AALTs\ bs\ {\isacharparenleft}{\kern0pt}map\ {\isacharparenleft}{\kern0pt}fuse\ bs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}\ {\isasymleadsto}\ AALTs\ {\isacharparenleft}{\kern0pt}bs{\isacharat}{\kern0pt}bs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ rs{\isachardoublequoteclose}\isanewline
+\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}AALTs\ {\isacharparenleft}{\kern0pt}bs{\isacharat}{\kern0pt}bs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ rs\ {\isasymleadsto}\ AALTs\ bs\ {\isacharparenleft}{\kern0pt}map\ {\isacharparenleft}{\kern0pt}fuse\ bs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}AALTs\ bs\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isasymleadsto}\ AZERO{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}AALTs\ bs\ {\isacharbrackleft}{\kern0pt}r{\isacharbrackright}{\kern0pt}\ {\isasymleadsto}\ fuse\ bs\ r{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}erase\ a{\isadigit{1}}\ {\isacharequal}{\kern0pt}\ erase\ a{\isadigit{2}}\ {\isasymLongrightarrow}\ AALTs\ bs\ {\isacharparenleft}{\kern0pt}rsa{\isacharat}{\kern0pt}{\isacharbrackleft}{\kern0pt}a{\isadigit{1}}{\isacharbrackright}{\kern0pt}{\isacharat}{\kern0pt}rsb{\isacharat}{\kern0pt}{\isacharbrackleft}{\kern0pt}a{\isadigit{2}}{\isacharbrackright}{\kern0pt}{\isacharat}{\kern0pt}rsc{\isacharparenright}{\kern0pt}\ {\isasymleadsto}\ AALTs\ bs\ {\isacharparenleft}{\kern0pt}rsa{\isacharat}{\kern0pt}{\isacharbrackleft}{\kern0pt}a{\isadigit{1}}{\isacharbrackright}{\kern0pt}{\isacharat}{\kern0pt}rsb{\isacharat}{\kern0pt}rsc{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isanewline
+\isacommand{inductive}\isamarkupfalse%
+\ rrewrites{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}arexp\ {\isasymRightarrow}\ arexp\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ {\isacharparenleft}{\kern0pt}{\isachardoublequoteopen}{\isacharunderscore}{\kern0pt}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ {\isacharunderscore}{\kern0pt}{\isachardoublequoteclose}\ {\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isacharcomma}{\kern0pt}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isacharbrackright}{\kern0pt}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isakeyword{where}\ \isanewline
+rs{\isadigit{1}}{\isacharbrackleft}{\kern0pt}intro{\isacharcomma}{\kern0pt}\ simp{\isacharbrackright}{\kern0pt}{\isacharcolon}{\kern0pt}{\isachardoublequoteopen}r\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ r{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ rs{\isadigit{2}}{\isacharbrackleft}{\kern0pt}intro{\isacharbrackright}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymlbrakk}r{\isadigit{1}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ r{\isadigit{2}}{\isacharsemicolon}{\kern0pt}\ r{\isadigit{2}}\ {\isasymleadsto}\ r{\isadigit{3}}{\isasymrbrakk}\ {\isasymLongrightarrow}\ r{\isadigit{1}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ r{\isadigit{3}}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isacommand{inductive}\isamarkupfalse%
+\ srewrites{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}arexp\ list\ {\isasymRightarrow}\ arexp\ list\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ {\isacharparenleft}{\kern0pt}{\isachardoublequoteopen}\ {\isacharunderscore}{\kern0pt}\ s{\isasymleadsto}{\isacharasterisk}{\kern0pt}\ {\isacharunderscore}{\kern0pt}{\isachardoublequoteclose}\ {\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isacharcomma}{\kern0pt}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isacharbrackright}{\kern0pt}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isakeyword{where}\isanewline
+ss{\isadigit{1}}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ s{\isasymleadsto}{\isacharasterisk}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}ss{\isadigit{2}}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymlbrakk}r\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ r{\isacharprime}{\kern0pt}{\isacharsemicolon}{\kern0pt}\ rs\ s{\isasymleadsto}{\isacharasterisk}{\kern0pt}\ rs{\isacharprime}{\kern0pt}{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\kern0pt}r{\isacharhash}{\kern0pt}rs{\isacharparenright}{\kern0pt}\ s{\isasymleadsto}{\isacharasterisk}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r{\isacharprime}{\kern0pt}{\isacharhash}{\kern0pt}rs{\isacharprime}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ {\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}r{\isadigit{1}}\ {\isasymleadsto}\ r{\isadigit{2}}\ {\isasymLongrightarrow}\ r{\isadigit{1}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ r{\isadigit{2}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ rrewrites{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}\ rrewrites{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}\ \isacommand{by}\isamarkupfalse%
+\ blast%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\ \isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ real{\isacharunderscore}{\kern0pt}trans{\isacharcolon}{\kern0pt}\ \isanewline
+\ \ \isakeyword{assumes}\ a{\isadigit{1}}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}r{\isadigit{1}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ r{\isadigit{2}}{\isachardoublequoteclose}\ \ \isakeyword{and}\ a{\isadigit{2}}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}r{\isadigit{2}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ r{\isadigit{3}}{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}r{\isadigit{1}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ r{\isadigit{3}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ a{\isadigit{2}}\ a{\isadigit{1}}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ r{\isadigit{2}}\ r{\isadigit{3}}\ arbitrary{\isacharcolon}{\kern0pt}\ r{\isadigit{1}}\ rule{\isacharcolon}{\kern0pt}\ rrewrites{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\ \isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}auto{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ \ many{\isacharunderscore}{\kern0pt}steps{\isacharunderscore}{\kern0pt}later{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymlbrakk}r{\isadigit{1}}\ {\isasymleadsto}\ r{\isadigit{2}}{\isacharsemicolon}{\kern0pt}\ r{\isadigit{2}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ r{\isadigit{3}}\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ r{\isadigit{1}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ r{\isadigit{3}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}meson\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ real{\isacharunderscore}{\kern0pt}trans{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ contextrewrites{\isadigit{1}}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}r\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ r{\isacharprime}{\kern0pt}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}r{\isacharhash}{\kern0pt}rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}r{\isacharprime}{\kern0pt}{\isacharhash}{\kern0pt}rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ r\ r{\isacharprime}{\kern0pt}\ rule{\isacharcolon}{\kern0pt}\ rrewrites{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ append{\isacharunderscore}{\kern0pt}Cons\ append{\isacharunderscore}{\kern0pt}Nil\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{6}}{\isacharparenright}{\kern0pt}\ rs{\isadigit{2}}{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ contextrewrites{\isadigit{2}}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}r\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ r{\isacharprime}{\kern0pt}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs{\isadigit{1}}{\isacharat}{\kern0pt}{\isacharbrackleft}{\kern0pt}r{\isacharbrackright}{\kern0pt}{\isacharat}{\kern0pt}rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs{\isadigit{1}}{\isacharat}{\kern0pt}{\isacharbrackleft}{\kern0pt}r{\isacharprime}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isacharat}{\kern0pt}rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ r\ r{\isacharprime}{\kern0pt}\ rule{\isacharcolon}{\kern0pt}\ rrewrites{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{6}}{\isacharparenright}{\kern0pt}\ \isacommand{by}\isamarkupfalse%
+\ blast%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ srewrites{\isacharunderscore}{\kern0pt}alt{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}rs{\isadigit{1}}\ s{\isasymleadsto}{\isacharasterisk}{\kern0pt}\ rs{\isadigit{2}}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs{\isacharat}{\kern0pt}rs{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs{\isacharat}{\kern0pt}rs{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\isanewline
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ rs{\isadigit{1}}\ rs{\isadigit{2}}\ arbitrary{\isacharcolon}{\kern0pt}\ bs\ rs\ rule{\isacharcolon}{\kern0pt}\ srewrites{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}drule{\isacharunderscore}{\kern0pt}tac\ x\ {\isacharequal}{\kern0pt}\ {\isachardoublequoteopen}bs{\isachardoublequoteclose}\ \isakeyword{in}\ meta{\isacharunderscore}{\kern0pt}spec{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}drule{\isacharunderscore}{\kern0pt}tac\ x\ {\isacharequal}{\kern0pt}\ {\isachardoublequoteopen}rsa{\isacharat}{\kern0pt}{\isacharbrackleft}{\kern0pt}r{\isacharprime}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isachardoublequoteclose}\ \isakeyword{in}\ meta{\isacharunderscore}{\kern0pt}spec{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ real{\isacharunderscore}{\kern0pt}trans{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{prefer}\isamarkupfalse%
+\ {\isadigit{2}}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}assumption{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}drule\ contextrewrites{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ auto\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{corollary}\isamarkupfalse%
+\ srewrites{\isacharunderscore}{\kern0pt}alt{\isadigit{1}}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}rs{\isadigit{1}}\ s{\isasymleadsto}{\isacharasterisk}{\kern0pt}\ rs{\isadigit{2}}\ {\isasymLongrightarrow}\ AALTs\ bs\ rs{\isadigit{1}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ AALTs\ bs\ rs{\isadigit{2}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ append{\isachardot}{\kern0pt}left{\isacharunderscore}{\kern0pt}neutral\ srewrites{\isacharunderscore}{\kern0pt}alt{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ star{\isacharunderscore}{\kern0pt}seq{\isacharcolon}{\kern0pt}\ \ {\isachardoublequoteopen}r{\isadigit{1}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ r{\isadigit{2}}\ {\isasymLongrightarrow}\ ASEQ\ bs\ r{\isadigit{1}}\ r{\isadigit{3}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ ASEQ\ bs\ r{\isadigit{2}}\ r{\isadigit{3}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ r{\isadigit{1}}\ r{\isadigit{2}}\ arbitrary{\isacharcolon}{\kern0pt}\ r{\isadigit{3}}\ rule{\isacharcolon}{\kern0pt}\ rrewrites{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}erule\ rrewrites{\isachardot}{\kern0pt}cases{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ rs{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}assumption{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ assumption%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ star{\isacharunderscore}{\kern0pt}seq{\isadigit{2}}{\isacharcolon}{\kern0pt}\ \ {\isachardoublequoteopen}r{\isadigit{3}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ r{\isadigit{4}}\ {\isasymLongrightarrow}\ ASEQ\ bs\ r{\isadigit{1}}\ r{\isadigit{3}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ ASEQ\ bs\ r{\isadigit{1}}\ r{\isadigit{4}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ r{\isadigit{3}}\ r{\isadigit{4}}\ arbitrary{\isacharcolon}{\kern0pt}\ r{\isadigit{1}}\ rule{\isacharcolon}{\kern0pt}\ rrewrites{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ auto\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{5}}{\isacharparenright}{\kern0pt}\ \isacommand{by}\isamarkupfalse%
+\ blast%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ continuous{\isacharunderscore}{\kern0pt}rewrite{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymlbrakk}r{\isadigit{1}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ AZERO{\isasymrbrakk}\ {\isasymLongrightarrow}\ ASEQ\ bs{\isadigit{1}}\ r{\isadigit{1}}\ r{\isadigit{2}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ AZERO{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induction\ ra{\isasymequiv}{\isachardoublequoteopen}r{\isadigit{1}}{\isachardoublequoteclose}\ rb{\isasymequiv}{\isachardoublequoteopen}AZERO{\isachardoublequoteclose}\ arbitrary{\isacharcolon}{\kern0pt}\ bs{\isadigit{1}}\ r{\isadigit{1}}\ r{\isadigit{2}}\ rule{\isacharcolon}{\kern0pt}\ rrewrites{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}meson\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}\ rrewrites{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}\ star{\isacharunderscore}{\kern0pt}seq{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\ \ \isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ bsimp{\isacharunderscore}{\kern0pt}aalts{\isacharunderscore}{\kern0pt}simpcases{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}AONE\ bs\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ {\isacharparenleft}{\kern0pt}bsimp\ {\isacharparenleft}{\kern0pt}AONE\ bs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\ \ {\isachardoublequoteopen}AZERO\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ bsimp\ AZERO{\isachardoublequoteclose}\ {\isachardoublequoteopen}ACHAR\ bs\ c\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ {\isacharparenleft}{\kern0pt}bsimp\ {\isacharparenleft}{\kern0pt}ACHAR\ bs\ c{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ rrewrites{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ rrewrites{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ rrewrites{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ trivialbsimpsrewrites{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymlbrakk}{\isasymAnd}x{\isachardot}{\kern0pt}\ x\ {\isasymin}\ set\ rs\ {\isasymLongrightarrow}\ x\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ f\ x\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ rs\ s{\isasymleadsto}{\isacharasterisk}{\kern0pt}\ {\isacharparenleft}{\kern0pt}map\ f\ rs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\isanewline
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induction\ rs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ ss{\isadigit{1}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ insert{\isacharunderscore}{\kern0pt}iff\ list{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isadigit{5}}{\isacharparenright}{\kern0pt}\ list{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{9}}{\isacharparenright}{\kern0pt}\ srewrites{\isachardot}{\kern0pt}simps{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ bsimp{\isacharunderscore}{\kern0pt}AALTsrewrites{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}AALTs\ bs{\isadigit{1}}\ rs\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ bsimp{\isacharunderscore}{\kern0pt}AALTs\ bs{\isadigit{1}}\ rs{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induction\ rs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ \ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}rs\ {\isacharequal}{\kern0pt}\ Nil{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isadigit{2}}{\isacharparenright}{\kern0pt}\ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}length\ {\isacharparenleft}{\kern0pt}a{\isacharhash}{\kern0pt}rs{\isacharparenright}{\kern0pt}\ {\isachargreater}{\kern0pt}\ {\isadigit{1}}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ bsimp{\isacharunderscore}{\kern0pt}AALTs{\isacharunderscore}{\kern0pt}qq{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\ \isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{inductive}\isamarkupfalse%
+\ frewrites{\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}arexp\ list\ {\isasymRightarrow}\ arexp\ list\ {\isasymRightarrow}\ bool{\isachardoublequoteclose}\ {\isacharparenleft}{\kern0pt}{\isachardoublequoteopen}\ {\isacharunderscore}{\kern0pt}\ f{\isasymleadsto}{\isacharasterisk}{\kern0pt}\ {\isacharunderscore}{\kern0pt}{\isachardoublequoteclose}\ {\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isacharcomma}{\kern0pt}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isacharbrackright}{\kern0pt}\ {\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isakeyword{where}\isanewline
+fs{\isadigit{1}}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ f{\isasymleadsto}{\isacharasterisk}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}fs{\isadigit{2}}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymlbrakk}rs\ f{\isasymleadsto}{\isacharasterisk}{\kern0pt}\ rs{\isacharprime}{\kern0pt}{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\kern0pt}AZERO{\isacharhash}{\kern0pt}rs{\isacharparenright}{\kern0pt}\ f{\isasymleadsto}{\isacharasterisk}{\kern0pt}\ rs{\isacharprime}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}fs{\isadigit{3}}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymlbrakk}rs\ f{\isasymleadsto}{\isacharasterisk}{\kern0pt}\ rs{\isacharprime}{\kern0pt}{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\kern0pt}{\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharhash}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}\ f{\isasymleadsto}{\isacharasterisk}{\kern0pt}\ {\isacharparenleft}{\kern0pt}{\isacharparenleft}{\kern0pt}map\ {\isacharparenleft}{\kern0pt}fuse\ bs{\isacharparenright}{\kern0pt}\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isacharat}{\kern0pt}\ rs{\isacharprime}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}fs{\isadigit{4}}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymlbrakk}rs\ f{\isasymleadsto}{\isacharasterisk}{\kern0pt}\ rs{\isacharprime}{\kern0pt}{\isacharsemicolon}{\kern0pt}nonalt\ r{\isacharsemicolon}{\kern0pt}\ nonazero\ r{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\kern0pt}r{\isacharhash}{\kern0pt}rs{\isacharparenright}{\kern0pt}\ f{\isasymleadsto}{\isacharasterisk}{\kern0pt}\ {\isacharparenleft}{\kern0pt}r{\isacharhash}{\kern0pt}rs{\isacharprime}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ flts{\isacharunderscore}{\kern0pt}prepend{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymlbrakk}nonalt\ a{\isacharsemicolon}{\kern0pt}\ nonazero\ a{\isasymrbrakk}\ {\isasymLongrightarrow}\ flts\ {\isacharparenleft}{\kern0pt}a{\isacharhash}{\kern0pt}rs{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ a\ {\isacharhash}{\kern0pt}\ {\isacharparenleft}{\kern0pt}flts\ rs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ append{\isacharunderscore}{\kern0pt}Cons\ append{\isacharunderscore}{\kern0pt}Nil\ flts{\isacharunderscore}{\kern0pt}single{\isadigit{1}}\ k{\isadigit{0}}{\isadigit{0}}{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ fltsfrewrites{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}rs\ f{\isasymleadsto}{\isacharasterisk}{\kern0pt}\ {\isacharparenleft}{\kern0pt}flts\ rs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induction\ rs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ fs{\isadigit{1}}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}a\ {\isacharequal}{\kern0pt}\ AZERO{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \ \isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ fs{\isadigit{2}}\ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}{\isasymexists}bs\ rs{\isachardot}{\kern0pt}\ a\ {\isacharequal}{\kern0pt}\ AALTs\ bs\ rs{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}erule\ exE{\isacharparenright}{\kern0pt}{\isacharplus}{\kern0pt}\isanewline
+\ \ \ \isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ fs{\isadigit{3}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ flts{\isacharunderscore}{\kern0pt}prepend{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ nonalt{\isachardot}{\kern0pt}elims{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{prefer}\isamarkupfalse%
+\ {\isadigit{2}}\isanewline
+\ \ \isacommand{thm}\isamarkupfalse%
+\ nonalt{\isachardot}{\kern0pt}elims\isanewline
+\ \ \ \isanewline
+\ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\ \ \ \isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ bbbbs{\isadigit{1}}\ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ nonalt{\isachardot}{\kern0pt}simps{\isacharparenright}{\kern0pt}{\isacharplus}{\kern0pt}\isanewline
+\ \ \ \isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}meson\ nonazero{\isachardot}{\kern0pt}elims{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}meson\ fs{\isadigit{4}}\ nonalt{\isachardot}{\kern0pt}elims{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}\ nonazero{\isachardot}{\kern0pt}elims{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ rrewrite{\isadigit{0}}away{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}AALTs\ bs\ {\isacharparenleft}{\kern0pt}\ AZERO\ {\isacharhash}{\kern0pt}\ rsb{\isacharparenright}{\kern0pt}\ {\isasymleadsto}\ AALTs\ bs\ rsb{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ append{\isacharunderscore}{\kern0pt}Nil\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{7}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ frewritesaalts{\isacharcolon}{\kern0pt}{\isachardoublequoteopen}rs\ f{\isasymleadsto}{\isacharasterisk}{\kern0pt}\ rs{\isacharprime}{\kern0pt}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs{\isadigit{1}}{\isacharat}{\kern0pt}rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs{\isadigit{1}}{\isacharat}{\kern0pt}rs{\isacharprime}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ rs\ rs{\isacharprime}{\kern0pt}\ arbitrary{\isacharcolon}{\kern0pt}\ bs\ rs{\isadigit{1}}\ rule{\isacharcolon}{\kern0pt}frewrites{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}drule{\isacharunderscore}{\kern0pt}tac\ x\ {\isacharequal}{\kern0pt}\ {\isachardoublequoteopen}bs{\isachardoublequoteclose}\ \isakeyword{in}\ meta{\isacharunderscore}{\kern0pt}spec{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}drule{\isacharunderscore}{\kern0pt}tac\ x\ {\isacharequal}{\kern0pt}\ {\isachardoublequoteopen}rs{\isadigit{1}}\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}AZERO{\isacharbrackright}{\kern0pt}{\isachardoublequoteclose}\ \isakeyword{in}\ meta{\isacharunderscore}{\kern0pt}spec{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ real{\isacharunderscore}{\kern0pt}trans{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{7}}{\isacharparenright}{\kern0pt}\ \isacommand{apply}\isamarkupfalse%
+\ presburger\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}drule{\isacharunderscore}{\kern0pt}tac\ x\ {\isacharequal}{\kern0pt}\ {\isachardoublequoteopen}bsa{\isachardoublequoteclose}\ \isakeyword{in}\ meta{\isacharunderscore}{\kern0pt}spec{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}drule{\isacharunderscore}{\kern0pt}tac\ x\ {\isacharequal}{\kern0pt}\ {\isachardoublequoteopen}rs{\isadigit{1}}a\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}AALTs\ bs\ rs{\isadigit{1}}{\isacharbrackright}{\kern0pt}{\isachardoublequoteclose}\ \isakeyword{in}\ meta{\isacharunderscore}{\kern0pt}spec{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ real{\isacharunderscore}{\kern0pt}trans{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{8}}{\isacharparenright}{\kern0pt}\ \isacommand{apply}\isamarkupfalse%
+\ presburger\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}drule{\isacharunderscore}{\kern0pt}tac\ x\ {\isacharequal}{\kern0pt}\ {\isachardoublequoteopen}bs{\isachardoublequoteclose}\ \isakeyword{in}\ meta{\isacharunderscore}{\kern0pt}spec{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}drule{\isacharunderscore}{\kern0pt}tac\ x\ {\isacharequal}{\kern0pt}\ {\isachardoublequoteopen}rs{\isadigit{1}}{\isacharat}{\kern0pt}{\isacharbrackleft}{\kern0pt}r{\isacharbrackright}{\kern0pt}{\isachardoublequoteclose}\ \isakeyword{in}\ meta{\isacharunderscore}{\kern0pt}spec{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ real{\isacharunderscore}{\kern0pt}trans{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ auto\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ fltsrewrites{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}\ \ AALTs\ bs{\isadigit{1}}\ rs\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ AALTs\ bs{\isadigit{1}}\ {\isacharparenleft}{\kern0pt}flts\ rs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induction\ rs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}a\ {\isacharequal}{\kern0pt}\ AZERO{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ append{\isacharunderscore}{\kern0pt}Nil\ flts{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}\ many{\isacharunderscore}{\kern0pt}steps{\isacharunderscore}{\kern0pt}later\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{7}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\isanewline
+\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}{\isasymexists}bs{\isadigit{2}}\ rs{\isadigit{2}}{\isachardot}{\kern0pt}\ a\ {\isacharequal}{\kern0pt}\ AALTs\ bs{\isadigit{2}}\ rs{\isadigit{2}}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}erule\ exE{\isacharparenright}{\kern0pt}{\isacharplus}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ flts{\isachardot}{\kern0pt}simps{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{prefer}\isamarkupfalse%
+\ {\isadigit{2}}\isanewline
+\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ flts{\isacharunderscore}{\kern0pt}prepend{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}meson\ nonalt{\isachardot}{\kern0pt}elims{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}meson\ nonazero{\isachardot}{\kern0pt}elims{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}{\isacharparenleft}{\kern0pt}a{\isacharhash}{\kern0pt}rs{\isacharparenright}{\kern0pt}\ f{\isasymleadsto}{\isacharasterisk}{\kern0pt}\ {\isacharparenleft}{\kern0pt}a{\isacharhash}{\kern0pt}flts\ rs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ append{\isacharunderscore}{\kern0pt}Nil\ frewritesaalts{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}meson\ fltsfrewrites\ fs{\isadigit{4}}\ nonalt{\isachardot}{\kern0pt}elims{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}\ nonazero{\isachardot}{\kern0pt}elims{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ append{\isacharunderscore}{\kern0pt}Cons\ append{\isacharunderscore}{\kern0pt}Nil\ fltsfrewrites\ frewritesaalts\ k{\isadigit{0}}{\isadigit{0}}\ k{\isadigit{0}}a{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ alts{\isacharunderscore}{\kern0pt}simpalts{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymAnd}bs{\isadigit{1}}\ rs{\isachardot}{\kern0pt}\ {\isacharparenleft}{\kern0pt}{\isasymAnd}x{\isachardot}{\kern0pt}\ x\ {\isasymin}\ set\ rs\ {\isasymLongrightarrow}\ x\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ bsimp\ x{\isacharparenright}{\kern0pt}\ {\isasymLongrightarrow}\ \isanewline
+AALTs\ bs{\isadigit{1}}\ rs\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ AALTs\ bs{\isadigit{1}}\ {\isacharparenleft}{\kern0pt}map\ bsimp\ rs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}\ rs\ s{\isasymleadsto}{\isacharasterisk}{\kern0pt}\ \ {\isacharparenleft}{\kern0pt}map\ bsimp\ rs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{prefer}\isamarkupfalse%
+\ {\isadigit{2}}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ trivialbsimpsrewrites\ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ srewrites{\isacharunderscore}{\kern0pt}alt{\isadigit{1}}\ \isacommand{by}\isamarkupfalse%
+\ auto%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ threelistsappend{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}rsa{\isacharat}{\kern0pt}a{\isacharhash}{\kern0pt}rsb\ {\isacharequal}{\kern0pt}\ {\isacharparenleft}{\kern0pt}rsa{\isacharat}{\kern0pt}{\isacharbrackleft}{\kern0pt}a{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharat}{\kern0pt}rsb{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+\ auto\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{fun}\isamarkupfalse%
+\ distinctByAcc\ {\isacharcolon}{\kern0pt}{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isacharprime}{\kern0pt}a\ list\ {\isasymRightarrow}\ {\isacharparenleft}{\kern0pt}{\isacharprime}{\kern0pt}a\ {\isasymRightarrow}\ {\isacharprime}{\kern0pt}b{\isacharparenright}{\kern0pt}\ {\isasymRightarrow}\ {\isacharprime}{\kern0pt}b\ set\ {\isasymRightarrow}\ {\isacharprime}{\kern0pt}b\ set{\isachardoublequoteclose}\isanewline
+\ \ \isakeyword{where}\isanewline
+\ \ {\isachardoublequoteopen}distinctByAcc\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ f\ acc\ {\isacharequal}{\kern0pt}\ acc{\isachardoublequoteclose}\isanewline
+{\isacharbar}{\kern0pt}\ {\isachardoublequoteopen}distinctByAcc\ {\isacharparenleft}{\kern0pt}x{\isacharhash}{\kern0pt}xs{\isacharparenright}{\kern0pt}\ f\ acc\ {\isacharequal}{\kern0pt}\ \isanewline
+\ \ \ \ \ {\isacharparenleft}{\kern0pt}if\ {\isacharparenleft}{\kern0pt}f\ x{\isacharparenright}{\kern0pt}\ {\isasymin}\ acc\ then\ distinctByAcc\ xs\ f\ acc\ \isanewline
+\ \ \ \ \ \ else\ \ {\isacharparenleft}{\kern0pt}distinctByAcc\ xs\ f\ {\isacharparenleft}{\kern0pt}{\isacharbraceleft}{\kern0pt}f\ x{\isacharbraceright}{\kern0pt}\ {\isasymunion}\ acc{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ dB{\isacharunderscore}{\kern0pt}single{\isacharunderscore}{\kern0pt}step{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}distinctBy\ {\isacharparenleft}{\kern0pt}a{\isacharhash}{\kern0pt}rs{\isacharparenright}{\kern0pt}\ f\ {\isacharbraceleft}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ a\ {\isacharhash}{\kern0pt}\ distinctBy\ rs\ f\ {\isacharbraceleft}{\kern0pt}f\ a{\isacharbraceright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ somewhereInside{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}r\ {\isasymin}\ set\ rs\ {\isasymLongrightarrow}\ {\isasymexists}rs{\isadigit{1}}\ rs{\isadigit{2}}{\isachardot}{\kern0pt}\ rs\ {\isacharequal}{\kern0pt}\ rs{\isadigit{1}}{\isacharat}{\kern0pt}{\isacharbrackleft}{\kern0pt}r{\isacharbrackright}{\kern0pt}{\isacharat}{\kern0pt}rs{\isadigit{2}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ split{\isacharunderscore}{\kern0pt}list\ \isacommand{by}\isamarkupfalse%
+\ fastforce%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ somewhereMapInside{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}f\ r\ {\isasymin}\ f\ {\isacharbackquote}{\kern0pt}\ set\ rs\ {\isasymLongrightarrow}\ {\isasymexists}rs{\isadigit{1}}\ rs{\isadigit{2}}\ a{\isachardot}{\kern0pt}\ rs\ {\isacharequal}{\kern0pt}\ rs{\isadigit{1}}{\isacharat}{\kern0pt}{\isacharbrackleft}{\kern0pt}a{\isacharbrackright}{\kern0pt}{\isacharat}{\kern0pt}rs{\isadigit{2}}\ {\isasymand}\ f\ a\ {\isacharequal}{\kern0pt}\ f\ r{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+\ auto\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ split{\isacharunderscore}{\kern0pt}list{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ alts{\isacharunderscore}{\kern0pt}dBrewrites{\isacharunderscore}{\kern0pt}withFront{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}\ AALTs\ bs\ {\isacharparenleft}{\kern0pt}rsa\ {\isacharat}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ AALTs\ bs\ {\isacharparenleft}{\kern0pt}rsa\ {\isacharat}{\kern0pt}\ distinctBy\ rs\ erase\ {\isacharparenleft}{\kern0pt}erase\ {\isacharbackquote}{\kern0pt}\ set\ rsa{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induction\ rs\ arbitrary{\isacharcolon}{\kern0pt}\ rsa{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}drule{\isacharunderscore}{\kern0pt}tac\ x\ {\isacharequal}{\kern0pt}\ {\isachardoublequoteopen}rsa{\isacharat}{\kern0pt}{\isacharbrackleft}{\kern0pt}a{\isacharbrackright}{\kern0pt}{\isachardoublequoteclose}\ \isakeyword{in}\ meta{\isacharunderscore}{\kern0pt}spec{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ threelistsappend{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ real{\isacharunderscore}{\kern0pt}trans{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}a\ {\isasymin}\ set\ rsa{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}drule\ somewhereInside{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}erule\ exE{\isacharparenright}{\kern0pt}{\isacharplus}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}\ AALTs\ bs\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}{\kern0pt}rs{\isadigit{1}}\ {\isacharat}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ a\ {\isacharhash}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ rs{\isadigit{2}}\ {\isacharat}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ a\ {\isacharhash}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ distinctBy\ rs\ erase\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}{\kern0pt}insert\ {\isacharparenleft}{\kern0pt}erase\ a{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}{\kern0pt}erase\ {\isacharbackquote}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}{\kern0pt}set\ rs{\isadigit{1}}\ {\isasymunion}\ set\ rs{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isasymleadsto}\ AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs{\isadigit{1}}{\isacharat}{\kern0pt}\ a\ {\isacharhash}{\kern0pt}\ rs{\isadigit{2}}\ {\isacharat}{\kern0pt}\ \ distinctBy\ rs\ erase\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}{\kern0pt}insert\ {\isacharparenleft}{\kern0pt}erase\ a{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}{\kern0pt}erase\ {\isacharbackquote}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}{\kern0pt}set\ rs{\isadigit{1}}\ {\isasymunion}\ set\ rs{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{prefer}\isamarkupfalse%
+\ {\isadigit{2}}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isadigit{3}}{\isacharparenright}{\kern0pt}\ \isacommand{apply}\isamarkupfalse%
+\ force\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ \isacommand{apply}\isamarkupfalse%
+\ force\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}erase\ {\isacharbackquote}{\kern0pt}\ set\ {\isacharparenleft}{\kern0pt}rsa\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}a{\isacharbrackright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ insert\ {\isacharparenleft}{\kern0pt}erase\ a{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}erase\ {\isacharbackquote}{\kern0pt}\ set\ rsa{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{prefer}\isamarkupfalse%
+\ {\isadigit{2}}\isanewline
+\ \ \ \ \isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}erase\ a\ {\isasymin}\ erase\ {\isacharbackquote}{\kern0pt}set\ rsa{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rsa\ {\isacharat}{\kern0pt}\ a\ {\isacharhash}{\kern0pt}\ distinctBy\ rs\ erase\ {\isacharparenleft}{\kern0pt}insert\ {\isacharparenleft}{\kern0pt}erase\ a{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}erase\ {\isacharbackquote}{\kern0pt}\ set\ rsa{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isasymleadsto}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ AALTs\ bs\ {\isacharparenleft}{\kern0pt}rsa\ {\isacharat}{\kern0pt}\ distinctBy\ rs\ erase\ {\isacharparenleft}{\kern0pt}insert\ {\isacharparenleft}{\kern0pt}erase\ a{\isacharparenright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}erase\ {\isacharbackquote}{\kern0pt}\ set\ rsa{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ force\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}smt\ {\isacharparenleft}{\kern0pt}verit{\isacharcomma}{\kern0pt}\ ccfv{\isacharunderscore}{\kern0pt}threshold{\isacharparenright}{\kern0pt}\ append{\isacharunderscore}{\kern0pt}Cons\ append{\isacharunderscore}{\kern0pt}assoc\ append{\isacharunderscore}{\kern0pt}self{\isacharunderscore}{\kern0pt}conv{\isadigit{2}}\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isadigit{3}}{\isacharparenright}{\kern0pt}\ same{\isacharunderscore}{\kern0pt}append{\isacharunderscore}{\kern0pt}eq\ somewhereMapInside{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ force%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\ \isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ alts{\isacharunderscore}{\kern0pt}dBrewrites{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}AALTs\ bs\ rs\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ AALTs\ bs\ {\isacharparenleft}{\kern0pt}distinctBy\ rs\ erase\ {\isacharbraceleft}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induction\ rs{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ alts{\isacharunderscore}{\kern0pt}dBrewrites{\isacharunderscore}{\kern0pt}withFront\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ append{\isacharunderscore}{\kern0pt}Nil\ dB{\isacharunderscore}{\kern0pt}single{\isacharunderscore}{\kern0pt}step\ empty{\isacharunderscore}{\kern0pt}set\ image{\isacharunderscore}{\kern0pt}empty{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isanewline
+\ \ \isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ bsimp{\isacharunderscore}{\kern0pt}rewrite{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}\ {\isacharparenleft}{\kern0pt}rrewrites\ r\ {\isacharparenleft}{\kern0pt}\ bsimp\ r{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induction\ r\ rule{\isacharcolon}{\kern0pt}\ bsimp{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}bsimp\ r{\isadigit{1}}\ {\isacharequal}{\kern0pt}\ AZERO{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ continuous{\isacharunderscore}{\kern0pt}rewrite\ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}{\isasymexists}bs{\isachardot}{\kern0pt}\ bsimp\ r{\isadigit{1}}\ {\isacharequal}{\kern0pt}\ AONE\ bs{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}erule\ exE{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ bsimp{\isacharunderscore}{\kern0pt}ASEQ{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}meson\ real{\isacharunderscore}{\kern0pt}trans\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}\ rrewrites{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}\ star{\isacharunderscore}{\kern0pt}seq\ star{\isacharunderscore}{\kern0pt}seq{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}smt\ {\isacharparenleft}{\kern0pt}verit{\isacharcomma}{\kern0pt}\ best{\isacharparenright}{\kern0pt}\ bsimp{\isacharunderscore}{\kern0pt}ASEQ{\isadigit{0}}\ bsimp{\isacharunderscore}{\kern0pt}ASEQ{\isadigit{1}}\ real{\isacharunderscore}{\kern0pt}trans\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}\ rs{\isadigit{2}}\ star{\isacharunderscore}{\kern0pt}seq\ star{\isacharunderscore}{\kern0pt}seq{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \isacommand{defer}\isamarkupfalse%
+\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ bsimp{\isacharunderscore}{\kern0pt}aalts{\isacharunderscore}{\kern0pt}simpcases{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}\ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ auto\isanewline
+\isanewline
+\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}AALTs\ bs{\isadigit{1}}\ rs\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ AALTs\ bs{\isadigit{1}}\ {\isacharparenleft}{\kern0pt}map\ bsimp\ rs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}AALTs\ bs{\isadigit{1}}\ {\isacharparenleft}{\kern0pt}map\ bsimp\ rs{\isacharparenright}{\kern0pt}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ AALTs\ bs{\isadigit{1}}\ {\isacharparenleft}{\kern0pt}flts\ {\isacharparenleft}{\kern0pt}map\ bsimp\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}AALTs\ bs{\isadigit{1}}\ {\isacharparenleft}{\kern0pt}flts\ {\isacharparenleft}{\kern0pt}map\ bsimp\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ AALTs\ bs{\isadigit{1}}\ {\isacharparenleft}{\kern0pt}distinctBy\ {\isacharparenleft}{\kern0pt}flts\ {\isacharparenleft}{\kern0pt}map\ bsimp\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ erase\ {\isacharbraceleft}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}AALTs\ bs{\isadigit{1}}\ {\isacharparenleft}{\kern0pt}distinctBy\ {\isacharparenleft}{\kern0pt}flts\ {\isacharparenleft}{\kern0pt}map\ bsimp\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ erase\ {\isacharbraceleft}{\kern0pt}{\isacharbraceright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ bsimp{\isacharunderscore}{\kern0pt}AALTs\ bs{\isadigit{1}}\ {\isacharparenleft}{\kern0pt}distinctBy\ {\isacharparenleft}{\kern0pt}flts\ {\isacharparenleft}{\kern0pt}map\ bsimp\ rs{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ erase\ {\isacharbraceleft}{\kern0pt}{\isacharbraceright}{\kern0pt}\ {\isacharparenright}{\kern0pt}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isanewline
+\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}meson\ real{\isacharunderscore}{\kern0pt}trans{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}meson\ bsimp{\isacharunderscore}{\kern0pt}AALTsrewrites{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}meson\ alts{\isacharunderscore}{\kern0pt}dBrewrites{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ fltsrewrites\ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ alts{\isacharunderscore}{\kern0pt}simpalts\ \isacommand{by}\isamarkupfalse%
+\ force%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ rewritenullable{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymlbrakk}r{\isadigit{1}}\ {\isasymleadsto}\ r{\isadigit{2}}{\isacharsemicolon}{\kern0pt}\ bnullable\ r{\isadigit{1}}\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ bnullable\ r{\isadigit{2}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induction\ r{\isadigit{1}}\ r{\isadigit{2}}\ rule{\isacharcolon}{\kern0pt}\ rrewrite{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}{\isacharplus}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isacharunderscore}{\kern0pt}fuse{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{4}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ UnCI\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isacharunderscore}{\kern0pt}fuse\ imageI{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isacharunderscore}{\kern0pt}fuse{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isacharunderscore}{\kern0pt}fuse{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{5}}{\isacharparenright}{\kern0pt}\ erase{\isacharunderscore}{\kern0pt}fuse{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isanewline
+\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}smt\ {\isacharparenleft}{\kern0pt}z{\isadigit{3}}{\isacharparenright}{\kern0pt}\ Un{\isacharunderscore}{\kern0pt}iff\ bnullable{\isacharunderscore}{\kern0pt}correctness\ insert{\isacharunderscore}{\kern0pt}iff\ list{\isachardot}{\kern0pt}set{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}\ qq{\isadigit{3}}\ set{\isacharunderscore}{\kern0pt}append{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ rewrite{\isacharunderscore}{\kern0pt}non{\isacharunderscore}{\kern0pt}nullable{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymlbrakk}r{\isadigit{1}}\ {\isasymleadsto}\ r{\isadigit{2}}{\isacharsemicolon}{\kern0pt}\ {\isasymnot}bnullable\ r{\isadigit{1}}\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isasymnot}bnullable\ r{\isadigit{2}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induction\ r{\isadigit{1}}\ r{\isadigit{2}}\ rule{\isacharcolon}{\kern0pt}\ rrewrite{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ auto\ \isanewline
+\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isacharunderscore}{\kern0pt}fuse{\isacharparenright}{\kern0pt}{\isacharplus}{\kern0pt}\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ rewritesnullable{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymlbrakk}\ r{\isadigit{1}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ r{\isadigit{2}}{\isacharsemicolon}{\kern0pt}\ bnullable\ r{\isadigit{1}}\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ bnullable\ r{\isadigit{2}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induction\ r{\isadigit{1}}\ r{\isadigit{2}}\ rule{\isacharcolon}{\kern0pt}\ rrewrites{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ rewritenullable{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ nonbnullable{\isacharunderscore}{\kern0pt}lists{\isacharunderscore}{\kern0pt}concat{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}\ {\isasymlbrakk}\ {\isasymnot}\ {\isacharparenleft}{\kern0pt}{\isasymexists}r{\isadigit{0}}{\isasymin}set\ rs{\isadigit{1}}{\isachardot}{\kern0pt}\ bnullable\ r{\isadigit{0}}{\isacharparenright}{\kern0pt}{\isacharsemicolon}{\kern0pt}\ {\isasymnot}\ bnullable\ r{\isacharsemicolon}{\kern0pt}\ {\isasymnot}\ {\isacharparenleft}{\kern0pt}{\isasymexists}r{\isadigit{0}}{\isasymin}set\ rs{\isadigit{2}}{\isachardot}{\kern0pt}\ bnullable\ r{\isadigit{0}}{\isacharparenright}{\kern0pt}{\isasymrbrakk}\ {\isasymLongrightarrow}\ \isanewline
+{\isasymnot}{\isacharparenleft}{\kern0pt}{\isasymexists}r{\isadigit{0}}\ {\isasymin}\ {\isacharparenleft}{\kern0pt}set\ {\isacharparenleft}{\kern0pt}rs{\isadigit{1}}{\isacharat}{\kern0pt}{\isacharbrackleft}{\kern0pt}r{\isacharbrackright}{\kern0pt}{\isacharat}{\kern0pt}rs{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardot}{\kern0pt}\ bnullable\ r{\isadigit{0}}\ {\isacharparenright}{\kern0pt}\ {\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\ \ \isacommand{done}\isamarkupfalse%
+%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ nomember{\isacharunderscore}{\kern0pt}bnullable{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymlbrakk}\ {\isasymnot}\ {\isacharparenleft}{\kern0pt}{\isasymexists}r{\isadigit{0}}{\isasymin}set\ rs{\isadigit{1}}{\isachardot}{\kern0pt}\ bnullable\ r{\isadigit{0}}{\isacharparenright}{\kern0pt}{\isacharsemicolon}{\kern0pt}\ {\isasymnot}\ bnullable\ r{\isacharsemicolon}{\kern0pt}\ {\isasymnot}\ {\isacharparenleft}{\kern0pt}{\isasymexists}r{\isadigit{0}}{\isasymin}set\ rs{\isadigit{2}}{\isachardot}{\kern0pt}\ bnullable\ r{\isadigit{0}}{\isacharparenright}{\kern0pt}{\isasymrbrakk}\isanewline
+\ {\isasymLongrightarrow}\ {\isasymnot}bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs{\isadigit{1}}\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}r{\isacharbrackright}{\kern0pt}\ {\isacharat}{\kern0pt}\ rs{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ nonbnullable{\isacharunderscore}{\kern0pt}lists{\isacharunderscore}{\kern0pt}concat\ qq{\isadigit{3}}\ \isacommand{by}\isamarkupfalse%
+\ presburger%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ bnullable{\isacharunderscore}{\kern0pt}segment{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}\ bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs{\isadigit{1}}{\isacharat}{\kern0pt}{\isacharbrackleft}{\kern0pt}r{\isacharbrackright}{\kern0pt}{\isacharat}{\kern0pt}rs{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isasymLongrightarrow}\ bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isasymor}\ bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isasymor}\ bnullable\ r{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}{\isasymexists}r{\isadigit{0}}{\isasymin}set\ rs{\isadigit{1}}{\isachardot}{\kern0pt}\ \ bnullable\ r{\isadigit{0}}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ qq{\isadigit{3}}\ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}bnullable\ r{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}{\isasymexists}r{\isadigit{0}}{\isasymin}set\ rs{\isadigit{2}}{\isachardot}{\kern0pt}\ \ bnullable\ r{\isadigit{0}}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ bnullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}\ \isacommand{apply}\isamarkupfalse%
+\ presburger\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}False{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ nomember{\isacharunderscore}{\kern0pt}bnullable\ \isacommand{by}\isamarkupfalse%
+\ blast%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\ \ \isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ bnullablewhichbmkeps{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymlbrakk}bnullable\ \ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs{\isadigit{1}}{\isacharat}{\kern0pt}{\isacharbrackleft}{\kern0pt}r{\isacharbrackright}{\kern0pt}{\isacharat}{\kern0pt}rs{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharsemicolon}{\kern0pt}\ {\isasymnot}\ bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharsemicolon}{\kern0pt}\ bnullable\ r\ {\isasymrbrakk}\isanewline
+\ {\isasymLongrightarrow}\ bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs{\isadigit{1}}{\isacharat}{\kern0pt}{\isacharbrackleft}{\kern0pt}r{\isacharbrackright}{\kern0pt}{\isacharat}{\kern0pt}rs{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bs\ {\isacharat}{\kern0pt}\ {\isacharparenleft}{\kern0pt}bmkeps\ r{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ qq{\isadigit{2}}\ bnullable{\isacharunderscore}{\kern0pt}Hdbmkeps{\isacharunderscore}{\kern0pt}Hd\ \isacommand{by}\isamarkupfalse%
+\ force%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ rrewrite{\isacharunderscore}{\kern0pt}nbnullable{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymlbrakk}\ r{\isadigit{1}}\ {\isasymleadsto}\ r{\isadigit{2}}\ {\isacharsemicolon}{\kern0pt}\ {\isasymnot}\ bnullable\ r{\isadigit{1}}\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isasymnot}bnullable\ r{\isadigit{2}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induction\ rule{\isacharcolon}{\kern0pt}\ rrewrite{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isacharunderscore}{\kern0pt}fuse{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isacharunderscore}{\kern0pt}fuse{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isacharunderscore}{\kern0pt}fuse{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isacharunderscore}{\kern0pt}fuse{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isacharunderscore}{\kern0pt}fuse{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}meson\ rewrite{\isacharunderscore}{\kern0pt}non{\isacharunderscore}{\kern0pt}nullable\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ spillbmkepslistr{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs{\isadigit{1}}\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ {\isasymLongrightarrow}\ bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}AALTs\ bs{\isadigit{1}}\ rs{\isadigit{1}}\ {\isacharhash}{\kern0pt}\ rsb{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}\ map\ {\isacharparenleft}{\kern0pt}fuse\ bs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ rs{\isadigit{1}}\ {\isacharat}{\kern0pt}\ rsb{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ bnullable{\isacharunderscore}{\kern0pt}Hdbmkeps{\isacharunderscore}{\kern0pt}Hd{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ bmkeps{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}\ k{\isadigit{0}}a\ list{\isachardot}{\kern0pt}set{\isacharunderscore}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}\ qq{\isadigit{1}}\ qq{\isadigit{4}}\ qs{\isadigit{3}}{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ third{\isacharunderscore}{\kern0pt}segment{\isacharunderscore}{\kern0pt}bnullable{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymlbrakk}bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs{\isadigit{1}}{\isacharat}{\kern0pt}rs{\isadigit{2}}{\isacharat}{\kern0pt}rs{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharsemicolon}{\kern0pt}\ {\isasymnot}bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharsemicolon}{\kern0pt}\ {\isasymnot}bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isasymrbrakk}\ {\isasymLongrightarrow}\ \isanewline
+bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ \isanewline
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ append{\isachardot}{\kern0pt}left{\isacharunderscore}{\kern0pt}neutral\ append{\isacharunderscore}{\kern0pt}Cons\ bnullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}\ bnullable{\isacharunderscore}{\kern0pt}segment\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{7}}{\isacharparenright}{\kern0pt}\ rrewrite{\isacharunderscore}{\kern0pt}nbnullable{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ third{\isacharunderscore}{\kern0pt}segment{\isacharunderscore}{\kern0pt}bmkeps{\isacharcolon}{\kern0pt}\ \ {\isachardoublequoteopen}{\isasymlbrakk}bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs{\isadigit{1}}{\isacharat}{\kern0pt}rs{\isadigit{2}}{\isacharat}{\kern0pt}rs{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharsemicolon}{\kern0pt}\ {\isasymnot}bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharsemicolon}{\kern0pt}\ {\isasymnot}bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isasymrbrakk}\ {\isasymLongrightarrow}\ \isanewline
+bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs{\isadigit{1}}{\isacharat}{\kern0pt}rs{\isadigit{2}}{\isacharat}{\kern0pt}rs{\isadigit{3}}{\isacharparenright}{\kern0pt}\ {\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}{\isasymforall}r\ {\isasymin}\ set\ {\isacharparenleft}{\kern0pt}rs{\isadigit{1}}{\isacharat}{\kern0pt}rs{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardot}{\kern0pt}\ {\isasymnot}bnullable\ r{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rs{\isadigit{1}}{\isacharat}{\kern0pt}rs{\isadigit{2}}{\isacharat}{\kern0pt}rs{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}{\isacharparenleft}{\kern0pt}rs{\isadigit{1}}{\isacharat}{\kern0pt}rs{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharat}{\kern0pt}rs{\isadigit{3}}{\isacharparenright}{\kern0pt}\ {\isacharparenright}{\kern0pt}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ qq{\isadigit{2}}\ qq{\isadigit{3}}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ append{\isachardot}{\kern0pt}assoc{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ append{\isachardot}{\kern0pt}assoc\ in{\isacharunderscore}{\kern0pt}set{\isacharunderscore}{\kern0pt}conv{\isacharunderscore}{\kern0pt}decomp\ r{\isadigit{2}}\ third{\isacharunderscore}{\kern0pt}segment{\isacharunderscore}{\kern0pt}bnullable{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ third{\isacharunderscore}{\kern0pt}segment{\isacharunderscore}{\kern0pt}bnullable\ \isacommand{by}\isamarkupfalse%
+\ blast%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ rewrite{\isacharunderscore}{\kern0pt}bmkepsalt{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}\ {\isasymlbrakk}bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rsa\ {\isacharat}{\kern0pt}\ AALTs\ bs{\isadigit{1}}\ rs{\isadigit{1}}\ {\isacharhash}{\kern0pt}\ rsb{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isacharsemicolon}{\kern0pt}\ bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rsa\ {\isacharat}{\kern0pt}\ map\ {\isacharparenleft}{\kern0pt}fuse\ bs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ rs{\isadigit{1}}\ {\isacharat}{\kern0pt}\ rsb{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isasymrbrakk}\isanewline
+\ \ \ \ \ \ \ {\isasymLongrightarrow}\ bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rsa\ {\isacharat}{\kern0pt}\ AALTs\ bs{\isadigit{1}}\ rs{\isadigit{1}}\ {\isacharhash}{\kern0pt}\ rsb{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rsa\ {\isacharat}{\kern0pt}\ map\ {\isacharparenleft}{\kern0pt}fuse\ bs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ rs{\isadigit{1}}\ {\isacharat}{\kern0pt}\ rsb{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rsa{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ qq{\isadigit{1}}\ \isacommand{apply}\isamarkupfalse%
+\ force\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs{\isadigit{1}}\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ qq{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ r{\isadigit{2}}\ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\ \ \isanewline
+\ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ list{\isachardot}{\kern0pt}set{\isacharunderscore}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}smt\ {\isacharparenleft}{\kern0pt}verit{\isacharcomma}{\kern0pt}\ ccfv{\isacharunderscore}{\kern0pt}threshold{\isacharparenright}{\kern0pt}\ append{\isacharunderscore}{\kern0pt}eq{\isacharunderscore}{\kern0pt}append{\isacharunderscore}{\kern0pt}conv{\isadigit{2}}\ list{\isachardot}{\kern0pt}set{\isacharunderscore}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}\ qq{\isadigit{2}}\ qq{\isadigit{3}}\ rewritenullable\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{8}}{\isacharparenright}{\kern0pt}\ self{\isacharunderscore}{\kern0pt}append{\isacharunderscore}{\kern0pt}conv{\isadigit{2}}\ spillbmkepslistr{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\isanewline
+\ \ \isacommand{thm}\isamarkupfalse%
+\ qq{\isadigit{1}}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}bmkeps\ \ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rsa\ {\isacharat}{\kern0pt}\ AALTs\ bs{\isadigit{1}}\ rs{\isadigit{1}}\ {\isacharhash}{\kern0pt}\ rsb{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bmkeps\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rsb{\isacharparenright}{\kern0pt}\ {\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{prefer}\isamarkupfalse%
+\ {\isadigit{2}}\isanewline
+\ \ \isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ append{\isacharunderscore}{\kern0pt}Cons\ append{\isacharunderscore}{\kern0pt}Nil\ bnullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}\ bnullable{\isacharunderscore}{\kern0pt}segment\ rewritenullable\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isadigit{1}}{\isacharparenright}{\kern0pt}\ third{\isacharunderscore}{\kern0pt}segment{\isacharunderscore}{\kern0pt}bmkeps{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ bnullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}\ rewrite{\isacharunderscore}{\kern0pt}non{\isacharunderscore}{\kern0pt}nullable\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isadigit{0}}{\isacharparenright}{\kern0pt}\ third{\isacharunderscore}{\kern0pt}segment{\isacharunderscore}{\kern0pt}bmkeps{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ rewrite{\isacharunderscore}{\kern0pt}bmkeps{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymlbrakk}\ r{\isadigit{1}}\ {\isasymleadsto}\ r{\isadigit{2}}{\isacharsemicolon}{\kern0pt}\ {\isacharparenleft}{\kern0pt}bnullable\ r{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isasymrbrakk}\ {\isasymLongrightarrow}\ bmkeps\ r{\isadigit{1}}\ {\isacharequal}{\kern0pt}\ bmkeps\ r{\isadigit{2}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\isanewline
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}frule\ rewritenullable{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induction\ r{\isadigit{1}}\ r{\isadigit{2}}\ rule{\isacharcolon}{\kern0pt}\ rrewrite{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ bnullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}\ bnullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{5}}{\isacharparenright}{\kern0pt}\ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ b{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}frule\ bnullable{\isacharunderscore}{\kern0pt}segment{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ qq{\isadigit{1}}\ \isacommand{apply}\isamarkupfalse%
+\ force\isanewline
+\ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}bnullable\ r{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ bnullablewhichbmkeps\ rewritenullable\ \isacommand{apply}\isamarkupfalse%
+\ presburger\isanewline
+\ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}bnullable\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ rs{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}{\isasymnot}\ bnullable\ r{\isacharprime}{\kern0pt}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ qq{\isadigit{2}}\ r{\isadigit{1}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ rrewrite{\isacharunderscore}{\kern0pt}nbnullable\ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\isanewline
+\ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ flts{\isacharunderscore}{\kern0pt}append\ qs{\isadigit{3}}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}meson\ rewrite{\isacharunderscore}{\kern0pt}bmkepsalt{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ bnullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}\ q{\isadigit{3}}a\ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ q{\isadigit{3}}a{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ bnullable{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}\ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ b{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}smt\ {\isacharparenleft}{\kern0pt}z{\isadigit{3}}{\isacharparenright}{\kern0pt}\ Un{\isacharunderscore}{\kern0pt}iff\ bnullable{\isacharunderscore}{\kern0pt}correctness\ erase{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{5}}{\isacharparenright}{\kern0pt}\ qq{\isadigit{1}}\ qq{\isadigit{2}}\ qq{\isadigit{3}}\ set{\isacharunderscore}{\kern0pt}append{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ rewrites{\isacharunderscore}{\kern0pt}bmkeps{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}{\isasymlbrakk}\ {\isacharparenleft}{\kern0pt}r{\isadigit{1}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharsemicolon}{\kern0pt}\ {\isacharparenleft}{\kern0pt}bnullable\ r{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isasymrbrakk}\ {\isasymLongrightarrow}\ bmkeps\ r{\isadigit{1}}\ {\isacharequal}{\kern0pt}\ bmkeps\ r{\isadigit{2}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induction\ r{\isadigit{1}}\ r{\isadigit{2}}\ rule{\isacharcolon}{\kern0pt}\ rrewrites{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}bnullable\ r{\isadigit{2}}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{prefer}\isamarkupfalse%
+\ {\isadigit{2}}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}metis\ rewritesnullable{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}bmkeps\ r{\isadigit{1}}\ {\isacharequal}{\kern0pt}\ bmkeps\ r{\isadigit{2}}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{prefer}\isamarkupfalse%
+\ {\isadigit{2}}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ fastforce\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ rewrite{\isacharunderscore}{\kern0pt}bmkeps\ \isacommand{by}\isamarkupfalse%
+\ presburger%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{thm}\isamarkupfalse%
+\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ alts{\isacharunderscore}{\kern0pt}rewrite{\isacharunderscore}{\kern0pt}front{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}r\ {\isasymleadsto}\ r{\isacharprime}{\kern0pt}\ {\isasymLongrightarrow}\ AALTs\ bs\ {\isacharparenleft}{\kern0pt}r\ {\isacharhash}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}\ {\isasymleadsto}\ AALTs\ bs\ {\isacharparenleft}{\kern0pt}r{\isacharprime}{\kern0pt}\ {\isacharhash}{\kern0pt}\ rs{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ append{\isacharunderscore}{\kern0pt}Cons\ append{\isacharunderscore}{\kern0pt}Nil\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{6}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ alt{\isacharunderscore}{\kern0pt}rewrite{\isacharunderscore}{\kern0pt}front{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}r\ {\isasymleadsto}\ r{\isacharprime}{\kern0pt}\ {\isasymLongrightarrow}\ AALT\ bs\ r\ r{\isadigit{2}}\ {\isasymleadsto}\ AALT\ bs\ r{\isacharprime}{\kern0pt}\ r{\isadigit{2}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ alts{\isacharunderscore}{\kern0pt}rewrite{\isacharunderscore}{\kern0pt}front\ \isacommand{by}\isamarkupfalse%
+\ blast%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ to{\isacharunderscore}{\kern0pt}zero{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}alt{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}\ AALT\ bs\ {\isacharparenleft}{\kern0pt}ASEQ\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ AZERO\ r{\isacharparenright}{\kern0pt}\ r{\isadigit{2}}\ {\isasymleadsto}\ \ AALT\ bs\ AZERO\ r{\isadigit{2}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ alts{\isacharunderscore}{\kern0pt}rewrite{\isacharunderscore}{\kern0pt}front\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ alt{\isacharunderscore}{\kern0pt}remove{\isadigit{0}}{\isacharunderscore}{\kern0pt}front{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}\ AALT\ bs\ AZERO\ r\ {\isasymleadsto}\ AALTs\ bs\ {\isacharbrackleft}{\kern0pt}r{\isacharbrackright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ rrewrite{\isadigit{0}}away{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ alt{\isacharunderscore}{\kern0pt}rewrites{\isacharunderscore}{\kern0pt}back{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}r{\isadigit{2}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ r{\isadigit{2}}{\isacharprime}{\kern0pt}\ {\isasymLongrightarrow}AALT\ bs\ r{\isadigit{1}}\ r{\isadigit{2}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ AALT\ bs\ r{\isadigit{1}}\ r{\isadigit{2}}{\isacharprime}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induction\ r{\isadigit{2}}\ r{\isadigit{2}}{\isacharprime}{\kern0pt}\ arbitrary{\isacharcolon}{\kern0pt}\ bs\ rule{\isacharcolon}{\kern0pt}\ rrewrites{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}meson\ rs{\isadigit{1}}\ rs{\isadigit{2}}\ srewrites{\isacharunderscore}{\kern0pt}alt{\isadigit{1}}\ ss{\isadigit{1}}\ ss{\isadigit{2}}{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ rewrite{\isacharunderscore}{\kern0pt}fuse{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}\ r{\isadigit{2}}\ {\isasymleadsto}\ r{\isadigit{3}}\ {\isasymLongrightarrow}\ fuse\ bs\ r{\isadigit{2}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ fuse\ bs\ r{\isadigit{3}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induction\ r{\isadigit{2}}\ r{\isadigit{3}}\ arbitrary{\isacharcolon}{\kern0pt}\ bs\ rule{\isacharcolon}{\kern0pt}\ rrewrite{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ auto\isanewline
+\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ continuous{\isacharunderscore}{\kern0pt}rewrite{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ fuse{\isacharunderscore}{\kern0pt}append\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{3}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ star{\isacharunderscore}{\kern0pt}seq\ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ star{\isacharunderscore}{\kern0pt}seq{\isadigit{2}}\ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ contextrewrites{\isadigit{2}}\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \isanewline
+\ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{7}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{8}}{\isacharparenright}{\kern0pt}\ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ append{\isacharunderscore}{\kern0pt}assoc\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{9}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ append{\isacharunderscore}{\kern0pt}assoc\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isadigit{0}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ fuse{\isacharunderscore}{\kern0pt}append\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isadigit{3}}{\isacharparenright}{\kern0pt}\ \isacommand{by}\isamarkupfalse%
+\ auto%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\ \ \isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ rewrites{\isacharunderscore}{\kern0pt}fuse{\isacharcolon}{\kern0pt}\ \ {\isachardoublequoteopen}r{\isadigit{2}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ r{\isadigit{2}}{\isacharprime}{\kern0pt}\ {\isasymLongrightarrow}\ \ {\isacharparenleft}{\kern0pt}fuse\ bs{\isadigit{1}}\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ \ {\isacharparenleft}{\kern0pt}fuse\ bs{\isadigit{1}}\ r{\isadigit{2}}{\isacharprime}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induction\ r{\isadigit{2}}\ r{\isadigit{2}}{\isacharprime}{\kern0pt}\ arbitrary{\isacharcolon}{\kern0pt}\ bs{\isadigit{1}}\ rule{\isacharcolon}{\kern0pt}\ rrewrites{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}meson\ real{\isacharunderscore}{\kern0pt}trans\ rewrite{\isacharunderscore}{\kern0pt}fuse{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ \ bder{\isacharunderscore}{\kern0pt}fuse{\isacharunderscore}{\kern0pt}list{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}\ map\ {\isacharparenleft}{\kern0pt}bder\ c\ {\isasymcirc}\ fuse\ bs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ rs{\isadigit{1}}\ {\isacharequal}{\kern0pt}\ map\ {\isacharparenleft}{\kern0pt}fuse\ bs{\isadigit{1}}\ {\isasymcirc}\ bder\ c{\isacharparenright}{\kern0pt}\ rs{\isadigit{1}}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induction\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ bder{\isacharunderscore}{\kern0pt}fuse{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ rewrite{\isacharunderscore}{\kern0pt}der{\isacharunderscore}{\kern0pt}altmiddle{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}bder\ c\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rsa\ {\isacharat}{\kern0pt}\ AALTs\ bs{\isadigit{1}}\ rs{\isadigit{1}}\ {\isacharhash}{\kern0pt}\ rsb{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ bder\ c\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rsa\ {\isacharat}{\kern0pt}\ map\ {\isacharparenleft}{\kern0pt}fuse\ bs{\isadigit{1}}{\isacharparenright}{\kern0pt}\ rs{\isadigit{1}}\ {\isacharat}{\kern0pt}\ rsb{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ bder{\isacharunderscore}{\kern0pt}fuse{\isacharunderscore}{\kern0pt}list{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ many{\isacharunderscore}{\kern0pt}steps{\isacharunderscore}{\kern0pt}later{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{8}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ fastforce%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ lock{\isacharunderscore}{\kern0pt}step{\isacharunderscore}{\kern0pt}der{\isacharunderscore}{\kern0pt}removal{\isacharcolon}{\kern0pt}\ \isanewline
+\ \ \isakeyword{shows}\ {\isachardoublequoteopen}\ erase\ a{\isadigit{1}}\ {\isacharequal}{\kern0pt}\ erase\ a{\isadigit{2}}\ {\isasymLongrightarrow}\ \isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ bder\ c\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rsa\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}a{\isadigit{1}}{\isacharbrackright}{\kern0pt}\ {\isacharat}{\kern0pt}\ rsb\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}a{\isadigit{2}}{\isacharbrackright}{\kern0pt}\ {\isacharat}{\kern0pt}\ rsc{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ \isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ bder\ c\ {\isacharparenleft}{\kern0pt}AALTs\ bs\ {\isacharparenleft}{\kern0pt}rsa\ {\isacharat}{\kern0pt}\ {\isacharbrackleft}{\kern0pt}a{\isadigit{1}}{\isacharbrackright}{\kern0pt}\ {\isacharat}{\kern0pt}\ rsb\ {\isacharat}{\kern0pt}\ rsc{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isadigit{3}}{\isacharparenright}{\kern0pt}\ \isacommand{by}\isamarkupfalse%
+\ auto%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ rewrite{\isacharunderscore}{\kern0pt}after{\isacharunderscore}{\kern0pt}der{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}r{\isadigit{1}}\ {\isasymleadsto}\ r{\isadigit{2}}\ {\isasymLongrightarrow}\ {\isacharparenleft}{\kern0pt}bder\ c\ r{\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ {\isacharparenleft}{\kern0pt}bder\ c\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induction\ r{\isadigit{1}}\ r{\isadigit{2}}\ arbitrary{\isacharcolon}{\kern0pt}\ c\ rule{\isacharcolon}{\kern0pt}\ rrewrite{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}meson\ contextrewrites{\isadigit{1}}\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isadigit{1}}{\isacharparenright}{\kern0pt}\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}\ rrewrite{\isadigit{0}}away\ rs{\isadigit{2}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}simp{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ many{\isacharunderscore}{\kern0pt}steps{\isacharunderscore}{\kern0pt}later{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ to{\isacharunderscore}{\kern0pt}zero{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}alt{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ many{\isacharunderscore}{\kern0pt}steps{\isacharunderscore}{\kern0pt}later{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ alt{\isacharunderscore}{\kern0pt}remove{\isadigit{0}}{\isacharunderscore}{\kern0pt}front{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ many{\isacharunderscore}{\kern0pt}steps{\isacharunderscore}{\kern0pt}later{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ bder{\isacharunderscore}{\kern0pt}fuse\ fuse{\isacharunderscore}{\kern0pt}append\ rs{\isadigit{1}}\ \isacommand{apply}\isamarkupfalse%
+\ presburger\isanewline
+\ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}bnullable\ r{\isadigit{1}}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{prefer}\isamarkupfalse%
+\ {\isadigit{2}}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}{\isasymnot}bnullable\ r{\isadigit{2}}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \isacommand{prefer}\isamarkupfalse%
+\ {\isadigit{2}}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ rewrite{\isacharunderscore}{\kern0pt}non{\isacharunderscore}{\kern0pt}nullable\ \isacommand{apply}\isamarkupfalse%
+\ presburger\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp{\isacharplus}{\kern0pt}\isanewline
+\ \ \isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ star{\isacharunderscore}{\kern0pt}seq\ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}bnullable\ r{\isadigit{2}}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp{\isacharplus}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}bmkeps\ r{\isadigit{1}}\ {\isacharequal}{\kern0pt}\ bmkeps\ r{\isadigit{2}}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{prefer}\isamarkupfalse%
+\ {\isadigit{2}}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ rewrite{\isacharunderscore}{\kern0pt}bmkeps\ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ contextrewrites{\isadigit{1}}\ star{\isacharunderscore}{\kern0pt}seq\ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ rewritenullable\ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}case{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}bnullable\ r{\isadigit{1}}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subgoal{\isacharunderscore}{\kern0pt}tac\ {\isachardoublequoteopen}ASEQ\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}bder\ c\ r{\isadigit{1}}{\isacharparenright}{\kern0pt}\ r{\isadigit{3}}\ {\isasymleadsto}\ ASEQ\ {\isacharbrackleft}{\kern0pt}{\isacharbrackright}{\kern0pt}\ {\isacharparenleft}{\kern0pt}bder\ c\ r{\isadigit{1}}{\isacharparenright}{\kern0pt}\ r{\isadigit{4}}{\isachardoublequoteclose}{\isacharparenright}{\kern0pt}\ \isanewline
+\ \ \ \ \ \ \ \ \ \ \ \isacommand{prefer}\isamarkupfalse%
+\ {\isadigit{2}}\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{5}}{\isacharparenright}{\kern0pt}\ \isacommand{apply}\isamarkupfalse%
+\ blast\isanewline
+\ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ many{\isacharunderscore}{\kern0pt}steps{\isacharunderscore}{\kern0pt}later{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ alt{\isacharunderscore}{\kern0pt}rewrite{\isacharunderscore}{\kern0pt}front{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ assumption\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}meson\ alt{\isacharunderscore}{\kern0pt}rewrites{\isacharunderscore}{\kern0pt}back\ rewrites{\isacharunderscore}{\kern0pt}fuse{\isacharparenright}{\kern0pt}\ \isanewline
+\isanewline
+\ \ \ \ \ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{5}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ contextrewrites{\isadigit{2}}\ \isacommand{apply}\isamarkupfalse%
+\ force\isanewline
+\isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{7}}{\isacharparenright}{\kern0pt}\ \isacommand{apply}\isamarkupfalse%
+\ force\isanewline
+\ \ \isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ rewrite{\isacharunderscore}{\kern0pt}der{\isacharunderscore}{\kern0pt}altmiddle\ \isacommand{apply}\isamarkupfalse%
+\ auto{\isacharbrackleft}{\kern0pt}{\isadigit{1}}{\isacharbrackright}{\kern0pt}\isanewline
+\ \ \isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ bder{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}\ bder{\isacharunderscore}{\kern0pt}fuse{\isacharunderscore}{\kern0pt}list\ map{\isacharunderscore}{\kern0pt}map\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{9}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ List{\isachardot}{\kern0pt}map{\isachardot}{\kern0pt}compositionality\ bder{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}\ bder{\isacharunderscore}{\kern0pt}fuse{\isacharunderscore}{\kern0pt}list\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isadigit{0}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ r{\isacharunderscore}{\kern0pt}in{\isacharunderscore}{\kern0pt}rstar\ rrewrite{\isachardot}{\kern0pt}intros{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isadigit{1}}{\isacharparenright}{\kern0pt}{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ bder{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{4}}{\isacharparenright}{\kern0pt}\ bder{\isacharunderscore}{\kern0pt}bsimp{\isacharunderscore}{\kern0pt}AALTs\ bsimp{\isacharunderscore}{\kern0pt}AALTs{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}\ bsimp{\isacharunderscore}{\kern0pt}AALTsrewrites{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \isanewline
+\ \ \isacommand{using}\isamarkupfalse%
+\ lock{\isacharunderscore}{\kern0pt}step{\isacharunderscore}{\kern0pt}der{\isacharunderscore}{\kern0pt}removal\ \isacommand{by}\isamarkupfalse%
+\ auto%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ rewrites{\isacharunderscore}{\kern0pt}after{\isacharunderscore}{\kern0pt}der{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}\ \ r{\isadigit{1}}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ r{\isadigit{2}}\ \ {\isasymLongrightarrow}\ \ {\isacharparenleft}{\kern0pt}bder\ c\ r{\isadigit{1}}{\isacharparenright}{\kern0pt}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ {\isacharparenleft}{\kern0pt}bder\ c\ r{\isadigit{2}}{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induction\ r{\isadigit{1}}\ r{\isadigit{2}}\ rule{\isacharcolon}{\kern0pt}\ rrewrites{\isachardot}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}rule\ rs{\isadigit{1}}{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}meson\ real{\isacharunderscore}{\kern0pt}trans\ rewrite{\isacharunderscore}{\kern0pt}after{\isacharunderscore}{\kern0pt}der{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\ \ \isanewline
+\isanewline
+\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ central{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}\ {\isacharparenleft}{\kern0pt}bders\ r\ s{\isacharparenright}{\kern0pt}\ {\isasymleadsto}{\isacharasterisk}{\kern0pt}\ \ {\isacharparenleft}{\kern0pt}bders{\isacharunderscore}{\kern0pt}simp\ r\ s{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\ \isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}induct\ s\ arbitrary{\isacharcolon}{\kern0pt}\ r\ rule{\isacharcolon}{\kern0pt}\ rev{\isacharunderscore}{\kern0pt}induct{\isacharparenright}{\kern0pt}\isanewline
+\isanewline
+\ \ \ \isacommand{apply}\isamarkupfalse%
+\ simp\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ bders{\isacharunderscore}{\kern0pt}append{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{apply}\isamarkupfalse%
+{\isacharparenleft}{\kern0pt}subst\ bders{\isacharunderscore}{\kern0pt}simp{\isacharunderscore}{\kern0pt}append{\isacharparenright}{\kern0pt}\isanewline
+\ \ \isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}metis\ bders{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}\ bders{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}\ bders{\isacharunderscore}{\kern0pt}simp{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{1}}{\isacharparenright}{\kern0pt}\ bders{\isacharunderscore}{\kern0pt}simp{\isachardot}{\kern0pt}simps{\isacharparenleft}{\kern0pt}{\isadigit{2}}{\isacharparenright}{\kern0pt}\ bsimp{\isacharunderscore}{\kern0pt}rewrite\ real{\isacharunderscore}{\kern0pt}trans\ rewrites{\isacharunderscore}{\kern0pt}after{\isacharunderscore}{\kern0pt}der{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isanewline
+\isacommand{thm}\isamarkupfalse%
+\ arexp{\isachardot}{\kern0pt}induct\isanewline
+\isanewline
+\isacommand{lemma}\isamarkupfalse%
+\ quasi{\isacharunderscore}{\kern0pt}main{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}bnullable\ {\isacharparenleft}{\kern0pt}bders\ r\ s{\isacharparenright}{\kern0pt}\ {\isasymLongrightarrow}\ bmkeps\ {\isacharparenleft}{\kern0pt}bders\ r\ s{\isacharparenright}{\kern0pt}\ {\isacharequal}{\kern0pt}\ bmkeps\ {\isacharparenleft}{\kern0pt}bders{\isacharunderscore}{\kern0pt}simp\ r\ s{\isacharparenright}{\kern0pt}{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ central\ rewrites{\isacharunderscore}{\kern0pt}bmkeps\ \isacommand{by}\isamarkupfalse%
+\ blast%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isacommand{theorem}\isamarkupfalse%
+\ main{\isacharunderscore}{\kern0pt}main{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}blexer\ r\ s\ {\isacharequal}{\kern0pt}\ blexer{\isacharunderscore}{\kern0pt}simp\ r\ s{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{by}\isamarkupfalse%
+\ {\isacharparenleft}{\kern0pt}simp\ add{\isacharcolon}{\kern0pt}\ b{\isadigit{4}}\ blexer{\isacharunderscore}{\kern0pt}def\ blexer{\isacharunderscore}{\kern0pt}simp{\isacharunderscore}{\kern0pt}def\ quasi{\isacharunderscore}{\kern0pt}main{\isacharparenright}{\kern0pt}%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{theorem}\isamarkupfalse%
+\ blexersimp{\isacharunderscore}{\kern0pt}correctness{\isacharcolon}{\kern0pt}\ {\isachardoublequoteopen}blexer{\isacharunderscore}{\kern0pt}simp\ r\ s{\isacharequal}{\kern0pt}\ lexer\ r\ s{\isachardoublequoteclose}\isanewline
+%
+\isadelimproof
+\ \ %
+\endisadelimproof
+%
+\isatagproof
+\isacommand{using}\isamarkupfalse%
+\ blexer{\isacharunderscore}{\kern0pt}correctness\ main{\isacharunderscore}{\kern0pt}main\ \isacommand{by}\isamarkupfalse%
+\ auto%
+\endisatagproof
+{\isafoldproof}%
+%
+\isadelimproof
+\isanewline
+%
+\endisadelimproof
+\isanewline
+\isanewline
+\isacommand{unused{\isacharunderscore}{\kern0pt}thms}\isamarkupfalse%
+\isanewline
+\isanewline
+%
+\isadelimtheory
+\isanewline
+%
+\endisadelimtheory
+%
+\isatagtheory
+\isacommand{end}\isamarkupfalse%
+%
+\endisatagtheory
+{\isafoldtheory}%
+%
+\isadelimtheory
+%
+\endisadelimtheory
+%
+\end{isabellebody}%
+\endinput
+%:%file=~/Dropbox/Workspace/journalpaper/lexing/thys2/SizeBound.thy%:%
+%:%6=1%:%
+%:%11=2%:%
+%:%12=2%:%
+%:%13=3%:%
+%:%14=4%:%
+%:%28=6%:%
+%:%38=8%:%
+%:%39=8%:%
+%:%40=9%:%
+%:%41=10%:%
+%:%42=10%:%
+%:%43=11%:%
+%:%44=12%:%
+%:%45=13%:%
+%:%46=14%:%
+%:%47=15%:%
+%:%48=16%:%
+%:%49=17%:%
+%:%50=18%:%
+%:%51=19%:%
+%:%52=20%:%
+%:%53=21%:%
+%:%54=21%:%
+%:%55=22%:%
+%:%56=23%:%
+%:%57=24%:%
+%:%58=25%:%
+%:%59=26%:%
+%:%60=26%:%
+%:%61=27%:%
+%:%62=28%:%
+%:%63=29%:%
+%:%64=30%:%
+%:%65=31%:%
+%:%66=32%:%
+%:%67=33%:%
+%:%68=34%:%
+%:%69=35%:%
+%:%70=36%:%
+%:%71=37%:%
+%:%72=38%:%
+%:%73=39%:%
+%:%75=41%:%
+%:%82=42%:%
+%:%83=42%:%
+%:%88=42%:%
+%:%91=43%:%
+%:%92=44%:%
+%:%93=44%:%
+%:%94=45%:%
+%:%95=46%:%
+%:%102=47%:%
+%:%103=47%:%
+%:%104=48%:%
+%:%105=48%:%
+%:%106=49%:%
+%:%107=49%:%
+%:%108=50%:%
+%:%109=50%:%
+%:%110=50%:%
+%:%111=51%:%
+%:%112=51%:%
+%:%117=51%:%
+%:%120=52%:%
+%:%121=53%:%
+%:%122=53%:%
+%:%129=54%:%
+%:%130=54%:%
+%:%131=55%:%
+%:%132=55%:%
+%:%133=56%:%
+%:%134=56%:%
+%:%139=56%:%
+%:%142=57%:%
+%:%143=58%:%
+%:%144=58%:%
+%:%145=59%:%
+%:%146=60%:%
+%:%147=61%:%
+%:%148=62%:%
+%:%149=63%:%
+%:%150=64%:%
+%:%151=64%:%
+%:%152=65%:%
+%:%153=66%:%
+%:%156=67%:%
+%:%160=67%:%
+%:%161=67%:%
+%:%162=68%:%
+%:%163=68%:%
+%:%164=69%:%
+%:%165=69%:%
+%:%166=70%:%
+%:%172=70%:%
+%:%175=71%:%
+%:%176=72%:%
+%:%177=72%:%
+%:%178=73%:%
+%:%179=74%:%
+%:%186=75%:%
+%:%187=75%:%
+%:%188=76%:%
+%:%189=76%:%
+%:%190=77%:%
+%:%191=77%:%
+%:%192=78%:%
+%:%193=78%:%
+%:%194=78%:%
+%:%199=78%:%
+%:%202=79%:%
+%:%203=80%:%
+%:%204=80%:%
+%:%205=81%:%
+%:%206=82%:%
+%:%209=83%:%
+%:%213=83%:%
+%:%214=83%:%
+%:%215=83%:%
+%:%216=84%:%
+%:%217=84%:%
+%:%231=87%:%
+%:%241=89%:%
+%:%242=89%:%
+%:%243=90%:%
+%:%244=91%:%
+%:%245=92%:%
+%:%246=93%:%
+%:%247=94%:%
+%:%248=95%:%
+%:%249=96%:%
+%:%250=97%:%
+%:%251=97%:%
+%:%252=98%:%
+%:%253=99%:%
+%:%254=100%:%
+%:%255=100%:%
+%:%256=101%:%
+%:%257=102%:%
+%:%258=103%:%
+%:%259=104%:%
+%:%260=105%:%
+%:%261=106%:%
+%:%262=107%:%
+%:%263=108%:%
+%:%264=108%:%
+%:%265=109%:%
+%:%266=110%:%
+%:%267=111%:%
+%:%268=112%:%
+%:%269=113%:%
+%:%270=114%:%
+%:%271=115%:%
+%:%272=116%:%
+%:%273=117%:%
+%:%274=118%:%
+%:%275=119%:%
+%:%276=120%:%
+%:%277=121%:%
+%:%278=122%:%
+%:%279=123%:%
+%:%280=123%:%
+%:%281=124%:%
+%:%282=125%:%
+%:%283=126%:%
+%:%284=127%:%
+%:%285=128%:%
+%:%286=129%:%
+%:%287=129%:%
+%:%288=130%:%
+%:%289=131%:%
+%:%290=132%:%
+%:%291=133%:%
+%:%292=134%:%
+%:%293=135%:%
+%:%294=136%:%
+%:%295=137%:%
+%:%296=138%:%
+%:%297=139%:%
+%:%298=140%:%
+%:%299=141%:%
+%:%300=142%:%
+%:%301=143%:%
+%:%302=144%:%
+%:%303=145%:%
+%:%304=145%:%
+%:%305=146%:%
+%:%306=147%:%
+%:%307=148%:%
+%:%308=149%:%
+%:%309=150%:%
+%:%310=151%:%
+%:%311=152%:%
+%:%312=153%:%
+%:%313=153%:%
+%:%314=154%:%
+%:%317=155%:%
+%:%321=155%:%
+%:%322=155%:%
+%:%323=156%:%
+%:%324=156%:%
+%:%325=157%:%
+%:%331=157%:%
+%:%334=158%:%
+%:%335=159%:%
+%:%336=160%:%
+%:%337=160%:%
+%:%338=161%:%
+%:%339=162%:%
+%:%340=163%:%
+%:%341=164%:%
+%:%342=165%:%
+%:%343=166%:%
+%:%344=167%:%
+%:%345=168%:%
+%:%346=169%:%
+%:%347=170%:%
+%:%348=170%:%
+%:%349=171%:%
+%:%350=172%:%
+%:%351=173%:%
+%:%352=174%:%
+%:%353=175%:%
+%:%354=176%:%
+%:%355=177%:%
+%:%356=178%:%
+%:%357=179%:%
+%:%358=180%:%
+%:%359=181%:%
+%:%360=182%:%
+%:%361=183%:%
+%:%362=183%:%
+%:%363=184%:%
+%:%364=185%:%
+%:%365=186%:%
+%:%366=187%:%
+%:%367=188%:%
+%:%368=189%:%
+%:%369=190%:%
+%:%370=191%:%
+%:%371=192%:%
+%:%372=193%:%
+%:%373=193%:%
+%:%374=194%:%
+%:%375=195%:%
+%:%376=196%:%
+%:%377=197%:%
+%:%378=198%:%
+%:%379=199%:%
+%:%380=200%:%
+%:%381=201%:%
+%:%382=202%:%
+%:%383=203%:%
+%:%384=203%:%
+%:%385=204%:%
+%:%386=205%:%
+%:%387=206%:%
+%:%388=207%:%
+%:%389=208%:%
+%:%390=209%:%
+%:%391=210%:%
+%:%394=213%:%
+%:%395=214%:%
+%:%396=215%:%
+%:%397=216%:%
+%:%398=217%:%
+%:%399=217%:%
+%:%400=218%:%
+%:%401=219%:%
+%:%402=220%:%
+%:%403=221%:%
+%:%404=222%:%
+%:%405=223%:%
+%:%406=223%:%
+%:%407=224%:%
+%:%410=225%:%
+%:%414=225%:%
+%:%415=225%:%
+%:%416=226%:%
+%:%417=226%:%
+%:%418=227%:%
+%:%424=227%:%
+%:%427=228%:%
+%:%428=229%:%
+%:%429=229%:%
+%:%430=230%:%
+%:%433=231%:%
+%:%437=231%:%
+%:%438=231%:%
+%:%439=232%:%
+%:%440=232%:%
+%:%441=233%:%
+%:%447=233%:%
+%:%450=234%:%
+%:%451=235%:%
+%:%452=235%:%
+%:%453=236%:%
+%:%456=237%:%
+%:%460=237%:%
+%:%461=237%:%
+%:%462=238%:%
+%:%463=238%:%
+%:%464=239%:%
+%:%470=239%:%
+%:%473=240%:%
+%:%474=241%:%
+%:%475=241%:%
+%:%476=242%:%
+%:%477=243%:%
+%:%478=243%:%
+%:%479=244%:%
+%:%482=245%:%
+%:%486=245%:%
+%:%487=245%:%
+%:%488=246%:%
+%:%489=246%:%
+%:%490=247%:%
+%:%496=247%:%
+%:%499=248%:%
+%:%500=249%:%
+%:%501=249%:%
+%:%502=250%:%
+%:%505=251%:%
+%:%509=251%:%
+%:%510=251%:%
+%:%511=252%:%
+%:%512=252%:%
+%:%513=253%:%
+%:%519=253%:%
+%:%522=254%:%
+%:%523=255%:%
+%:%524=255%:%
+%:%525=256%:%
+%:%528=257%:%
+%:%532=257%:%
+%:%533=257%:%
+%:%534=258%:%
+%:%535=258%:%
+%:%536=259%:%
+%:%542=259%:%
+%:%545=260%:%
+%:%546=261%:%
+%:%547=261%:%
+%:%548=262%:%
+%:%549=263%:%
+%:%552=264%:%
+%:%556=264%:%
+%:%557=264%:%
+%:%558=265%:%
+%:%559=265%:%
+%:%560=266%:%
+%:%561=266%:%
+%:%562=267%:%
+%:%568=267%:%
+%:%571=268%:%
+%:%572=269%:%
+%:%573=270%:%
+%:%574=270%:%
+%:%575=271%:%
+%:%576=272%:%
+%:%579=273%:%
+%:%583=273%:%
+%:%584=273%:%
+%:%585=274%:%
+%:%586=274%:%
+%:%587=275%:%
+%:%588=275%:%
+%:%589=276%:%
+%:%590=276%:%
+%:%591=277%:%
+%:%592=277%:%
+%:%593=278%:%
+%:%594=278%:%
+%:%595=279%:%
+%:%596=279%:%
+%:%597=280%:%
+%:%598=280%:%
+%:%599=281%:%
+%:%600=281%:%
+%:%601=282%:%
+%:%602=283%:%
+%:%603=283%:%
+%:%604=284%:%
+%:%605=284%:%
+%:%606=285%:%
+%:%607=285%:%
+%:%608=286%:%
+%:%609=286%:%
+%:%610=287%:%
+%:%611=287%:%
+%:%612=288%:%
+%:%613=288%:%
+%:%614=289%:%
+%:%615=289%:%
+%:%616=290%:%
+%:%617=290%:%
+%:%618=291%:%
+%:%619=291%:%
+%:%620=292%:%
+%:%626=292%:%
+%:%629=293%:%
+%:%630=294%:%
+%:%631=294%:%
+%:%632=295%:%
+%:%633=296%:%
+%:%636=297%:%
+%:%640=297%:%
+%:%641=297%:%
+%:%642=298%:%
+%:%643=298%:%
+%:%648=298%:%
+%:%651=299%:%
+%:%652=300%:%
+%:%653=301%:%
+%:%654=301%:%
+%:%655=302%:%
+%:%656=303%:%
+%:%659=304%:%
+%:%663=304%:%
+%:%664=304%:%
+%:%665=305%:%
+%:%666=305%:%
+%:%667=306%:%
+%:%668=306%:%
+%:%669=307%:%
+%:%675=307%:%
+%:%678=308%:%
+%:%679=309%:%
+%:%680=310%:%
+%:%681=310%:%
+%:%682=311%:%
+%:%683=312%:%
+%:%686=313%:%
+%:%690=313%:%
+%:%691=313%:%
+%:%692=314%:%
+%:%693=314%:%
+%:%698=314%:%
+%:%701=315%:%
+%:%702=316%:%
+%:%703=316%:%
+%:%704=317%:%
+%:%705=318%:%
+%:%708=319%:%
+%:%712=319%:%
+%:%713=319%:%
+%:%714=320%:%
+%:%715=320%:%
+%:%716=321%:%
+%:%717=321%:%
+%:%718=322%:%
+%:%724=322%:%
+%:%727=323%:%
+%:%728=324%:%
+%:%729=324%:%
+%:%730=325%:%
+%:%731=326%:%
+%:%734=327%:%
+%:%738=327%:%
+%:%739=327%:%
+%:%740=328%:%
+%:%741=328%:%
+%:%742=329%:%
+%:%743=329%:%
+%:%744=330%:%
+%:%750=330%:%
+%:%753=331%:%
+%:%754=332%:%
+%:%755=332%:%
+%:%756=333%:%
+%:%757=334%:%
+%:%758=335%:%
+%:%759=336%:%
+%:%762=337%:%
+%:%766=337%:%
+%:%767=337%:%
+%:%768=338%:%
+%:%769=338%:%
+%:%770=339%:%
+%:%771=339%:%
+%:%772=340%:%
+%:%773=340%:%
+%:%774=341%:%
+%:%775=341%:%
+%:%776=341%:%
+%:%777=342%:%
+%:%778=342%:%
+%:%779=343%:%
+%:%780=343%:%
+%:%781=344%:%
+%:%782=344%:%
+%:%783=345%:%
+%:%784=345%:%
+%:%785=346%:%
+%:%786=346%:%
+%:%787=347%:%
+%:%788=347%:%
+%:%789=348%:%
+%:%790=348%:%
+%:%791=349%:%
+%:%792=349%:%
+%:%793=350%:%
+%:%794=350%:%
+%:%795=351%:%
+%:%796=351%:%
+%:%797=352%:%
+%:%798=352%:%
+%:%799=353%:%
+%:%800=353%:%
+%:%801=354%:%
+%:%802=354%:%
+%:%803=355%:%
+%:%804=355%:%
+%:%805=356%:%
+%:%806=356%:%
+%:%807=357%:%
+%:%808=357%:%
+%:%809=358%:%
+%:%810=358%:%
+%:%811=359%:%
+%:%812=359%:%
+%:%813=360%:%
+%:%814=360%:%
+%:%815=361%:%
+%:%816=361%:%
+%:%817=362%:%
+%:%818=362%:%
+%:%819=363%:%
+%:%825=363%:%
+%:%828=364%:%
+%:%829=365%:%
+%:%830=366%:%
+%:%831=366%:%
+%:%832=367%:%
+%:%833=368%:%
+%:%834=369%:%
+%:%837=370%:%
+%:%841=370%:%
+%:%842=370%:%
+%:%843=371%:%
+%:%844=371%:%
+%:%845=372%:%
+%:%846=372%:%
+%:%847=373%:%
+%:%848=373%:%
+%:%849=374%:%
+%:%850=374%:%
+%:%851=375%:%
+%:%852=375%:%
+%:%853=376%:%
+%:%854=376%:%
+%:%855=377%:%
+%:%856=377%:%
+%:%857=378%:%
+%:%858=378%:%
+%:%859=379%:%
+%:%860=379%:%
+%:%861=380%:%
+%:%862=380%:%
+%:%863=381%:%
+%:%864=381%:%
+%:%865=382%:%
+%:%866=382%:%
+%:%867=383%:%
+%:%868=383%:%
+%:%869=384%:%
+%:%870=384%:%
+%:%871=385%:%
+%:%872=385%:%
+%:%873=386%:%
+%:%874=386%:%
+%:%875=387%:%
+%:%876=387%:%
+%:%877=388%:%
+%:%878=388%:%
+%:%879=389%:%
+%:%880=389%:%
+%:%881=390%:%
+%:%882=390%:%
+%:%883=390%:%
+%:%884=391%:%
+%:%885=391%:%
+%:%886=392%:%
+%:%887=392%:%
+%:%888=393%:%
+%:%889=393%:%
+%:%890=394%:%
+%:%891=394%:%
+%:%892=394%:%
+%:%893=395%:%
+%:%894=395%:%
+%:%895=396%:%
+%:%896=396%:%
+%:%897=397%:%
+%:%898=397%:%
+%:%899=398%:%
+%:%900=398%:%
+%:%901=399%:%
+%:%902=399%:%
+%:%903=399%:%
+%:%904=400%:%
+%:%905=400%:%
+%:%906=401%:%
+%:%907=401%:%
+%:%908=401%:%
+%:%913=401%:%
+%:%916=402%:%
+%:%917=403%:%
+%:%918=403%:%
+%:%919=404%:%
+%:%920=405%:%
+%:%923=406%:%
+%:%927=406%:%
+%:%928=406%:%
+%:%929=407%:%
+%:%930=407%:%
+%:%931=408%:%
+%:%932=408%:%
+%:%933=409%:%
+%:%934=409%:%
+%:%935=410%:%
+%:%936=410%:%
+%:%937=411%:%
+%:%938=411%:%
+%:%939=412%:%
+%:%940=412%:%
+%:%941=413%:%
+%:%942=413%:%
+%:%943=414%:%
+%:%944=414%:%
+%:%945=415%:%
+%:%946=415%:%
+%:%947=416%:%
+%:%953=416%:%
+%:%956=417%:%
+%:%957=418%:%
+%:%958=418%:%
+%:%959=419%:%
+%:%960=420%:%
+%:%963=421%:%
+%:%967=421%:%
+%:%968=421%:%
+%:%969=422%:%
+%:%970=422%:%
+%:%971=423%:%
+%:%972=423%:%
+%:%973=424%:%
+%:%974=424%:%
+%:%975=425%:%
+%:%976=425%:%
+%:%977=426%:%
+%:%978=426%:%
+%:%979=427%:%
+%:%980=427%:%
+%:%981=428%:%
+%:%982=428%:%
+%:%983=429%:%
+%:%984=429%:%
+%:%985=430%:%
+%:%986=430%:%
+%:%987=431%:%
+%:%988=431%:%
+%:%989=432%:%
+%:%990=432%:%
+%:%991=433%:%
+%:%992=433%:%
+%:%993=434%:%
+%:%994=434%:%
+%:%995=435%:%
+%:%996=435%:%
+%:%997=436%:%
+%:%998=436%:%
+%:%999=437%:%
+%:%1000=437%:%
+%:%1001=438%:%
+%:%1002=438%:%
+%:%1003=439%:%
+%:%1004=439%:%
+%:%1005=440%:%
+%:%1006=440%:%
+%:%1007=441%:%
+%:%1008=441%:%
+%:%1009=442%:%
+%:%1010=442%:%
+%:%1011=443%:%
+%:%1012=443%:%
+%:%1013=444%:%
+%:%1014=444%:%
+%:%1015=445%:%
+%:%1016=445%:%
+%:%1017=446%:%
+%:%1018=446%:%
+%:%1019=447%:%
+%:%1020=447%:%
+%:%1021=448%:%
+%:%1022=448%:%
+%:%1023=449%:%
+%:%1024=449%:%
+%:%1025=450%:%
+%:%1026=450%:%
+%:%1027=451%:%
+%:%1028=451%:%
+%:%1029=452%:%
+%:%1030=452%:%
+%:%1031=452%:%
+%:%1032=453%:%
+%:%1033=453%:%
+%:%1034=454%:%
+%:%1035=454%:%
+%:%1036=455%:%
+%:%1037=455%:%
+%:%1038=456%:%
+%:%1039=456%:%
+%:%1040=457%:%
+%:%1041=457%:%
+%:%1042=458%:%
+%:%1043=458%:%
+%:%1044=458%:%
+%:%1045=459%:%
+%:%1046=459%:%
+%:%1047=460%:%
+%:%1048=460%:%
+%:%1049=461%:%
+%:%1050=461%:%
+%:%1051=462%:%
+%:%1052=462%:%
+%:%1053=463%:%
+%:%1054=463%:%
+%:%1055=464%:%
+%:%1056=464%:%
+%:%1057=465%:%
+%:%1058=465%:%
+%:%1059=465%:%
+%:%1060=466%:%
+%:%1061=466%:%
+%:%1062=467%:%
+%:%1063=467%:%
+%:%1064=468%:%
+%:%1065=468%:%
+%:%1066=469%:%
+%:%1067=469%:%
+%:%1068=470%:%
+%:%1069=470%:%
+%:%1070=471%:%
+%:%1071=471%:%
+%:%1072=472%:%
+%:%1073=472%:%
+%:%1074=473%:%
+%:%1075=473%:%
+%:%1076=474%:%
+%:%1077=474%:%
+%:%1078=475%:%
+%:%1079=475%:%
+%:%1080=476%:%
+%:%1081=476%:%
+%:%1082=477%:%
+%:%1083=477%:%
+%:%1084=478%:%
+%:%1090=478%:%
+%:%1093=479%:%
+%:%1094=480%:%
+%:%1095=481%:%
+%:%1096=482%:%
+%:%1097=482%:%
+%:%1098=483%:%
+%:%1099=484%:%
+%:%1102=485%:%
+%:%1106=485%:%
+%:%1107=485%:%
+%:%1108=486%:%
+%:%1109=486%:%
+%:%1110=487%:%
+%:%1111=487%:%
+%:%1112=488%:%
+%:%1113=488%:%
+%:%1114=488%:%
+%:%1115=489%:%
+%:%1116=489%:%
+%:%1117=489%:%
+%:%1118=489%:%
+%:%1119=490%:%
+%:%1120=490%:%
+%:%1121=490%:%
+%:%1122=491%:%
+%:%1123=491%:%
+%:%1124=491%:%
+%:%1125=492%:%
+%:%1126=492%:%
+%:%1127=492%:%
+%:%1128=493%:%
+%:%1129=493%:%
+%:%1130=494%:%
+%:%1131=494%:%
+%:%1132=495%:%
+%:%1133=495%:%
+%:%1134=496%:%
+%:%1135=496%:%
+%:%1136=497%:%
+%:%1137=497%:%
+%:%1138=498%:%
+%:%1139=498%:%
+%:%1140=498%:%
+%:%1141=499%:%
+%:%1142=499%:%
+%:%1143=499%:%
+%:%1144=500%:%
+%:%1145=500%:%
+%:%1146=501%:%
+%:%1147=501%:%
+%:%1148=502%:%
+%:%1149=502%:%
+%:%1150=503%:%
+%:%1151=503%:%
+%:%1152=503%:%
+%:%1153=504%:%
+%:%1154=504%:%
+%:%1155=504%:%
+%:%1156=505%:%
+%:%1157=505%:%
+%:%1158=505%:%
+%:%1159=506%:%
+%:%1160=506%:%
+%:%1161=506%:%
+%:%1162=507%:%
+%:%1163=507%:%
+%:%1164=507%:%
+%:%1165=508%:%
+%:%1166=508%:%
+%:%1167=509%:%
+%:%1173=509%:%
+%:%1176=510%:%
+%:%1177=511%:%
+%:%1178=512%:%
+%:%1179=512%:%
+%:%1180=513%:%
+%:%1181=514%:%
+%:%1182=515%:%
+%:%1183=516%:%
+%:%1184=517%:%
+%:%1185=517%:%
+%:%1186=518%:%
+%:%1187=519%:%
+%:%1188=520%:%
+%:%1189=521%:%
+%:%1190=521%:%
+%:%1191=522%:%
+%:%1198=523%:%
+%:%1199=523%:%
+%:%1200=524%:%
+%:%1201=524%:%
+%:%1202=524%:%
+%:%1203=525%:%
+%:%1204=525%:%
+%:%1205=526%:%
+%:%1206=526%:%
+%:%1207=527%:%
+%:%1208=527%:%
+%:%1209=528%:%
+%:%1210=528%:%
+%:%1211=528%:%
+%:%1212=529%:%
+%:%1213=529%:%
+%:%1214=530%:%
+%:%1215=530%:%
+%:%1216=530%:%
+%:%1217=531%:%
+%:%1218=531%:%
+%:%1219=532%:%
+%:%1220=532%:%
+%:%1221=533%:%
+%:%1222=533%:%
+%:%1223=533%:%
+%:%1224=534%:%
+%:%1225=534%:%
+%:%1226=534%:%
+%:%1227=535%:%
+%:%1228=535%:%
+%:%1229=535%:%
+%:%1230=536%:%
+%:%1231=536%:%
+%:%1232=536%:%
+%:%1233=537%:%
+%:%1234=537%:%
+%:%1235=537%:%
+%:%1236=538%:%
+%:%1237=538%:%
+%:%1238=539%:%
+%:%1239=539%:%
+%:%1240=539%:%
+%:%1241=540%:%
+%:%1242=540%:%
+%:%1243=541%:%
+%:%1244=541%:%
+%:%1245=542%:%
+%:%1246=542%:%
+%:%1247=543%:%
+%:%1248=543%:%
+%:%1249=544%:%
+%:%1255=544%:%
+%:%1258=545%:%
+%:%1259=546%:%
+%:%1260=547%:%
+%:%1261=547%:%
+%:%1262=548%:%
+%:%1263=549%:%
+%:%1264=550%:%
+%:%1266=552%:%
+%:%1267=553%:%
+%:%1268=554%:%
+%:%1269=555%:%
+%:%1270=556%:%
+%:%1271=557%:%
+%:%1272=557%:%
+%:%1273=558%:%
+%:%1274=559%:%
+%:%1275=560%:%
+%:%1276=561%:%
+%:%1277=562%:%
+%:%1278=563%:%
+%:%1279=564%:%
+%:%1280=565%:%
+%:%1281=566%:%
+%:%1282=567%:%
+%:%1283=567%:%
+%:%1284=568%:%
+%:%1285=569%:%
+%:%1286=570%:%
+%:%1287=571%:%
+%:%1288=572%:%
+%:%1289=573%:%
+%:%1290=574%:%
+%:%1291=575%:%
+%:%1292=576%:%
+%:%1293=576%:%
+%:%1294=577%:%
+%:%1295=578%:%
+%:%1296=579%:%
+%:%1297=580%:%
+%:%1298=581%:%
+%:%1299=582%:%
+%:%1300=583%:%
+%:%1301=584%:%
+%:%1302=584%:%
+%:%1303=585%:%
+%:%1304=586%:%
+%:%1305=587%:%
+%:%1306=588%:%
+%:%1307=589%:%
+%:%1308=590%:%
+%:%1309=591%:%
+%:%1310=591%:%
+%:%1311=592%:%
+%:%1312=593%:%
+%:%1313=594%:%
+%:%1314=595%:%
+%:%1315=596%:%
+%:%1316=597%:%
+%:%1317=598%:%
+%:%1318=599%:%
+%:%1319=600%:%
+%:%1320=600%:%
+%:%1321=601%:%
+%:%1322=602%:%
+%:%1323=603%:%
+%:%1324=604%:%
+%:%1325=605%:%
+%:%1326=606%:%
+%:%1327=606%:%
+%:%1328=607%:%
+%:%1329=608%:%
+%:%1330=609%:%
+%:%1331=610%:%
+%:%1332=610%:%
+%:%1333=611%:%
+%:%1334=612%:%
+%:%1335=612%:%
+%:%1336=613%:%
+%:%1339=614%:%
+%:%1343=614%:%
+%:%1344=614%:%
+%:%1345=615%:%
+%:%1346=615%:%
+%:%1347=616%:%
+%:%1348=616%:%
+%:%1349=617%:%
+%:%1355=617%:%
+%:%1358=618%:%
+%:%1359=619%:%
+%:%1360=620%:%
+%:%1361=621%:%
+%:%1362=622%:%
+%:%1363=623%:%
+%:%1364=624%:%
+%:%1365=625%:%
+%:%1366=625%:%
+%:%1367=626%:%
+%:%1370=627%:%
+%:%1374=627%:%
+%:%1375=627%:%
+%:%1376=628%:%
+%:%1377=628%:%
+%:%1378=629%:%
+%:%1379=629%:%
+%:%1384=629%:%
+%:%1387=630%:%
+%:%1388=631%:%
+%:%1389=631%:%
+%:%1390=632%:%
+%:%1393=633%:%
+%:%1397=633%:%
+%:%1398=633%:%
+%:%1399=634%:%
+%:%1400=634%:%
+%:%1401=635%:%
+%:%1407=635%:%
+%:%1410=636%:%
+%:%1411=637%:%
+%:%1412=637%:%
+%:%1413=638%:%
+%:%1416=639%:%
+%:%1420=639%:%
+%:%1421=639%:%
+%:%1422=640%:%
+%:%1423=640%:%
+%:%1424=641%:%
+%:%1425=641%:%
+%:%1426=642%:%
+%:%1427=642%:%
+%:%1428=643%:%
+%:%1429=643%:%
+%:%1430=644%:%
+%:%1431=644%:%
+%:%1432=645%:%
+%:%1438=645%:%
+%:%1441=646%:%
+%:%1442=647%:%
+%:%1443=647%:%
+%:%1444=648%:%
+%:%1447=649%:%
+%:%1451=649%:%
+%:%1452=649%:%
+%:%1453=650%:%
+%:%1454=650%:%
+%:%1455=651%:%
+%:%1456=651%:%
+%:%1457=652%:%
+%:%1458=652%:%
+%:%1459=652%:%
+%:%1460=653%:%
+%:%1461=653%:%
+%:%1466=653%:%
+%:%1469=654%:%
+%:%1470=655%:%
+%:%1471=655%:%
+%:%1472=656%:%
+%:%1475=657%:%
+%:%1479=657%:%
+%:%1480=657%:%
+%:%1481=658%:%
+%:%1482=658%:%
+%:%1483=659%:%
+%:%1484=659%:%
+%:%1485=660%:%
+%:%1486=660%:%
+%:%1491=660%:%
+%:%1494=661%:%
+%:%1495=662%:%
+%:%1496=662%:%
+%:%1497=663%:%
+%:%1500=664%:%
+%:%1504=664%:%
+%:%1505=664%:%
+%:%1510=664%:%
+%:%1513=665%:%
+%:%1514=666%:%
+%:%1515=666%:%
+%:%1516=667%:%
+%:%1519=668%:%
+%:%1523=668%:%
+%:%1524=668%:%
+%:%1525=669%:%
+%:%1526=669%:%
+%:%1527=670%:%
+%:%1528=670%:%
+%:%1529=671%:%
+%:%1530=671%:%
+%:%1531=672%:%
+%:%1532=672%:%
+%:%1533=673%:%
+%:%1534=673%:%
+%:%1535=674%:%
+%:%1536=674%:%
+%:%1537=675%:%
+%:%1538=675%:%
+%:%1539=676%:%
+%:%1540=676%:%
+%:%1541=677%:%
+%:%1542=677%:%
+%:%1543=678%:%
+%:%1544=678%:%
+%:%1545=679%:%
+%:%1546=679%:%
+%:%1547=680%:%
+%:%1548=680%:%
+%:%1549=681%:%
+%:%1550=681%:%
+%:%1551=682%:%
+%:%1552=682%:%
+%:%1553=683%:%
+%:%1554=683%:%
+%:%1555=684%:%
+%:%1556=684%:%
+%:%1557=685%:%
+%:%1558=685%:%
+%:%1559=686%:%
+%:%1560=686%:%
+%:%1561=686%:%
+%:%1566=686%:%
+%:%1569=687%:%
+%:%1570=688%:%
+%:%1571=688%:%
+%:%1572=689%:%
+%:%1575=690%:%
+%:%1579=690%:%
+%:%1580=690%:%
+%:%1581=691%:%
+%:%1582=691%:%
+%:%1583=692%:%
+%:%1589=692%:%
+%:%1592=693%:%
+%:%1593=694%:%
+%:%1594=695%:%
+%:%1595=696%:%
+%:%1596=696%:%
+%:%1597=697%:%
+%:%1598=698%:%
+%:%1601=699%:%
+%:%1605=699%:%
+%:%1606=699%:%
+%:%1607=700%:%
+%:%1608=700%:%
+%:%1609=701%:%
+%:%1610=701%:%
+%:%1611=702%:%
+%:%1617=702%:%
+%:%1620=703%:%
+%:%1621=704%:%
+%:%1622=704%:%
+%:%1623=705%:%
+%:%1626=706%:%
+%:%1630=706%:%
+%:%1631=706%:%
+%:%1632=707%:%
+%:%1633=707%:%
+%:%1634=708%:%
+%:%1640=708%:%
+%:%1643=709%:%
+%:%1644=710%:%
+%:%1645=711%:%
+%:%1646=711%:%
+%:%1647=712%:%
+%:%1650=713%:%
+%:%1654=713%:%
+%:%1655=713%:%
+%:%1656=714%:%
+%:%1657=714%:%
+%:%1658=715%:%
+%:%1659=715%:%
+%:%1660=716%:%
+%:%1661=716%:%
+%:%1662=717%:%
+%:%1663=717%:%
+%:%1664=718%:%
+%:%1665=718%:%
+%:%1666=719%:%
+%:%1667=719%:%
+%:%1672=719%:%
+%:%1675=720%:%
+%:%1676=721%:%
+%:%1677=722%:%
+%:%1678=722%:%
+%:%1679=723%:%
+%:%1680=724%:%
+%:%1683=725%:%
+%:%1687=725%:%
+%:%1688=725%:%
+%:%1693=725%:%
+%:%1696=726%:%
+%:%1697=727%:%
+%:%1698=728%:%
+%:%1699=728%:%
+%:%1700=729%:%
+%:%1703=730%:%
+%:%1707=730%:%
+%:%1708=730%:%
+%:%1713=730%:%
+%:%1716=731%:%
+%:%1717=732%:%
+%:%1718=733%:%
+%:%1719=733%:%
+%:%1720=734%:%
+%:%1721=735%:%
+%:%1724=736%:%
+%:%1728=736%:%
+%:%1729=736%:%
+%:%1730=737%:%
+%:%1731=737%:%
+%:%1732=738%:%
+%:%1733=738%:%
+%:%1734=739%:%
+%:%1735=739%:%
+%:%1736=740%:%
+%:%1737=740%:%
+%:%1742=740%:%
+%:%1745=741%:%
+%:%1746=742%:%
+%:%1747=742%:%
+%:%1748=743%:%
+%:%1749=744%:%
+%:%1752=745%:%
+%:%1756=745%:%
+%:%1757=745%:%
+%:%1758=746%:%
+%:%1759=746%:%
+%:%1760=747%:%
+%:%1761=747%:%
+%:%1762=748%:%
+%:%1763=748%:%
+%:%1764=749%:%
+%:%1765=749%:%
+%:%1770=749%:%
+%:%1773=750%:%
+%:%1774=751%:%
+%:%1775=751%:%
+%:%1776=752%:%
+%:%1779=753%:%
+%:%1783=753%:%
+%:%1784=753%:%
+%:%1785=754%:%
+%:%1786=754%:%
+%:%1787=755%:%
+%:%1788=755%:%
+%:%1789=756%:%
+%:%1795=756%:%
+%:%1798=757%:%
+%:%1799=758%:%
+%:%1800=759%:%
+%:%1801=760%:%
+%:%1802=761%:%
+%:%1803=762%:%
+%:%1804=762%:%
+%:%1805=763%:%
+%:%1806=764%:%
+%:%1807=765%:%
+%:%1808=766%:%
+%:%1809=767%:%
+%:%1810=768%:%
+%:%1811=769%:%
+%:%1812=770%:%
+%:%1813=770%:%
+%:%1814=771%:%
+%:%1817=772%:%
+%:%1821=772%:%
+%:%1822=772%:%
+%:%1823=773%:%
+%:%1824=773%:%
+%:%1825=774%:%
+%:%1831=774%:%
+%:%1834=775%:%
+%:%1835=776%:%
+%:%1836=776%:%
+%:%1837=777%:%
+%:%1840=778%:%
+%:%1844=778%:%
+%:%1845=778%:%
+%:%1846=779%:%
+%:%1847=779%:%
+%:%1848=780%:%
+%:%1849=780%:%
+%:%1854=780%:%
+%:%1857=781%:%
+%:%1858=782%:%
+%:%1859=782%:%
+%:%1860=783%:%
+%:%1863=784%:%
+%:%1867=784%:%
+%:%1868=784%:%
+%:%1869=785%:%
+%:%1875=785%:%
+%:%1878=786%:%
+%:%1879=787%:%
+%:%1880=788%:%
+%:%1881=789%:%
+%:%1882=790%:%
+%:%1883=791%:%
+%:%1884=792%:%
+%:%1885=793%:%
+%:%1886=794%:%
+%:%1887=794%:%
+%:%1888=795%:%
+%:%1889=796%:%
+%:%1892=797%:%
+%:%1896=797%:%
+%:%1897=797%:%
+%:%1898=798%:%
+%:%1899=798%:%
+%:%1900=799%:%
+%:%1901=799%:%
+%:%1902=800%:%
+%:%1903=800%:%
+%:%1904=801%:%
+%:%1905=801%:%
+%:%1906=802%:%
+%:%1912=802%:%
+%:%1915=803%:%
+%:%1916=804%:%
+%:%1917=805%:%
+%:%1918=806%:%
+%:%1919=806%:%
+%:%1920=807%:%
+%:%1923=808%:%
+%:%1927=808%:%
+%:%1928=808%:%
+%:%1929=808%:%
+%:%1934=808%:%
+%:%1937=809%:%
+%:%1938=810%:%
+%:%1939=811%:%
+%:%1940=812%:%
+%:%1941=813%:%
+%:%1942=814%:%
+%:%1943=814%:%
+%:%1944=815%:%
+%:%1947=816%:%
+%:%1951=816%:%
+%:%1952=816%:%
+%:%1953=817%:%
+%:%1954=817%:%
+%:%1955=818%:%
+%:%1956=818%:%
+%:%1957=819%:%
+%:%1958=819%:%
+%:%1959=820%:%
+%:%1960=820%:%
+%:%1961=821%:%
+%:%1962=821%:%
+%:%1963=822%:%
+%:%1964=822%:%
+%:%1965=823%:%
+%:%1966=823%:%
+%:%1967=824%:%
+%:%1973=824%:%
+%:%1976=825%:%
+%:%1977=826%:%
+%:%1978=826%:%
+%:%1979=827%:%
+%:%1980=828%:%
+%:%1981=829%:%
+%:%1982=830%:%
+%:%1983=831%:%
+%:%1984=832%:%
+%:%1985=832%:%
+%:%1986=833%:%
+%:%1987=834%:%
+%:%1990=835%:%
+%:%1994=835%:%
+%:%1995=835%:%
+%:%1996=836%:%
+%:%1997=836%:%
+%:%1998=837%:%
+%:%1999=837%:%
+%:%2000=838%:%
+%:%2006=838%:%
+%:%2009=839%:%
+%:%2010=840%:%
+%:%2011=841%:%
+%:%2012=842%:%
+%:%2013=842%:%
+%:%2014=843%:%
+%:%2015=844%:%
+%:%2018=845%:%
+%:%2022=845%:%
+%:%2023=845%:%
+%:%2024=846%:%
+%:%2025=846%:%
+%:%2026=847%:%
+%:%2027=847%:%
+%:%2028=848%:%
+%:%2029=848%:%
+%:%2030=849%:%
+%:%2031=849%:%
+%:%2032=850%:%
+%:%2033=850%:%
+%:%2034=851%:%
+%:%2035=851%:%
+%:%2036=852%:%
+%:%2037=852%:%
+%:%2038=853%:%
+%:%2039=853%:%
+%:%2040=854%:%
+%:%2041=854%:%
+%:%2042=855%:%
+%:%2043=855%:%
+%:%2044=856%:%
+%:%2045=856%:%
+%:%2046=857%:%
+%:%2047=857%:%
+%:%2048=858%:%
+%:%2054=858%:%
+%:%2057=859%:%
+%:%2058=860%:%
+%:%2059=860%:%
+%:%2060=861%:%
+%:%2061=862%:%
+%:%2064=863%:%
+%:%2068=863%:%
+%:%2069=863%:%
+%:%2070=864%:%
+%:%2071=864%:%
+%:%2072=865%:%
+%:%2073=865%:%
+%:%2074=866%:%
+%:%2075=866%:%
+%:%2080=866%:%
+%:%2083=867%:%
+%:%2084=868%:%
+%:%2085=869%:%
+%:%2086=869%:%
+%:%2087=870%:%
+%:%2088=871%:%
+%:%2091=872%:%
+%:%2095=872%:%
+%:%2096=872%:%
+%:%2097=873%:%
+%:%2098=873%:%
+%:%2099=874%:%
+%:%2100=874%:%
+%:%2101=875%:%
+%:%2102=875%:%
+%:%2103=876%:%
+%:%2104=876%:%
+%:%2105=877%:%
+%:%2106=877%:%
+%:%2107=878%:%
+%:%2108=878%:%
+%:%2109=879%:%
+%:%2110=879%:%
+%:%2111=880%:%
+%:%2112=880%:%
+%:%2113=881%:%
+%:%2114=881%:%
+%:%2115=882%:%
+%:%2116=882%:%
+%:%2117=883%:%
+%:%2118=883%:%
+%:%2119=884%:%
+%:%2120=884%:%
+%:%2121=885%:%
+%:%2122=885%:%
+%:%2123=886%:%
+%:%2124=886%:%
+%:%2125=887%:%
+%:%2126=887%:%
+%:%2127=888%:%
+%:%2128=888%:%
+%:%2129=889%:%
+%:%2130=889%:%
+%:%2131=890%:%
+%:%2132=890%:%
+%:%2133=891%:%
+%:%2134=891%:%
+%:%2135=892%:%
+%:%2136=892%:%
+%:%2137=893%:%
+%:%2138=893%:%
+%:%2139=894%:%
+%:%2140=894%:%
+%:%2141=895%:%
+%:%2142=895%:%
+%:%2143=896%:%
+%:%2144=896%:%
+%:%2145=897%:%
+%:%2146=897%:%
+%:%2147=898%:%
+%:%2148=898%:%
+%:%2149=899%:%
+%:%2150=899%:%
+%:%2151=900%:%
+%:%2152=900%:%
+%:%2153=901%:%
+%:%2154=901%:%
+%:%2155=902%:%
+%:%2156=902%:%
+%:%2157=903%:%
+%:%2158=903%:%
+%:%2159=904%:%
+%:%2160=904%:%
+%:%2161=905%:%
+%:%2162=905%:%
+%:%2163=906%:%
+%:%2164=906%:%
+%:%2165=907%:%
+%:%2166=907%:%
+%:%2167=908%:%
+%:%2168=908%:%
+%:%2169=909%:%
+%:%2170=909%:%
+%:%2171=910%:%
+%:%2172=910%:%
+%:%2173=911%:%
+%:%2174=911%:%
+%:%2175=912%:%
+%:%2176=912%:%
+%:%2177=913%:%
+%:%2178=913%:%
+%:%2179=914%:%
+%:%2180=914%:%
+%:%2181=915%:%
+%:%2182=915%:%
+%:%2183=916%:%
+%:%2184=916%:%
+%:%2185=917%:%
+%:%2186=917%:%
+%:%2187=918%:%
+%:%2188=918%:%
+%:%2189=919%:%
+%:%2190=919%:%
+%:%2191=920%:%
+%:%2192=920%:%
+%:%2193=921%:%
+%:%2194=921%:%
+%:%2195=922%:%
+%:%2196=922%:%
+%:%2197=923%:%
+%:%2198=923%:%
+%:%2199=924%:%
+%:%2200=924%:%
+%:%2201=925%:%
+%:%2202=925%:%
+%:%2203=926%:%
+%:%2204=926%:%
+%:%2205=927%:%
+%:%2206=927%:%
+%:%2207=928%:%
+%:%2208=928%:%
+%:%2209=929%:%
+%:%2210=929%:%
+%:%2211=930%:%
+%:%2212=930%:%
+%:%2213=931%:%
+%:%2214=931%:%
+%:%2215=931%:%
+%:%2220=931%:%
+%:%2223=932%:%
+%:%2224=933%:%
+%:%2225=934%:%
+%:%2226=935%:%
+%:%2227=936%:%
+%:%2228=936%:%
+%:%2229=937%:%
+%:%2232=938%:%
+%:%2236=938%:%
+%:%2237=938%:%
+%:%2238=939%:%
+%:%2239=939%:%
+%:%2240=940%:%
+%:%2246=940%:%
+%:%2249=941%:%
+%:%2250=942%:%
+%:%2251=943%:%
+%:%2252=943%:%
+%:%2253=944%:%
+%:%2254=945%:%
+%:%2255=946%:%
+%:%2256=947%:%
+%:%2257=948%:%
+%:%2258=949%:%
+%:%2259=950%:%
+%:%2261=952%:%
+%:%2262=953%:%
+%:%2263=954%:%
+%:%2264=955%:%
+%:%2265=956%:%
+%:%2266=957%:%
+%:%2267=958%:%
+%:%2268=959%:%
+%:%2269=960%:%
+%:%2270=961%:%
+%:%2271=962%:%
+%:%2272=963%:%
+%:%2273=964%:%
+%:%2274=965%:%
+%:%2275=966%:%
+%:%2276=967%:%
+%:%2277=967%:%
+%:%2278=968%:%
+%:%2279=969%:%
+%:%2282=970%:%
+%:%2286=970%:%
+%:%2287=970%:%
+%:%2288=971%:%
+%:%2294=971%:%
+%:%2297=972%:%
+%:%2298=973%:%
+%:%2299=974%:%
+%:%2300=975%:%
+%:%2301=975%:%
+%:%2302=976%:%
+%:%2305=977%:%
+%:%2309=977%:%
+%:%2310=977%:%
+%:%2311=978%:%
+%:%2312=978%:%
+%:%2313=979%:%
+%:%2314=979%:%
+%:%2315=980%:%
+%:%2316=980%:%
+%:%2317=981%:%
+%:%2318=981%:%
+%:%2319=982%:%
+%:%2325=982%:%
+%:%2328=983%:%
+%:%2329=984%:%
+%:%2330=985%:%
+%:%2331=986%:%
+%:%2332=986%:%
+%:%2333=987%:%
+%:%2334=988%:%
+%:%2337=989%:%
+%:%2341=989%:%
+%:%2342=989%:%
+%:%2343=990%:%
+%:%2344=990%:%
+%:%2345=991%:%
+%:%2351=991%:%
+%:%2354=992%:%
+%:%2355=993%:%
+%:%2356=994%:%
+%:%2357=995%:%
+%:%2358=996%:%
+%:%2359=997%:%
+%:%2360=998%:%
+%:%2361=999%:%
+%:%2362=1000%:%
+%:%2363=1000%:%
+%:%2364=1001%:%
+%:%2365=1002%:%
+%:%2366=1003%:%
+%:%2367=1004%:%
+%:%2368=1005%:%
+%:%2369=1006%:%
+%:%2370=1007%:%
+%:%2371=1008%:%
+%:%2372=1009%:%
+%:%2373=1010%:%
+%:%2374=1011%:%
+%:%2375=1012%:%
+%:%2376=1014%:%
+%:%2377=1015%:%
+%:%2378=1016%:%
+%:%2379=1017%:%
+%:%2380=1018%:%
+%:%2381=1019%:%
+%:%2382=1020%:%
+%:%2383=1021%:%
+%:%2384=1021%:%
+%:%2385=1022%:%
+%:%2386=1023%:%
+%:%2387=1024%:%
+%:%2388=1025%:%
+%:%2389=1026%:%
+%:%2390=1026%:%
+%:%2391=1027%:%
+%:%2392=1028%:%
+%:%2393=1029%:%
+%:%2394=1032%:%
+%:%2395=1033%:%
+%:%2396=1034%:%
+%:%2397=1035%:%
+%:%2398=1036%:%
+%:%2399=1036%:%
+%:%2402=1037%:%
+%:%2406=1037%:%
+%:%2407=1037%:%
+%:%2408=1037%:%
+%:%2413=1037%:%
+%:%2416=1038%:%
+%:%2417=1039%:%
+%:%2418=1039%:%
+%:%2419=1040%:%
+%:%2420=1041%:%
+%:%2423=1042%:%
+%:%2427=1042%:%
+%:%2428=1042%:%
+%:%2429=1043%:%
+%:%2430=1043%:%
+%:%2431=1044%:%
+%:%2432=1044%:%
+%:%2433=1045%:%
+%:%2439=1045%:%
+%:%2442=1046%:%
+%:%2443=1047%:%
+%:%2444=1048%:%
+%:%2445=1048%:%
+%:%2448=1049%:%
+%:%2452=1049%:%
+%:%2453=1049%:%
+%:%2458=1049%:%
+%:%2461=1050%:%
+%:%2462=1051%:%
+%:%2463=1052%:%
+%:%2464=1052%:%
+%:%2467=1053%:%
+%:%2471=1053%:%
+%:%2472=1053%:%
+%:%2473=1054%:%
+%:%2474=1054%:%
+%:%2475=1055%:%
+%:%2476=1055%:%
+%:%2481=1055%:%
+%:%2484=1056%:%
+%:%2485=1057%:%
+%:%2486=1058%:%
+%:%2487=1058%:%
+%:%2490=1059%:%
+%:%2494=1059%:%
+%:%2495=1059%:%
+%:%2496=1060%:%
+%:%2497=1060%:%
+%:%2498=1061%:%
+%:%2499=1061%:%
+%:%2500=1061%:%
+%:%2505=1061%:%
+%:%2508=1062%:%
+%:%2509=1063%:%
+%:%2510=1064%:%
+%:%2511=1065%:%
+%:%2512=1065%:%
+%:%2515=1066%:%
+%:%2516=1067%:%
+%:%2520=1067%:%
+%:%2521=1067%:%
+%:%2522=1068%:%
+%:%2523=1068%:%
+%:%2524=1069%:%
+%:%2525=1069%:%
+%:%2526=1070%:%
+%:%2527=1070%:%
+%:%2528=1071%:%
+%:%2529=1071%:%
+%:%2530=1072%:%
+%:%2531=1072%:%
+%:%2532=1073%:%
+%:%2533=1073%:%
+%:%2534=1074%:%
+%:%2535=1074%:%
+%:%2536=1075%:%
+%:%2537=1075%:%
+%:%2538=1076%:%
+%:%2539=1076%:%
+%:%2540=1077%:%
+%:%2546=1077%:%
+%:%2549=1078%:%
+%:%2550=1079%:%
+%:%2551=1080%:%
+%:%2552=1080%:%
+%:%2555=1081%:%
+%:%2559=1081%:%
+%:%2560=1081%:%
+%:%2565=1081%:%
+%:%2568=1082%:%
+%:%2569=1083%:%
+%:%2570=1084%:%
+%:%2571=1084%:%
+%:%2574=1085%:%
+%:%2578=1085%:%
+%:%2579=1085%:%
+%:%2580=1086%:%
+%:%2581=1086%:%
+%:%2582=1087%:%
+%:%2583=1087%:%
+%:%2584=1088%:%
+%:%2585=1088%:%
+%:%2586=1089%:%
+%:%2587=1089%:%
+%:%2588=1090%:%
+%:%2589=1090%:%
+%:%2590=1091%:%
+%:%2591=1091%:%
+%:%2592=1092%:%
+%:%2593=1092%:%
+%:%2594=1093%:%
+%:%2595=1093%:%
+%:%2596=1094%:%
+%:%2597=1094%:%
+%:%2598=1095%:%
+%:%2599=1095%:%
+%:%2604=1095%:%
+%:%2607=1096%:%
+%:%2608=1097%:%
+%:%2609=1097%:%
+%:%2612=1098%:%
+%:%2616=1098%:%
+%:%2617=1098%:%
+%:%2618=1099%:%
+%:%2619=1099%:%
+%:%2620=1100%:%
+%:%2621=1100%:%
+%:%2622=1100%:%
+%:%2627=1100%:%
+%:%2630=1101%:%
+%:%2631=1102%:%
+%:%2632=1103%:%
+%:%2633=1103%:%
+%:%2636=1104%:%
+%:%2640=1104%:%
+%:%2641=1104%:%
+%:%2642=1105%:%
+%:%2643=1105%:%
+%:%2644=1106%:%
+%:%2645=1107%:%
+%:%2646=1107%:%
+%:%2651=1107%:%
+%:%2654=1108%:%
+%:%2655=1109%:%
+%:%2656=1110%:%
+%:%2657=1111%:%
+%:%2658=1111%:%
+%:%2661=1112%:%
+%:%2665=1112%:%
+%:%2666=1112%:%
+%:%2667=1113%:%
+%:%2668=1113%:%
+%:%2669=1114%:%
+%:%2670=1114%:%
+%:%2675=1114%:%
+%:%2678=1115%:%
+%:%2679=1116%:%
+%:%2680=1116%:%
+%:%2683=1117%:%
+%:%2684=1118%:%
+%:%2688=1118%:%
+%:%2689=1118%:%
+%:%2690=1119%:%
+%:%2691=1119%:%
+%:%2692=1120%:%
+%:%2693=1120%:%
+%:%2694=1121%:%
+%:%2695=1121%:%
+%:%2700=1121%:%
+%:%2703=1122%:%
+%:%2704=1123%:%
+%:%2705=1124%:%
+%:%2706=1124%:%
+%:%2709=1125%:%
+%:%2713=1125%:%
+%:%2714=1125%:%
+%:%2715=1126%:%
+%:%2716=1126%:%
+%:%2717=1127%:%
+%:%2718=1127%:%
+%:%2719=1128%:%
+%:%2720=1128%:%
+%:%2721=1129%:%
+%:%2722=1129%:%
+%:%2723=1130%:%
+%:%2724=1130%:%
+%:%2725=1131%:%
+%:%2726=1131%:%
+%:%2727=1131%:%
+%:%2728=1132%:%
+%:%2729=1132%:%
+%:%2730=1133%:%
+%:%2731=1133%:%
+%:%2732=1134%:%
+%:%2733=1134%:%
+%:%2734=1135%:%
+%:%2740=1135%:%
+%:%2743=1136%:%
+%:%2744=1137%:%
+%:%2745=1137%:%
+%:%2746=1138%:%
+%:%2747=1139%:%
+%:%2748=1140%:%
+%:%2749=1141%:%
+%:%2750=1142%:%
+%:%2751=1143%:%
+%:%2752=1144%:%
+%:%2753=1145%:%
+%:%2754=1146%:%
+%:%2755=1147%:%
+%:%2756=1148%:%
+%:%2757=1148%:%
+%:%2760=1149%:%
+%:%2764=1149%:%
+%:%2765=1149%:%
+%:%2770=1149%:%
+%:%2773=1150%:%
+%:%2774=1151%:%
+%:%2775=1151%:%
+%:%2778=1152%:%
+%:%2782=1152%:%
+%:%2783=1152%:%
+%:%2784=1153%:%
+%:%2785=1153%:%
+%:%2786=1154%:%
+%:%2787=1154%:%
+%:%2788=1155%:%
+%:%2789=1156%:%
+%:%2790=1156%:%
+%:%2791=1157%:%
+%:%2792=1158%:%
+%:%2793=1159%:%
+%:%2794=1159%:%
+%:%2795=1159%:%
+%:%2796=1160%:%
+%:%2797=1160%:%
+%:%2798=1161%:%
+%:%2799=1161%:%
+%:%2800=1162%:%
+%:%2801=1163%:%
+%:%2802=1163%:%
+%:%2803=1164%:%
+%:%2804=1164%:%
+%:%2805=1165%:%
+%:%2806=1165%:%
+%:%2807=1166%:%
+%:%2808=1166%:%
+%:%2809=1167%:%
+%:%2810=1167%:%
+%:%2811=1168%:%
+%:%2812=1169%:%
+%:%2813=1169%:%
+%:%2814=1170%:%
+%:%2815=1171%:%
+%:%2816=1171%:%
+%:%2817=1171%:%
+%:%2818=1172%:%
+%:%2819=1172%:%
+%:%2820=1173%:%
+%:%2821=1174%:%
+%:%2822=1174%:%
+%:%2823=1175%:%
+%:%2824=1176%:%
+%:%2825=1176%:%
+%:%2830=1176%:%
+%:%2833=1177%:%
+%:%2834=1178%:%
+%:%2835=1179%:%
+%:%2836=1179%:%
+%:%2839=1180%:%
+%:%2843=1180%:%
+%:%2844=1180%:%
+%:%2849=1180%:%
+%:%2852=1181%:%
+%:%2853=1182%:%
+%:%2854=1183%:%
+%:%2855=1183%:%
+%:%2858=1184%:%
+%:%2862=1184%:%
+%:%2863=1184%:%
+%:%2864=1185%:%
+%:%2865=1185%:%
+%:%2866=1186%:%
+%:%2867=1186%:%
+%:%2868=1187%:%
+%:%2869=1187%:%
+%:%2870=1188%:%
+%:%2871=1188%:%
+%:%2872=1189%:%
+%:%2873=1189%:%
+%:%2874=1190%:%
+%:%2875=1190%:%
+%:%2876=1190%:%
+%:%2877=1191%:%
+%:%2878=1191%:%
+%:%2879=1192%:%
+%:%2880=1192%:%
+%:%2881=1193%:%
+%:%2882=1193%:%
+%:%2883=1194%:%
+%:%2884=1194%:%
+%:%2885=1195%:%
+%:%2886=1195%:%
+%:%2887=1195%:%
+%:%2888=1196%:%
+%:%2889=1196%:%
+%:%2890=1197%:%
+%:%2891=1197%:%
+%:%2892=1198%:%
+%:%2893=1198%:%
+%:%2894=1199%:%
+%:%2895=1199%:%
+%:%2896=1200%:%
+%:%2897=1200%:%
+%:%2898=1201%:%
+%:%2904=1201%:%
+%:%2907=1202%:%
+%:%2908=1203%:%
+%:%2909=1203%:%
+%:%2912=1204%:%
+%:%2916=1204%:%
+%:%2917=1204%:%
+%:%2918=1205%:%
+%:%2919=1205%:%
+%:%2920=1206%:%
+%:%2921=1206%:%
+%:%2922=1207%:%
+%:%2923=1207%:%
+%:%2924=1208%:%
+%:%2925=1209%:%
+%:%2926=1210%:%
+%:%2927=1211%:%
+%:%2928=1211%:%
+%:%2929=1212%:%
+%:%2930=1212%:%
+%:%2931=1213%:%
+%:%2932=1213%:%
+%:%2933=1214%:%
+%:%2934=1214%:%
+%:%2935=1215%:%
+%:%2936=1216%:%
+%:%2937=1216%:%
+%:%2938=1217%:%
+%:%2939=1218%:%
+%:%2940=1218%:%
+%:%2941=1219%:%
+%:%2942=1220%:%
+%:%2943=1220%:%
+%:%2944=1221%:%
+%:%2945=1221%:%
+%:%2946=1222%:%
+%:%2947=1222%:%
+%:%2948=1223%:%
+%:%2949=1223%:%
+%:%2950=1224%:%
+%:%2951=1224%:%
+%:%2956=1224%:%
+%:%2959=1225%:%
+%:%2960=1226%:%
+%:%2961=1226%:%
+%:%2962=1227%:%
+%:%2965=1228%:%
+%:%2969=1228%:%
+%:%2970=1228%:%
+%:%2971=1229%:%
+%:%2972=1229%:%
+%:%2973=1230%:%
+%:%2974=1230%:%
+%:%2975=1230%:%
+%:%2976=1231%:%
+%:%2977=1231%:%
+%:%2978=1231%:%
+%:%2983=1231%:%
+%:%2986=1232%:%
+%:%2987=1233%:%
+%:%2988=1234%:%
+%:%2989=1234%:%
+%:%2992=1235%:%
+%:%2996=1235%:%
+%:%2997=1235%:%
+%:%2998=1236%:%
+%:%3004=1236%:%
+%:%3007=1237%:%
+%:%3008=1238%:%
+%:%3009=1238%:%
+%:%3010=1239%:%
+%:%3011=1240%:%
+%:%3012=1241%:%
+%:%3014=1243%:%
+%:%3015=1244%:%
+%:%3016=1245%:%
+%:%3017=1245%:%
+%:%3020=1246%:%
+%:%3024=1246%:%
+%:%3025=1246%:%
+%:%3026=1247%:%
+%:%3032=1247%:%
+%:%3035=1248%:%
+%:%3036=1249%:%
+%:%3037=1249%:%
+%:%3040=1250%:%
+%:%3044=1250%:%
+%:%3045=1250%:%
+%:%3046=1250%:%
+%:%3051=1250%:%
+%:%3054=1251%:%
+%:%3055=1252%:%
+%:%3056=1252%:%
+%:%3059=1253%:%
+%:%3063=1253%:%
+%:%3064=1253%:%
+%:%3065=1254%:%
+%:%3066=1254%:%
+%:%3071=1254%:%
+%:%3074=1255%:%
+%:%3075=1256%:%
+%:%3076=1256%:%
+%:%3079=1257%:%
+%:%3083=1257%:%
+%:%3084=1257%:%
+%:%3085=1258%:%
+%:%3086=1258%:%
+%:%3087=1259%:%
+%:%3088=1259%:%
+%:%3089=1260%:%
+%:%3090=1260%:%
+%:%3091=1261%:%
+%:%3092=1261%:%
+%:%3093=1262%:%
+%:%3094=1262%:%
+%:%3095=1263%:%
+%:%3096=1263%:%
+%:%3097=1264%:%
+%:%3098=1264%:%
+%:%3099=1265%:%
+%:%3100=1265%:%
+%:%3101=1266%:%
+%:%3102=1266%:%
+%:%3103=1267%:%
+%:%3104=1267%:%
+%:%3105=1268%:%
+%:%3106=1268%:%
+%:%3117=1279%:%
+%:%3118=1280%:%
+%:%3119=1280%:%
+%:%3120=1281%:%
+%:%3121=1281%:%
+%:%3122=1281%:%
+%:%3123=1282%:%
+%:%3124=1282%:%
+%:%3125=1282%:%
+%:%3126=1283%:%
+%:%3127=1283%:%
+%:%3128=1284%:%
+%:%3129=1284%:%
+%:%3130=1285%:%
+%:%3131=1286%:%
+%:%3132=1286%:%
+%:%3133=1287%:%
+%:%3134=1287%:%
+%:%3135=1288%:%
+%:%3136=1289%:%
+%:%3137=1289%:%
+%:%3138=1290%:%
+%:%3139=1290%:%
+%:%3140=1291%:%
+%:%3141=1292%:%
+%:%3142=1292%:%
+%:%3143=1293%:%
+%:%3144=1293%:%
+%:%3145=1294%:%
+%:%3146=1294%:%
+%:%3151=1294%:%
+%:%3154=1295%:%
+%:%3155=1296%:%
+%:%3156=1297%:%
+%:%3157=1298%:%
+%:%3158=1298%:%
+%:%3161=1299%:%
+%:%3165=1299%:%
+%:%3166=1299%:%
+%:%3167=1300%:%
+%:%3168=1300%:%
+%:%3169=1301%:%
+%:%3170=1301%:%
+%:%3171=1302%:%
+%:%3172=1302%:%
+%:%3173=1303%:%
+%:%3174=1303%:%
+%:%3179=1303%:%
+%:%3182=1304%:%
+%:%3183=1305%:%
+%:%3184=1306%:%
+%:%3185=1307%:%
+%:%3186=1308%:%
+%:%3187=1309%:%
+%:%3188=1310%:%
+%:%3189=1310%:%
+%:%3192=1311%:%
+%:%3196=1311%:%
+%:%3197=1311%:%
+%:%3198=1312%:%
+%:%3199=1312%:%
+%:%3200=1313%:%
+%:%3201=1313%:%
+%:%3202=1314%:%
+%:%3203=1314%:%
+%:%3204=1315%:%
+%:%3205=1315%:%
+%:%3206=1315%:%
+%:%3207=1316%:%
+%:%3208=1316%:%
+%:%3209=1317%:%
+%:%3210=1317%:%
+%:%3211=1318%:%
+%:%3212=1318%:%
+%:%3213=1319%:%
+%:%3214=1319%:%
+%:%3215=1320%:%
+%:%3216=1320%:%
+%:%3217=1321%:%
+%:%3218=1321%:%
+%:%3219=1322%:%
+%:%3220=1322%:%
+%:%3221=1323%:%
+%:%3222=1323%:%
+%:%3223=1323%:%
+%:%3224=1324%:%
+%:%3225=1324%:%
+%:%3226=1325%:%
+%:%3227=1325%:%
+%:%3228=1326%:%
+%:%3229=1326%:%
+%:%3230=1327%:%
+%:%3231=1328%:%
+%:%3232=1328%:%
+%:%3233=1329%:%
+%:%3234=1330%:%
+%:%3235=1331%:%
+%:%3236=1331%:%
+%:%3237=1332%:%
+%:%3238=1332%:%
+%:%3239=1333%:%
+%:%3240=1333%:%
+%:%3241=1334%:%
+%:%3242=1334%:%
+%:%3243=1335%:%
+%:%3244=1336%:%
+%:%3245=1337%:%
+%:%3246=1337%:%
+%:%3247=1338%:%
+%:%3248=1339%:%
+%:%3249=1339%:%
+%:%3250=1340%:%
+%:%3251=1341%:%
+%:%3252=1341%:%
+%:%3253=1342%:%
+%:%3254=1343%:%
+%:%3255=1343%:%
+%:%3256=1343%:%
+%:%3257=1344%:%
+%:%3258=1345%:%
+%:%3259=1345%:%
+%:%3260=1345%:%
+%:%3265=1345%:%
+%:%3268=1346%:%
+%:%3269=1347%:%
+%:%3270=1348%:%
+%:%3271=1348%:%
+%:%3274=1349%:%
+%:%3278=1349%:%
+%:%3279=1349%:%
+%:%3280=1350%:%
+%:%3281=1350%:%
+%:%3282=1351%:%
+%:%3283=1351%:%
+%:%3284=1352%:%
+%:%3285=1352%:%
+%:%3286=1353%:%
+%:%3287=1353%:%
+%:%3288=1354%:%
+%:%3289=1354%:%
+%:%3290=1355%:%
+%:%3291=1355%:%
+%:%3292=1356%:%
+%:%3293=1356%:%
+%:%3294=1357%:%
+%:%3295=1357%:%
+%:%3296=1358%:%
+%:%3297=1358%:%
+%:%3298=1359%:%
+%:%3299=1359%:%
+%:%3300=1360%:%
+%:%3301=1361%:%
+%:%3302=1361%:%
+%:%3303=1362%:%
+%:%3304=1363%:%
+%:%3305=1364%:%
+%:%3306=1364%:%
+%:%3311=1364%:%
+%:%3314=1365%:%
+%:%3315=1366%:%
+%:%3316=1366%:%
+%:%3319=1367%:%
+%:%3323=1367%:%
+%:%3324=1367%:%
+%:%3325=1368%:%
+%:%3326=1368%:%
+%:%3327=1369%:%
+%:%3328=1369%:%
+%:%3329=1370%:%
+%:%3335=1370%:%
+%:%3338=1371%:%
+%:%3339=1372%:%
+%:%3340=1373%:%
+%:%3341=1373%:%
+%:%3344=1374%:%
+%:%3348=1374%:%
+%:%3349=1374%:%
+%:%3350=1375%:%
+%:%3351=1375%:%
+%:%3352=1376%:%
+%:%3353=1376%:%
+%:%3354=1377%:%
+%:%3355=1377%:%
+%:%3356=1378%:%
+%:%3357=1378%:%
+%:%3358=1379%:%
+%:%3364=1379%:%
+%:%3367=1380%:%
+%:%3368=1381%:%
+%:%3369=1381%:%
+%:%3370=1382%:%
+%:%3373=1383%:%
+%:%3377=1383%:%
+%:%3378=1383%:%
+%:%3379=1384%:%
+%:%3380=1384%:%
+%:%3381=1385%:%
+%:%3387=1385%:%
+%:%3390=1386%:%
+%:%3391=1387%:%
+%:%3392=1388%:%
+%:%3393=1389%:%
+%:%3394=1389%:%
+%:%3395=1390%:%
+%:%3398=1391%:%
+%:%3402=1391%:%
+%:%3403=1391%:%
+%:%3404=1391%:%
+%:%3409=1391%:%
+%:%3412=1392%:%
+%:%3413=1393%:%
+%:%3414=1393%:%
+%:%3417=1394%:%
+%:%3421=1394%:%
+%:%3422=1394%:%
+%:%3423=1395%:%
+%:%3424=1396%:%
+%:%3425=1396%:%
+%:%3426=1396%:%
+%:%3427=1397%:%
+%:%3428=1397%:%
+%:%3429=1398%:%
+%:%3430=1399%:%
+%:%3431=1399%:%
+%:%3432=1400%:%
+%:%3433=1400%:%
+%:%3434=1401%:%
+%:%3435=1402%:%
+%:%3436=1402%:%
+%:%3437=1402%:%
+%:%3438=1403%:%
+%:%3439=1403%:%
+%:%3440=1404%:%
+%:%3441=1405%:%
+%:%3442=1405%:%
+%:%3443=1406%:%
+%:%3444=1407%:%
+%:%3445=1407%:%
+%:%3446=1407%:%
+%:%3451=1407%:%
+%:%3454=1408%:%
+%:%3455=1409%:%
+%:%3456=1410%:%
+%:%3457=1411%:%
+%:%3458=1411%:%
+%:%3459=1412%:%
+%:%3462=1413%:%
+%:%3466=1413%:%
+%:%3467=1413%:%
+%:%3468=1413%:%
+%:%3473=1413%:%
+%:%3476=1414%:%
+%:%3477=1415%:%
+%:%3478=1415%:%
+%:%3481=1416%:%
+%:%3485=1416%:%
+%:%3486=1416%:%
+%:%3487=1417%:%
+%:%3488=1417%:%
+%:%3489=1418%:%
+%:%3490=1418%:%
+%:%3491=1419%:%
+%:%3492=1419%:%
+%:%3493=1420%:%
+%:%3494=1420%:%
+%:%3495=1421%:%
+%:%3496=1421%:%
+%:%3497=1422%:%
+%:%3498=1422%:%
+%:%3499=1423%:%
+%:%3500=1423%:%
+%:%3501=1424%:%
+%:%3502=1424%:%
+%:%3503=1425%:%
+%:%3504=1425%:%
+%:%3505=1426%:%
+%:%3506=1426%:%
+%:%3507=1427%:%
+%:%3508=1427%:%
+%:%3509=1428%:%
+%:%3510=1428%:%
+%:%3511=1429%:%
+%:%3512=1429%:%
+%:%3513=1430%:%
+%:%3514=1430%:%
+%:%3515=1431%:%
+%:%3516=1431%:%
+%:%3517=1432%:%
+%:%3518=1432%:%
+%:%3519=1433%:%
+%:%3520=1434%:%
+%:%3521=1434%:%
+%:%3522=1435%:%
+%:%3523=1436%:%
+%:%3524=1436%:%
+%:%3529=1436%:%
+%:%3532=1437%:%
+%:%3533=1438%:%
+%:%3534=1439%:%
+%:%3535=1440%:%
+%:%3536=1441%:%
+%:%3537=1441%:%
+%:%3538=1442%:%
+%:%3541=1443%:%
+%:%3545=1443%:%
+%:%3546=1443%:%
+%:%3547=1444%:%
+%:%3548=1445%:%
+%:%3549=1445%:%
+%:%3550=1446%:%
+%:%3551=1446%:%
+%:%3556=1446%:%
+%:%3559=1447%:%
+%:%3560=1448%:%
+%:%3561=1448%:%
+%:%3562=1449%:%
+%:%3565=1450%:%
+%:%3566=1451%:%
+%:%3570=1451%:%
+%:%3571=1451%:%
+%:%3576=1451%:%
+%:%3579=1452%:%
+%:%3580=1453%:%
+%:%3581=1454%:%
+%:%3582=1454%:%
+%:%3583=1455%:%
+%:%3586=1456%:%
+%:%3590=1456%:%
+%:%3591=1456%:%
+%:%3592=1457%:%
+%:%3593=1457%:%
+%:%3594=1458%:%
+%:%3595=1458%:%
+%:%3596=1459%:%
+%:%3597=1459%:%
+%:%3598=1460%:%
+%:%3599=1461%:%
+%:%3600=1461%:%
+%:%3601=1462%:%
+%:%3602=1463%:%
+%:%3603=1463%:%
+%:%3604=1464%:%
+%:%3605=1465%:%
+%:%3606=1465%:%
+%:%3607=1465%:%
+%:%3612=1465%:%
+%:%3615=1466%:%
+%:%3616=1467%:%
+%:%3617=1468%:%
+%:%3618=1468%:%
+%:%3619=1469%:%
+%:%3622=1470%:%
+%:%3626=1470%:%
+%:%3627=1470%:%
+%:%3628=1471%:%
+%:%3629=1472%:%
+%:%3630=1472%:%
+%:%3631=1472%:%
+%:%3632=1473%:%
+%:%3633=1473%:%
+%:%3634=1474%:%
+%:%3635=1474%:%
+%:%3636=1475%:%
+%:%3637=1476%:%
+%:%3638=1477%:%
+%:%3639=1477%:%
+%:%3640=1477%:%
+%:%3641=1478%:%
+%:%3642=1479%:%
+%:%3643=1479%:%
+%:%3644=1480%:%
+%:%3645=1480%:%
+%:%3646=1481%:%
+%:%3647=1482%:%
+%:%3648=1483%:%
+%:%3649=1483%:%
+%:%3650=1484%:%
+%:%3651=1484%:%
+%:%3652=1485%:%
+%:%3653=1485%:%
+%:%3654=1486%:%
+%:%3655=1487%:%
+%:%3656=1487%:%
+%:%3657=1488%:%
+%:%3658=1489%:%
+%:%3659=1489%:%
+%:%3664=1489%:%
+%:%3667=1490%:%
+%:%3668=1491%:%
+%:%3669=1492%:%
+%:%3670=1493%:%
+%:%3671=1493%:%
+%:%3674=1494%:%
+%:%3675=1495%:%
+%:%3679=1495%:%
+%:%3680=1495%:%
+%:%3681=1496%:%
+%:%3682=1496%:%
+%:%3683=1497%:%
+%:%3684=1497%:%
+%:%3685=1498%:%
+%:%3686=1498%:%
+%:%3687=1499%:%
+%:%3688=1499%:%
+%:%3689=1499%:%
+%:%3690=1500%:%
+%:%3691=1500%:%
+%:%3692=1501%:%
+%:%3693=1501%:%
+%:%3694=1502%:%
+%:%3695=1502%:%
+%:%3696=1503%:%
+%:%3697=1503%:%
+%:%3698=1504%:%
+%:%3699=1504%:%
+%:%3700=1505%:%
+%:%3701=1505%:%
+%:%3702=1505%:%
+%:%3703=1506%:%
+%:%3704=1506%:%
+%:%3705=1507%:%
+%:%3706=1507%:%
+%:%3707=1507%:%
+%:%3708=1508%:%
+%:%3709=1508%:%
+%:%3710=1509%:%
+%:%3711=1509%:%
+%:%3712=1510%:%
+%:%3713=1510%:%
+%:%3714=1511%:%
+%:%3715=1512%:%
+%:%3716=1512%:%
+%:%3717=1512%:%
+%:%3718=1513%:%
+%:%3719=1514%:%
+%:%3720=1514%:%
+%:%3721=1515%:%
+%:%3722=1515%:%
+%:%3723=1516%:%
+%:%3724=1517%:%
+%:%3725=1517%:%
+%:%3726=1518%:%
+%:%3727=1519%:%
+%:%3728=1519%:%
+%:%3729=1519%:%
+%:%3730=1520%:%
+%:%3731=1521%:%
+%:%3732=1521%:%
+%:%3733=1522%:%
+%:%3734=1523%:%
+%:%3735=1523%:%
+%:%3736=1523%:%
+%:%3737=1524%:%
+%:%3738=1525%:%
+%:%3739=1525%:%
+%:%3740=1526%:%
+%:%3741=1527%:%
+%:%3742=1527%:%
+%:%3747=1527%:%
+%:%3750=1528%:%
+%:%3751=1529%:%
+%:%3752=1530%:%
+%:%3753=1531%:%
+%:%3754=1531%:%
+%:%3757=1532%:%
+%:%3761=1532%:%
+%:%3762=1532%:%
+%:%3763=1533%:%
+%:%3764=1533%:%
+%:%3765=1534%:%
+%:%3766=1534%:%
+%:%3767=1535%:%
+%:%3768=1535%:%
+%:%3769=1536%:%
+%:%3770=1536%:%
+%:%3771=1537%:%
+%:%3772=1537%:%
+%:%3773=1538%:%
+%:%3774=1538%:%
+%:%3775=1539%:%
+%:%3776=1539%:%
+%:%3777=1540%:%
+%:%3778=1540%:%
+%:%3779=1540%:%
+%:%3784=1540%:%
+%:%3787=1541%:%
+%:%3788=1542%:%
+%:%3789=1543%:%
+%:%3790=1543%:%
+%:%3791=1544%:%
+%:%3792=1544%:%
+%:%3795=1545%:%
+%:%3799=1545%:%
+%:%3800=1545%:%
+%:%3805=1545%:%
+%:%3808=1546%:%
+%:%3809=1547%:%
+%:%3810=1547%:%
+%:%3813=1548%:%
+%:%3817=1548%:%
+%:%3818=1548%:%
+%:%3819=1548%:%
+%:%3824=1548%:%
+%:%3827=1549%:%
+%:%3828=1550%:%
+%:%3829=1550%:%
+%:%3832=1551%:%
+%:%3836=1551%:%
+%:%3837=1551%:%
+%:%3842=1551%:%
+%:%3845=1552%:%
+%:%3846=1553%:%
+%:%3847=1553%:%
+%:%3850=1554%:%
+%:%3854=1554%:%
+%:%3855=1554%:%
+%:%3860=1554%:%
+%:%3863=1555%:%
+%:%3864=1556%:%
+%:%3865=1556%:%
+%:%3868=1557%:%
+%:%3872=1557%:%
+%:%3873=1557%:%
+%:%3874=1558%:%
+%:%3875=1558%:%
+%:%3876=1559%:%
+%:%3877=1559%:%
+%:%3882=1559%:%
+%:%3885=1560%:%
+%:%3886=1561%:%
+%:%3887=1561%:%
+%:%3890=1562%:%
+%:%3894=1562%:%
+%:%3895=1562%:%
+%:%3896=1563%:%
+%:%3897=1563%:%
+%:%3898=1564%:%
+%:%3899=1565%:%
+%:%3900=1565%:%
+%:%3901=1566%:%
+%:%3902=1567%:%
+%:%3903=1567%:%
+%:%3904=1568%:%
+%:%3905=1569%:%
+%:%3906=1569%:%
+%:%3907=1570%:%
+%:%3908=1571%:%
+%:%3909=1571%:%
+%:%3910=1571%:%
+%:%3911=1572%:%
+%:%3912=1573%:%
+%:%3913=1573%:%
+%:%3914=1573%:%
+%:%3915=1574%:%
+%:%3916=1575%:%
+%:%3917=1575%:%
+%:%3918=1575%:%
+%:%3919=1576%:%
+%:%3920=1577%:%
+%:%3921=1577%:%
+%:%3922=1578%:%
+%:%3923=1579%:%
+%:%3924=1579%:%
+%:%3925=1579%:%
+%:%3926=1580%:%
+%:%3927=1581%:%
+%:%3928=1581%:%
+%:%3929=1582%:%
+%:%3930=1583%:%
+%:%3931=1583%:%
+%:%3932=1584%:%
+%:%3933=1585%:%
+%:%3934=1585%:%
+%:%3935=1586%:%
+%:%3936=1587%:%
+%:%3937=1587%:%
+%:%3938=1588%:%
+%:%3939=1589%:%
+%:%3940=1589%:%
+%:%3941=1589%:%
+%:%3946=1589%:%
+%:%3949=1590%:%
+%:%3950=1591%:%
+%:%3951=1592%:%
+%:%3952=1593%:%
+%:%3953=1593%:%
+%:%3956=1594%:%
+%:%3960=1594%:%
+%:%3961=1594%:%
+%:%3962=1595%:%
+%:%3963=1595%:%
+%:%3964=1596%:%
+%:%3965=1596%:%
+%:%3970=1596%:%
+%:%3973=1597%:%
+%:%3974=1598%:%
+%:%3975=1598%:%
+%:%3978=1599%:%
+%:%3982=1599%:%
+%:%3983=1599%:%
+%:%3984=1600%:%
+%:%3985=1600%:%
+%:%3986=1601%:%
+%:%3987=1601%:%
+%:%3992=1601%:%
+%:%3995=1602%:%
+%:%3996=1603%:%
+%:%3997=1604%:%
+%:%3998=1605%:%
+%:%3999=1605%:%
+%:%4002=1606%:%
+%:%4006=1606%:%
+%:%4007=1606%:%
+%:%4008=1607%:%
+%:%4009=1607%:%
+%:%4010=1608%:%
+%:%4011=1608%:%
+%:%4012=1609%:%
+%:%4013=1609%:%
+%:%4014=1610%:%
+%:%4015=1610%:%
+%:%4016=1611%:%
+%:%4017=1612%:%
+%:%4018=1612%:%
+%:%4023=1612%:%
+%:%4026=1613%:%
+%:%4027=1614%:%
+%:%4028=1614%:%
+%:%4029=1615%:%
+%:%4031=1617%:%
+%:%4034=1618%:%
+%:%4038=1618%:%
+%:%4039=1618%:%
+%:%4040=1619%:%
+%:%4041=1620%:%
+%:%4042=1620%:%
+%:%4043=1620%:%
+%:%4048=1620%:%
+%:%4051=1621%:%
+%:%4052=1622%:%
+%:%4053=1622%:%
+%:%4056=1623%:%
+%:%4060=1623%:%
+%:%4061=1623%:%
+%:%4062=1624%:%
+%:%4063=1625%:%
+%:%4064=1625%:%
+%:%4065=1626%:%
+%:%4066=1626%:%
+%:%4067=1627%:%
+%:%4068=1628%:%
+%:%4069=1628%:%
+%:%4070=1629%:%
+%:%4071=1629%:%
+%:%4072=1630%:%
+%:%4073=1630%:%
+%:%4074=1631%:%
+%:%4075=1631%:%
+%:%4076=1632%:%
+%:%4077=1632%:%
+%:%4078=1633%:%
+%:%4079=1633%:%
+%:%4080=1634%:%
+%:%4081=1634%:%
+%:%4082=1635%:%
+%:%4083=1635%:%
+%:%4084=1636%:%
+%:%4085=1636%:%
+%:%4086=1636%:%
+%:%4087=1637%:%
+%:%4088=1637%:%
+%:%4089=1638%:%
+%:%4090=1638%:%
+%:%4091=1639%:%
+%:%4092=1639%:%
+%:%4093=1640%:%
+%:%4094=1640%:%
+%:%4095=1641%:%
+%:%4096=1641%:%
+%:%4097=1641%:%
+%:%4098=1642%:%
+%:%4099=1642%:%
+%:%4100=1643%:%
+%:%4101=1644%:%
+%:%4102=1644%:%
+%:%4103=1644%:%
+%:%4104=1645%:%
+%:%4105=1645%:%
+%:%4106=1646%:%
+%:%4107=1646%:%
+%:%4108=1647%:%
+%:%4109=1647%:%
+%:%4110=1648%:%
+%:%4111=1648%:%
+%:%4112=1649%:%
+%:%4113=1649%:%
+%:%4114=1649%:%
+%:%4115=1650%:%
+%:%4116=1650%:%
+%:%4117=1650%:%
+%:%4118=1651%:%
+%:%4119=1651%:%
+%:%4120=1651%:%
+%:%4121=1652%:%
+%:%4122=1652%:%
+%:%4123=1653%:%
+%:%4124=1653%:%
+%:%4125=1654%:%
+%:%4126=1654%:%
+%:%4127=1655%:%
+%:%4128=1655%:%
+%:%4129=1656%:%
+%:%4130=1656%:%
+%:%4131=1656%:%
+%:%4132=1657%:%
+%:%4133=1657%:%
+%:%4134=1658%:%
+%:%4135=1658%:%
+%:%4136=1659%:%
+%:%4137=1659%:%
+%:%4138=1660%:%
+%:%4139=1660%:%
+%:%4140=1661%:%
+%:%4141=1662%:%
+%:%4142=1662%:%
+%:%4143=1663%:%
+%:%4144=1664%:%
+%:%4145=1664%:%
+%:%4146=1664%:%
+%:%4147=1665%:%
+%:%4148=1666%:%
+%:%4149=1666%:%
+%:%4150=1666%:%
+%:%4151=1667%:%
+%:%4152=1668%:%
+%:%4153=1668%:%
+%:%4154=1668%:%
+%:%4155=1669%:%
+%:%4156=1670%:%
+%:%4157=1670%:%
+%:%4158=1671%:%
+%:%4159=1672%:%
+%:%4160=1672%:%
+%:%4161=1673%:%
+%:%4162=1674%:%
+%:%4163=1674%:%
+%:%4164=1675%:%
+%:%4165=1676%:%
+%:%4166=1676%:%
+%:%4167=1677%:%
+%:%4168=1678%:%
+%:%4169=1679%:%
+%:%4170=1679%:%
+%:%4171=1679%:%
+%:%4176=1679%:%
+%:%4179=1680%:%
+%:%4180=1681%:%
+%:%4181=1682%:%
+%:%4182=1683%:%
+%:%4183=1683%:%
+%:%4186=1684%:%
+%:%4190=1684%:%
+%:%4191=1684%:%
+%:%4192=1685%:%
+%:%4193=1685%:%
+%:%4194=1686%:%
+%:%4195=1686%:%
+%:%4200=1686%:%
+%:%4203=1687%:%
+%:%4204=1688%:%
+%:%4205=1689%:%
+%:%4206=1690%:%
+%:%4207=1691%:%
+%:%4208=1691%:%
+%:%4211=1692%:%
+%:%4215=1692%:%
+%:%4216=1692%:%
+%:%4217=1693%:%
+%:%4218=1694%:%
+%:%4219=1694%:%
+%:%4220=1695%:%
+%:%4221=1695%:%
+%:%4222=1696%:%
+%:%4223=1696%:%
+%:%4224=1697%:%
+%:%4225=1697%:%
+%:%4230=1697%:%
+%:%4233=1698%:%
+%:%4234=1699%:%
+%:%4235=1700%:%
+%:%4236=1701%:%
+%:%4237=1701%:%
+%:%4238=1702%:%
+%:%4239=1703%:%
+%:%4240=1703%:%
+%:%4243=1704%:%
+%:%4247=1704%:%
+%:%4248=1704%:%
+%:%4249=1704%:%
+%:%4254=1704%:%
+%:%4257=1705%:%
+%:%4258=1706%:%
+%:%4259=1706%:%
+%:%4262=1707%:%
+%:%4266=1707%:%
+%:%4267=1707%:%
+%:%4272=1707%:%
+%:%4275=1708%:%
+%:%4276=1709%:%
+%:%4277=1710%:%
+%:%4278=1710%:%
+%:%4281=1711%:%
+%:%4285=1711%:%
+%:%4286=1711%:%
+%:%4287=1711%:%
+%:%4292=1711%:%
+%:%4295=1712%:%
+%:%4296=1713%:%
+%:%4297=1714%:%
+%:%4298=1714%:%
+%:%4299=1715%:%
+%:%4302=1716%:%
+%:%4307=1717%:%
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys2/Journal/comment.sty	Mon Nov 01 10:40:21 2021 +0000
@@ -0,0 +1,278 @@
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+% Comment.sty   version 3.6, October 1999
+%
+% Purpose:
+% selectively in/exclude pieces of text: the user can define new
+% comment versions, and each is controlled separately.
+% Special comments can be defined where the user specifies the
+% action that is to be taken with each comment line.
+%
+% Author
+%    Victor Eijkhout
+%    Department of Computer Science
+%    University of Tennessee
+%    107 Ayres Hall
+%    Knoxville TN 37996
+%    USA
+%
+%    victor@eijkhout.net
+%
+% This program is free software; you can redistribute it and/or
+% modify it under the terms of the GNU General Public License
+% as published by the Free Software Foundation; either version 2
+% of the License, or (at your option) any later version.
+% 
+% This program is distributed in the hope that it will be useful,
+% but WITHOUT ANY WARRANTY; without even the implied warranty of
+% MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE.  See the
+% GNU General Public License for more details.
+%
+% For a copy of the GNU General Public License, write to the 
+% Free Software Foundation, Inc.,
+% 59 Temple Place - Suite 330, Boston, MA  02111-1307, USA,
+% or find it on the net, for instance at
+% http://www.gnu.org/copyleft/gpl.html
+%
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+% This style can be used with plain TeX or LaTeX, and probably
+% most other packages too.
+%
+% Usage: all text included between
+%    \comment ... \endcomment
+% or \begin{comment} ... \end{comment}
+% is discarded. 
+%
+% The opening and closing commands should appear on a line
+% of their own. No starting spaces, nothing after it.
+% This environment should work with arbitrary amounts
+% of comment, and the comment can be arbitrary text.
+%
+% Other `comment' environments are defined by
+% and are selected/deselected with
+% \includecomment{versiona}
+% \excludecoment{versionb}
+%
+% These environments are used as
+% \versiona ... \endversiona
+% or \begin{versiona} ... \end{versiona}
+% with the opening and closing commands again on a line of 
+% their own.
+%
+% LaTeX users note: for an included comment, the
+% \begin and \end lines act as if they don't exist.
+% In particular, they don't imply grouping, so assignments 
+% &c are not local.
+%
+% Special comments are defined as
+% \specialcomment{name}{before commands}{after commands}
+% where the second and third arguments are executed before
+% and after each comment block. You can use this for global
+% formatting commands.
+% To keep definitions &c local, you can include \begingroup
+% in the `before commands' and \endgroup in the `after commands'.
+% ex:
+% \specialcomment{smalltt}
+%     {\begingroup\ttfamily\footnotesize}{\endgroup}
+% You do *not* have to do an additional
+% \includecomment{smalltt}
+% To remove 'smalltt' blocks, give \excludecomment{smalltt}
+% after the definition.
+%
+% Processing comments can apply processing to each line.
+% \processcomment{name}{each-line commands}%
+%    {before commands}{after commands}
+% By defining a control sequence 
+% \def\Thiscomment##1{...} in the before commands the user can
+% specify what is to be done with each comment line.
+% BUG this does not work quite yet BUG
+%
+% Trick for short in/exclude macros (such as \maybe{this snippet}):
+%\includecomment{cond}
+%\newcommand{\maybe}[1]{}
+%\begin{cond}
+%\renewcommand{\maybe}[1]{#1}
+%\end{cond}
+%
+% Basic approach of the implementation:
+% to comment something out, scoop up  every line in verbatim mode
+% as macro argument, then throw it away.
+% For inclusions, in LaTeX the block is written out to
+% a file \CommentCutFile (default "comment.cut"), which is
+% then included.
+% In plain TeX (and other formats) both the opening and
+% closing comands are defined as noop.
+%
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+% Changes in version 3.1
+% - updated author's address
+% - cleaned up some code
+% - trailing contents on \begin{env} line is always discarded
+%  even if you've done \includecomment{env}
+% - comments no longer define grouping!! you can even
+%   \includecomment{env}
+%   \begin{env}
+%   \begin{itemize}
+%   \end{env}
+%  Isn't that something ...
+% - included comments are written to file and input again.
+% Changes in 3.2
+% - \specialcomment brought up to date (thanks to Ivo Welch).
+% Changes in 3.3
+% - updated author's address again
+% - parametrised \CommentCutFile
+% Changes in 3.4
+% - added GNU public license
+% - added \processcomment, because Ivo's fix (above) brought an
+%   inconsistency to light.
+% Changes in 3.5
+% - corrected typo in header.
+% - changed author email
+% - corrected \specialcomment yet again.
+% - fixed excludecomment of an earlier defined environment.
+% Changes in 3.6
+% - The 'cut' file is now written more verbatim, using \meaning;
+%   some people reported having trouble with ISO latin 1, or umlaute.sty.
+% - removed some \newif statements.
+%   Has this suddenly become \outer again?
+%
+% Known bugs:
+% - excludecomment leads to one superfluous space
+% - processcomment leads to a superfluous line break
+%
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+
+\def\makeinnocent#1{\catcode`#1=12 }
+\def\csarg#1#2{\expandafter#1\csname#2\endcsname}
+\def\latexname{lplain}\def\latexename{LaTeX2e}
+\newwrite\CommentStream
+\def\CommentCutFile{comment.cut}
+
+\def\ProcessComment#1% start it all of
+   {\begingroup
+    \def\CurrentComment{#1}%
+    \let\do\makeinnocent \dospecials 
+    \makeinnocent\^^L% and whatever other special cases
+    \endlinechar`\^^M \catcode`\^^M=12 \xComment}
+%\def\ProcessCommentWithArg#1#2% to be used in \leveledcomment
+%   {\begingroup
+%    \def\CurrentComment{#1}%
+%    \let\do\makeinnocent \dospecials 
+%    \makeinnocent\^^L% and whatever other special cases
+%    \endlinechar`\^^M \catcode`\^^M=12 \xComment}
+{\catcode`\^^M=12 \endlinechar=-1 %
+ \gdef\xComment#1^^M{%
+    \expandafter\ProcessCommentLine}
+ \gdef\ProcessCommentLine#1^^M{\def\test{#1}
+      \csarg\ifx{End\CurrentComment Test}\test
+          \edef\next{\noexpand\EndOfComment{\CurrentComment}}%
+      \else \ThisComment{#1}\let\next\ProcessCommentLine
+      \fi \next}
+}
+
+\def\CSstringmeaning#1{\expandafter\CSgobblearrow\meaning#1}
+\def\CSstringcsnoescape#1{\expandafter\CSgobbleescape\string#1}
+{\escapechar-1
+\expandafter\expandafter\expandafter\gdef
+  \expandafter\expandafter\expandafter\CSgobblearrow
+    \expandafter\string\csname macro:->\endcsname{}
+}
+\def\CSgobbleescape#1{\ifnum`\\=`#1 \else #1\fi}
+\def\WriteCommentLine#1{\def\CStmp{#1}%
+    \immediate\write\CommentStream{\CSstringmeaning\CStmp}}
+
+% 3.1 change: in LaTeX and LaTeX2e prevent grouping
+\if 0%
+\ifx\fmtname\latexename 
+    0%
+\else \ifx\fmtname\latexname 
+          0%
+      \else 
+          1%
+\fi   \fi
+%%%%
+%%%% definitions for LaTeX
+%%%%
+\def\AfterIncludedComment
+   {\immediate\closeout\CommentStream
+    \input{\CommentCutFile}\relax
+    }%
+\def\TossComment{\immediate\closeout\CommentStream}
+\def\BeforeIncludedComment
+   {\immediate\openout\CommentStream=\CommentCutFile
+    \let\ThisComment\WriteCommentLine}
+\def\includecomment
+ #1{\message{Include comment '#1'}%
+    \csarg\let{After#1Comment}\AfterIncludedComment
+    \csarg\def{#1}{\BeforeIncludedComment
+        \ProcessComment{#1}}%
+    \CommentEndDef{#1}}
+\long\def\specialcomment
+ #1#2#3{\message{Special comment '#1'}%
+    % note: \AfterIncludedComment does \input, so #2 goes here!
+    \csarg\def{After#1Comment}{#2\AfterIncludedComment#3}%
+    \csarg\def{#1}{\BeforeIncludedComment\relax
+          \ProcessComment{#1}}%
+    \CommentEndDef{#1}}
+\long\def\processcomment
+ #1#2#3#4{\message{Lines-Processing comment '#1'}%
+    \csarg\def{After#1Comment}{#3\AfterIncludedComment#4}%
+    \csarg\def{#1}{\BeforeIncludedComment#2\relax
+          \ProcessComment{#1}}%
+    \CommentEndDef{#1}}
+\def\leveledcomment
+ #1#2{\message{Include comment '#1' up to level '#2'}%
+    %\csname #1IsLeveledCommenttrue\endcsname
+    \csarg\let{After#1Comment}\AfterIncludedComment
+    \csarg\def{#1}{\BeforeIncludedComment
+        \ProcessCommentWithArg{#1}}%
+    \CommentEndDef{#1}}
+\else 
+%%%%
+%%%%plain TeX and other formats
+%%%%
+\def\includecomment
+ #1{\message{Including comment '#1'}%
+    \csarg\def{#1}{}%
+    \csarg\def{end#1}{}}
+\long\def\specialcomment
+ #1#2#3{\message{Special comment '#1'}%
+    \csarg\def{#1}{\def\ThisComment{}\def\AfterComment{#3}#2%
+           \ProcessComment{#1}}%
+    \CommentEndDef{#1}}
+\fi
+
+%%%%
+%%%% general definition of skipped comment
+%%%%
+\def\excludecomment
+ #1{\message{Excluding comment '#1'}%
+    \csarg\def{#1}{\let\AfterComment\relax
+           \def\ThisComment####1{}\ProcessComment{#1}}%
+    \csarg\let{After#1Comment}\TossComment
+    \CommentEndDef{#1}}
+
+\if 0%
+\ifx\fmtname\latexename 
+    0%
+\else \ifx\fmtname\latexname 
+          0%
+      \else 
+          1%
+\fi   \fi
+% latex & latex2e:
+\def\EndOfComment#1{\endgroup\end{#1}%
+    \csname After#1Comment\endcsname}
+\def\CommentEndDef#1{{\escapechar=-1\relax
+    \csarg\xdef{End#1Test}{\string\\end\string\{#1\string\}}%
+    }}
+\else
+% plain & other
+\def\EndOfComment#1{\endgroup\AfterComment}
+\def\CommentEndDef#1{{\escapechar=-1\relax
+    \csarg\xdef{End#1Test}{\string\\end#1}%
+    }}
+\fi
+
+\excludecomment{comment}
+
+\endinput
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys2/Journal/document/llncs.cls	Mon Nov 01 10:40:21 2021 +0000
@@ -0,0 +1,1189 @@
+% LLNCS DOCUMENT CLASS -- version 2.13 (28-Jan-2002)
+% Springer Verlag LaTeX2e support for Lecture Notes in Computer Science
+%
+%%
+%% \CharacterTable
+%%  {Upper-case    \A\B\C\D\E\F\G\H\I\J\K\L\M\N\O\P\Q\R\S\T\U\V\W\X\Y\Z
+%%   Lower-case    \a\b\c\d\e\f\g\h\i\j\k\l\m\n\o\p\q\r\s\t\u\v\w\x\y\z
+%%   Digits        \0\1\2\3\4\5\6\7\8\9
+%%   Exclamation   \!     Double quote  \"     Hash (number) \#
+%%   Dollar        \$     Percent       \%     Ampersand     \&
+%%   Acute accent  \'     Left paren    \(     Right paren   \)
+%%   Asterisk      \*     Plus          \+     Comma         \,
+%%   Minus         \-     Point         \.     Solidus       \/
+%%   Colon         \:     Semicolon     \;     Less than     \<
+%%   Equals        \=     Greater than  \>     Question mark \?
+%%   Commercial at \@     Left bracket  \[     Backslash     \\
+%%   Right bracket \]     Circumflex    \^     Underscore    \_
+%%   Grave accent  \`     Left brace    \{     Vertical bar  \|
+%%   Right brace   \}     Tilde         \~}
+%%
+\NeedsTeXFormat{LaTeX2e}[1995/12/01]
+\ProvidesClass{llncs}[2002/01/28 v2.13
+^^J LaTeX document class for Lecture Notes in Computer Science]
+% Options
+\let\if@envcntreset\iffalse
+\DeclareOption{envcountreset}{\let\if@envcntreset\iftrue}
+\DeclareOption{citeauthoryear}{\let\citeauthoryear=Y}
+\DeclareOption{oribibl}{\let\oribibl=Y}
+\let\if@custvec\iftrue
+\DeclareOption{orivec}{\let\if@custvec\iffalse}
+\let\if@envcntsame\iffalse
+\DeclareOption{envcountsame}{\let\if@envcntsame\iftrue}
+\let\if@envcntsect\iffalse
+\DeclareOption{envcountsect}{\let\if@envcntsect\iftrue}
+\let\if@runhead\iffalse
+\DeclareOption{runningheads}{\let\if@runhead\iftrue}
+
+\let\if@openbib\iffalse
+\DeclareOption{openbib}{\let\if@openbib\iftrue}
+
+% languages
+\let\switcht@@therlang\relax
+\def\ds@deutsch{\def\switcht@@therlang{\switcht@deutsch}}
+\def\ds@francais{\def\switcht@@therlang{\switcht@francais}}
+
+\DeclareOption*{\PassOptionsToClass{\CurrentOption}{article}}
+
+\ProcessOptions
+
+\LoadClass[twoside]{article}
+\RequirePackage{multicol} % needed for the list of participants, index
+
+\setlength{\textwidth}{12.2cm}
+\setlength{\textheight}{19.3cm}
+\renewcommand\@pnumwidth{2em}
+\renewcommand\@tocrmarg{3.5em}
+%
+\def\@dottedtocline#1#2#3#4#5{%
+  \ifnum #1>\c@tocdepth \else
+    \vskip \z@ \@plus.2\p@
+    {\leftskip #2\relax \rightskip \@tocrmarg \advance\rightskip by 0pt plus 2cm
+               \parfillskip -\rightskip \pretolerance=10000
+     \parindent #2\relax\@afterindenttrue
+     \interlinepenalty\@M
+     \leavevmode
+     \@tempdima #3\relax
+     \advance\leftskip \@tempdima \null\nobreak\hskip -\leftskip
+     {#4}\nobreak
+     \leaders\hbox{$\m@th
+        \mkern \@dotsep mu\hbox{.}\mkern \@dotsep
+        mu$}\hfill
+     \nobreak
+     \hb@xt@\@pnumwidth{\hfil\normalfont \normalcolor #5}%
+     \par}%
+  \fi}
+%
+\def\switcht@albion{%
+\def\abstractname{Abstract.}
+\def\ackname{Acknowledgement.}
+\def\andname{and}
+\def\lastandname{\unskip, and}
+\def\appendixname{Appendix}
+\def\chaptername{Chapter}
+\def\claimname{Claim}
+\def\conjecturename{Conjecture}
+\def\contentsname{Table of Contents}
+\def\corollaryname{Corollary}
+\def\definitionname{Definition}
+\def\examplename{Example}
+\def\exercisename{Exercise}
+\def\figurename{Fig.}
+\def\keywordname{{\bf Key words:}}
+\def\indexname{Index}
+\def\lemmaname{Lemma}
+\def\contriblistname{List of Contributors}
+\def\listfigurename{List of Figures}
+\def\listtablename{List of Tables}
+\def\mailname{{\it Correspondence to\/}:}
+\def\noteaddname{Note added in proof}
+\def\notename{Note}
+\def\partname{Part}
+\def\problemname{Problem}
+\def\proofname{Proof}
+\def\propertyname{Property}
+\def\propositionname{Proposition}
+\def\questionname{Question}
+\def\remarkname{Remark}
+\def\seename{see}
+\def\solutionname{Solution}
+\def\subclassname{{\it Subject Classifications\/}:}
+\def\tablename{Table}
+\def\theoremname{Theorem}}
+\switcht@albion
+% Names of theorem like environments are already defined
+% but must be translated if another language is chosen
+%
+% French section
+\def\switcht@francais{%\typeout{On parle francais.}%
+ \def\abstractname{R\'esum\'e.}%
+ \def\ackname{Remerciements.}%
+ \def\andname{et}%
+ \def\lastandname{ et}%
+ \def\appendixname{Appendice}
+ \def\chaptername{Chapitre}%
+ \def\claimname{Pr\'etention}%
+ \def\conjecturename{Hypoth\`ese}%
+ \def\contentsname{Table des mati\`eres}%
+ \def\corollaryname{Corollaire}%
+ \def\definitionname{D\'efinition}%
+ \def\examplename{Exemple}%
+ \def\exercisename{Exercice}%
+ \def\figurename{Fig.}%
+ \def\keywordname{{\bf Mots-cl\'e:}}
+ \def\indexname{Index}
+ \def\lemmaname{Lemme}%
+ \def\contriblistname{Liste des contributeurs}
+ \def\listfigurename{Liste des figures}%
+ \def\listtablename{Liste des tables}%
+ \def\mailname{{\it Correspondence to\/}:}
+ \def\noteaddname{Note ajout\'ee \`a l'\'epreuve}%
+ \def\notename{Remarque}%
+ \def\partname{Partie}%
+ \def\problemname{Probl\`eme}%
+ \def\proofname{Preuve}%
+ \def\propertyname{Caract\'eristique}%
+%\def\propositionname{Proposition}%
+ \def\questionname{Question}%
+ \def\remarkname{Remarque}%
+ \def\seename{voir}
+ \def\solutionname{Solution}%
+ \def\subclassname{{\it Subject Classifications\/}:}
+ \def\tablename{Tableau}%
+ \def\theoremname{Th\'eor\`eme}%
+}
+%
+% German section
+\def\switcht@deutsch{%\typeout{Man spricht deutsch.}%
+ \def\abstractname{Zusammenfassung.}%
+ \def\ackname{Danksagung.}%
+ \def\andname{und}%
+ \def\lastandname{ und}%
+ \def\appendixname{Anhang}%
+ \def\chaptername{Kapitel}%
+ \def\claimname{Behauptung}%
+ \def\conjecturename{Hypothese}%
+ \def\contentsname{Inhaltsverzeichnis}%
+ \def\corollaryname{Korollar}%
+%\def\definitionname{Definition}%
+ \def\examplename{Beispiel}%
+ \def\exercisename{\"Ubung}%
+ \def\figurename{Abb.}%
+ \def\keywordname{{\bf Schl\"usselw\"orter:}}
+ \def\indexname{Index}
+%\def\lemmaname{Lemma}%
+ \def\contriblistname{Mitarbeiter}
+ \def\listfigurename{Abbildungsverzeichnis}%
+ \def\listtablename{Tabellenverzeichnis}%
+ \def\mailname{{\it Correspondence to\/}:}
+ \def\noteaddname{Nachtrag}%
+ \def\notename{Anmerkung}%
+ \def\partname{Teil}%
+%\def\problemname{Problem}%
+ \def\proofname{Beweis}%
+ \def\propertyname{Eigenschaft}%
+%\def\propositionname{Proposition}%
+ \def\questionname{Frage}%
+ \def\remarkname{Anmerkung}%
+ \def\seename{siehe}
+ \def\solutionname{L\"osung}%
+ \def\subclassname{{\it Subject Classifications\/}:}
+ \def\tablename{Tabelle}%
+%\def\theoremname{Theorem}%
+}
+
+% Ragged bottom for the actual page
+\def\thisbottomragged{\def\@textbottom{\vskip\z@ plus.0001fil
+\global\let\@textbottom\relax}}
+
+\renewcommand\small{%
+   \@setfontsize\small\@ixpt{11}%
+   \abovedisplayskip 8.5\p@ \@plus3\p@ \@minus4\p@
+   \abovedisplayshortskip \z@ \@plus2\p@
+   \belowdisplayshortskip 4\p@ \@plus2\p@ \@minus2\p@
+   \def\@listi{\leftmargin\leftmargini
+               \parsep 0\p@ \@plus1\p@ \@minus\p@
+               \topsep 8\p@ \@plus2\p@ \@minus4\p@
+               \itemsep0\p@}%
+   \belowdisplayskip \abovedisplayskip
+}
+
+\frenchspacing
+\widowpenalty=10000
+\clubpenalty=10000
+
+\setlength\oddsidemargin   {63\p@}
+\setlength\evensidemargin  {63\p@}
+\setlength\marginparwidth  {90\p@}
+
+\setlength\headsep   {16\p@}
+
+\setlength\footnotesep{7.7\p@}
+\setlength\textfloatsep{8mm\@plus 2\p@ \@minus 4\p@}
+\setlength\intextsep   {8mm\@plus 2\p@ \@minus 2\p@}
+
+\setcounter{secnumdepth}{2}
+
+\newcounter {chapter}
+\renewcommand\thechapter      {\@arabic\c@chapter}
+
+\newif\if@mainmatter \@mainmattertrue
+\newcommand\frontmatter{\cleardoublepage
+            \@mainmatterfalse\pagenumbering{Roman}}
+\newcommand\mainmatter{\cleardoublepage
+       \@mainmattertrue\pagenumbering{arabic}}
+\newcommand\backmatter{\if@openright\cleardoublepage\else\clearpage\fi
+      \@mainmatterfalse}
+
+\renewcommand\part{\cleardoublepage
+                 \thispagestyle{empty}%
+                 \if@twocolumn
+                     \onecolumn
+                     \@tempswatrue
+                   \else
+                     \@tempswafalse
+                 \fi
+                 \null\vfil
+                 \secdef\@part\@spart}
+
+\def\@part[#1]#2{%
+    \ifnum \c@secnumdepth >-2\relax
+      \refstepcounter{part}%
+      \addcontentsline{toc}{part}{\thepart\hspace{1em}#1}%
+    \else
+      \addcontentsline{toc}{part}{#1}%
+    \fi
+    \markboth{}{}%
+    {\centering
+     \interlinepenalty \@M
+     \normalfont
+     \ifnum \c@secnumdepth >-2\relax
+       \huge\bfseries \partname~\thepart
+       \par
+       \vskip 20\p@
+     \fi
+     \Huge \bfseries #2\par}%
+    \@endpart}
+\def\@spart#1{%
+    {\centering
+     \interlinepenalty \@M
+     \normalfont
+     \Huge \bfseries #1\par}%
+    \@endpart}
+\def\@endpart{\vfil\newpage
+              \if@twoside
+                \null
+                \thispagestyle{empty}%
+                \newpage
+              \fi
+              \if@tempswa
+                \twocolumn
+              \fi}
+
+\newcommand\chapter{\clearpage
+                    \thispagestyle{empty}%
+                    \global\@topnum\z@
+                    \@afterindentfalse
+                    \secdef\@chapter\@schapter}
+\def\@chapter[#1]#2{\ifnum \c@secnumdepth >\m@ne
+                       \if@mainmatter
+                         \refstepcounter{chapter}%
+                         \typeout{\@chapapp\space\thechapter.}%
+                         \addcontentsline{toc}{chapter}%
+                                  {\protect\numberline{\thechapter}#1}%
+                       \else
+                         \addcontentsline{toc}{chapter}{#1}%
+                       \fi
+                    \else
+                      \addcontentsline{toc}{chapter}{#1}%
+                    \fi
+                    \chaptermark{#1}%
+                    \addtocontents{lof}{\protect\addvspace{10\p@}}%
+                    \addtocontents{lot}{\protect\addvspace{10\p@}}%
+                    \if@twocolumn
+                      \@topnewpage[\@makechapterhead{#2}]%
+                    \else
+                      \@makechapterhead{#2}%
+                      \@afterheading
+                    \fi}
+\def\@makechapterhead#1{%
+% \vspace*{50\p@}%
+  {\centering
+    \ifnum \c@secnumdepth >\m@ne
+      \if@mainmatter
+        \large\bfseries \@chapapp{} \thechapter
+        \par\nobreak
+        \vskip 20\p@
+      \fi
+    \fi
+    \interlinepenalty\@M
+    \Large \bfseries #1\par\nobreak
+    \vskip 40\p@
+  }}
+\def\@schapter#1{\if@twocolumn
+                   \@topnewpage[\@makeschapterhead{#1}]%
+                 \else
+                   \@makeschapterhead{#1}%
+                   \@afterheading
+                 \fi}
+\def\@makeschapterhead#1{%
+% \vspace*{50\p@}%
+  {\centering
+    \normalfont
+    \interlinepenalty\@M
+    \Large \bfseries  #1\par\nobreak
+    \vskip 40\p@
+  }}
+
+\renewcommand\section{\@startsection{section}{1}{\z@}%
+                       {-18\p@ \@plus -4\p@ \@minus -4\p@}%
+                       {12\p@ \@plus 4\p@ \@minus 4\p@}%
+                       {\normalfont\large\bfseries\boldmath
+                        \rightskip=\z@ \@plus 8em\pretolerance=10000 }}
+\renewcommand\subsection{\@startsection{subsection}{2}{\z@}%
+                       {-18\p@ \@plus -4\p@ \@minus -4\p@}%
+                       {8\p@ \@plus 4\p@ \@minus 4\p@}%
+                       {\normalfont\normalsize\bfseries\boldmath
+                        \rightskip=\z@ \@plus 8em\pretolerance=10000 }}
+\renewcommand\subsubsection{\@startsection{subsubsection}{3}{\z@}%
+                       {-18\p@ \@plus -4\p@ \@minus -4\p@}%
+                       {-0.5em \@plus -0.22em \@minus -0.1em}%
+                       {\normalfont\normalsize\bfseries\boldmath}}
+\renewcommand\paragraph{\@startsection{paragraph}{4}{\z@}%
+                       {-12\p@ \@plus -4\p@ \@minus -4\p@}%
+                       {-0.5em \@plus -0.22em \@minus -0.1em}%
+                       {\normalfont\normalsize\itshape}}
+\renewcommand\subparagraph[1]{\typeout{LLNCS warning: You should not use
+                  \string\subparagraph\space with this class}\vskip0.5cm
+You should not use \verb|\subparagraph| with this class.\vskip0.5cm}
+
+\DeclareMathSymbol{\Gamma}{\mathalpha}{letters}{"00}
+\DeclareMathSymbol{\Delta}{\mathalpha}{letters}{"01}
+\DeclareMathSymbol{\Theta}{\mathalpha}{letters}{"02}
+\DeclareMathSymbol{\Lambda}{\mathalpha}{letters}{"03}
+\DeclareMathSymbol{\Xi}{\mathalpha}{letters}{"04}
+\DeclareMathSymbol{\Pi}{\mathalpha}{letters}{"05}
+\DeclareMathSymbol{\Sigma}{\mathalpha}{letters}{"06}
+\DeclareMathSymbol{\Upsilon}{\mathalpha}{letters}{"07}
+\DeclareMathSymbol{\Phi}{\mathalpha}{letters}{"08}
+\DeclareMathSymbol{\Psi}{\mathalpha}{letters}{"09}
+\DeclareMathSymbol{\Omega}{\mathalpha}{letters}{"0A}
+
+\let\footnotesize\small
+
+\if@custvec
+\def\vec#1{\mathchoice{\mbox{\boldmath$\displaystyle#1$}}
+{\mbox{\boldmath$\textstyle#1$}}
+{\mbox{\boldmath$\scriptstyle#1$}}
+{\mbox{\boldmath$\scriptscriptstyle#1$}}}
+\fi
+
+\def\squareforqed{\hbox{\rlap{$\sqcap$}$\sqcup$}}
+\def\qed{\ifmmode\squareforqed\else{\unskip\nobreak\hfil
+\penalty50\hskip1em\null\nobreak\hfil\squareforqed
+\parfillskip=0pt\finalhyphendemerits=0\endgraf}\fi}
+
+\def\getsto{\mathrel{\mathchoice {\vcenter{\offinterlineskip
+\halign{\hfil
+$\displaystyle##$\hfil\cr\gets\cr\to\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr\gets
+\cr\to\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr\gets
+\cr\to\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
+\gets\cr\to\cr}}}}}
+\def\lid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil
+$\displaystyle##$\hfil\cr<\cr\noalign{\vskip1.2pt}=\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr<\cr
+\noalign{\vskip1.2pt}=\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr<\cr
+\noalign{\vskip1pt}=\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
+<\cr
+\noalign{\vskip0.9pt}=\cr}}}}}
+\def\gid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil
+$\displaystyle##$\hfil\cr>\cr\noalign{\vskip1.2pt}=\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr>\cr
+\noalign{\vskip1.2pt}=\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr>\cr
+\noalign{\vskip1pt}=\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
+>\cr
+\noalign{\vskip0.9pt}=\cr}}}}}
+\def\grole{\mathrel{\mathchoice {\vcenter{\offinterlineskip
+\halign{\hfil
+$\displaystyle##$\hfil\cr>\cr\noalign{\vskip-1pt}<\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr
+>\cr\noalign{\vskip-1pt}<\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr
+>\cr\noalign{\vskip-0.8pt}<\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
+>\cr\noalign{\vskip-0.3pt}<\cr}}}}}
+\def\bbbr{{\rm I\!R}} %reelle Zahlen
+\def\bbbm{{\rm I\!M}}
+\def\bbbn{{\rm I\!N}} %natuerliche Zahlen
+\def\bbbf{{\rm I\!F}}
+\def\bbbh{{\rm I\!H}}
+\def\bbbk{{\rm I\!K}}
+\def\bbbp{{\rm I\!P}}
+\def\bbbone{{\mathchoice {\rm 1\mskip-4mu l} {\rm 1\mskip-4mu l}
+{\rm 1\mskip-4.5mu l} {\rm 1\mskip-5mu l}}}
+\def\bbbc{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm C$}\hbox{\hbox
+to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
+{\setbox0=\hbox{$\textstyle\rm C$}\hbox{\hbox
+to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptstyle\rm C$}\hbox{\hbox
+to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptscriptstyle\rm C$}\hbox{\hbox
+to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}}}
+\def\bbbq{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm
+Q$}\hbox{\raise
+0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}}
+{\setbox0=\hbox{$\textstyle\rm Q$}\hbox{\raise
+0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptstyle\rm Q$}\hbox{\raise
+0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptscriptstyle\rm Q$}\hbox{\raise
+0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}}}
+\def\bbbt{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm
+T$}\hbox{\hbox to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
+{\setbox0=\hbox{$\textstyle\rm T$}\hbox{\hbox
+to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptstyle\rm T$}\hbox{\hbox
+to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptscriptstyle\rm T$}\hbox{\hbox
+to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}}}
+\def\bbbs{{\mathchoice
+{\setbox0=\hbox{$\displaystyle     \rm S$}\hbox{\raise0.5\ht0\hbox
+to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox
+to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}}
+{\setbox0=\hbox{$\textstyle        \rm S$}\hbox{\raise0.5\ht0\hbox
+to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox
+to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptstyle      \rm S$}\hbox{\raise0.5\ht0\hbox
+to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox
+to0pt{\kern0.5\wd0\vrule height0.45\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptscriptstyle\rm S$}\hbox{\raise0.5\ht0\hbox
+to0pt{\kern0.4\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox
+to0pt{\kern0.55\wd0\vrule height0.45\ht0\hss}\box0}}}}
+\def\bbbz{{\mathchoice {\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}}
+{\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}}
+{\hbox{$\mathsf\scriptstyle Z\kern-0.3em Z$}}
+{\hbox{$\mathsf\scriptscriptstyle Z\kern-0.2em Z$}}}}
+
+\let\ts\,
+
+\setlength\leftmargini  {17\p@}
+\setlength\leftmargin    {\leftmargini}
+\setlength\leftmarginii  {\leftmargini}
+\setlength\leftmarginiii {\leftmargini}
+\setlength\leftmarginiv  {\leftmargini}
+\setlength  \labelsep  {.5em}
+\setlength  \labelwidth{\leftmargini}
+\addtolength\labelwidth{-\labelsep}
+
+\def\@listI{\leftmargin\leftmargini
+            \parsep 0\p@ \@plus1\p@ \@minus\p@
+            \topsep 8\p@ \@plus2\p@ \@minus4\p@
+            \itemsep0\p@}
+\let\@listi\@listI
+\@listi
+\def\@listii {\leftmargin\leftmarginii
+              \labelwidth\leftmarginii
+              \advance\labelwidth-\labelsep
+              \topsep    0\p@ \@plus2\p@ \@minus\p@}
+\def\@listiii{\leftmargin\leftmarginiii
+              \labelwidth\leftmarginiii
+              \advance\labelwidth-\labelsep
+              \topsep    0\p@ \@plus\p@\@minus\p@
+              \parsep    \z@
+              \partopsep \p@ \@plus\z@ \@minus\p@}
+
+\renewcommand\labelitemi{\normalfont\bfseries --}
+\renewcommand\labelitemii{$\m@th\bullet$}
+
+\setlength\arraycolsep{1.4\p@}
+\setlength\tabcolsep{1.4\p@}
+
+\def\tableofcontents{\chapter*{\contentsname\@mkboth{{\contentsname}}%
+                                                    {{\contentsname}}}
+ \def\authcount##1{\setcounter{auco}{##1}\setcounter{@auth}{1}}
+ \def\lastand{\ifnum\value{auco}=2\relax
+                 \unskip{} \andname\
+              \else
+                 \unskip \lastandname\
+              \fi}%
+ \def\and{\stepcounter{@auth}\relax
+          \ifnum\value{@auth}=\value{auco}%
+             \lastand
+          \else
+             \unskip,
+          \fi}%
+ \@starttoc{toc}\if@restonecol\twocolumn\fi}
+
+\def\l@part#1#2{\addpenalty{\@secpenalty}%
+   \addvspace{2em plus\p@}%  % space above part line
+   \begingroup
+     \parindent \z@
+     \rightskip \z@ plus 5em
+     \hrule\vskip5pt
+     \large               % same size as for a contribution heading
+     \bfseries\boldmath   % set line in boldface
+     \leavevmode          % TeX command to enter horizontal mode.
+     #1\par
+     \vskip5pt
+     \hrule
+     \vskip1pt
+     \nobreak             % Never break after part entry
+   \endgroup}
+
+\def\@dotsep{2}
+
+\def\hyperhrefextend{\ifx\hyper@anchor\@undefined\else
+{chapter.\thechapter}\fi}
+
+\def\addnumcontentsmark#1#2#3{%
+\addtocontents{#1}{\protect\contentsline{#2}{\protect\numberline
+                     {\thechapter}#3}{\thepage}\hyperhrefextend}}
+\def\addcontentsmark#1#2#3{%
+\addtocontents{#1}{\protect\contentsline{#2}{#3}{\thepage}\hyperhrefextend}}
+\def\addcontentsmarkwop#1#2#3{%
+\addtocontents{#1}{\protect\contentsline{#2}{#3}{0}\hyperhrefextend}}
+
+\def\@adcmk[#1]{\ifcase #1 \or
+\def\@gtempa{\addnumcontentsmark}%
+  \or    \def\@gtempa{\addcontentsmark}%
+  \or    \def\@gtempa{\addcontentsmarkwop}%
+  \fi\@gtempa{toc}{chapter}}
+\def\addtocmark{\@ifnextchar[{\@adcmk}{\@adcmk[3]}}
+
+\def\l@chapter#1#2{\addpenalty{-\@highpenalty}
+ \vskip 1.0em plus 1pt \@tempdima 1.5em \begingroup
+ \parindent \z@ \rightskip \@tocrmarg
+ \advance\rightskip by 0pt plus 2cm
+ \parfillskip -\rightskip \pretolerance=10000
+ \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
+ {\large\bfseries\boldmath#1}\ifx0#2\hfil\null
+ \else
+      \nobreak
+      \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
+      \@dotsep mu$}\hfill
+      \nobreak\hbox to\@pnumwidth{\hss #2}%
+ \fi\par
+ \penalty\@highpenalty \endgroup}
+
+\def\l@title#1#2{\addpenalty{-\@highpenalty}
+ \addvspace{8pt plus 1pt}
+ \@tempdima \z@
+ \begingroup
+ \parindent \z@ \rightskip \@tocrmarg
+ \advance\rightskip by 0pt plus 2cm
+ \parfillskip -\rightskip \pretolerance=10000
+ \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
+ #1\nobreak
+ \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
+ \@dotsep mu$}\hfill
+ \nobreak\hbox to\@pnumwidth{\hss #2}\par
+ \penalty\@highpenalty \endgroup}
+
+\def\l@author#1#2{\addpenalty{\@highpenalty}
+ \@tempdima=\z@ %15\p@
+ \begingroup
+ \parindent \z@ \rightskip \@tocrmarg
+ \advance\rightskip by 0pt plus 2cm
+ \pretolerance=10000
+ \leavevmode \advance\leftskip\@tempdima %\hskip -\leftskip
+ \textit{#1}\par
+ \penalty\@highpenalty \endgroup}
+
+%\setcounter{tocdepth}{0}
+\newdimen\tocchpnum
+\newdimen\tocsecnum
+\newdimen\tocsectotal
+\newdimen\tocsubsecnum
+\newdimen\tocsubsectotal
+\newdimen\tocsubsubsecnum
+\newdimen\tocsubsubsectotal
+\newdimen\tocparanum
+\newdimen\tocparatotal
+\newdimen\tocsubparanum
+\tocchpnum=\z@            % no chapter numbers
+\tocsecnum=15\p@          % section 88. plus 2.222pt
+\tocsubsecnum=23\p@       % subsection 88.8 plus 2.222pt
+\tocsubsubsecnum=27\p@    % subsubsection 88.8.8 plus 1.444pt
+\tocparanum=35\p@         % paragraph 88.8.8.8 plus 1.666pt
+\tocsubparanum=43\p@      % subparagraph 88.8.8.8.8 plus 1.888pt
+\def\calctocindent{%
+\tocsectotal=\tocchpnum
+\advance\tocsectotal by\tocsecnum
+\tocsubsectotal=\tocsectotal
+\advance\tocsubsectotal by\tocsubsecnum
+\tocsubsubsectotal=\tocsubsectotal
+\advance\tocsubsubsectotal by\tocsubsubsecnum
+\tocparatotal=\tocsubsubsectotal
+\advance\tocparatotal by\tocparanum}
+\calctocindent
+
+\def\l@section{\@dottedtocline{1}{\tocchpnum}{\tocsecnum}}
+\def\l@subsection{\@dottedtocline{2}{\tocsectotal}{\tocsubsecnum}}
+\def\l@subsubsection{\@dottedtocline{3}{\tocsubsectotal}{\tocsubsubsecnum}}
+\def\l@paragraph{\@dottedtocline{4}{\tocsubsubsectotal}{\tocparanum}}
+\def\l@subparagraph{\@dottedtocline{5}{\tocparatotal}{\tocsubparanum}}
+
+\def\listoffigures{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
+ \fi\section*{\listfigurename\@mkboth{{\listfigurename}}{{\listfigurename}}}
+ \@starttoc{lof}\if@restonecol\twocolumn\fi}
+\def\l@figure{\@dottedtocline{1}{0em}{1.5em}}
+
+\def\listoftables{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
+ \fi\section*{\listtablename\@mkboth{{\listtablename}}{{\listtablename}}}
+ \@starttoc{lot}\if@restonecol\twocolumn\fi}
+\let\l@table\l@figure
+
+\renewcommand\listoffigures{%
+    \section*{\listfigurename
+      \@mkboth{\listfigurename}{\listfigurename}}%
+    \@starttoc{lof}%
+    }
+
+\renewcommand\listoftables{%
+    \section*{\listtablename
+      \@mkboth{\listtablename}{\listtablename}}%
+    \@starttoc{lot}%
+    }
+
+\ifx\oribibl\undefined
+\ifx\citeauthoryear\undefined
+\renewenvironment{thebibliography}[1]
+     {\section*{\refname}
+      \def\@biblabel##1{##1.}
+      \small
+      \list{\@biblabel{\@arabic\c@enumiv}}%
+           {\settowidth\labelwidth{\@biblabel{#1}}%
+            \leftmargin\labelwidth
+            \advance\leftmargin\labelsep
+            \if@openbib
+              \advance\leftmargin\bibindent
+              \itemindent -\bibindent
+              \listparindent \itemindent
+              \parsep \z@
+            \fi
+            \usecounter{enumiv}%
+            \let\p@enumiv\@empty
+            \renewcommand\theenumiv{\@arabic\c@enumiv}}%
+      \if@openbib
+        \renewcommand\newblock{\par}%
+      \else
+        \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}%
+      \fi
+      \sloppy\clubpenalty4000\widowpenalty4000%
+      \sfcode`\.=\@m}
+     {\def\@noitemerr
+       {\@latex@warning{Empty `thebibliography' environment}}%
+      \endlist}
+\def\@lbibitem[#1]#2{\item[{[#1]}\hfill]\if@filesw
+     {\let\protect\noexpand\immediate
+     \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
+\newcount\@tempcntc
+\def\@citex[#1]#2{\if@filesw\immediate\write\@auxout{\string\citation{#2}}\fi
+  \@tempcnta\z@\@tempcntb\m@ne\def\@citea{}\@cite{\@for\@citeb:=#2\do
+    {\@ifundefined
+       {b@\@citeb}{\@citeo\@tempcntb\m@ne\@citea\def\@citea{,}{\bfseries
+        ?}\@warning
+       {Citation `\@citeb' on page \thepage \space undefined}}%
+    {\setbox\z@\hbox{\global\@tempcntc0\csname b@\@citeb\endcsname\relax}%
+     \ifnum\@tempcntc=\z@ \@citeo\@tempcntb\m@ne
+       \@citea\def\@citea{,}\hbox{\csname b@\@citeb\endcsname}%
+     \else
+      \advance\@tempcntb\@ne
+      \ifnum\@tempcntb=\@tempcntc
+      \else\advance\@tempcntb\m@ne\@citeo
+      \@tempcnta\@tempcntc\@tempcntb\@tempcntc\fi\fi}}\@citeo}{#1}}
+\def\@citeo{\ifnum\@tempcnta>\@tempcntb\else
+               \@citea\def\@citea{,\,\hskip\z@skip}%
+               \ifnum\@tempcnta=\@tempcntb\the\@tempcnta\else
+               {\advance\@tempcnta\@ne\ifnum\@tempcnta=\@tempcntb \else
+                \def\@citea{--}\fi
+      \advance\@tempcnta\m@ne\the\@tempcnta\@citea\the\@tempcntb}\fi\fi}
+\else
+\renewenvironment{thebibliography}[1]
+     {\section*{\refname}
+      \small
+      \list{}%
+           {\settowidth\labelwidth{}%
+            \leftmargin\parindent
+            \itemindent=-\parindent
+            \labelsep=\z@
+            \if@openbib
+              \advance\leftmargin\bibindent
+              \itemindent -\bibindent
+              \listparindent \itemindent
+              \parsep \z@
+            \fi
+            \usecounter{enumiv}%
+            \let\p@enumiv\@empty
+            \renewcommand\theenumiv{}}%
+      \if@openbib
+        \renewcommand\newblock{\par}%
+      \else
+        \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}%
+      \fi
+      \sloppy\clubpenalty4000\widowpenalty4000%
+      \sfcode`\.=\@m}
+     {\def\@noitemerr
+       {\@latex@warning{Empty `thebibliography' environment}}%
+      \endlist}
+      \def\@cite#1{#1}%
+      \def\@lbibitem[#1]#2{\item[]\if@filesw
+        {\def\protect##1{\string ##1\space}\immediate
+      \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
+   \fi
+\else
+\@cons\@openbib@code{\noexpand\small}
+\fi
+
+\def\idxquad{\hskip 10\p@}% space that divides entry from number
+
+\def\@idxitem{\par\hangindent 10\p@}
+
+\def\subitem{\par\setbox0=\hbox{--\enspace}% second order
+                \noindent\hangindent\wd0\box0}% index entry
+
+\def\subsubitem{\par\setbox0=\hbox{--\,--\enspace}% third
+                \noindent\hangindent\wd0\box0}% order index entry
+
+\def\indexspace{\par \vskip 10\p@ plus5\p@ minus3\p@\relax}
+
+\renewenvironment{theindex}
+               {\@mkboth{\indexname}{\indexname}%
+                \thispagestyle{empty}\parindent\z@
+                \parskip\z@ \@plus .3\p@\relax
+                \let\item\par
+                \def\,{\relax\ifmmode\mskip\thinmuskip
+                             \else\hskip0.2em\ignorespaces\fi}%
+                \normalfont\small
+                \begin{multicols}{2}[\@makeschapterhead{\indexname}]%
+                }
+                {\end{multicols}}
+
+\renewcommand\footnoterule{%
+  \kern-3\p@
+  \hrule\@width 2truecm
+  \kern2.6\p@}
+  \newdimen\fnindent
+  \fnindent1em
+\long\def\@makefntext#1{%
+    \parindent \fnindent%
+    \leftskip \fnindent%
+    \noindent
+    \llap{\hb@xt@1em{\hss\@makefnmark\ }}\ignorespaces#1}
+
+\long\def\@makecaption#1#2{%
+  \vskip\abovecaptionskip
+  \sbox\@tempboxa{{\bfseries #1.} #2}%
+  \ifdim \wd\@tempboxa >\hsize
+    {\bfseries #1.} #2\par
+  \else
+    \global \@minipagefalse
+    \hb@xt@\hsize{\hfil\box\@tempboxa\hfil}%
+  \fi
+  \vskip\belowcaptionskip}
+
+\def\fps@figure{htbp}
+\def\fnum@figure{\figurename\thinspace\thefigure}
+\def \@floatboxreset {%
+        \reset@font
+        \small
+        \@setnobreak
+        \@setminipage
+}
+\def\fps@table{htbp}
+\def\fnum@table{\tablename~\thetable}
+\renewenvironment{table}
+               {\setlength\abovecaptionskip{0\p@}%
+                \setlength\belowcaptionskip{10\p@}%
+                \@float{table}}
+               {\end@float}
+\renewenvironment{table*}
+               {\setlength\abovecaptionskip{0\p@}%
+                \setlength\belowcaptionskip{10\p@}%
+                \@dblfloat{table}}
+               {\end@dblfloat}
+
+\long\def\@caption#1[#2]#3{\par\addcontentsline{\csname
+  ext@#1\endcsname}{#1}{\protect\numberline{\csname
+  the#1\endcsname}{\ignorespaces #2}}\begingroup
+    \@parboxrestore
+    \@makecaption{\csname fnum@#1\endcsname}{\ignorespaces #3}\par
+  \endgroup}
+
+% LaTeX does not provide a command to enter the authors institute
+% addresses. The \institute command is defined here.
+
+\newcounter{@inst}
+\newcounter{@auth}
+\newcounter{auco}
+\newdimen\instindent
+\newbox\authrun
+\newtoks\authorrunning
+\newtoks\tocauthor
+\newbox\titrun
+\newtoks\titlerunning
+\newtoks\toctitle
+
+\def\clearheadinfo{\gdef\@author{No Author Given}%
+                   \gdef\@title{No Title Given}%
+                   \gdef\@subtitle{}%
+                   \gdef\@institute{No Institute Given}%
+                   \gdef\@thanks{}%
+                   \global\titlerunning={}\global\authorrunning={}%
+                   \global\toctitle={}\global\tocauthor={}}
+
+\def\institute#1{\gdef\@institute{#1}}
+
+\def\institutename{\par
+ \begingroup
+ \parskip=\z@
+ \parindent=\z@
+ \setcounter{@inst}{1}%
+ \def\and{\par\stepcounter{@inst}%
+ \noindent$^{\the@inst}$\enspace\ignorespaces}%
+ \setbox0=\vbox{\def\thanks##1{}\@institute}%
+ \ifnum\c@@inst=1\relax
+   \gdef\fnnstart{0}%
+ \else
+   \xdef\fnnstart{\c@@inst}%
+   \setcounter{@inst}{1}%
+   \noindent$^{\the@inst}$\enspace
+ \fi
+ \ignorespaces
+ \@institute\par
+ \endgroup}
+
+\def\@fnsymbol#1{\ensuremath{\ifcase#1\or\star\or{\star\star}\or
+   {\star\star\star}\or \dagger\or \ddagger\or
+   \mathchar "278\or \mathchar "27B\or \|\or **\or \dagger\dagger
+   \or \ddagger\ddagger \else\@ctrerr\fi}}
+
+\def\inst#1{\unskip$^{#1}$}
+\def\fnmsep{\unskip$^,$}
+\def\email#1{{\tt#1}}
+\AtBeginDocument{\@ifundefined{url}{\def\url#1{#1}}{}%
+\@ifpackageloaded{babel}{%
+\@ifundefined{extrasenglish}{}{\addto\extrasenglish{\switcht@albion}}%
+\@ifundefined{extrasfrenchb}{}{\addto\extrasfrenchb{\switcht@francais}}%
+\@ifundefined{extrasgerman}{}{\addto\extrasgerman{\switcht@deutsch}}%
+}{\switcht@@therlang}%
+}
+\def\homedir{\~{ }}
+
+\def\subtitle#1{\gdef\@subtitle{#1}}
+\clearheadinfo
+
+\renewcommand\maketitle{\newpage
+  \refstepcounter{chapter}%
+  \stepcounter{section}%
+  \setcounter{section}{0}%
+  \setcounter{subsection}{0}%
+  \setcounter{figure}{0}
+  \setcounter{table}{0}
+  \setcounter{equation}{0}
+  \setcounter{footnote}{0}%
+  \begingroup
+    \parindent=\z@
+    \renewcommand\thefootnote{\@fnsymbol\c@footnote}%
+    \if@twocolumn
+      \ifnum \col@number=\@ne
+        \@maketitle
+      \else
+        \twocolumn[\@maketitle]%
+      \fi
+    \else
+      \newpage
+      \global\@topnum\z@   % Prevents figures from going at top of page.
+      \@maketitle
+    \fi
+    \thispagestyle{empty}\@thanks
+%
+    \def\\{\unskip\ \ignorespaces}\def\inst##1{\unskip{}}%
+    \def\thanks##1{\unskip{}}\def\fnmsep{\unskip}%
+    \instindent=\hsize
+    \advance\instindent by-\headlineindent
+%    \if!\the\toctitle!\addcontentsline{toc}{title}{\@title}\else
+%       \addcontentsline{toc}{title}{\the\toctitle}\fi
+    \if@runhead
+       \if!\the\titlerunning!\else
+         \edef\@title{\the\titlerunning}%
+       \fi
+       \global\setbox\titrun=\hbox{\small\rm\unboldmath\ignorespaces\@title}%
+       \ifdim\wd\titrun>\instindent
+          \typeout{Title too long for running head. Please supply}%
+          \typeout{a shorter form with \string\titlerunning\space prior to
+                   \string\maketitle}%
+          \global\setbox\titrun=\hbox{\small\rm
+          Title Suppressed Due to Excessive Length}%
+       \fi
+       \xdef\@title{\copy\titrun}%
+    \fi
+%
+    \if!\the\tocauthor!\relax
+      {\def\and{\noexpand\protect\noexpand\and}%
+      \protected@xdef\toc@uthor{\@author}}%
+    \else
+      \def\\{\noexpand\protect\noexpand\newline}%
+      \protected@xdef\scratch{\the\tocauthor}%
+      \protected@xdef\toc@uthor{\scratch}%
+    \fi
+%    \addcontentsline{toc}{author}{\toc@uthor}%
+    \if@runhead
+       \if!\the\authorrunning!
+         \value{@inst}=\value{@auth}%
+         \setcounter{@auth}{1}%
+       \else
+         \edef\@author{\the\authorrunning}%
+       \fi
+       \global\setbox\authrun=\hbox{\small\unboldmath\@author\unskip}%
+       \ifdim\wd\authrun>\instindent
+          \typeout{Names of authors too long for running head. Please supply}%
+          \typeout{a shorter form with \string\authorrunning\space prior to
+                   \string\maketitle}%
+          \global\setbox\authrun=\hbox{\small\rm
+          Authors Suppressed Due to Excessive Length}%
+       \fi
+       \xdef\@author{\copy\authrun}%
+       \markboth{\@author}{\@title}%
+     \fi
+  \endgroup
+  \setcounter{footnote}{\fnnstart}%
+  \clearheadinfo}
+%
+\def\@maketitle{\newpage
+ \markboth{}{}%
+ \def\lastand{\ifnum\value{@inst}=2\relax
+                 \unskip{} \andname\
+              \else
+                 \unskip \lastandname\
+              \fi}%
+ \def\and{\stepcounter{@auth}\relax
+          \ifnum\value{@auth}=\value{@inst}%
+             \lastand
+          \else
+             \unskip,
+          \fi}%
+ \begin{center}%
+ \let\newline\\
+ {\Large \bfseries\boldmath
+  \pretolerance=10000
+  \@title \par}\vskip .8cm
+\if!\@subtitle!\else {\large \bfseries\boldmath
+  \vskip -.65cm
+  \pretolerance=10000
+  \@subtitle \par}\vskip .8cm\fi
+ \setbox0=\vbox{\setcounter{@auth}{1}\def\and{\stepcounter{@auth}}%
+ \def\thanks##1{}\@author}%
+ \global\value{@inst}=\value{@auth}%
+ \global\value{auco}=\value{@auth}%
+ \setcounter{@auth}{1}%
+{\lineskip .5em
+\noindent\ignorespaces
+\@author\vskip.35cm}
+ {\small\institutename}
+ \end{center}%
+ }
+
+% definition of the "\spnewtheorem" command.
+%
+% Usage:
+%
+%     \spnewtheorem{env_nam}{caption}[within]{cap_font}{body_font}
+% or  \spnewtheorem{env_nam}[numbered_like]{caption}{cap_font}{body_font}
+% or  \spnewtheorem*{env_nam}{caption}{cap_font}{body_font}
+%
+% New is "cap_font" and "body_font". It stands for
+% fontdefinition of the caption and the text itself.
+%
+% "\spnewtheorem*" gives a theorem without number.
+%
+% A defined spnewthoerem environment is used as described
+% by Lamport.
+%
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+
+\def\@thmcountersep{}
+\def\@thmcounterend{.}
+
+\def\spnewtheorem{\@ifstar{\@sthm}{\@Sthm}}
+
+% definition of \spnewtheorem with number
+
+\def\@spnthm#1#2{%
+  \@ifnextchar[{\@spxnthm{#1}{#2}}{\@spynthm{#1}{#2}}}
+\def\@Sthm#1{\@ifnextchar[{\@spothm{#1}}{\@spnthm{#1}}}
+
+\def\@spxnthm#1#2[#3]#4#5{\expandafter\@ifdefinable\csname #1\endcsname
+   {\@definecounter{#1}\@addtoreset{#1}{#3}%
+   \expandafter\xdef\csname the#1\endcsname{\expandafter\noexpand
+     \csname the#3\endcsname \noexpand\@thmcountersep \@thmcounter{#1}}%
+   \expandafter\xdef\csname #1name\endcsname{#2}%
+   \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#4}{#5}}%
+                              \global\@namedef{end#1}{\@endtheorem}}}
+
+\def\@spynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
+   {\@definecounter{#1}%
+   \expandafter\xdef\csname the#1\endcsname{\@thmcounter{#1}}%
+   \expandafter\xdef\csname #1name\endcsname{#2}%
+   \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#3}{#4}}%
+                               \global\@namedef{end#1}{\@endtheorem}}}
+
+\def\@spothm#1[#2]#3#4#5{%
+  \@ifundefined{c@#2}{\@latexerr{No theorem environment `#2' defined}\@eha}%
+  {\expandafter\@ifdefinable\csname #1\endcsname
+  {\global\@namedef{the#1}{\@nameuse{the#2}}%
+  \expandafter\xdef\csname #1name\endcsname{#3}%
+  \global\@namedef{#1}{\@spthm{#2}{\csname #1name\endcsname}{#4}{#5}}%
+  \global\@namedef{end#1}{\@endtheorem}}}}
+
+\def\@spthm#1#2#3#4{\topsep 7\p@ \@plus2\p@ \@minus4\p@
+\refstepcounter{#1}%
+\@ifnextchar[{\@spythm{#1}{#2}{#3}{#4}}{\@spxthm{#1}{#2}{#3}{#4}}}
+
+\def\@spxthm#1#2#3#4{\@spbegintheorem{#2}{\csname the#1\endcsname}{#3}{#4}%
+                    \ignorespaces}
+
+\def\@spythm#1#2#3#4[#5]{\@spopargbegintheorem{#2}{\csname
+       the#1\endcsname}{#5}{#3}{#4}\ignorespaces}
+
+\def\@spbegintheorem#1#2#3#4{\trivlist
+                 \item[\hskip\labelsep{#3#1\ #2\@thmcounterend}]#4}
+
+\def\@spopargbegintheorem#1#2#3#4#5{\trivlist
+      \item[\hskip\labelsep{#4#1\ #2}]{#4(#3)\@thmcounterend\ }#5}
+
+% definition of \spnewtheorem* without number
+
+\def\@sthm#1#2{\@Ynthm{#1}{#2}}
+
+\def\@Ynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
+   {\global\@namedef{#1}{\@Thm{\csname #1name\endcsname}{#3}{#4}}%
+    \expandafter\xdef\csname #1name\endcsname{#2}%
+    \global\@namedef{end#1}{\@endtheorem}}}
+
+\def\@Thm#1#2#3{\topsep 7\p@ \@plus2\p@ \@minus4\p@
+\@ifnextchar[{\@Ythm{#1}{#2}{#3}}{\@Xthm{#1}{#2}{#3}}}
+
+\def\@Xthm#1#2#3{\@Begintheorem{#1}{#2}{#3}\ignorespaces}
+
+\def\@Ythm#1#2#3[#4]{\@Opargbegintheorem{#1}
+       {#4}{#2}{#3}\ignorespaces}
+
+\def\@Begintheorem#1#2#3{#3\trivlist
+                           \item[\hskip\labelsep{#2#1\@thmcounterend}]}
+
+\def\@Opargbegintheorem#1#2#3#4{#4\trivlist
+      \item[\hskip\labelsep{#3#1}]{#3(#2)\@thmcounterend\ }}
+
+\if@envcntsect
+   \def\@thmcountersep{.}
+   \spnewtheorem{theorem}{Theorem}[section]{\bfseries}{\itshape}
+\else
+   \spnewtheorem{theorem}{Theorem}{\bfseries}{\itshape}
+   \if@envcntreset
+      \@addtoreset{theorem}{section}
+   \else
+      \@addtoreset{theorem}{chapter}
+   \fi
+\fi
+
+%definition of divers theorem environments
+\spnewtheorem*{claim}{Claim}{\itshape}{\rmfamily}
+\spnewtheorem*{proof}{Proof}{\itshape}{\rmfamily}
+\if@envcntsame % alle Umgebungen wie Theorem.
+   \def\spn@wtheorem#1#2#3#4{\@spothm{#1}[theorem]{#2}{#3}{#4}}
+\else % alle Umgebungen mit eigenem Zaehler
+   \if@envcntsect % mit section numeriert
+      \def\spn@wtheorem#1#2#3#4{\@spxnthm{#1}{#2}[section]{#3}{#4}}
+   \else % nicht mit section numeriert
+      \if@envcntreset
+         \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
+                                   \@addtoreset{#1}{section}}
+      \else
+         \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
+                                   \@addtoreset{#1}{chapter}}%
+      \fi
+   \fi
+\fi
+\spn@wtheorem{case}{Case}{\itshape}{\rmfamily}
+\spn@wtheorem{conjecture}{Conjecture}{\itshape}{\rmfamily}
+\spn@wtheorem{corollary}{Corollary}{\bfseries}{\itshape}
+\spn@wtheorem{definition}{Definition}{\bfseries}{\itshape}
+\spn@wtheorem{example}{Example}{\itshape}{\rmfamily}
+\spn@wtheorem{exercise}{Exercise}{\itshape}{\rmfamily}
+\spn@wtheorem{lemma}{Lemma}{\bfseries}{\itshape}
+\spn@wtheorem{note}{Note}{\itshape}{\rmfamily}
+\spn@wtheorem{problem}{Problem}{\itshape}{\rmfamily}
+\spn@wtheorem{property}{Property}{\itshape}{\rmfamily}
+\spn@wtheorem{proposition}{Proposition}{\bfseries}{\itshape}
+\spn@wtheorem{question}{Question}{\itshape}{\rmfamily}
+\spn@wtheorem{solution}{Solution}{\itshape}{\rmfamily}
+\spn@wtheorem{remark}{Remark}{\itshape}{\rmfamily}
+
+\def\@takefromreset#1#2{%
+    \def\@tempa{#1}%
+    \let\@tempd\@elt
+    \def\@elt##1{%
+        \def\@tempb{##1}%
+        \ifx\@tempa\@tempb\else
+            \@addtoreset{##1}{#2}%
+        \fi}%
+    \expandafter\expandafter\let\expandafter\@tempc\csname cl@#2\endcsname
+    \expandafter\def\csname cl@#2\endcsname{}%
+    \@tempc
+    \let\@elt\@tempd}
+
+\def\theopargself{\def\@spopargbegintheorem##1##2##3##4##5{\trivlist
+      \item[\hskip\labelsep{##4##1\ ##2}]{##4##3\@thmcounterend\ }##5}
+                  \def\@Opargbegintheorem##1##2##3##4{##4\trivlist
+      \item[\hskip\labelsep{##3##1}]{##3##2\@thmcounterend\ }}
+      }
+
+\renewenvironment{abstract}{%
+      \list{}{\advance\topsep by0.35cm\relax\small
+      \leftmargin=1cm
+      \labelwidth=\z@
+      \listparindent=\z@
+      \itemindent\listparindent
+      \rightmargin\leftmargin}\item[\hskip\labelsep
+                                    \bfseries\abstractname]}
+    {\endlist}
+
+\newdimen\headlineindent             % dimension for space between
+\headlineindent=1.166cm              % number and text of headings.
+
+\def\ps@headings{\let\@mkboth\@gobbletwo
+   \let\@oddfoot\@empty\let\@evenfoot\@empty
+   \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}%
+                  \leftmark\hfil}
+   \def\@oddhead{\normalfont\small\hfil\rightmark\hspace{\headlineindent}%
+                 \llap{\thepage}}
+   \def\chaptermark##1{}%
+   \def\sectionmark##1{}%
+   \def\subsectionmark##1{}}
+
+\def\ps@titlepage{\let\@mkboth\@gobbletwo
+   \let\@oddfoot\@empty\let\@evenfoot\@empty
+   \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}%
+                  \hfil}
+   \def\@oddhead{\normalfont\small\hfil\hspace{\headlineindent}%
+                 \llap{\thepage}}
+   \def\chaptermark##1{}%
+   \def\sectionmark##1{}%
+   \def\subsectionmark##1{}}
+
+\if@runhead\ps@headings\else
+\ps@empty\fi
+
+\setlength\arraycolsep{1.4\p@}
+\setlength\tabcolsep{1.4\p@}
+
+\endinput
+%end of file llncs.cls
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys2/Journal/document/root.bib	Mon Nov 01 10:40:21 2021 +0000
@@ -0,0 +1,340 @@
+@article{HosoyaVouillonPierce2005,
+  author = {H.~Hosoya and J.~Vouillon and B.~C.~Pierce},
+  title = {{R}egular {E}xpression {T}ypes for {XML}},
+  journal = {ACM Transactions on Programming Languages and Systems (TOPLAS)},
+  year = {2005},
+  volume = 27,
+  number = 1,
+  pages = {46--90}
+}
+
+
+@Misc{POSIX,
+  title =     {{T}he {O}pen {G}roup {B}ase {S}pecification {I}ssue 6 {IEEE} {S}td 1003.1 2004 {E}dition},
+  year =      {2004},
+  note =    {\url{http://pubs.opengroup.org/onlinepubs/009695399/basedefs/xbd_chap09.html}}
+}
+
+
+@InProceedings{AusafDyckhoffUrban2016,
+  author =   {F.~Ausaf and R.~Dyckhoff and C.~Urban},
+  title = 	 {{POSIX} {L}exing with {D}erivatives of {R}egular {E}xpressions ({P}roof {P}earl)},
+  year = 	 {2016},
+  booktitle = {Proc.~of the 7th International Conference on Interactive Theorem Proving (ITP)},
+  volume = 	 {9807},
+  series = 	 {LNCS},
+  pages =        {69--86}
+}
+
+@article{aduAFP16,
+  author =   {F.~Ausaf and R.~Dyckhoff and C.~Urban},
+  title =    {{POSIX} {L}exing with {D}erivatives of {R}egular {E}xpressions},
+  journal =  {Archive of Formal Proofs},
+  year =     2016,
+  note =     {\url{http://www.isa-afp.org/entries/Posix-Lexing.shtml}, Formal proof development},
+  ISSN =     {2150-914x}
+}
+
+
+@TechReport{CrashCourse2014,
+  author =       {N.~B.~B.~Grathwohl and F.~Henglein and U.~T.~Rasmussen},
+  title =        {{A} {C}rash-{C}ourse in {R}egular {E}xpression {P}arsing and {R}egular 
+                  {E}xpressions as {T}ypes},
+  institution =  {University of Copenhagen},
+  year =         {2014},
+  annote =       {draft report}
+}
+
+@inproceedings{Sulzmann2014,
+  author    = {M.~Sulzmann and K.~Lu},
+  title     = {{POSIX} {R}egular {E}xpression {P}arsing with {D}erivatives},
+  booktitle = {Proc.~of the 12th International Conference on Functional and Logic Programming (FLOPS)},
+  pages     = {203--220},
+  year      = {2014},
+  volume =    {8475},
+  series =    {LNCS}
+}
+
+@inproceedings{Sulzmann2014b,
+  author    = {M.~Sulzmann and P.~van Steenhoven},
+  title     = {{A} {F}lexible and {E}fficient {ML} {L}exer {T}ool {B}ased on {E}xtended {R}egular
+               {E}xpression {S}ubmatching},
+  booktitle = {Proc.~of the 23rd International Conference on Compiler Construction (CC)},
+  pages     = {174--191},
+  year      = {2014},
+  volume    = {8409},
+  series =    {LNCS}
+}
+
+@book{Pierce2015,
+  author = {B.~C.~Pierce and C.~Casinghino and M.~Gaboardi and
+            M.~Greenberg and C.~Hri\c{t}cu and 
+            V.~Sj\"{o}berg and B.~Yorgey},
+  title = {{S}oftware {F}oundations},
+  year = {2015},
+  publisher = {Electronic textbook},
+  note = {\url{http://www.cis.upenn.edu/~bcpierce/sf}}
+}
+
+@Misc{Kuklewicz,
+  author = 	 {C.~Kuklewicz},
+  title = 	 {{R}egex {P}osix},
+  howpublished = "\url{https://wiki.haskell.org/Regex_Posix}"
+}
+
+@article{Vansummeren2006,
+  author = {S.~Vansummeren},
+  title = {{T}ype {I}nference for {U}nique {P}attern {M}atching},
+  year = {2006},
+  journal = {ACM Transactions on Programming Languages and Systems},
+  volume = {28},
+  number = {3},
+  pages = {389--428}
+}
+
+@InProceedings{Asperti12,
+  author =       {A.~Asperti},
+  title =        {{A} {C}ompact {P}roof of {D}ecidability for {R}egular {E}xpression {E}quivalence},
+  booktitle = {Proc.~of the 3rd International Conference on Interactive Theorem Proving (ITP)},
+  pages =     {283--298},
+  year =      {2012},
+  volume =    {7406},
+  series =    {LNCS}
+}
+
+@inproceedings{Frisch2004,
+  author    = {A.~Frisch and L.~Cardelli},
+  title     = {{G}reedy {R}egular {E}xpression {M}atching},
+  booktitle = {Proc.~of the 31st International Conference on Automata, Languages and Programming (ICALP)},
+  pages     = {618--629},
+  year      = {2004},
+  volume =    {3142},
+  series =    {LNCS}
+}
+
+@ARTICLE{Antimirov95,
+    author = {V.~Antimirov},
+    title = {{P}artial {D}erivatives of {R}egular {E}xpressions and 
+     {F}inite {A}utomata {C}onstructions},
+    journal = {Theoretical Computer Science},
+    year = {1995},
+    volume = {155},
+    pages = {291--319}
+}
+
+@inproceedings{Nipkow98,
+ author={T.~Nipkow},
+ title={{V}erified {L}exical {A}nalysis},
+ booktitle={Proc.~of the 11th International Conference on Theorem Proving in Higher Order Logics (TPHOLs)},
+ series={LNCS},
+ volume=1479,
+ pages={1--15},
+ year=1998
+}
+
+@article{Brzozowski1964,
+  author    = {J.~A.~Brzozowski},
+  title     = {{D}erivatives of {R}egular {E}xpressions},
+  journal   = {Journal of the {ACM}},
+  volume    = {11},
+  number    = {4},
+  pages     = {481--494},
+  year      = {1964}
+}
+
+@article{Leroy2009,
+  author = {X.~Leroy},
+  title = {{F}ormal {V}erification of a {R}ealistic {C}ompiler},
+  journal = {Communications of the ACM},
+  year = 2009,
+  volume = 52,
+  number = 7,
+  pages = {107--115}
+}
+
+@InProceedings{Paulson2015,
+  author = 	 {L.~C.~Paulson},
+  title = 	 {{A} {F}ormalisation of {F}inite {A}utomata {U}sing {H}ereditarily {F}inite {S}ets},
+  booktitle = {Proc.~of the 25th International Conference on Automated Deduction (CADE)},
+  pages = 	 {231--245},
+  year = 	 {2015},
+  volume = 	 {9195},
+  series = 	 {LNAI}
+}
+
+@Article{Wu2014,
+  author = 	 {C.~Wu and X.~Zhang and C.~Urban},
+  title = 	 {{A} {F}ormalisation of the {M}yhill-{N}erode {T}heorem based on {R}egular {E}xpressions},
+  journal = 	 {Journal of Automatic Reasoning},
+  year = 	 {2014},
+  volume = 	 {52},
+  number = 	 {4},
+  pages = 	 {451--480}
+}
+
+@InProceedings{Regehr2011,
+  author = 	 {X.~Yang and Y.~Chen and E.~Eide and J.~Regehr},
+  title = 	 {{F}inding and {U}nderstanding {B}ugs in {C} {C}ompilers},
+  pages = 	 {283--294},
+  year = 	 {2011},
+  booktitle =    {Proc.~of the 32nd ACM SIGPLAN Conference on Programming Language Design and 
+                  Implementation (PLDI)}
+}
+
+@Article{Norrish2014,
+  author = 	 {A.~Barthwal and M.~Norrish},
+  title = 	 {{A} {M}echanisation of {S}ome {C}ontext-{F}ree {L}anguage {T}heory in {HOL4}},
+  journal          = {Journal of Computer and System Sciences},
+  year = 	 {2014},
+  volume = 	 {80},
+  number = 	 {2},
+  pages = 	 {346--362}
+}
+
+@Article{Thompson1968,
+ author = {K.~Thompson},
+ title = {{P}rogramming {T}echniques: {R}egular {E}xpression {S}earch {A}lgorithm},
+ journal = {Communications of the ACM},
+ issue_date = {June 1968},
+ volume = {11},
+ number = {6},
+ year = {1968},
+ pages = {419--422}
+}
+
+@article{Owens2009,
+  author    = {S.~Owens and J.~H.~Reppy and A.~Turon},
+  title     = {{R}egular-{E}xpression {D}erivatives {R}e-{E}xamined},
+  journal   = {Journal of Functinal Programming},
+  volume    = {19},
+  number    = {2},
+  pages     = {173--190},
+  year      = {2009}
+}
+
+@inproceedings{Sulzmann2015,
+  author    = {M.~Sulzmann and P.~Thiemann},
+  title     = {{D}erivatives for {R}egular {S}huffle {E}xpressions},
+  booktitle = {Proc.~of the 9th International Conference on Language and Automata Theory 
+               and Applications (LATA)},
+  pages     = {275--286},
+  year      = {2015},
+  volume = 	 {8977},
+  series = 	 {LNCS}
+}
+
+@inproceedings{Chen2012,
+  author    = {H.~Chen and S.~Yu},
+  title     = {{D}erivatives of {R}egular {E}xpressions and an {A}pplication},
+  booktitle = {Proc.~in the International Workshop on Theoretical
+               Computer Science (WTCS)},
+  pages     = {343--356},
+  year      = {2012},
+  volume = 	 {7160},
+  series = 	 {LNCS}
+}
+
+@article{Krauss2011,
+  author={A.~Krauss and T.~Nipkow},
+  title={{P}roof {P}earl: {R}egular {E}xpression {E}quivalence and {R}elation {A}lgebra},
+  journal={Journal of Automated Reasoning},
+  volume=49,
+  pages={95--106},
+  year=2012
+}
+
+@InProceedings{Traytel2015,
+  author =	{D.~Traytel},
+  title =	{{A} {C}oalgebraic {D}ecision {P}rocedure for {WS1S}},
+  booktitle =	{Proc.~of the 24th Annual Conference on Computer Science Logic (CSL)},
+  pages =	{487--503},
+  series =	{LIPIcs},
+  year =	{2015},
+  volume =	{41}
+}
+
+@inproceedings{Traytel2013,
+  author={D.~Traytel and T.~Nipkow},
+  title={{A} {V}erified {D}ecision {P}rocedure for {MSO} on 
+         {W}ords {B}ased on {D}erivatives of {R}egular {E}xpressions},
+  booktitle={Proc.~of the 18th ACM SIGPLAN International Conference on Functional Programming (ICFP)},
+  pages={3-12},
+  year=2013
+}
+
+@InProceedings{Coquand2012,
+  author =       {T.~Coquand and V.~Siles},
+  title =        {{A} {D}ecision {P}rocedure for {R}egular {E}xpression {E}quivalence in {T}ype {T}heory},
+  booktitle = {Proc.~of the 1st International Conference on Certified Programs and Proofs (CPP)},
+  pages =     {119--134},
+  year =      {2011},
+  volume =    {7086},
+  series =    {LNCS}
+}
+
+@InProceedings{Almeidaetal10,
+  author =       {J.~B.~Almeida and N.~Moriera and D.~Pereira and S.~M.~de Sousa},
+  title =        {{P}artial {D}erivative {A}utomata {F}ormalized in {C}oq},
+  booktitle =    {Proc.~of the 15th International Conference on Implementation
+                  and Application of Automata (CIAA)},
+  pages =        {59-68},
+  year =         {2010},
+  volume =       {6482},
+  series =       {LNCS}
+}
+
+@article{Owens2008,
+  author    = {S.~Owens and K.~Slind},
+  title     = {{A}dapting {F}unctional {P}rograms to {H}igher {O}rder {L}ogic},
+  journal   = {Higher-Order and Symbolic Computation},
+  volume    = {21},
+  number    = {4},
+  year      = {2008},
+  pages     = {377--409}
+}
+
+
+@article{Owens2,
+  author    = {S.~Owens and K.~Slind},
+  title     = {Adapting functional programs to higher order logic},
+  journal   = {Higher-Order and Symbolic Computation},
+  volume    = {21},
+  number    = {4},
+  pages     = {377--409},
+  year      = {2008},
+  url       = {http://dx.doi.org/10.1007/s10990-008-9038-0},
+  doi       = {10.1007/s10990-008-9038-0},
+  timestamp = {Wed, 16 Dec 2009 13:51:02 +0100},
+  biburl    = {http://dblp.uni-trier.de/rec/bib/journals/lisp/OwensS08},
+  bibsource = {dblp computer science bibliography, http://dblp.org}
+}
+
+@misc{PCRE,
+  title = "{PCRE - Perl Compatible Regular Expressions}",
+  url = {http://www.pcre.org},
+}
+
+
+
+@InProceedings{OkuiSuzuki2010,
+  author =       {S.~Okui and T.~Suzuki},
+  title =        {{D}isambiguation in {R}egular {E}xpression {M}atching via
+                  {P}osition {A}utomata with {A}ugmented {T}ransitions},
+  booktitle = {Proc.~of the 15th International Conference on Implementation
+                  and Application of Automata (CIAA)},
+  year =      {2010},
+  volume =    {6482},
+  series =    {LNCS},
+  pages =     {231--240}
+}
+
+
+
+@TechReport{OkuiSuzukiTech,
+  author =       {S.~Okui and T.~Suzuki},
+  title =        {{D}isambiguation in {R}egular {E}xpression {M}atching via
+                  {P}osition {A}utomata with {A}ugmented {T}ransitions},
+  institution =  {University of Aizu},
+  year =         {2013}
+}
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys2/Journal/document/root.tex	Mon Nov 01 10:40:21 2021 +0000
@@ -0,0 +1,101 @@
+\documentclass[runningheads]{llncs}
+\usepackage{times}
+\usepackage{isabelle}
+\usepackage{isabellesym}
+\usepackage{amsmath}
+\usepackage{amssymb}
+\usepackage{mathpartir}
+\usepackage{tikz}
+\usepackage{pgf}
+\usetikzlibrary{positioning}
+\usepackage{pdfsetup}
+\usepackage{stmaryrd}
+\usepackage{url}
+\usepackage{color}
+\usepackage[safe]{tipa}
+\usepackage[sc]{mathpazo}
+\usepackage{fontspec}
+%\setmainfont[Ligatures=TeX]{Palatino Linotype}
+
+
+\titlerunning{POSIX Lexing with Derivatives of Regular Expressions}
+
+\urlstyle{rm}
+\isabellestyle{it}
+\renewcommand{\isastyleminor}{\it}% 
+\renewcommand{\isastyle}{\normalsize\it}%
+
+
+\def\dn{\,\stackrel{\mbox{\scriptsize def}}{=}\,}
+\renewcommand{\isasymequiv}{$\dn$}
+\renewcommand{\isasymemptyset}{$\varnothing$}
+\renewcommand{\isacharunderscore}{\mbox{$\_\!\_$}}
+\renewcommand{\isasymiota}{\makebox[0mm]{${}^{\prime}$}}
+\renewcommand{\isasymin}{\ensuremath{\,\in\,}}
+
+
+\def\Brz{Brzozowski}
+\def\der{\backslash}
+\newtheorem{falsehood}{Falsehood}
+\newtheorem{conject}{Conjecture}
+
+\begin{document}
+\renewcommand{\thefootnote}{$\star$} \footnotetext[1]{This paper is a
+  revised and expanded version of \cite{AusafDyckhoffUrban2016}.
+  Compared with that paper we give a second definition for POSIX
+  values introduced by Okui Suzuki \cite{OkuiSuzuki2010,OkuiSuzukiTech}
+  and prove that it is
+  equivalent to our original one. This
+  second definition is based on an ordering of values and very similar to,
+  but not equivalent with, the
+  definition given by Sulzmann and Lu~\cite{Sulzmann2014}.
+  The advantage of the definition based on the
+  ordering is that it implements more directly the informal rules from the
+  POSIX standard.
+  We also prove Sulzmann \& Lu's conjecture that their bitcoded version
+  of the POSIX algorithm is correct. Furthermore we extend our results to additional constructors of regular
+  expressions.}  \renewcommand{\thefootnote}{\arabic{footnote}}
+
+
+\title{POSIX {L}exing with {D}erivatives of {R}egular {E}xpressions}
+\author{Fahad Ausaf\inst{1} \and Roy Dyckhoff\inst{2} \and Christian Urban\inst{3}}
+\institute{King's College London\\
+		\email{fahad.ausaf@icloud.com}
+\and University of St Andrews\\
+		\email{roy.dyckhoff@st-andrews.ac.uk}
+\and King's College London\\
+		\email{christian.urban@kcl.ac.uk}}
+\maketitle
+
+\begin{abstract}
+Brzozowski introduced the notion of derivatives for regular
+expressions. They can be used for a very simple regular expression
+matching algorithm.  Sulzmann and Lu cleverly extended this algorithm
+in order to deal with POSIX matching, which is the underlying
+disambiguation strategy for regular expressions needed in lexers.
+Their algorithm generates POSIX values which encode the information of
+\emph{how} a regular expression matches a string---that is, which part
+of the string is matched by which part of the regular expression.  In
+this paper we give our inductive definition of what a POSIX value is
+and show $(i)$ that such a value is unique (for given regular
+expression and string being matched) and $(ii)$ that Sulzmann and Lu's
+algorithm always generates such a value (provided that the regular
+expression matches the string). We show that $(iii)$ our inductive
+definition of a POSIX value is equivalent to an alternative definition
+by Okui and Suzuki which identifies POSIX values as least elements
+according to an ordering of values.  We also prove the correctness of
+Sulzmann's bitcoded version of the POSIX matching algorithm and extend the
+results to additional constructors for regular expressions.  \smallskip
+
+{\bf Keywords:} POSIX matching, Derivatives of Regular Expressions,
+Isabelle/HOL
+\end{abstract}
+
+\input{session}
+
+\end{document}
+
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: t
+%%% End:
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys2/Journal/isabelle.sty	Mon Nov 01 10:40:21 2021 +0000
@@ -0,0 +1,267 @@
+%%
+%% macros for Isabelle generated LaTeX output
+%%
+
+%%% Simple document preparation (based on theory token language and symbols)
+
+% isabelle environments
+
+\newcommand{\isabellecontext}{UNKNOWN}
+\newcommand{\setisabellecontext}[1]{\def\isabellecontext{#1}}
+
+\newcommand{\isastyle}{\UNDEF}
+\newcommand{\isastylett}{\UNDEF}
+\newcommand{\isastyleminor}{\UNDEF}
+\newcommand{\isastyleminortt}{\UNDEF}
+\newcommand{\isastylescript}{\UNDEF}
+\newcommand{\isastyletext}{\normalsize\normalfont\rmfamily}
+\newcommand{\isastyletxt}{\normalfont\rmfamily}
+\newcommand{\isastylecmt}{\normalfont\rmfamily}
+
+\newcommand{\isaspacing}{%
+  \sfcode 42 1000 % .
+  \sfcode 63 1000 % ?
+  \sfcode 33 1000 % !
+  \sfcode 58 1000 % :
+  \sfcode 59 1000 % ;
+  \sfcode 44 1000 % ,
+}
+
+%symbol markup -- \emph achieves decent spacing via italic corrections
+\newcommand{\isamath}[1]{\emph{$#1$}}
+\newcommand{\isatext}[1]{\emph{#1}}
+\DeclareRobustCommand{\isascriptstyle}{\def\isamath##1{##1}\def\isatext##1{\mbox{\isaspacing\isastylescript##1}}}
+\newcommand{\isactrlsub}[1]{\emph{\isascriptstyle${}\sb{#1}$}}
+\newcommand{\isactrlsup}[1]{\emph{\isascriptstyle${}\sp{#1}$}}
+\DeclareRobustCommand{\isactrlbsub}{\emph\bgroup\math{}\sb\bgroup\mbox\bgroup\isaspacing\isastylescript}
+\DeclareRobustCommand{\isactrlesub}{\egroup\egroup\endmath\egroup}
+\DeclareRobustCommand{\isactrlbsup}{\emph\bgroup\math{}\sp\bgroup\mbox\bgroup\isaspacing\isastylescript}
+\DeclareRobustCommand{\isactrlesup}{\egroup\egroup\endmath\egroup}
+\newcommand{\isactrlbold}[1]{{\bfseries\upshape\boldmath#1}}
+
+\newenvironment{isaantiq}{{\isacharat\isacharbraceleft}}{{\isacharbraceright}}
+
+\newdimen\isa@parindent\newdimen\isa@parskip
+
+\newenvironment{isabellebody}{%
+\isamarkuptrue\par%
+\isa@parindent\parindent\parindent0pt%
+\isa@parskip\parskip\parskip0pt%
+\isaspacing\isastyle}{\par}
+
+\newenvironment{isabellebodytt}{%
+\isamarkuptrue\par%
+\isa@parindent\parindent\parindent0pt%
+\isa@parskip\parskip\parskip0pt%
+\isaspacing\isastylett}{\par}
+
+\newenvironment{isabelle}
+{\begin{trivlist}\begin{isabellebody}\item\relax}
+{\end{isabellebody}\end{trivlist}}
+
+\newenvironment{isabellett}
+{\begin{trivlist}\begin{isabellebodytt}\item\relax}
+{\end{isabellebodytt}\end{trivlist}}
+
+\newcommand{\isa}[1]{\emph{\isaspacing\isastyleminor #1}}
+\newcommand{\isatt}[1]{\emph{\isaspacing\isastyleminortt #1}}
+
+\newcommand{\isaindent}[1]{\hphantom{#1}}
+\newcommand{\isanewline}{\mbox{}\par\mbox{}}
+\newcommand{\isasep}{}
+\newcommand{\isadigit}[1]{#1}
+
+\newcommand{\isachardefaults}{%
+\def\isacharbell{\isamath{\bigbox}}%requires stmaryrd
+\chardef\isacharbang=`\!%
+\chardef\isachardoublequote=`\"%
+\chardef\isachardoublequoteopen=`\"%
+\chardef\isachardoublequoteclose=`\"%
+\chardef\isacharhash=`\#%
+\chardef\isachardollar=`\$%
+\chardef\isacharpercent=`\%%
+\chardef\isacharampersand=`\&%
+\chardef\isacharprime=`\'%
+\chardef\isacharparenleft=`\(%
+\chardef\isacharparenright=`\)%
+\chardef\isacharasterisk=`\*%
+\chardef\isacharplus=`\+%
+\chardef\isacharcomma=`\,%
+\chardef\isacharminus=`\-%
+\chardef\isachardot=`\.%
+\chardef\isacharslash=`\/%
+\chardef\isacharcolon=`\:%
+\chardef\isacharsemicolon=`\;%
+\chardef\isacharless=`\<%
+\chardef\isacharequal=`\=%
+\chardef\isachargreater=`\>%
+\chardef\isacharquery=`\?%
+\chardef\isacharat=`\@%
+\chardef\isacharbrackleft=`\[%
+\chardef\isacharbackslash=`\\%
+\chardef\isacharbrackright=`\]%
+\chardef\isacharcircum=`\^%
+\chardef\isacharunderscore=`\_%
+\def\isacharunderscorekeyword{\_}%
+\chardef\isacharbackquote=`\`%
+\chardef\isacharbackquoteopen=`\`%
+\chardef\isacharbackquoteclose=`\`%
+\chardef\isacharbraceleft=`\{%
+\chardef\isacharbar=`\|%
+\chardef\isacharbraceright=`\}%
+\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$}}}%
+}
+
+
+% keyword and section markup
+
+\newcommand{\isakeyword}[1]
+{\emph{\normalfont\bfseries\def\isachardot{.}\def\isacharunderscore{\isacharunderscorekeyword}%
+\def\isacharbraceleft{\{}\def\isacharbraceright{\}}#1}}
+\newcommand{\isacommand}[1]{\isakeyword{#1}}
+
+\newcommand{\isakeywordcontrol}[1]
+{\emph{\normalfont\bfseries\itshape\def\isacharunderscore{\isacharunderscorekeyword}#1\,}}
+
+\newcommand{\isamarkupchapter}[1]{\chapter{#1}}
+\newcommand{\isamarkupsection}[1]{\section{#1}}
+\newcommand{\isamarkupsubsection}[1]{\subsection{#1}}
+\newcommand{\isamarkupsubsubsection}[1]{\subsubsection{#1}}
+\newcommand{\isamarkupparagraph}[1]{\paragraph{#1}}
+\newcommand{\isamarkupsubparagraph}[1]{\subparagraph{#1}}
+
+\newif\ifisamarkup
+\newcommand{\isabeginpar}{\par\ifisamarkup\relax\else\medskip\fi}
+\newcommand{\isaendpar}{\par\medskip}
+\newenvironment{isapar}{\parindent\isa@parindent\parskip\isa@parskip\isabeginpar}{\isaendpar}
+\newenvironment{isamarkuptext}{\par\isastyletext\begin{isapar}}{\end{isapar}}
+\newenvironment{isamarkuptxt}{\par\isastyletxt\begin{isapar}}{\end{isapar}}
+\newcommand{\isamarkupcmt}[1]{{\isastylecmt--- #1}}
+
+
+% styles
+
+\def\isabellestyle#1{\csname isabellestyle#1\endcsname}
+
+\newcommand{\isabellestyledefault}{%
+\def\isastyle{\small\normalfont\ttfamily\slshape}%
+\def\isastylett{\small\normalfont\ttfamily}%
+\def\isastyleminor{\small\normalfont\ttfamily\slshape}%
+\def\isastyleminortt{\small\normalfont\ttfamily}%
+\def\isastylescript{\footnotesize\normalfont\ttfamily\slshape}%
+\isachardefaults%
+}
+\isabellestyledefault
+
+\newcommand{\isabellestylett}{%
+\def\isastyle{\small\normalfont\ttfamily}%
+\def\isastylett{\small\normalfont\ttfamily}%
+\def\isastyleminor{\small\normalfont\ttfamily}%
+\def\isastyleminortt{\small\normalfont\ttfamily}%
+\def\isastylescript{\footnotesize\normalfont\ttfamily}%
+\isachardefaults%
+}
+
+\newcommand{\isabellestyleit}{%
+\def\isastyle{\small\normalfont\itshape}%
+\def\isastylett{\small\normalfont\ttfamily}%
+\def\isastyleminor{\normalfont\itshape}%
+\def\isastyleminortt{\normalfont\ttfamily}%
+\def\isastylescript{\footnotesize\normalfont\itshape}%
+\isachardefaults%
+\def\isacharunderscorekeyword{\mbox{-}}%
+\def\isacharbang{\isamath{!}}%
+\def\isachardoublequote{}%
+\def\isachardoublequoteopen{}%
+\def\isachardoublequoteclose{}%
+\def\isacharhash{\isamath{\#}}%
+\def\isachardollar{\isamath{\$}}%
+\def\isacharpercent{\isamath{\%}}%
+\def\isacharampersand{\isamath{\&}}%
+\def\isacharprime{\isamath{\mskip2mu{'}\mskip-2mu}}%
+\def\isacharparenleft{\isamath{(}}%
+\def\isacharparenright{\isamath{)}}%
+\def\isacharasterisk{\isamath{*}}%
+\def\isacharplus{\isamath{+}}%
+\def\isacharcomma{\isamath{\mathord,}}%
+\def\isacharminus{\isamath{-}}%
+\def\isachardot{\isamath{\mathord.}}%
+\def\isacharslash{\isamath{/}}%
+\def\isacharcolon{\isamath{\mathord:}}%
+\def\isacharsemicolon{\isamath{\mathord;}}%
+\def\isacharless{\isamath{<}}%
+\def\isacharequal{\isamath{=}}%
+\def\isachargreater{\isamath{>}}%
+\def\isacharat{\isamath{@}}%
+\def\isacharbrackleft{\isamath{[}}%
+\def\isacharbackslash{\isamath{\backslash}}%
+\def\isacharbrackright{\isamath{]}}%
+\def\isacharunderscore{\mbox{-}}%
+\def\isacharbraceleft{\isamath{\{}}%
+\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$}}}%
+}
+
+\newcommand{\isabellestyleliteral}{%
+\isabellestyleit%
+\def\isacharunderscore{\_}%
+\def\isacharunderscorekeyword{\_}%
+\chardef\isacharbackquoteopen=`\`%
+\chardef\isacharbackquoteclose=`\`%
+}
+
+\newcommand{\isabellestyleliteralunderscore}{%
+\isabellestyleliteral%
+\def\isacharunderscore{\textunderscore}%
+\def\isacharunderscorekeyword{\textunderscore}%
+}
+
+\newcommand{\isabellestylesl}{%
+\isabellestyleit%
+\def\isastyle{\small\normalfont\slshape}%
+\def\isastylett{\small\normalfont\ttfamily}%
+\def\isastyleminor{\normalfont\slshape}%
+\def\isastyleminortt{\normalfont\ttfamily}%
+\def\isastylescript{\footnotesize\normalfont\slshape}%
+}
+
+
+% cancel text
+
+\usepackage[normalem]{ulem}
+\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
+
+\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}}{}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys2/Journal/isabellesym.sty	Mon Nov 01 10:40:21 2021 +0000
@@ -0,0 +1,427 @@
+%%
+%% definitions of standard Isabelle symbols
+%%
+
+\newcommand{\isasymzero}{\isamath{\mathbf{0}}}  %requires amssymb
+\newcommand{\isasymone}{\isamath{\mathbf{1}}}  %requires amssymb
+\newcommand{\isasymtwo}{\isamath{\mathbf{2}}}  %requires amssymb
+\newcommand{\isasymthree}{\isamath{\mathbf{3}}}  %requires amssymb
+\newcommand{\isasymfour}{\isamath{\mathbf{4}}}  %requires amssymb
+\newcommand{\isasymfive}{\isamath{\mathbf{5}}}  %requires amssymb
+\newcommand{\isasymsix}{\isamath{\mathbf{6}}}  %requires amssymb
+\newcommand{\isasymseven}{\isamath{\mathbf{7}}}  %requires amssymb
+\newcommand{\isasymeight}{\isamath{\mathbf{8}}}  %requires amssymb
+\newcommand{\isasymnine}{\isamath{\mathbf{9}}}  %requires amssymb
+\newcommand{\isasymA}{\isamath{\mathcal{A}}}
+\newcommand{\isasymB}{\isamath{\mathcal{B}}}
+\newcommand{\isasymC}{\isamath{\mathcal{C}}}
+\newcommand{\isasymD}{\isamath{\mathcal{D}}}
+\newcommand{\isasymE}{\isamath{\mathcal{E}}}
+\newcommand{\isasymF}{\isamath{\mathcal{F}}}
+\newcommand{\isasymG}{\isamath{\mathcal{G}}}
+\newcommand{\isasymH}{\isamath{\mathcal{H}}}
+\newcommand{\isasymI}{\isamath{\mathcal{I}}}
+\newcommand{\isasymJ}{\isamath{\mathcal{J}}}
+\newcommand{\isasymK}{\isamath{\mathcal{K}}}
+\newcommand{\isasymL}{\isamath{\mathcal{L}}}
+\newcommand{\isasymM}{\isamath{\mathcal{M}}}
+\newcommand{\isasymN}{\isamath{\mathcal{N}}}
+\newcommand{\isasymO}{\isamath{\mathcal{O}}}
+\newcommand{\isasymP}{\isamath{\mathcal{P}}}
+\newcommand{\isasymQ}{\isamath{\mathcal{Q}}}
+\newcommand{\isasymR}{\isamath{\mathcal{R}}}
+\newcommand{\isasymS}{\isamath{\mathcal{S}}}
+\newcommand{\isasymT}{\isamath{\mathcal{T}}}
+\newcommand{\isasymU}{\isamath{\mathcal{U}}}
+\newcommand{\isasymV}{\isamath{\mathcal{V}}}
+\newcommand{\isasymW}{\isamath{\mathcal{W}}}
+\newcommand{\isasymX}{\isamath{\mathcal{X}}}
+\newcommand{\isasymY}{\isamath{\mathcal{Y}}}
+\newcommand{\isasymZ}{\isamath{\mathcal{Z}}}
+\newcommand{\isasyma}{\isamath{\mathrm{a}}}
+\newcommand{\isasymb}{\isamath{\mathrm{b}}}
+\newcommand{\isasymc}{\isamath{\mathrm{c}}}
+\newcommand{\isasymd}{\isamath{\mathrm{d}}}
+\newcommand{\isasyme}{\isamath{\mathrm{e}}}
+\newcommand{\isasymf}{\isamath{\mathrm{f}}}
+\newcommand{\isasymg}{\isamath{\mathrm{g}}}
+\newcommand{\isasymh}{\isamath{\mathrm{h}}}
+\newcommand{\isasymi}{\isamath{\mathrm{i}}}
+\newcommand{\isasymj}{\isamath{\mathrm{j}}}
+\newcommand{\isasymk}{\isamath{\mathrm{k}}}
+\newcommand{\isasyml}{\isamath{\mathrm{l}}}
+\newcommand{\isasymm}{\isamath{\mathrm{m}}}
+\newcommand{\isasymn}{\isamath{\mathrm{n}}}
+\newcommand{\isasymo}{\isamath{\mathrm{o}}}
+\newcommand{\isasymp}{\isamath{\mathrm{p}}}
+\newcommand{\isasymq}{\isamath{\mathrm{q}}}
+\newcommand{\isasymr}{\isamath{\mathrm{r}}}
+\newcommand{\isasyms}{\isamath{\mathrm{s}}}
+\newcommand{\isasymt}{\isamath{\mathrm{t}}}
+\newcommand{\isasymu}{\isamath{\mathrm{u}}}
+\newcommand{\isasymv}{\isamath{\mathrm{v}}}
+\newcommand{\isasymw}{\isamath{\mathrm{w}}}
+\newcommand{\isasymx}{\isamath{\mathrm{x}}}
+\newcommand{\isasymy}{\isamath{\mathrm{y}}}
+\newcommand{\isasymz}{\isamath{\mathrm{z}}}
+\newcommand{\isasymAA}{\isamath{\mathfrak{A}}}  %requires eufrak
+\newcommand{\isasymBB}{\isamath{\mathfrak{B}}}  %requires eufrak
+\newcommand{\isasymCC}{\isamath{\mathfrak{C}}}  %requires eufrak
+\newcommand{\isasymDD}{\isamath{\mathfrak{D}}}  %requires eufrak
+\newcommand{\isasymEE}{\isamath{\mathfrak{E}}}  %requires eufrak
+\newcommand{\isasymFF}{\isamath{\mathfrak{F}}}  %requires eufrak
+\newcommand{\isasymGG}{\isamath{\mathfrak{G}}}  %requires eufrak
+\newcommand{\isasymHH}{\isamath{\mathfrak{H}}}  %requires eufrak
+\newcommand{\isasymII}{\isamath{\mathfrak{I}}}  %requires eufrak
+\newcommand{\isasymJJ}{\isamath{\mathfrak{J}}}  %requires eufrak
+\newcommand{\isasymKK}{\isamath{\mathfrak{K}}}  %requires eufrak
+\newcommand{\isasymLL}{\isamath{\mathfrak{L}}}  %requires eufrak
+\newcommand{\isasymMM}{\isamath{\mathfrak{M}}}  %requires eufrak
+\newcommand{\isasymNN}{\isamath{\mathfrak{N}}}  %requires eufrak
+\newcommand{\isasymOO}{\isamath{\mathfrak{O}}}  %requires eufrak
+\newcommand{\isasymPP}{\isamath{\mathfrak{P}}}  %requires eufrak
+\newcommand{\isasymQQ}{\isamath{\mathfrak{Q}}}  %requires eufrak
+\newcommand{\isasymRR}{\isamath{\mathfrak{R}}}  %requires eufrak
+\newcommand{\isasymSS}{\isamath{\mathfrak{S}}}  %requires eufrak
+\newcommand{\isasymTT}{\isamath{\mathfrak{T}}}  %requires eufrak
+\newcommand{\isasymUU}{\isamath{\mathfrak{U}}}  %requires eufrak
+\newcommand{\isasymVV}{\isamath{\mathfrak{V}}}  %requires eufrak
+\newcommand{\isasymWW}{\isamath{\mathfrak{W}}}  %requires eufrak
+\newcommand{\isasymXX}{\isamath{\mathfrak{X}}}  %requires eufrak
+\newcommand{\isasymYY}{\isamath{\mathfrak{Y}}}  %requires eufrak
+\newcommand{\isasymZZ}{\isamath{\mathfrak{Z}}}  %requires eufrak
+\newcommand{\isasymaa}{\isamath{\mathfrak{a}}}  %requires eufrak
+\newcommand{\isasymbb}{\isamath{\mathfrak{b}}}  %requires eufrak
+\newcommand{\isasymcc}{\isamath{\mathfrak{c}}}  %requires eufrak
+\newcommand{\isasymdd}{\isamath{\mathfrak{d}}}  %requires eufrak
+\newcommand{\isasymee}{\isamath{\mathfrak{e}}}  %requires eufrak
+\newcommand{\isasymff}{\isamath{\mathfrak{f}}}  %requires eufrak
+\newcommand{\isasymgg}{\isamath{\mathfrak{g}}}  %requires eufrak
+\newcommand{\isasymhh}{\isamath{\mathfrak{h}}}  %requires eufrak
+\newcommand{\isasymii}{\isamath{\mathfrak{i}}}  %requires eufrak
+\newcommand{\isasymjj}{\isamath{\mathfrak{j}}}  %requires eufrak
+\newcommand{\isasymkk}{\isamath{\mathfrak{k}}}  %requires eufrak
+\newcommand{\isasymll}{\isamath{\mathfrak{l}}}  %requires eufrak
+\newcommand{\isasymmm}{\isamath{\mathfrak{m}}}  %requires eufrak
+\newcommand{\isasymnn}{\isamath{\mathfrak{n}}}  %requires eufrak
+\newcommand{\isasymoo}{\isamath{\mathfrak{o}}}  %requires eufrak
+\newcommand{\isasympp}{\isamath{\mathfrak{p}}}  %requires eufrak
+\newcommand{\isasymqq}{\isamath{\mathfrak{q}}}  %requires eufrak
+\newcommand{\isasymrr}{\isamath{\mathfrak{r}}}  %requires eufrak
+\newcommand{\isasymss}{\isamath{\mathfrak{s}}}  %requires eufrak
+\newcommand{\isasymtt}{\isamath{\mathfrak{t}}}  %requires eufrak
+\newcommand{\isasymuu}{\isamath{\mathfrak{u}}}  %requires eufrak
+\newcommand{\isasymvv}{\isamath{\mathfrak{v}}}  %requires eufrak
+\newcommand{\isasymww}{\isamath{\mathfrak{w}}}  %requires eufrak
+\newcommand{\isasymxx}{\isamath{\mathfrak{x}}}  %requires eufrak
+\newcommand{\isasymyy}{\isamath{\mathfrak{y}}}  %requires eufrak
+\newcommand{\isasymzz}{\isamath{\mathfrak{z}}}  %requires eufrak
+\newcommand{\isasymalpha}{\isamath{\alpha}}
+\newcommand{\isasymbeta}{\isamath{\beta}}
+\newcommand{\isasymgamma}{\isamath{\gamma}}
+\newcommand{\isasymdelta}{\isamath{\delta}}
+\newcommand{\isasymepsilon}{\isamath{\varepsilon}}
+\newcommand{\isasymzeta}{\isamath{\zeta}}
+\newcommand{\isasymeta}{\isamath{\eta}}
+\newcommand{\isasymtheta}{\isamath{\vartheta}}
+\newcommand{\isasymiota}{\isamath{\iota}}
+\newcommand{\isasymkappa}{\isamath{\kappa}}
+\newcommand{\isasymlambda}{\isamath{\lambda}}
+\newcommand{\isasymmu}{\isamath{\mu}}
+\newcommand{\isasymnu}{\isamath{\nu}}
+\newcommand{\isasymxi}{\isamath{\xi}}
+\newcommand{\isasympi}{\isamath{\pi}}
+\newcommand{\isasymrho}{\isamath{\varrho}}
+\newcommand{\isasymsigma}{\isamath{\sigma}}
+\newcommand{\isasymtau}{\isamath{\tau}}
+\newcommand{\isasymupsilon}{\isamath{\upsilon}}
+\newcommand{\isasymphi}{\isamath{\varphi}}
+\newcommand{\isasymchi}{\isamath{\chi}}
+\newcommand{\isasympsi}{\isamath{\psi}}
+\newcommand{\isasymomega}{\isamath{\omega}}
+\newcommand{\isasymGamma}{\isamath{\Gamma}}
+\newcommand{\isasymDelta}{\isamath{\Delta}}
+\newcommand{\isasymTheta}{\isamath{\Theta}}
+\newcommand{\isasymLambda}{\isamath{\Lambda}}
+\newcommand{\isasymXi}{\isamath{\Xi}}
+\newcommand{\isasymPi}{\isamath{\Pi}}
+\newcommand{\isasymSigma}{\isamath{\Sigma}}
+\newcommand{\isasymUpsilon}{\isamath{\Upsilon}}
+\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{\isasymleftarrow}{\isamath{\leftarrow}}
+\newcommand{\isasymrightarrow}{\isamath{\rightarrow}}
+\newcommand{\isasymlongleftarrow}{\isamath{\longleftarrow}}
+\newcommand{\isasymlongrightarrow}{\isamath{\longrightarrow}}
+\newcommand{\isasymlonglongleftarrow}{\isamath{\xleftarrow{\hphantom{AAA}}}}  %requires amsmath
+\newcommand{\isasymlonglongrightarrow}{\isamath{\xrightarrow{\hphantom{AAA}}}}  %requires amsmath
+\newcommand{\isasymlonglonglongleftarrow}{\isamath{\xleftarrow{\hphantom{AAAA}}}}  %requires amsmath
+\newcommand{\isasymlonglonglongrightarrow}{\isamath{\xrightarrow{\hphantom{AAAA}}}}  %requires amsmath
+\newcommand{\isasymLeftarrow}{\isamath{\Leftarrow}}
+\newcommand{\isasymRightarrow}{\isamath{\Rightarrow}}
+\newcommand{\isasymLongleftarrow}{\isamath{\Longleftarrow}}
+\newcommand{\isasymLongrightarrow}{\isamath{\Longrightarrow}}
+\newcommand{\isasymLleftarrow}{\isamath{\Lleftarrow}}  %requires amssymb
+\newcommand{\isasymRrightarrow}{\isamath{\Rrightarrow}}  %requires amssymb
+\newcommand{\isasymleftrightarrow}{\isamath{\leftrightarrow}}
+\newcommand{\isasymLeftrightarrow}{\isamath{\Leftrightarrow}}
+\newcommand{\isasymlongleftrightarrow}{\isamath{\longleftrightarrow}}
+\newcommand{\isasymLongleftrightarrow}{\isamath{\Longleftrightarrow}}
+\newcommand{\isasymmapsto}{\isamath{\mapsto}}
+\newcommand{\isasymlongmapsto}{\isamath{\longmapsto}}
+\newcommand{\isasymmidarrow}{\isamath{\relbar}}
+\newcommand{\isasymMidarrow}{\isamath{\Relbar}}
+\newcommand{\isasymhookleftarrow}{\isamath{\hookleftarrow}}
+\newcommand{\isasymhookrightarrow}{\isamath{\hookrightarrow}}
+\newcommand{\isasymleftharpoondown}{\isamath{\leftharpoondown}}
+\newcommand{\isasymrightharpoondown}{\isamath{\rightharpoondown}}
+\newcommand{\isasymleftharpoonup}{\isamath{\leftharpoonup}}
+\newcommand{\isasymrightharpoonup}{\isamath{\rightharpoonup}}
+\newcommand{\isasymrightleftharpoons}{\isamath{\rightleftharpoons}}
+\newcommand{\isasymleadsto}{\isamath{\leadsto}}  %requires amssymb
+\newcommand{\isasymdownharpoonleft}{\isamath{\downharpoonleft}}  %requires amssymb
+\newcommand{\isasymdownharpoonright}{\isamath{\downharpoonright}}  %requires amssymb
+\newcommand{\isasymupharpoonleft}{\isamath{\upharpoonleft}}  %requires amssymb
+\newcommand{\isasymupharpoonright}{\isamath{\upharpoonright}}  %requires amssymb
+\newcommand{\isasymrestriction}{\isamath{\restriction}}  %requires amssymb
+\newcommand{\isasymColon}{\isamath{\mathrel{::}}}
+\newcommand{\isasymup}{\isamath{\uparrow}}
+\newcommand{\isasymUp}{\isamath{\Uparrow}}
+\newcommand{\isasymdown}{\isamath{\downarrow}}
+\newcommand{\isasymDown}{\isamath{\Downarrow}}
+\newcommand{\isasymupdown}{\isamath{\updownarrow}}
+\newcommand{\isasymUpdown}{\isamath{\Updownarrow}}
+\newcommand{\isasymlangle}{\isamath{\langle}}
+\newcommand{\isasymrangle}{\isamath{\rangle}}
+\newcommand{\isasymllangle}{\isamath{\langle\mskip-5mu\langle}}
+\newcommand{\isasymrrangle}{\isamath{\rangle\mskip-5mu\rangle}}
+\newcommand{\isasymlceil}{\isamath{\lceil}}
+\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{\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{\isasymbottom}{\isamath{\bot}}
+\newcommand{\isasymtop}{\isamath{\top}}
+\newcommand{\isasymand}{\isamath{\wedge}}
+\newcommand{\isasymAnd}{\isamath{\bigwedge}}
+\newcommand{\isasymor}{\isamath{\vee}}
+\newcommand{\isasymOr}{\isamath{\bigvee}}
+\newcommand{\isasymforall}{\isamath{\forall\,}}
+\newcommand{\isasymexists}{\isamath{\exists\,}}
+\newcommand{\isasymnot}{\isamath{\neg}}
+\newcommand{\isasymnexists}{\isamath{\nexists\,}}  %requires amssymb
+\newcommand{\isasymcircle}{\isamath{\ocircle}}  %requires wasysym
+\newcommand{\isasymbox}{\isamath{\Box}}  %requires amssymb
+\newcommand{\isasymdiamond}{\isamath{\Diamond}}  %requires amssymb
+\newcommand{\isasymdiamondop}{\isamath{\diamond}}
+\newcommand{\isasymsurd}{\isamath{\surd}}
+\newcommand{\isasymturnstile}{\isamath{\vdash}}
+\newcommand{\isasymTurnstile}{\isamath{\models}}
+\newcommand{\isasymtturnstile}{\isamath{\vdash\!\!\!\vdash}}
+\newcommand{\isasymTTurnstile}{\isamath{\mid\!\models}}
+\newcommand{\isasymstileturn}{\isamath{\dashv}}
+\newcommand{\isasymle}{\isamath{\le}}
+\newcommand{\isasymge}{\isamath{\ge}}
+\newcommand{\isasymlless}{\isamath{\ll}}
+\newcommand{\isasymggreater}{\isamath{\gg}}
+\newcommand{\isasymlesssim}{\isamath{\lesssim}}  %requires amssymb
+\newcommand{\isasymgreatersim}{\isamath{\gtrsim}}  %requires amssymb
+\newcommand{\isasymlessapprox}{\isamath{\lessapprox}}  %requires amssymb
+\newcommand{\isasymgreaterapprox}{\isamath{\gtrapprox}}  %requires amssymb
+\newcommand{\isasymin}{\isamath{\in}}
+\newcommand{\isasymnotin}{\isamath{\notin}}
+\newcommand{\isasymsubset}{\isamath{\subset}}
+\newcommand{\isasymsupset}{\isamath{\supset}}
+\newcommand{\isasymsubseteq}{\isamath{\subseteq}}
+\newcommand{\isasymsupseteq}{\isamath{\supseteq}}
+\newcommand{\isasymsqsubset}{\isamath{\sqsubset}}  %requires amssymb
+\newcommand{\isasymsqsupset}{\isamath{\sqsupset}}  %requires amssymb
+\newcommand{\isasymsqsubseteq}{\isamath{\sqsubseteq}}
+\newcommand{\isasymsqsupseteq}{\isamath{\sqsupseteq}}
+\newcommand{\isasyminter}{\isamath{\cap}}
+\newcommand{\isasymInter}{\isamath{\bigcap\,}}
+\newcommand{\isasymunion}{\isamath{\cup}}
+\newcommand{\isasymUnion}{\isamath{\bigcup\,}}
+\newcommand{\isasymsqunion}{\isamath{\sqcup}}
+\newcommand{\isasymSqunion}{\isamath{\bigsqcup\,}}
+\newcommand{\isasymsqinter}{\isamath{\sqcap}}
+\newcommand{\isasymSqinter}{\isamath{\bigsqcap\,}}  %requires stmaryrd
+\newcommand{\isasymsetminus}{\isamath{\setminus}}
+\newcommand{\isasympropto}{\isamath{\propto}}
+\newcommand{\isasymuplus}{\isamath{\uplus}}
+\newcommand{\isasymUplus}{\isamath{\biguplus\,}}
+\newcommand{\isasymnoteq}{\isamath{\not=}}
+\newcommand{\isasymsim}{\isamath{\sim}}
+\newcommand{\isasymdoteq}{\isamath{\doteq}}
+\newcommand{\isasymsimeq}{\isamath{\simeq}}
+\newcommand{\isasymapprox}{\isamath{\approx}}
+\newcommand{\isasymasymp}{\isamath{\asymp}}
+\newcommand{\isasymcong}{\isamath{\cong}}
+\newcommand{\isasymsmile}{\isamath{\smile}}
+\newcommand{\isasymequiv}{\isamath{\equiv}}
+\newcommand{\isasymfrown}{\isamath{\frown}}
+\newcommand{\isasymJoin}{\isamath{\Join}}  %requires amssymb
+\newcommand{\isasymbowtie}{\isamath{\bowtie}}
+\newcommand{\isasymprec}{\isamath{\prec}}
+\newcommand{\isasymsucc}{\isamath{\succ}}
+\newcommand{\isasympreceq}{\isamath{\preceq}}
+\newcommand{\isasymsucceq}{\isamath{\succeq}}
+\newcommand{\isasymparallel}{\isamath{\parallel}}
+\newcommand{\isasymbar}{\isamath{\mid}}
+\newcommand{\isasymbbar}{\isamath{[\mskip-1.5mu]}}
+\newcommand{\isasymplusminus}{\isamath{\pm}}
+\newcommand{\isasymminusplus}{\isamath{\mp}}
+\newcommand{\isasymtimes}{\isamath{\times}}
+\newcommand{\isasymdiv}{\isamath{\div}}
+\newcommand{\isasymcdot}{\isamath{\cdot}}
+\newcommand{\isasymsqdot}{\isamath{\sbox\z@{$\centerdot$}\ht\z@=.33333\ht\z@\vcenter{\box\z@}}}  %requires amssymb
+\newcommand{\isasymstar}{\isamath{\star}}
+\newcommand{\isasymbullet}{\boldmath\isamath{\mathchoice{\displaystyle{\cdot}}{\textstyle{\cdot}}{\scriptstyle{\bullet}}{\scriptscriptstyle{\bullet}}}}
+\newcommand{\isasymcirc}{\isamath{\circ}}
+\newcommand{\isasymdagger}{\isamath{\dagger}}
+\newcommand{\isasymddagger}{\isamath{\ddagger}}
+\newcommand{\isasymlhd}{\isamath{\lhd}}  %requires amssymb
+\newcommand{\isasymrhd}{\isamath{\rhd}}  %requires amssymb
+\newcommand{\isasymunlhd}{\isamath{\unlhd}}  %requires amssymb
+\newcommand{\isasymunrhd}{\isamath{\unrhd}}  %requires amssymb
+\newcommand{\isasymtriangleleft}{\isamath{\triangleleft}}
+\newcommand{\isasymtriangleright}{\isamath{\triangleright}}
+\newcommand{\isasymtriangle}{\isamath{\triangle}}
+\newcommand{\isasymtriangleq}{\isamath{\triangleq}}  %requires amssymb
+\newcommand{\isasymoplus}{\isamath{\oplus}}
+\newcommand{\isasymOplus}{\isamath{\bigoplus\,}}
+\newcommand{\isasymotimes}{\isamath{\otimes}}
+\newcommand{\isasymOtimes}{\isamath{\bigotimes\,}}
+\newcommand{\isasymodot}{\isamath{\odot}}
+\newcommand{\isasymOdot}{\isamath{\bigodot\,}}
+\newcommand{\isasymominus}{\isamath{\ominus}}
+\newcommand{\isasymoslash}{\isamath{\oslash}}
+\newcommand{\isasymdots}{\isamath{\dots}}
+\newcommand{\isasymcdots}{\isamath{\cdots}}
+\newcommand{\isasymSum}{\isamath{\sum\,}}
+\newcommand{\isasymProd}{\isamath{\prod\,}}
+\newcommand{\isasymCoprod}{\isamath{\coprod\,}}
+\newcommand{\isasyminfinity}{\isamath{\infty}}
+\newcommand{\isasymintegral}{\isamath{\int\,}}
+\newcommand{\isasymointegral}{\isamath{\oint\,}}
+\newcommand{\isasymclubsuit}{\isamath{\clubsuit}}
+\newcommand{\isasymdiamondsuit}{\isamath{\diamondsuit}}
+\newcommand{\isasymheartsuit}{\isamath{\heartsuit}}
+\newcommand{\isasymspadesuit}{\isamath{\spadesuit}}
+\newcommand{\isasymaleph}{\isamath{\aleph}}
+\newcommand{\isasymemptyset}{\isamath{\emptyset}}
+\newcommand{\isasymnabla}{\isamath{\nabla}}
+\newcommand{\isasympartial}{\isamath{\partial}}
+\newcommand{\isasymRe}{\isamath{\Re}}
+\newcommand{\isasymIm}{\isamath{\Im}}
+\newcommand{\isasymflat}{\isamath{\flat}}
+\newcommand{\isasymnatural}{\isamath{\natural}}
+\newcommand{\isasymsharp}{\isamath{\sharp}}
+\newcommand{\isasymangle}{\isamath{\angle}}
+\newcommand{\isasymcopyright}{\isatext{\normalfont\rmfamily\copyright}}
+\newcommand{\isasymregistered}{\isatext{\normalfont\rmfamily\textregistered}}
+\newcommand{\isasyminverse}{\isamath{{}^{-1}}}
+\newcommand{\isasymonequarter}{\isatext{\normalfont\rmfamily\textonequarter}}  %requires textcomp
+\newcommand{\isasymonehalf}{\isatext{\normalfont\rmfamily\textonehalf}}  %requires textcomp
+\newcommand{\isasymthreequarters}{\isatext{\normalfont\rmfamily\textthreequarters}}  %requires textcomp
+\newcommand{\isasymordfeminine}{\isatext{\normalfont\rmfamily\textordfeminine}}
+\newcommand{\isasymordmasculine}{\isatext{\normalfont\rmfamily\textordmasculine}}
+\newcommand{\isasymsection}{\isatext{\normalfont\rmfamily\S}}
+\newcommand{\isasymparagraph}{\isatext{\normalfont\rmfamily\P}}
+\newcommand{\isasymexclamdown}{\isatext{\normalfont\rmfamily\textexclamdown}}
+\newcommand{\isasymquestiondown}{\isatext{\normalfont\rmfamily\textquestiondown}}
+\newcommand{\isasymeuro}{\isatext{\euro}}  %requires eurosym
+\newcommand{\isasympounds}{\isamath{\pounds}}
+\newcommand{\isasymyen}{\isatext{\yen}}  %requires amssymb
+\newcommand{\isasymcent}{\isatext{\textcent}}  %requires textcomp
+\newcommand{\isasymcurrency}{\isatext{\textcurrency}} %requires textcomp
+\newcommand{\isasymdegree}{\isatext{\normalfont\rmfamily\textdegree}}  %requires textcomp
+\newcommand{\isasymhyphen}{\isatext{\normalfont\rmfamily-}}
+\newcommand{\isasymamalg}{\isamath{\amalg}}
+\newcommand{\isasymmho}{\isamath{\mho}}  %requires amssymb
+\newcommand{\isasymlozenge}{\isamath{\lozenge}}  %requires amssymb
+\newcommand{\isasymwp}{\isamath{\wp}}
+\newcommand{\isasymwrong}{\isamath{\wr}}
+\newcommand{\isasymacute}{\isatext{\'\relax}}
+\newcommand{\isasymindex}{\isatext{\i}}
+\newcommand{\isasymdieresis}{\isatext{\"\relax}}
+\newcommand{\isasymcedilla}{\isatext{\c\relax}}
+\newcommand{\isasymhungarumlaut}{\isatext{\H\relax}}
+\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$}}}
+\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{\isactrlmarker}{\isatext{\ding{48}}}  %requires pifont
+\newcommand{\isactrlassert}{\isakeywordcontrol{assert}}
+\newcommand{\isactrlcancel}{\isakeywordcontrol{cancel}}
+\newcommand{\isactrlbinding}{\isakeywordcontrol{binding}}
+\newcommand{\isactrlclass}{\isakeywordcontrol{class}}
+\newcommand{\isactrlclassUNDERSCOREsyntax}{\isakeywordcontrol{class{\isacharunderscore}syntax}}
+\newcommand{\isactrlcommandUNDERSCOREkeyword}{\isakeywordcontrol{command{\isacharunderscore}keyword}}
+\newcommand{\isactrlconst}{\isakeywordcontrol{const}}
+\newcommand{\isactrlconstUNDERSCOREabbrev}{\isakeywordcontrol{const{\isacharunderscore}abbrev}}
+\newcommand{\isactrlconstUNDERSCOREname}{\isakeywordcontrol{const{\isacharunderscore}name}}
+\newcommand{\isactrlconstUNDERSCOREsyntax}{\isakeywordcontrol{const{\isacharunderscore}syntax}}
+\newcommand{\isactrlcontext}{\isakeywordcontrol{context}}
+\newcommand{\isactrlcprop}{\isakeywordcontrol{cprop}}
+\newcommand{\isactrlcterm}{\isakeywordcontrol{cterm}}
+\newcommand{\isactrlctyp}{\isakeywordcontrol{ctyp}}
+\newcommand{\isactrldir}{\isakeywordcontrol{dir}}
+\newcommand{\isactrlfile}{\isakeywordcontrol{file}}
+\newcommand{\isactrlhere}{\isakeywordcontrol{here}}
+\newcommand{\isactrlkeyword}{\isakeywordcontrol{keyword}}
+\newcommand{\isactrllatex}{\isakeywordcontrol{latex}}
+\newcommand{\isactrllocale}{\isakeywordcontrol{locale}}
+\newcommand{\isactrlmakeUNDERSCOREstring}{\isakeywordcontrol{make{\isacharunderscore}string}}
+\newcommand{\isactrlmasterUNDERSCOREdir}{\isakeywordcontrol{master{\isacharunderscore}dir}}
+\newcommand{\isactrlmethod}{\isakeywordcontrol{method}}
+\newcommand{\isactrlnamedUNDERSCOREtheorems}{\isakeywordcontrol{named{\isacharunderscore}theorems}}
+\newcommand{\isactrlnonterminal}{\isakeywordcontrol{nonterminal}}
+\newcommand{\isactrloracleUNDERSCOREname}{\isakeywordcontrol{oracle{\isacharunderscore}name}}
+\newcommand{\isactrlpath}{\isakeywordcontrol{path}}
+\newcommand{\isactrlpathUNDERSCOREbinding}{\isakeywordcontrol{path{\isacharunderscore}binding}}
+\newcommand{\isactrlplugin}{\isakeywordcontrol{plugin}}
+\newcommand{\isactrlprint}{\isakeywordcontrol{print}}
+\newcommand{\isactrlprop}{\isakeywordcontrol{prop}}
+\newcommand{\isactrlscala}{\isakeywordcontrol{scala}}
+\newcommand{\isactrlscalaUNDERSCOREfunction}{\isakeywordcontrol{scala{\isacharunderscore}function}}
+\newcommand{\isactrlscalaUNDERSCOREmethod}{\isakeywordcontrol{scala{\isacharunderscore}method}}
+\newcommand{\isactrlscalaUNDERSCOREobject}{\isakeywordcontrol{scala{\isacharunderscore}object}}
+\newcommand{\isactrlscalaUNDERSCOREtype}{\isakeywordcontrol{scala{\isacharunderscore}type}}
+\newcommand{\isactrlsimproc}{\isakeywordcontrol{simproc}}
+\newcommand{\isactrlsort}{\isakeywordcontrol{sort}}
+\newcommand{\isactrlsyntaxUNDERSCOREconst}{\isakeywordcontrol{syntax{\isacharunderscore}const}}
+\newcommand{\isactrlsystemUNDERSCOREoption}{\isakeywordcontrol{system{\isacharunderscore}option}}
+\newcommand{\isactrlterm}{\isakeywordcontrol{term}}
+\newcommand{\isactrltheory}{\isakeywordcontrol{theory}}
+\newcommand{\isactrltheoryUNDERSCOREcontext}{\isakeywordcontrol{theory{\isacharunderscore}context}}
+\newcommand{\isactrltyp}{\isakeywordcontrol{typ}}
+\newcommand{\isactrltypeUNDERSCOREabbrev}{\isakeywordcontrol{type{\isacharunderscore}abbrev}}
+\newcommand{\isactrltypeUNDERSCOREname}{\isakeywordcontrol{type{\isacharunderscore}name}}
+\newcommand{\isactrltypeUNDERSCOREsyntax}{\isakeywordcontrol{type{\isacharunderscore}syntax}}
+\newcommand{\isactrlundefined}{\isakeywordcontrol{undefined}}
+
+\newcommand{\isactrlcode}{\isakeywordcontrol{code}}
+\newcommand{\isactrlcomputation}{\isakeywordcontrol{computation}}
+\newcommand{\isactrlcomputationUNDERSCOREconv}{\isakeywordcontrol{computation{\isacharunderscore}conv}}
+\newcommand{\isactrlcomputationUNDERSCOREcheck}{\isakeywordcontrol{computation{\isacharunderscore}check}}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys2/Journal/isabelletags.sty	Mon Nov 01 10:40:21 2021 +0000
@@ -0,0 +1,1 @@
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys2/Journal/llncs.cls	Mon Nov 01 10:40:21 2021 +0000
@@ -0,0 +1,1189 @@
+% LLNCS DOCUMENT CLASS -- version 2.13 (28-Jan-2002)
+% Springer Verlag LaTeX2e support for Lecture Notes in Computer Science
+%
+%%
+%% \CharacterTable
+%%  {Upper-case    \A\B\C\D\E\F\G\H\I\J\K\L\M\N\O\P\Q\R\S\T\U\V\W\X\Y\Z
+%%   Lower-case    \a\b\c\d\e\f\g\h\i\j\k\l\m\n\o\p\q\r\s\t\u\v\w\x\y\z
+%%   Digits        \0\1\2\3\4\5\6\7\8\9
+%%   Exclamation   \!     Double quote  \"     Hash (number) \#
+%%   Dollar        \$     Percent       \%     Ampersand     \&
+%%   Acute accent  \'     Left paren    \(     Right paren   \)
+%%   Asterisk      \*     Plus          \+     Comma         \,
+%%   Minus         \-     Point         \.     Solidus       \/
+%%   Colon         \:     Semicolon     \;     Less than     \<
+%%   Equals        \=     Greater than  \>     Question mark \?
+%%   Commercial at \@     Left bracket  \[     Backslash     \\
+%%   Right bracket \]     Circumflex    \^     Underscore    \_
+%%   Grave accent  \`     Left brace    \{     Vertical bar  \|
+%%   Right brace   \}     Tilde         \~}
+%%
+\NeedsTeXFormat{LaTeX2e}[1995/12/01]
+\ProvidesClass{llncs}[2002/01/28 v2.13
+^^J LaTeX document class for Lecture Notes in Computer Science]
+% Options
+\let\if@envcntreset\iffalse
+\DeclareOption{envcountreset}{\let\if@envcntreset\iftrue}
+\DeclareOption{citeauthoryear}{\let\citeauthoryear=Y}
+\DeclareOption{oribibl}{\let\oribibl=Y}
+\let\if@custvec\iftrue
+\DeclareOption{orivec}{\let\if@custvec\iffalse}
+\let\if@envcntsame\iffalse
+\DeclareOption{envcountsame}{\let\if@envcntsame\iftrue}
+\let\if@envcntsect\iffalse
+\DeclareOption{envcountsect}{\let\if@envcntsect\iftrue}
+\let\if@runhead\iffalse
+\DeclareOption{runningheads}{\let\if@runhead\iftrue}
+
+\let\if@openbib\iffalse
+\DeclareOption{openbib}{\let\if@openbib\iftrue}
+
+% languages
+\let\switcht@@therlang\relax
+\def\ds@deutsch{\def\switcht@@therlang{\switcht@deutsch}}
+\def\ds@francais{\def\switcht@@therlang{\switcht@francais}}
+
+\DeclareOption*{\PassOptionsToClass{\CurrentOption}{article}}
+
+\ProcessOptions
+
+\LoadClass[twoside]{article}
+\RequirePackage{multicol} % needed for the list of participants, index
+
+\setlength{\textwidth}{12.2cm}
+\setlength{\textheight}{19.3cm}
+\renewcommand\@pnumwidth{2em}
+\renewcommand\@tocrmarg{3.5em}
+%
+\def\@dottedtocline#1#2#3#4#5{%
+  \ifnum #1>\c@tocdepth \else
+    \vskip \z@ \@plus.2\p@
+    {\leftskip #2\relax \rightskip \@tocrmarg \advance\rightskip by 0pt plus 2cm
+               \parfillskip -\rightskip \pretolerance=10000
+     \parindent #2\relax\@afterindenttrue
+     \interlinepenalty\@M
+     \leavevmode
+     \@tempdima #3\relax
+     \advance\leftskip \@tempdima \null\nobreak\hskip -\leftskip
+     {#4}\nobreak
+     \leaders\hbox{$\m@th
+        \mkern \@dotsep mu\hbox{.}\mkern \@dotsep
+        mu$}\hfill
+     \nobreak
+     \hb@xt@\@pnumwidth{\hfil\normalfont \normalcolor #5}%
+     \par}%
+  \fi}
+%
+\def\switcht@albion{%
+\def\abstractname{Abstract.}
+\def\ackname{Acknowledgement.}
+\def\andname{and}
+\def\lastandname{\unskip, and}
+\def\appendixname{Appendix}
+\def\chaptername{Chapter}
+\def\claimname{Claim}
+\def\conjecturename{Conjecture}
+\def\contentsname{Table of Contents}
+\def\corollaryname{Corollary}
+\def\definitionname{Definition}
+\def\examplename{Example}
+\def\exercisename{Exercise}
+\def\figurename{Fig.}
+\def\keywordname{{\bf Key words:}}
+\def\indexname{Index}
+\def\lemmaname{Lemma}
+\def\contriblistname{List of Contributors}
+\def\listfigurename{List of Figures}
+\def\listtablename{List of Tables}
+\def\mailname{{\it Correspondence to\/}:}
+\def\noteaddname{Note added in proof}
+\def\notename{Note}
+\def\partname{Part}
+\def\problemname{Problem}
+\def\proofname{Proof}
+\def\propertyname{Property}
+\def\propositionname{Proposition}
+\def\questionname{Question}
+\def\remarkname{Remark}
+\def\seename{see}
+\def\solutionname{Solution}
+\def\subclassname{{\it Subject Classifications\/}:}
+\def\tablename{Table}
+\def\theoremname{Theorem}}
+\switcht@albion
+% Names of theorem like environments are already defined
+% but must be translated if another language is chosen
+%
+% French section
+\def\switcht@francais{%\typeout{On parle francais.}%
+ \def\abstractname{R\'esum\'e.}%
+ \def\ackname{Remerciements.}%
+ \def\andname{et}%
+ \def\lastandname{ et}%
+ \def\appendixname{Appendice}
+ \def\chaptername{Chapitre}%
+ \def\claimname{Pr\'etention}%
+ \def\conjecturename{Hypoth\`ese}%
+ \def\contentsname{Table des mati\`eres}%
+ \def\corollaryname{Corollaire}%
+ \def\definitionname{D\'efinition}%
+ \def\examplename{Exemple}%
+ \def\exercisename{Exercice}%
+ \def\figurename{Fig.}%
+ \def\keywordname{{\bf Mots-cl\'e:}}
+ \def\indexname{Index}
+ \def\lemmaname{Lemme}%
+ \def\contriblistname{Liste des contributeurs}
+ \def\listfigurename{Liste des figures}%
+ \def\listtablename{Liste des tables}%
+ \def\mailname{{\it Correspondence to\/}:}
+ \def\noteaddname{Note ajout\'ee \`a l'\'epreuve}%
+ \def\notename{Remarque}%
+ \def\partname{Partie}%
+ \def\problemname{Probl\`eme}%
+ \def\proofname{Preuve}%
+ \def\propertyname{Caract\'eristique}%
+%\def\propositionname{Proposition}%
+ \def\questionname{Question}%
+ \def\remarkname{Remarque}%
+ \def\seename{voir}
+ \def\solutionname{Solution}%
+ \def\subclassname{{\it Subject Classifications\/}:}
+ \def\tablename{Tableau}%
+ \def\theoremname{Th\'eor\`eme}%
+}
+%
+% German section
+\def\switcht@deutsch{%\typeout{Man spricht deutsch.}%
+ \def\abstractname{Zusammenfassung.}%
+ \def\ackname{Danksagung.}%
+ \def\andname{und}%
+ \def\lastandname{ und}%
+ \def\appendixname{Anhang}%
+ \def\chaptername{Kapitel}%
+ \def\claimname{Behauptung}%
+ \def\conjecturename{Hypothese}%
+ \def\contentsname{Inhaltsverzeichnis}%
+ \def\corollaryname{Korollar}%
+%\def\definitionname{Definition}%
+ \def\examplename{Beispiel}%
+ \def\exercisename{\"Ubung}%
+ \def\figurename{Abb.}%
+ \def\keywordname{{\bf Schl\"usselw\"orter:}}
+ \def\indexname{Index}
+%\def\lemmaname{Lemma}%
+ \def\contriblistname{Mitarbeiter}
+ \def\listfigurename{Abbildungsverzeichnis}%
+ \def\listtablename{Tabellenverzeichnis}%
+ \def\mailname{{\it Correspondence to\/}:}
+ \def\noteaddname{Nachtrag}%
+ \def\notename{Anmerkung}%
+ \def\partname{Teil}%
+%\def\problemname{Problem}%
+ \def\proofname{Beweis}%
+ \def\propertyname{Eigenschaft}%
+%\def\propositionname{Proposition}%
+ \def\questionname{Frage}%
+ \def\remarkname{Anmerkung}%
+ \def\seename{siehe}
+ \def\solutionname{L\"osung}%
+ \def\subclassname{{\it Subject Classifications\/}:}
+ \def\tablename{Tabelle}%
+%\def\theoremname{Theorem}%
+}
+
+% Ragged bottom for the actual page
+\def\thisbottomragged{\def\@textbottom{\vskip\z@ plus.0001fil
+\global\let\@textbottom\relax}}
+
+\renewcommand\small{%
+   \@setfontsize\small\@ixpt{11}%
+   \abovedisplayskip 8.5\p@ \@plus3\p@ \@minus4\p@
+   \abovedisplayshortskip \z@ \@plus2\p@
+   \belowdisplayshortskip 4\p@ \@plus2\p@ \@minus2\p@
+   \def\@listi{\leftmargin\leftmargini
+               \parsep 0\p@ \@plus1\p@ \@minus\p@
+               \topsep 8\p@ \@plus2\p@ \@minus4\p@
+               \itemsep0\p@}%
+   \belowdisplayskip \abovedisplayskip
+}
+
+\frenchspacing
+\widowpenalty=10000
+\clubpenalty=10000
+
+\setlength\oddsidemargin   {63\p@}
+\setlength\evensidemargin  {63\p@}
+\setlength\marginparwidth  {90\p@}
+
+\setlength\headsep   {16\p@}
+
+\setlength\footnotesep{7.7\p@}
+\setlength\textfloatsep{8mm\@plus 2\p@ \@minus 4\p@}
+\setlength\intextsep   {8mm\@plus 2\p@ \@minus 2\p@}
+
+\setcounter{secnumdepth}{2}
+
+\newcounter {chapter}
+\renewcommand\thechapter      {\@arabic\c@chapter}
+
+\newif\if@mainmatter \@mainmattertrue
+\newcommand\frontmatter{\cleardoublepage
+            \@mainmatterfalse\pagenumbering{Roman}}
+\newcommand\mainmatter{\cleardoublepage
+       \@mainmattertrue\pagenumbering{arabic}}
+\newcommand\backmatter{\if@openright\cleardoublepage\else\clearpage\fi
+      \@mainmatterfalse}
+
+\renewcommand\part{\cleardoublepage
+                 \thispagestyle{empty}%
+                 \if@twocolumn
+                     \onecolumn
+                     \@tempswatrue
+                   \else
+                     \@tempswafalse
+                 \fi
+                 \null\vfil
+                 \secdef\@part\@spart}
+
+\def\@part[#1]#2{%
+    \ifnum \c@secnumdepth >-2\relax
+      \refstepcounter{part}%
+      \addcontentsline{toc}{part}{\thepart\hspace{1em}#1}%
+    \else
+      \addcontentsline{toc}{part}{#1}%
+    \fi
+    \markboth{}{}%
+    {\centering
+     \interlinepenalty \@M
+     \normalfont
+     \ifnum \c@secnumdepth >-2\relax
+       \huge\bfseries \partname~\thepart
+       \par
+       \vskip 20\p@
+     \fi
+     \Huge \bfseries #2\par}%
+    \@endpart}
+\def\@spart#1{%
+    {\centering
+     \interlinepenalty \@M
+     \normalfont
+     \Huge \bfseries #1\par}%
+    \@endpart}
+\def\@endpart{\vfil\newpage
+              \if@twoside
+                \null
+                \thispagestyle{empty}%
+                \newpage
+              \fi
+              \if@tempswa
+                \twocolumn
+              \fi}
+
+\newcommand\chapter{\clearpage
+                    \thispagestyle{empty}%
+                    \global\@topnum\z@
+                    \@afterindentfalse
+                    \secdef\@chapter\@schapter}
+\def\@chapter[#1]#2{\ifnum \c@secnumdepth >\m@ne
+                       \if@mainmatter
+                         \refstepcounter{chapter}%
+                         \typeout{\@chapapp\space\thechapter.}%
+                         \addcontentsline{toc}{chapter}%
+                                  {\protect\numberline{\thechapter}#1}%
+                       \else
+                         \addcontentsline{toc}{chapter}{#1}%
+                       \fi
+                    \else
+                      \addcontentsline{toc}{chapter}{#1}%
+                    \fi
+                    \chaptermark{#1}%
+                    \addtocontents{lof}{\protect\addvspace{10\p@}}%
+                    \addtocontents{lot}{\protect\addvspace{10\p@}}%
+                    \if@twocolumn
+                      \@topnewpage[\@makechapterhead{#2}]%
+                    \else
+                      \@makechapterhead{#2}%
+                      \@afterheading
+                    \fi}
+\def\@makechapterhead#1{%
+% \vspace*{50\p@}%
+  {\centering
+    \ifnum \c@secnumdepth >\m@ne
+      \if@mainmatter
+        \large\bfseries \@chapapp{} \thechapter
+        \par\nobreak
+        \vskip 20\p@
+      \fi
+    \fi
+    \interlinepenalty\@M
+    \Large \bfseries #1\par\nobreak
+    \vskip 40\p@
+  }}
+\def\@schapter#1{\if@twocolumn
+                   \@topnewpage[\@makeschapterhead{#1}]%
+                 \else
+                   \@makeschapterhead{#1}%
+                   \@afterheading
+                 \fi}
+\def\@makeschapterhead#1{%
+% \vspace*{50\p@}%
+  {\centering
+    \normalfont
+    \interlinepenalty\@M
+    \Large \bfseries  #1\par\nobreak
+    \vskip 40\p@
+  }}
+
+\renewcommand\section{\@startsection{section}{1}{\z@}%
+                       {-18\p@ \@plus -4\p@ \@minus -4\p@}%
+                       {12\p@ \@plus 4\p@ \@minus 4\p@}%
+                       {\normalfont\large\bfseries\boldmath
+                        \rightskip=\z@ \@plus 8em\pretolerance=10000 }}
+\renewcommand\subsection{\@startsection{subsection}{2}{\z@}%
+                       {-18\p@ \@plus -4\p@ \@minus -4\p@}%
+                       {8\p@ \@plus 4\p@ \@minus 4\p@}%
+                       {\normalfont\normalsize\bfseries\boldmath
+                        \rightskip=\z@ \@plus 8em\pretolerance=10000 }}
+\renewcommand\subsubsection{\@startsection{subsubsection}{3}{\z@}%
+                       {-18\p@ \@plus -4\p@ \@minus -4\p@}%
+                       {-0.5em \@plus -0.22em \@minus -0.1em}%
+                       {\normalfont\normalsize\bfseries\boldmath}}
+\renewcommand\paragraph{\@startsection{paragraph}{4}{\z@}%
+                       {-12\p@ \@plus -4\p@ \@minus -4\p@}%
+                       {-0.5em \@plus -0.22em \@minus -0.1em}%
+                       {\normalfont\normalsize\itshape}}
+\renewcommand\subparagraph[1]{\typeout{LLNCS warning: You should not use
+                  \string\subparagraph\space with this class}\vskip0.5cm
+You should not use \verb|\subparagraph| with this class.\vskip0.5cm}
+
+\DeclareMathSymbol{\Gamma}{\mathalpha}{letters}{"00}
+\DeclareMathSymbol{\Delta}{\mathalpha}{letters}{"01}
+\DeclareMathSymbol{\Theta}{\mathalpha}{letters}{"02}
+\DeclareMathSymbol{\Lambda}{\mathalpha}{letters}{"03}
+\DeclareMathSymbol{\Xi}{\mathalpha}{letters}{"04}
+\DeclareMathSymbol{\Pi}{\mathalpha}{letters}{"05}
+\DeclareMathSymbol{\Sigma}{\mathalpha}{letters}{"06}
+\DeclareMathSymbol{\Upsilon}{\mathalpha}{letters}{"07}
+\DeclareMathSymbol{\Phi}{\mathalpha}{letters}{"08}
+\DeclareMathSymbol{\Psi}{\mathalpha}{letters}{"09}
+\DeclareMathSymbol{\Omega}{\mathalpha}{letters}{"0A}
+
+\let\footnotesize\small
+
+\if@custvec
+\def\vec#1{\mathchoice{\mbox{\boldmath$\displaystyle#1$}}
+{\mbox{\boldmath$\textstyle#1$}}
+{\mbox{\boldmath$\scriptstyle#1$}}
+{\mbox{\boldmath$\scriptscriptstyle#1$}}}
+\fi
+
+\def\squareforqed{\hbox{\rlap{$\sqcap$}$\sqcup$}}
+\def\qed{\ifmmode\squareforqed\else{\unskip\nobreak\hfil
+\penalty50\hskip1em\null\nobreak\hfil\squareforqed
+\parfillskip=0pt\finalhyphendemerits=0\endgraf}\fi}
+
+\def\getsto{\mathrel{\mathchoice {\vcenter{\offinterlineskip
+\halign{\hfil
+$\displaystyle##$\hfil\cr\gets\cr\to\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr\gets
+\cr\to\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr\gets
+\cr\to\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
+\gets\cr\to\cr}}}}}
+\def\lid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil
+$\displaystyle##$\hfil\cr<\cr\noalign{\vskip1.2pt}=\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr<\cr
+\noalign{\vskip1.2pt}=\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr<\cr
+\noalign{\vskip1pt}=\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
+<\cr
+\noalign{\vskip0.9pt}=\cr}}}}}
+\def\gid{\mathrel{\mathchoice {\vcenter{\offinterlineskip\halign{\hfil
+$\displaystyle##$\hfil\cr>\cr\noalign{\vskip1.2pt}=\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr>\cr
+\noalign{\vskip1.2pt}=\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr>\cr
+\noalign{\vskip1pt}=\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
+>\cr
+\noalign{\vskip0.9pt}=\cr}}}}}
+\def\grole{\mathrel{\mathchoice {\vcenter{\offinterlineskip
+\halign{\hfil
+$\displaystyle##$\hfil\cr>\cr\noalign{\vskip-1pt}<\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\textstyle##$\hfil\cr
+>\cr\noalign{\vskip-1pt}<\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptstyle##$\hfil\cr
+>\cr\noalign{\vskip-0.8pt}<\cr}}}
+{\vcenter{\offinterlineskip\halign{\hfil$\scriptscriptstyle##$\hfil\cr
+>\cr\noalign{\vskip-0.3pt}<\cr}}}}}
+\def\bbbr{{\rm I\!R}} %reelle Zahlen
+\def\bbbm{{\rm I\!M}}
+\def\bbbn{{\rm I\!N}} %natuerliche Zahlen
+\def\bbbf{{\rm I\!F}}
+\def\bbbh{{\rm I\!H}}
+\def\bbbk{{\rm I\!K}}
+\def\bbbp{{\rm I\!P}}
+\def\bbbone{{\mathchoice {\rm 1\mskip-4mu l} {\rm 1\mskip-4mu l}
+{\rm 1\mskip-4.5mu l} {\rm 1\mskip-5mu l}}}
+\def\bbbc{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm C$}\hbox{\hbox
+to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
+{\setbox0=\hbox{$\textstyle\rm C$}\hbox{\hbox
+to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptstyle\rm C$}\hbox{\hbox
+to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptscriptstyle\rm C$}\hbox{\hbox
+to0pt{\kern0.4\wd0\vrule height0.9\ht0\hss}\box0}}}}
+\def\bbbq{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm
+Q$}\hbox{\raise
+0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}}
+{\setbox0=\hbox{$\textstyle\rm Q$}\hbox{\raise
+0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.8\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptstyle\rm Q$}\hbox{\raise
+0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptscriptstyle\rm Q$}\hbox{\raise
+0.15\ht0\hbox to0pt{\kern0.4\wd0\vrule height0.7\ht0\hss}\box0}}}}
+\def\bbbt{{\mathchoice {\setbox0=\hbox{$\displaystyle\rm
+T$}\hbox{\hbox to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
+{\setbox0=\hbox{$\textstyle\rm T$}\hbox{\hbox
+to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptstyle\rm T$}\hbox{\hbox
+to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptscriptstyle\rm T$}\hbox{\hbox
+to0pt{\kern0.3\wd0\vrule height0.9\ht0\hss}\box0}}}}
+\def\bbbs{{\mathchoice
+{\setbox0=\hbox{$\displaystyle     \rm S$}\hbox{\raise0.5\ht0\hbox
+to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox
+to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}}
+{\setbox0=\hbox{$\textstyle        \rm S$}\hbox{\raise0.5\ht0\hbox
+to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\hbox
+to0pt{\kern0.55\wd0\vrule height0.5\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptstyle      \rm S$}\hbox{\raise0.5\ht0\hbox
+to0pt{\kern0.35\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox
+to0pt{\kern0.5\wd0\vrule height0.45\ht0\hss}\box0}}
+{\setbox0=\hbox{$\scriptscriptstyle\rm S$}\hbox{\raise0.5\ht0\hbox
+to0pt{\kern0.4\wd0\vrule height0.45\ht0\hss}\raise0.05\ht0\hbox
+to0pt{\kern0.55\wd0\vrule height0.45\ht0\hss}\box0}}}}
+\def\bbbz{{\mathchoice {\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}}
+{\hbox{$\mathsf\textstyle Z\kern-0.4em Z$}}
+{\hbox{$\mathsf\scriptstyle Z\kern-0.3em Z$}}
+{\hbox{$\mathsf\scriptscriptstyle Z\kern-0.2em Z$}}}}
+
+\let\ts\,
+
+\setlength\leftmargini  {17\p@}
+\setlength\leftmargin    {\leftmargini}
+\setlength\leftmarginii  {\leftmargini}
+\setlength\leftmarginiii {\leftmargini}
+\setlength\leftmarginiv  {\leftmargini}
+\setlength  \labelsep  {.5em}
+\setlength  \labelwidth{\leftmargini}
+\addtolength\labelwidth{-\labelsep}
+
+\def\@listI{\leftmargin\leftmargini
+            \parsep 0\p@ \@plus1\p@ \@minus\p@
+            \topsep 8\p@ \@plus2\p@ \@minus4\p@
+            \itemsep0\p@}
+\let\@listi\@listI
+\@listi
+\def\@listii {\leftmargin\leftmarginii
+              \labelwidth\leftmarginii
+              \advance\labelwidth-\labelsep
+              \topsep    0\p@ \@plus2\p@ \@minus\p@}
+\def\@listiii{\leftmargin\leftmarginiii
+              \labelwidth\leftmarginiii
+              \advance\labelwidth-\labelsep
+              \topsep    0\p@ \@plus\p@\@minus\p@
+              \parsep    \z@
+              \partopsep \p@ \@plus\z@ \@minus\p@}
+
+\renewcommand\labelitemi{\normalfont\bfseries --}
+\renewcommand\labelitemii{$\m@th\bullet$}
+
+\setlength\arraycolsep{1.4\p@}
+\setlength\tabcolsep{1.4\p@}
+
+\def\tableofcontents{\chapter*{\contentsname\@mkboth{{\contentsname}}%
+                                                    {{\contentsname}}}
+ \def\authcount##1{\setcounter{auco}{##1}\setcounter{@auth}{1}}
+ \def\lastand{\ifnum\value{auco}=2\relax
+                 \unskip{} \andname\
+              \else
+                 \unskip \lastandname\
+              \fi}%
+ \def\and{\stepcounter{@auth}\relax
+          \ifnum\value{@auth}=\value{auco}%
+             \lastand
+          \else
+             \unskip,
+          \fi}%
+ \@starttoc{toc}\if@restonecol\twocolumn\fi}
+
+\def\l@part#1#2{\addpenalty{\@secpenalty}%
+   \addvspace{2em plus\p@}%  % space above part line
+   \begingroup
+     \parindent \z@
+     \rightskip \z@ plus 5em
+     \hrule\vskip5pt
+     \large               % same size as for a contribution heading
+     \bfseries\boldmath   % set line in boldface
+     \leavevmode          % TeX command to enter horizontal mode.
+     #1\par
+     \vskip5pt
+     \hrule
+     \vskip1pt
+     \nobreak             % Never break after part entry
+   \endgroup}
+
+\def\@dotsep{2}
+
+\def\hyperhrefextend{\ifx\hyper@anchor\@undefined\else
+{chapter.\thechapter}\fi}
+
+\def\addnumcontentsmark#1#2#3{%
+\addtocontents{#1}{\protect\contentsline{#2}{\protect\numberline
+                     {\thechapter}#3}{\thepage}\hyperhrefextend}}
+\def\addcontentsmark#1#2#3{%
+\addtocontents{#1}{\protect\contentsline{#2}{#3}{\thepage}\hyperhrefextend}}
+\def\addcontentsmarkwop#1#2#3{%
+\addtocontents{#1}{\protect\contentsline{#2}{#3}{0}\hyperhrefextend}}
+
+\def\@adcmk[#1]{\ifcase #1 \or
+\def\@gtempa{\addnumcontentsmark}%
+  \or    \def\@gtempa{\addcontentsmark}%
+  \or    \def\@gtempa{\addcontentsmarkwop}%
+  \fi\@gtempa{toc}{chapter}}
+\def\addtocmark{\@ifnextchar[{\@adcmk}{\@adcmk[3]}}
+
+\def\l@chapter#1#2{\addpenalty{-\@highpenalty}
+ \vskip 1.0em plus 1pt \@tempdima 1.5em \begingroup
+ \parindent \z@ \rightskip \@tocrmarg
+ \advance\rightskip by 0pt plus 2cm
+ \parfillskip -\rightskip \pretolerance=10000
+ \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
+ {\large\bfseries\boldmath#1}\ifx0#2\hfil\null
+ \else
+      \nobreak
+      \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
+      \@dotsep mu$}\hfill
+      \nobreak\hbox to\@pnumwidth{\hss #2}%
+ \fi\par
+ \penalty\@highpenalty \endgroup}
+
+\def\l@title#1#2{\addpenalty{-\@highpenalty}
+ \addvspace{8pt plus 1pt}
+ \@tempdima \z@
+ \begingroup
+ \parindent \z@ \rightskip \@tocrmarg
+ \advance\rightskip by 0pt plus 2cm
+ \parfillskip -\rightskip \pretolerance=10000
+ \leavevmode \advance\leftskip\@tempdima \hskip -\leftskip
+ #1\nobreak
+ \leaders\hbox{$\m@th \mkern \@dotsep mu.\mkern
+ \@dotsep mu$}\hfill
+ \nobreak\hbox to\@pnumwidth{\hss #2}\par
+ \penalty\@highpenalty \endgroup}
+
+\def\l@author#1#2{\addpenalty{\@highpenalty}
+ \@tempdima=\z@ %15\p@
+ \begingroup
+ \parindent \z@ \rightskip \@tocrmarg
+ \advance\rightskip by 0pt plus 2cm
+ \pretolerance=10000
+ \leavevmode \advance\leftskip\@tempdima %\hskip -\leftskip
+ \textit{#1}\par
+ \penalty\@highpenalty \endgroup}
+
+%\setcounter{tocdepth}{0}
+\newdimen\tocchpnum
+\newdimen\tocsecnum
+\newdimen\tocsectotal
+\newdimen\tocsubsecnum
+\newdimen\tocsubsectotal
+\newdimen\tocsubsubsecnum
+\newdimen\tocsubsubsectotal
+\newdimen\tocparanum
+\newdimen\tocparatotal
+\newdimen\tocsubparanum
+\tocchpnum=\z@            % no chapter numbers
+\tocsecnum=15\p@          % section 88. plus 2.222pt
+\tocsubsecnum=23\p@       % subsection 88.8 plus 2.222pt
+\tocsubsubsecnum=27\p@    % subsubsection 88.8.8 plus 1.444pt
+\tocparanum=35\p@         % paragraph 88.8.8.8 plus 1.666pt
+\tocsubparanum=43\p@      % subparagraph 88.8.8.8.8 plus 1.888pt
+\def\calctocindent{%
+\tocsectotal=\tocchpnum
+\advance\tocsectotal by\tocsecnum
+\tocsubsectotal=\tocsectotal
+\advance\tocsubsectotal by\tocsubsecnum
+\tocsubsubsectotal=\tocsubsectotal
+\advance\tocsubsubsectotal by\tocsubsubsecnum
+\tocparatotal=\tocsubsubsectotal
+\advance\tocparatotal by\tocparanum}
+\calctocindent
+
+\def\l@section{\@dottedtocline{1}{\tocchpnum}{\tocsecnum}}
+\def\l@subsection{\@dottedtocline{2}{\tocsectotal}{\tocsubsecnum}}
+\def\l@subsubsection{\@dottedtocline{3}{\tocsubsectotal}{\tocsubsubsecnum}}
+\def\l@paragraph{\@dottedtocline{4}{\tocsubsubsectotal}{\tocparanum}}
+\def\l@subparagraph{\@dottedtocline{5}{\tocparatotal}{\tocsubparanum}}
+
+\def\listoffigures{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
+ \fi\section*{\listfigurename\@mkboth{{\listfigurename}}{{\listfigurename}}}
+ \@starttoc{lof}\if@restonecol\twocolumn\fi}
+\def\l@figure{\@dottedtocline{1}{0em}{1.5em}}
+
+\def\listoftables{\@restonecolfalse\if@twocolumn\@restonecoltrue\onecolumn
+ \fi\section*{\listtablename\@mkboth{{\listtablename}}{{\listtablename}}}
+ \@starttoc{lot}\if@restonecol\twocolumn\fi}
+\let\l@table\l@figure
+
+\renewcommand\listoffigures{%
+    \section*{\listfigurename
+      \@mkboth{\listfigurename}{\listfigurename}}%
+    \@starttoc{lof}%
+    }
+
+\renewcommand\listoftables{%
+    \section*{\listtablename
+      \@mkboth{\listtablename}{\listtablename}}%
+    \@starttoc{lot}%
+    }
+
+\ifx\oribibl\undefined
+\ifx\citeauthoryear\undefined
+\renewenvironment{thebibliography}[1]
+     {\section*{\refname}
+      \def\@biblabel##1{##1.}
+      \small
+      \list{\@biblabel{\@arabic\c@enumiv}}%
+           {\settowidth\labelwidth{\@biblabel{#1}}%
+            \leftmargin\labelwidth
+            \advance\leftmargin\labelsep
+            \if@openbib
+              \advance\leftmargin\bibindent
+              \itemindent -\bibindent
+              \listparindent \itemindent
+              \parsep \z@
+            \fi
+            \usecounter{enumiv}%
+            \let\p@enumiv\@empty
+            \renewcommand\theenumiv{\@arabic\c@enumiv}}%
+      \if@openbib
+        \renewcommand\newblock{\par}%
+      \else
+        \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}%
+      \fi
+      \sloppy\clubpenalty4000\widowpenalty4000%
+      \sfcode`\.=\@m}
+     {\def\@noitemerr
+       {\@latex@warning{Empty `thebibliography' environment}}%
+      \endlist}
+\def\@lbibitem[#1]#2{\item[{[#1]}\hfill]\if@filesw
+     {\let\protect\noexpand\immediate
+     \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
+\newcount\@tempcntc
+\def\@citex[#1]#2{\if@filesw\immediate\write\@auxout{\string\citation{#2}}\fi
+  \@tempcnta\z@\@tempcntb\m@ne\def\@citea{}\@cite{\@for\@citeb:=#2\do
+    {\@ifundefined
+       {b@\@citeb}{\@citeo\@tempcntb\m@ne\@citea\def\@citea{,}{\bfseries
+        ?}\@warning
+       {Citation `\@citeb' on page \thepage \space undefined}}%
+    {\setbox\z@\hbox{\global\@tempcntc0\csname b@\@citeb\endcsname\relax}%
+     \ifnum\@tempcntc=\z@ \@citeo\@tempcntb\m@ne
+       \@citea\def\@citea{,}\hbox{\csname b@\@citeb\endcsname}%
+     \else
+      \advance\@tempcntb\@ne
+      \ifnum\@tempcntb=\@tempcntc
+      \else\advance\@tempcntb\m@ne\@citeo
+      \@tempcnta\@tempcntc\@tempcntb\@tempcntc\fi\fi}}\@citeo}{#1}}
+\def\@citeo{\ifnum\@tempcnta>\@tempcntb\else
+               \@citea\def\@citea{,\,\hskip\z@skip}%
+               \ifnum\@tempcnta=\@tempcntb\the\@tempcnta\else
+               {\advance\@tempcnta\@ne\ifnum\@tempcnta=\@tempcntb \else
+                \def\@citea{--}\fi
+      \advance\@tempcnta\m@ne\the\@tempcnta\@citea\the\@tempcntb}\fi\fi}
+\else
+\renewenvironment{thebibliography}[1]
+     {\section*{\refname}
+      \small
+      \list{}%
+           {\settowidth\labelwidth{}%
+            \leftmargin\parindent
+            \itemindent=-\parindent
+            \labelsep=\z@
+            \if@openbib
+              \advance\leftmargin\bibindent
+              \itemindent -\bibindent
+              \listparindent \itemindent
+              \parsep \z@
+            \fi
+            \usecounter{enumiv}%
+            \let\p@enumiv\@empty
+            \renewcommand\theenumiv{}}%
+      \if@openbib
+        \renewcommand\newblock{\par}%
+      \else
+        \renewcommand\newblock{\hskip .11em \@plus.33em \@minus.07em}%
+      \fi
+      \sloppy\clubpenalty4000\widowpenalty4000%
+      \sfcode`\.=\@m}
+     {\def\@noitemerr
+       {\@latex@warning{Empty `thebibliography' environment}}%
+      \endlist}
+      \def\@cite#1{#1}%
+      \def\@lbibitem[#1]#2{\item[]\if@filesw
+        {\def\protect##1{\string ##1\space}\immediate
+      \write\@auxout{\string\bibcite{#2}{#1}}}\fi\ignorespaces}
+   \fi
+\else
+\@cons\@openbib@code{\noexpand\small}
+\fi
+
+\def\idxquad{\hskip 10\p@}% space that divides entry from number
+
+\def\@idxitem{\par\hangindent 10\p@}
+
+\def\subitem{\par\setbox0=\hbox{--\enspace}% second order
+                \noindent\hangindent\wd0\box0}% index entry
+
+\def\subsubitem{\par\setbox0=\hbox{--\,--\enspace}% third
+                \noindent\hangindent\wd0\box0}% order index entry
+
+\def\indexspace{\par \vskip 10\p@ plus5\p@ minus3\p@\relax}
+
+\renewenvironment{theindex}
+               {\@mkboth{\indexname}{\indexname}%
+                \thispagestyle{empty}\parindent\z@
+                \parskip\z@ \@plus .3\p@\relax
+                \let\item\par
+                \def\,{\relax\ifmmode\mskip\thinmuskip
+                             \else\hskip0.2em\ignorespaces\fi}%
+                \normalfont\small
+                \begin{multicols}{2}[\@makeschapterhead{\indexname}]%
+                }
+                {\end{multicols}}
+
+\renewcommand\footnoterule{%
+  \kern-3\p@
+  \hrule\@width 2truecm
+  \kern2.6\p@}
+  \newdimen\fnindent
+  \fnindent1em
+\long\def\@makefntext#1{%
+    \parindent \fnindent%
+    \leftskip \fnindent%
+    \noindent
+    \llap{\hb@xt@1em{\hss\@makefnmark\ }}\ignorespaces#1}
+
+\long\def\@makecaption#1#2{%
+  \vskip\abovecaptionskip
+  \sbox\@tempboxa{{\bfseries #1.} #2}%
+  \ifdim \wd\@tempboxa >\hsize
+    {\bfseries #1.} #2\par
+  \else
+    \global \@minipagefalse
+    \hb@xt@\hsize{\hfil\box\@tempboxa\hfil}%
+  \fi
+  \vskip\belowcaptionskip}
+
+\def\fps@figure{htbp}
+\def\fnum@figure{\figurename\thinspace\thefigure}
+\def \@floatboxreset {%
+        \reset@font
+        \small
+        \@setnobreak
+        \@setminipage
+}
+\def\fps@table{htbp}
+\def\fnum@table{\tablename~\thetable}
+\renewenvironment{table}
+               {\setlength\abovecaptionskip{0\p@}%
+                \setlength\belowcaptionskip{10\p@}%
+                \@float{table}}
+               {\end@float}
+\renewenvironment{table*}
+               {\setlength\abovecaptionskip{0\p@}%
+                \setlength\belowcaptionskip{10\p@}%
+                \@dblfloat{table}}
+               {\end@dblfloat}
+
+\long\def\@caption#1[#2]#3{\par\addcontentsline{\csname
+  ext@#1\endcsname}{#1}{\protect\numberline{\csname
+  the#1\endcsname}{\ignorespaces #2}}\begingroup
+    \@parboxrestore
+    \@makecaption{\csname fnum@#1\endcsname}{\ignorespaces #3}\par
+  \endgroup}
+
+% LaTeX does not provide a command to enter the authors institute
+% addresses. The \institute command is defined here.
+
+\newcounter{@inst}
+\newcounter{@auth}
+\newcounter{auco}
+\newdimen\instindent
+\newbox\authrun
+\newtoks\authorrunning
+\newtoks\tocauthor
+\newbox\titrun
+\newtoks\titlerunning
+\newtoks\toctitle
+
+\def\clearheadinfo{\gdef\@author{No Author Given}%
+                   \gdef\@title{No Title Given}%
+                   \gdef\@subtitle{}%
+                   \gdef\@institute{No Institute Given}%
+                   \gdef\@thanks{}%
+                   \global\titlerunning={}\global\authorrunning={}%
+                   \global\toctitle={}\global\tocauthor={}}
+
+\def\institute#1{\gdef\@institute{#1}}
+
+\def\institutename{\par
+ \begingroup
+ \parskip=\z@
+ \parindent=\z@
+ \setcounter{@inst}{1}%
+ \def\and{\par\stepcounter{@inst}%
+ \noindent$^{\the@inst}$\enspace\ignorespaces}%
+ \setbox0=\vbox{\def\thanks##1{}\@institute}%
+ \ifnum\c@@inst=1\relax
+   \gdef\fnnstart{0}%
+ \else
+   \xdef\fnnstart{\c@@inst}%
+   \setcounter{@inst}{1}%
+   \noindent$^{\the@inst}$\enspace
+ \fi
+ \ignorespaces
+ \@institute\par
+ \endgroup}
+
+\def\@fnsymbol#1{\ensuremath{\ifcase#1\or\star\or{\star\star}\or
+   {\star\star\star}\or \dagger\or \ddagger\or
+   \mathchar "278\or \mathchar "27B\or \|\or **\or \dagger\dagger
+   \or \ddagger\ddagger \else\@ctrerr\fi}}
+
+\def\inst#1{\unskip$^{#1}$}
+\def\fnmsep{\unskip$^,$}
+\def\email#1{{\tt#1}}
+\AtBeginDocument{\@ifundefined{url}{\def\url#1{#1}}{}%
+\@ifpackageloaded{babel}{%
+\@ifundefined{extrasenglish}{}{\addto\extrasenglish{\switcht@albion}}%
+\@ifundefined{extrasfrenchb}{}{\addto\extrasfrenchb{\switcht@francais}}%
+\@ifundefined{extrasgerman}{}{\addto\extrasgerman{\switcht@deutsch}}%
+}{\switcht@@therlang}%
+}
+\def\homedir{\~{ }}
+
+\def\subtitle#1{\gdef\@subtitle{#1}}
+\clearheadinfo
+
+\renewcommand\maketitle{\newpage
+  \refstepcounter{chapter}%
+  \stepcounter{section}%
+  \setcounter{section}{0}%
+  \setcounter{subsection}{0}%
+  \setcounter{figure}{0}
+  \setcounter{table}{0}
+  \setcounter{equation}{0}
+  \setcounter{footnote}{0}%
+  \begingroup
+    \parindent=\z@
+    \renewcommand\thefootnote{\@fnsymbol\c@footnote}%
+    \if@twocolumn
+      \ifnum \col@number=\@ne
+        \@maketitle
+      \else
+        \twocolumn[\@maketitle]%
+      \fi
+    \else
+      \newpage
+      \global\@topnum\z@   % Prevents figures from going at top of page.
+      \@maketitle
+    \fi
+    \thispagestyle{empty}\@thanks
+%
+    \def\\{\unskip\ \ignorespaces}\def\inst##1{\unskip{}}%
+    \def\thanks##1{\unskip{}}\def\fnmsep{\unskip}%
+    \instindent=\hsize
+    \advance\instindent by-\headlineindent
+%    \if!\the\toctitle!\addcontentsline{toc}{title}{\@title}\else
+%       \addcontentsline{toc}{title}{\the\toctitle}\fi
+    \if@runhead
+       \if!\the\titlerunning!\else
+         \edef\@title{\the\titlerunning}%
+       \fi
+       \global\setbox\titrun=\hbox{\small\rm\unboldmath\ignorespaces\@title}%
+       \ifdim\wd\titrun>\instindent
+          \typeout{Title too long for running head. Please supply}%
+          \typeout{a shorter form with \string\titlerunning\space prior to
+                   \string\maketitle}%
+          \global\setbox\titrun=\hbox{\small\rm
+          Title Suppressed Due to Excessive Length}%
+       \fi
+       \xdef\@title{\copy\titrun}%
+    \fi
+%
+    \if!\the\tocauthor!\relax
+      {\def\and{\noexpand\protect\noexpand\and}%
+      \protected@xdef\toc@uthor{\@author}}%
+    \else
+      \def\\{\noexpand\protect\noexpand\newline}%
+      \protected@xdef\scratch{\the\tocauthor}%
+      \protected@xdef\toc@uthor{\scratch}%
+    \fi
+%    \addcontentsline{toc}{author}{\toc@uthor}%
+    \if@runhead
+       \if!\the\authorrunning!
+         \value{@inst}=\value{@auth}%
+         \setcounter{@auth}{1}%
+       \else
+         \edef\@author{\the\authorrunning}%
+       \fi
+       \global\setbox\authrun=\hbox{\small\unboldmath\@author\unskip}%
+       \ifdim\wd\authrun>\instindent
+          \typeout{Names of authors too long for running head. Please supply}%
+          \typeout{a shorter form with \string\authorrunning\space prior to
+                   \string\maketitle}%
+          \global\setbox\authrun=\hbox{\small\rm
+          Authors Suppressed Due to Excessive Length}%
+       \fi
+       \xdef\@author{\copy\authrun}%
+       \markboth{\@author}{\@title}%
+     \fi
+  \endgroup
+  \setcounter{footnote}{\fnnstart}%
+  \clearheadinfo}
+%
+\def\@maketitle{\newpage
+ \markboth{}{}%
+ \def\lastand{\ifnum\value{@inst}=2\relax
+                 \unskip{} \andname\
+              \else
+                 \unskip \lastandname\
+              \fi}%
+ \def\and{\stepcounter{@auth}\relax
+          \ifnum\value{@auth}=\value{@inst}%
+             \lastand
+          \else
+             \unskip,
+          \fi}%
+ \begin{center}%
+ \let\newline\\
+ {\Large \bfseries\boldmath
+  \pretolerance=10000
+  \@title \par}\vskip .8cm
+\if!\@subtitle!\else {\large \bfseries\boldmath
+  \vskip -.65cm
+  \pretolerance=10000
+  \@subtitle \par}\vskip .8cm\fi
+ \setbox0=\vbox{\setcounter{@auth}{1}\def\and{\stepcounter{@auth}}%
+ \def\thanks##1{}\@author}%
+ \global\value{@inst}=\value{@auth}%
+ \global\value{auco}=\value{@auth}%
+ \setcounter{@auth}{1}%
+{\lineskip .5em
+\noindent\ignorespaces
+\@author\vskip.35cm}
+ {\small\institutename}
+ \end{center}%
+ }
+
+% definition of the "\spnewtheorem" command.
+%
+% Usage:
+%
+%     \spnewtheorem{env_nam}{caption}[within]{cap_font}{body_font}
+% or  \spnewtheorem{env_nam}[numbered_like]{caption}{cap_font}{body_font}
+% or  \spnewtheorem*{env_nam}{caption}{cap_font}{body_font}
+%
+% New is "cap_font" and "body_font". It stands for
+% fontdefinition of the caption and the text itself.
+%
+% "\spnewtheorem*" gives a theorem without number.
+%
+% A defined spnewthoerem environment is used as described
+% by Lamport.
+%
+%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+
+\def\@thmcountersep{}
+\def\@thmcounterend{.}
+
+\def\spnewtheorem{\@ifstar{\@sthm}{\@Sthm}}
+
+% definition of \spnewtheorem with number
+
+\def\@spnthm#1#2{%
+  \@ifnextchar[{\@spxnthm{#1}{#2}}{\@spynthm{#1}{#2}}}
+\def\@Sthm#1{\@ifnextchar[{\@spothm{#1}}{\@spnthm{#1}}}
+
+\def\@spxnthm#1#2[#3]#4#5{\expandafter\@ifdefinable\csname #1\endcsname
+   {\@definecounter{#1}\@addtoreset{#1}{#3}%
+   \expandafter\xdef\csname the#1\endcsname{\expandafter\noexpand
+     \csname the#3\endcsname \noexpand\@thmcountersep \@thmcounter{#1}}%
+   \expandafter\xdef\csname #1name\endcsname{#2}%
+   \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#4}{#5}}%
+                              \global\@namedef{end#1}{\@endtheorem}}}
+
+\def\@spynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
+   {\@definecounter{#1}%
+   \expandafter\xdef\csname the#1\endcsname{\@thmcounter{#1}}%
+   \expandafter\xdef\csname #1name\endcsname{#2}%
+   \global\@namedef{#1}{\@spthm{#1}{\csname #1name\endcsname}{#3}{#4}}%
+                               \global\@namedef{end#1}{\@endtheorem}}}
+
+\def\@spothm#1[#2]#3#4#5{%
+  \@ifundefined{c@#2}{\@latexerr{No theorem environment `#2' defined}\@eha}%
+  {\expandafter\@ifdefinable\csname #1\endcsname
+  {\global\@namedef{the#1}{\@nameuse{the#2}}%
+  \expandafter\xdef\csname #1name\endcsname{#3}%
+  \global\@namedef{#1}{\@spthm{#2}{\csname #1name\endcsname}{#4}{#5}}%
+  \global\@namedef{end#1}{\@endtheorem}}}}
+
+\def\@spthm#1#2#3#4{\topsep 7\p@ \@plus2\p@ \@minus4\p@
+\refstepcounter{#1}%
+\@ifnextchar[{\@spythm{#1}{#2}{#3}{#4}}{\@spxthm{#1}{#2}{#3}{#4}}}
+
+\def\@spxthm#1#2#3#4{\@spbegintheorem{#2}{\csname the#1\endcsname}{#3}{#4}%
+                    \ignorespaces}
+
+\def\@spythm#1#2#3#4[#5]{\@spopargbegintheorem{#2}{\csname
+       the#1\endcsname}{#5}{#3}{#4}\ignorespaces}
+
+\def\@spbegintheorem#1#2#3#4{\trivlist
+                 \item[\hskip\labelsep{#3#1\ #2\@thmcounterend}]#4}
+
+\def\@spopargbegintheorem#1#2#3#4#5{\trivlist
+      \item[\hskip\labelsep{#4#1\ #2}]{#4(#3)\@thmcounterend\ }#5}
+
+% definition of \spnewtheorem* without number
+
+\def\@sthm#1#2{\@Ynthm{#1}{#2}}
+
+\def\@Ynthm#1#2#3#4{\expandafter\@ifdefinable\csname #1\endcsname
+   {\global\@namedef{#1}{\@Thm{\csname #1name\endcsname}{#3}{#4}}%
+    \expandafter\xdef\csname #1name\endcsname{#2}%
+    \global\@namedef{end#1}{\@endtheorem}}}
+
+\def\@Thm#1#2#3{\topsep 7\p@ \@plus2\p@ \@minus4\p@
+\@ifnextchar[{\@Ythm{#1}{#2}{#3}}{\@Xthm{#1}{#2}{#3}}}
+
+\def\@Xthm#1#2#3{\@Begintheorem{#1}{#2}{#3}\ignorespaces}
+
+\def\@Ythm#1#2#3[#4]{\@Opargbegintheorem{#1}
+       {#4}{#2}{#3}\ignorespaces}
+
+\def\@Begintheorem#1#2#3{#3\trivlist
+                           \item[\hskip\labelsep{#2#1\@thmcounterend}]}
+
+\def\@Opargbegintheorem#1#2#3#4{#4\trivlist
+      \item[\hskip\labelsep{#3#1}]{#3(#2)\@thmcounterend\ }}
+
+\if@envcntsect
+   \def\@thmcountersep{.}
+   \spnewtheorem{theorem}{Theorem}[section]{\bfseries}{\itshape}
+\else
+   \spnewtheorem{theorem}{Theorem}{\bfseries}{\itshape}
+   \if@envcntreset
+      \@addtoreset{theorem}{section}
+   \else
+      \@addtoreset{theorem}{chapter}
+   \fi
+\fi
+
+%definition of divers theorem environments
+\spnewtheorem*{claim}{Claim}{\itshape}{\rmfamily}
+\spnewtheorem*{proof}{Proof}{\itshape}{\rmfamily}
+\if@envcntsame % alle Umgebungen wie Theorem.
+   \def\spn@wtheorem#1#2#3#4{\@spothm{#1}[theorem]{#2}{#3}{#4}}
+\else % alle Umgebungen mit eigenem Zaehler
+   \if@envcntsect % mit section numeriert
+      \def\spn@wtheorem#1#2#3#4{\@spxnthm{#1}{#2}[section]{#3}{#4}}
+   \else % nicht mit section numeriert
+      \if@envcntreset
+         \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
+                                   \@addtoreset{#1}{section}}
+      \else
+         \def\spn@wtheorem#1#2#3#4{\@spynthm{#1}{#2}{#3}{#4}
+                                   \@addtoreset{#1}{chapter}}%
+      \fi
+   \fi
+\fi
+\spn@wtheorem{case}{Case}{\itshape}{\rmfamily}
+\spn@wtheorem{conjecture}{Conjecture}{\itshape}{\rmfamily}
+\spn@wtheorem{corollary}{Corollary}{\bfseries}{\itshape}
+\spn@wtheorem{definition}{Definition}{\bfseries}{\itshape}
+\spn@wtheorem{example}{Example}{\itshape}{\rmfamily}
+\spn@wtheorem{exercise}{Exercise}{\itshape}{\rmfamily}
+\spn@wtheorem{lemma}{Lemma}{\bfseries}{\itshape}
+\spn@wtheorem{note}{Note}{\itshape}{\rmfamily}
+\spn@wtheorem{problem}{Problem}{\itshape}{\rmfamily}
+\spn@wtheorem{property}{Property}{\itshape}{\rmfamily}
+\spn@wtheorem{proposition}{Proposition}{\bfseries}{\itshape}
+\spn@wtheorem{question}{Question}{\itshape}{\rmfamily}
+\spn@wtheorem{solution}{Solution}{\itshape}{\rmfamily}
+\spn@wtheorem{remark}{Remark}{\itshape}{\rmfamily}
+
+\def\@takefromreset#1#2{%
+    \def\@tempa{#1}%
+    \let\@tempd\@elt
+    \def\@elt##1{%
+        \def\@tempb{##1}%
+        \ifx\@tempa\@tempb\else
+            \@addtoreset{##1}{#2}%
+        \fi}%
+    \expandafter\expandafter\let\expandafter\@tempc\csname cl@#2\endcsname
+    \expandafter\def\csname cl@#2\endcsname{}%
+    \@tempc
+    \let\@elt\@tempd}
+
+\def\theopargself{\def\@spopargbegintheorem##1##2##3##4##5{\trivlist
+      \item[\hskip\labelsep{##4##1\ ##2}]{##4##3\@thmcounterend\ }##5}
+                  \def\@Opargbegintheorem##1##2##3##4{##4\trivlist
+      \item[\hskip\labelsep{##3##1}]{##3##2\@thmcounterend\ }}
+      }
+
+\renewenvironment{abstract}{%
+      \list{}{\advance\topsep by0.35cm\relax\small
+      \leftmargin=1cm
+      \labelwidth=\z@
+      \listparindent=\z@
+      \itemindent\listparindent
+      \rightmargin\leftmargin}\item[\hskip\labelsep
+                                    \bfseries\abstractname]}
+    {\endlist}
+
+\newdimen\headlineindent             % dimension for space between
+\headlineindent=1.166cm              % number and text of headings.
+
+\def\ps@headings{\let\@mkboth\@gobbletwo
+   \let\@oddfoot\@empty\let\@evenfoot\@empty
+   \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}%
+                  \leftmark\hfil}
+   \def\@oddhead{\normalfont\small\hfil\rightmark\hspace{\headlineindent}%
+                 \llap{\thepage}}
+   \def\chaptermark##1{}%
+   \def\sectionmark##1{}%
+   \def\subsectionmark##1{}}
+
+\def\ps@titlepage{\let\@mkboth\@gobbletwo
+   \let\@oddfoot\@empty\let\@evenfoot\@empty
+   \def\@evenhead{\normalfont\small\rlap{\thepage}\hspace{\headlineindent}%
+                  \hfil}
+   \def\@oddhead{\normalfont\small\hfil\hspace{\headlineindent}%
+                 \llap{\thepage}}
+   \def\chaptermark##1{}%
+   \def\sectionmark##1{}%
+   \def\subsectionmark##1{}}
+
+\if@runhead\ps@headings\else
+\ps@empty\fi
+
+\setlength\arraycolsep{1.4\p@}
+\setlength\tabcolsep{1.4\p@}
+
+\endinput
+%end of file llncs.cls
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys2/Journal/pdfsetup.sty	Mon Nov 01 10:40:21 2021 +0000
@@ -0,0 +1,7 @@
+%%
+%% default hyperref setup (both for pdf and dvi output)
+%%
+
+\usepackage{color}
+\definecolor{linkcolor}{rgb}{0,0,0.5}
+\usepackage[colorlinks=true,linkcolor=linkcolor,citecolor=linkcolor,filecolor=linkcolor,urlcolor=linkcolor,pdfpagelabels]{hyperref}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys2/Journal/railsetup.sty	Mon Nov 01 10:40:21 2021 +0000
@@ -0,0 +1,1202 @@
+% rail.sty - style file to support railroad diagrams
+%
+% 09-Jul-90 L. Rooijakkers
+% 08-Oct-90 L. Rooijakkers	fixed centering bug when \rail@tmpc<0.
+% 07-Feb-91 L. Rooijakkers	added \railoptions command, indexing
+% 08-Feb-91 L. Rooijakkers	minor fixes
+% 28-Jun-94 K. Barthelmann	turned into LaTeX2e package
+% 08-Dec-96 K. Barthelmann	replaced \@writefile
+% 13-Dec-96 K. Barthelmann	cleanup
+% 22-Feb-98 K. Barthelmann	fixed catcodes of special characters
+% 18-Apr-98 K. Barthelmann	fixed \par handling
+% 19-May-98 J. Olsson		Added new macros to support arrow heads.
+% 26-Jul-98 K. Barthelmann	changed \par to output newlines
+% 02-May-11 M. Wenzel           default setup for Isabelle
+%
+% This style file needs to be used in conjunction with the 'rail'
+% program. Running LaTeX as 'latex file' produces file.rai, which should be
+% processed by Rail with 'rail file'. This produces file.rao, which will
+% be picked up by LaTeX on the next 'latex file' run.
+%
+% LaTeX will warn if there is no file.rao or it's out of date.
+%
+% The macros in this file thus consist of two parts: those that read and
+% write the .rai and .rao files, and those that do the actual formatting
+% of the railroad diagrams.
+
+\NeedsTeXFormat{LaTeX2e}
+\ProvidesPackage{rail}[1998/05/19]
+
+% railroad diagram formatting parameters (user level)
+% all of these are copied into their internal versions by \railinit
+%
+% \railunit : \unitlength within railroad diagrams
+% \railextra : extra length at outside of diagram
+% \railboxheight : height of ovals and frames
+% \railboxskip : vertical space between lines
+% \railboxleft : space to the left of a box
+% \railboxright : space to the right of a box
+% \railovalspace : extra space around contents of oval
+% \railframespace : extra space around contents of frame
+% \railtextleft : space to the left of text
+% \railtextright : space to the right of text
+% \railtextup : space to lift text up
+% \railjoinsize : circle size of join/split arcs
+% \railjoinadjust : space to adjust join
+%
+% \railnamesep : separator between name and rule body
+
+\newlength\railunit
+\newlength\railextra
+\newlength\railboxheight
+\newlength\railboxskip
+\newlength\railboxleft
+\newlength\railboxright
+\newlength\railovalspace
+\newlength\railframespace
+\newlength\railtextleft
+\newlength\railtextright
+\newlength\railtextup
+\newlength\railjoinsize
+\newlength\railjoinadjust
+\newlength\railnamesep
+
+% initialize the parameters
+
+\setlength\railunit{1sp}
+\setlength\railextra{4ex}
+\setlength\railboxleft{1ex}
+\setlength\railboxright{1ex}
+\setlength\railovalspace{2ex}
+\setlength\railframespace{2ex}
+\setlength\railtextleft{1ex}
+\setlength\railtextright{1ex}
+\setlength\railjoinadjust{0pt}
+\setlength\railnamesep{1ex}
+
+\DeclareOption{10pt}{
+  \setlength\railboxheight{16pt}
+  \setlength\railboxskip{24pt}
+  \setlength\railtextup{5pt}
+  \setlength\railjoinsize{16pt}
+}
+\DeclareOption{11pt}{
+  \setlength\railboxheight{16pt}
+  \setlength\railboxskip{24pt}
+  \setlength\railtextup{5pt}
+  \setlength\railjoinsize{16pt}
+}
+\DeclareOption{12pt}{
+  \setlength\railboxheight{20pt}
+  \setlength\railboxskip{28pt}
+  \setlength\railtextup{6pt}
+  \setlength\railjoinsize{20pt}
+}
+
+\ExecuteOptions{10pt}
+\ProcessOptions
+
+% internal versions of the formatting parameters
+%
+% \rail@extra   : \railextra
+% \rail@boxht   : \railboxheight
+% \rail@boxsp   : \railboxskip
+% \rail@boxlf   : \railboxleft
+% \rail@boxrt   : \railboxright
+% \rail@boxhht  : \railboxheight / 2
+% \rail@ovalsp  : \railovalspace
+% \rail@framesp : \railframespace
+% \rail@textlf  : \railtextleft
+% \rail@textrt  : \railtextright
+% \rail@textup  : \railtextup
+% \rail@joinsz  : \railjoinsize
+% \rail@joinhsz : \railjoinsize / 2
+% \rail@joinadj : \railjoinadjust
+%
+% \railinit : internalize all of the parameters.
+
+\newcount\rail@extra
+\newcount\rail@boxht
+\newcount\rail@boxsp
+\newcount\rail@boxlf
+\newcount\rail@boxrt
+\newcount\rail@boxhht
+\newcount\rail@ovalsp
+\newcount\rail@framesp
+\newcount\rail@textlf
+\newcount\rail@textrt
+\newcount\rail@textup
+\newcount\rail@joinsz
+\newcount\rail@joinhsz
+\newcount\rail@joinadj
+
+\newcommand\railinit{
+\rail@extra=\railextra
+\divide\rail@extra by \railunit
+\rail@boxht=\railboxheight
+\divide\rail@boxht by \railunit
+\rail@boxsp=\railboxskip
+\divide\rail@boxsp by \railunit
+\rail@boxlf=\railboxleft
+\divide\rail@boxlf by \railunit
+\rail@boxrt=\railboxright
+\divide\rail@boxrt by \railunit
+\rail@boxhht=\railboxheight
+\divide\rail@boxhht by \railunit
+\divide\rail@boxhht by 2
+\rail@ovalsp=\railovalspace
+\divide\rail@ovalsp by \railunit
+\rail@framesp=\railframespace
+\divide\rail@framesp by \railunit
+\rail@textlf=\railtextleft
+\divide\rail@textlf by \railunit
+\rail@textrt=\railtextright
+\divide\rail@textrt by \railunit
+\rail@textup=\railtextup
+\divide\rail@textup by \railunit
+\rail@joinsz=\railjoinsize
+\divide\rail@joinsz by \railunit
+\rail@joinhsz=\railjoinsize
+\divide\rail@joinhsz by \railunit
+\divide\rail@joinhsz by 2
+\rail@joinadj=\railjoinadjust
+\divide\rail@joinadj by \railunit
+}
+
+\AtBeginDocument{\railinit}
+
+% \rail@param : declarations for list environment
+%
+% \railparam{TEXT} : sets \rail@param to TEXT
+%
+% \rail@reserved : characters reserved for grammar
+
+\newcommand\railparam[1]{
+\def\rail@param{
+  \setlength\leftmargin{0pt}\setlength\rightmargin{0pt}
+  \setlength\labelwidth{0pt}\setlength\labelsep{0pt}
+  \setlength\itemindent{0pt}\setlength\listparindent{0pt}
+  #1
+}
+}
+\railparam{}
+
+\newtoks\rail@reserved
+\rail@reserved={:;|*+?[]()'"}
+
+% \rail@termfont : format setup for terminals
+%
+% \rail@nontfont : format setup for nonterminals
+%
+% \rail@annofont : format setup for annotations
+%
+% \rail@rulefont : format setup for rule names
+%
+% \rail@indexfont : format setup for index entry
+%
+% \railtermfont{TEXT} : set terminal format setup to TEXT
+%
+% \railnontermfont{TEXT} : set nonterminal format setup to TEXT
+%
+% \railannotatefont{TEXT} : set annotation format setup to TEXT
+%
+% \railnamefont{TEXT} : set rule name format setup to TEXT
+%
+% \railindexfont{TEXT} : set index entry format setup to TEXT
+
+\def\rail@termfont{\ttfamily\upshape}
+\def\rail@nontfont{\rmfamily\upshape}
+\def\rail@annofont{\rmfamily\itshape}
+\def\rail@namefont{\rmfamily\itshape}
+\def\rail@indexfont{\rmfamily\itshape}
+
+\newcommand\railtermfont[1]{
+\def\rail@termfont{#1}
+}
+
+\newcommand\railnontermfont[1]{
+\def\rail@nontfont{#1}
+}
+
+\newcommand\railannotatefont[1]{
+\def\rail@annofont{#1}
+}
+
+\newcommand\railnamefont[1]{
+\def\rail@namefont{#1}
+}
+
+\newcommand\railindexfont[1]{
+\def\rail@indexfont{#1}
+}
+
+% railroad read/write macros
+%
+% \begin{rail} TEXT \end{rail} : TEXT is written out to the .rai file,
+%                                as \rail@i{NR}{TEXT}. Then the matching
+%                                \rail@o{NR}{FMT} from the .rao file is
+%                                executed (if defined).
+%
+% \railoptions{OPTIONS} : OPTIONS are written out to the .rai file,
+%                         as \rail@p{OPTIONS}.
+%
+% \railterm{IDENT,IDENT,...} : format IDENT as terminals. writes out
+%                              \rail@t{IDENT} to the .rai file
+%
+% \railalias{IDENT}{TEXT} : format IDENT as TEXT. defines \rail@t@IDENT as
+%                           TEXT.
+%
+% \railtoken{IDENT}{TEXT} : abbreviates \railalias{IDENT}{TEXT}\railterm{IDENT}
+%                           (for backward compatibility)
+%
+% \rail@setcodes : guards special characters
+%
+% \rail@makeother{CHARACTER} : sets \catcode of CHARACTER to "other"
+%                              used inside a loop for \rail@setcodes
+%
+% \rail@nr : railroad diagram counter
+%
+% \ifrail@match : current \rail@i{NR}{TEXT} matches
+%
+% \rail@first : actions to be done first. read in .rao file,
+%               open .rai file if \@filesw true, undefine \rail@first.
+%               executed from \begin{rail}, \railoptions and \railterm.
+%
+% \rail@i{NR}{TEXT} : defines \rail@i@NR as TEXT. written to the .rai
+%                     file by \rail, read from the .rao file by
+%                     \rail@first
+%
+% \rail@t{IDENT} : tells Rail that IDENT is to be custom formatted,
+%                  written to the .rai file by \railterm.
+%
+% \rail@o{NR}{TEXT} : defines \rail@o@NR as TEXT, read from the .rao
+%                     file by \rail@first.
+%
+% \rail@p{OPTIONS} : pass options to rail, written to the .rai file by
+%                    \railoptions
+%
+% \rail@write{TEXT} : write TEXT to the .rai file
+%
+% \rail@warn : warn user for mismatching diagrams
+%
+% \rail@endwarn : either \relax or \rail@warn
+%
+% \ifrail@all : checked at the end of the document
+
+\def\rail@makeother#1{
+  \expandafter\catcode\expandafter`\csname\string #1\endcsname=12
+}
+
+\def\rail@setcodes{
+\let\par=\relax
+\let\\=\relax
+\expandafter\@tfor\expandafter\rail@symbol\expandafter:\expandafter=%
+  \the\rail@reserved
+\do{\expandafter\rail@makeother\rail@symbol}
+}
+
+\newcount\rail@nr
+
+\newif\ifrail@all
+\rail@alltrue
+
+\newif\ifrail@match
+
+\def\rail@first{
+\begingroup
+\makeatletter
+\rail@setcodes
+\InputIfFileExists{\jobname.rao}{}{\PackageInfo{rail}{No file \jobname.rao}}
+\makeatother
+\endgroup
+\if@filesw
+\newwrite\tf@rai
+\immediate\openout\tf@rai=\jobname.rai
+\fi
+\global\let\rail@first=\relax
+}
+
+\long\def\rail@body#1\end{
+{
+\newlinechar=`^^J
+\def\par{\string\par^^J}
+\rail@write{\string\rail@i{\number\rail@nr}{#1}}
+}
+\xdef\rail@i@{#1}
+\end
+}
+
+\newenvironment{rail}{
+\global\advance\rail@nr by 1
+\rail@first
+\begingroup
+\rail@setcodes
+\rail@body
+}{
+\endgroup
+\rail@matchtrue
+\@ifundefined{rail@o@\number\rail@nr}{\rail@matchfalse}{}
+\expandafter\ifx\csname rail@i@\number\rail@nr\endcsname\rail@i@
+\else
+\rail@matchfalse
+\fi
+\ifrail@match
+\csname rail@o@\number\rail@nr\endcsname
+\else
+\PackageWarning{rail}{Railroad diagram {\number\rail@nr} doesn't match}
+\global\let\rail@endwarn=\rail@warn
+\begin{list}{}{\rail@param}
+\rail@begin{1}{}
+\rail@setbox{\bfseries ???}
+\rail@oval
+\rail@end
+\end{list}
+\fi
+}
+
+\newcommand\railoptions[1]{
+\rail@first
+\rail@write{\string\rail@p{#1}}
+}
+
+\newcommand\railterm[1]{
+\rail@first
+\@for\rail@@:=#1\do{
+\rail@write{\string\rail@t{\rail@@}}
+}
+}
+
+\newcommand\railalias[2]{
+\expandafter\def\csname rail@t@#1\endcsname{#2}
+}
+
+\newcommand\railtoken[2]{\railalias{#1}{#2}\railterm{#1}}
+
+\long\def\rail@i#1#2{
+\expandafter\gdef\csname rail@i@#1\endcsname{#2}
+}
+
+\def\rail@o#1#2{
+\expandafter\gdef\csname rail@o@#1\endcsname{
+\begin{list}{}{\rail@param}
+#2
+\end{list}
+}
+}
+
+\def\rail@t#1{}
+
+\def\rail@p#1{}
+
+\long\def\rail@write#1{\@ifundefined{tf@rai}{}{\immediate\write\tf@rai{#1}}}
+
+\def\rail@warn{
+\PackageWarningNoLine{rail}{Railroad diagram(s) may have changed.
+                            Use 'rail' and rerun}
+}
+
+\let\rail@endwarn=\relax
+
+\AtEndDocument{\rail@endwarn}
+
+% index entry macro
+%
+% \rail@index{IDENT} : add index entry for IDENT
+
+\def\rail@index#1{
+\index{\rail@indexfont#1}
+}
+
+% railroad formatting primitives
+%
+% \rail@x : current x
+% \rail@y : current y
+% \rail@ex : current end x
+% \rail@sx : starting x for \rail@cr
+% \rail@rx : rightmost previous x for \rail@cr
+%
+% \rail@tmpa : temporary count
+% \rail@tmpb : temporary count
+% \rail@tmpc : temporary count
+%
+% \rail@put : put at (\rail@x,\rail@y)
+% \rail@vput : put vector at (\rail@x,\rail@y)
+%
+% \rail@eline : end line by drawing from \rail@ex to \rail@x
+%
+% \rail@vreline : end line by drawing a vector from \rail@x to \rail@ex
+%
+% \rail@vleline : end line by drawing a vector from \rail@ex to \rail@x
+%
+% \rail@sety{LEVEL} : set \rail@y to level LEVEL
+
+\newcount\rail@x
+\newcount\rail@y
+\newcount\rail@ex
+\newcount\rail@sx
+\newcount\rail@rx
+
+\newcount\rail@tmpa
+\newcount\rail@tmpb
+\newcount\rail@tmpc
+
+\def\rail@put{\put(\number\rail@x,\number\rail@y)}
+
+\def\rail@vput{\put(\number\rail@ex,\number\rail@y)}
+
+\def\rail@eline{
+\rail@tmpb=\rail@x
+\advance\rail@tmpb by -\rail@ex
+\rail@put{\line(-1,0){\number\rail@tmpb}}
+}
+
+\def\rail@vreline{
+\rail@tmpb=\rail@x
+\advance\rail@tmpb by -\rail@ex
+\rail@vput{\vector(1,0){\number\rail@tmpb}}
+}
+
+\def\rail@vleline{
+\rail@tmpb=\rail@x
+\advance\rail@tmpb by -\rail@ex
+\rail@put{\vector(-1,0){\number\rail@tmpb}}
+}
+
+\def\rail@sety#1{
+\rail@y=#1
+\multiply\rail@y by -\rail@boxsp
+\advance\rail@y by -\rail@boxht
+}
+
+% \rail@begin{HEIGHT}{NAME} : begin a railroad diagram of height HEIGHT
+%
+% \rail@end : end a railroad diagram
+%
+% \rail@expand{IDENT} : expand IDENT
+
+\def\rail@begin#1#2{
+\item
+\begin{minipage}[t]{\linewidth}
+\ifx\@empty#2\else
+{\rail@namefont \rail@expand{#2}}\\*[\railnamesep]
+\fi
+\unitlength=\railunit
+\rail@tmpa=#1
+\multiply\rail@tmpa by \rail@boxsp
+\begin{picture}(0,\number\rail@tmpa)(0,-\number\rail@tmpa)
+\rail@ex=0
+\rail@rx=0
+\rail@x=\rail@extra
+\rail@sx=\rail@x
+\rail@sety{0}
+}
+
+\def\rail@end{
+\advance\rail@x by \rail@extra
+\rail@eline
+\end{picture}
+\end{minipage}
+}
+
+\def\rail@vend{
+\advance\rail@x by \rail@extra
+\rail@vreline
+\end{picture}
+\end{minipage}
+}
+
+\def\rail@expand#1{\@ifundefined{rail@t@#1}{#1}{\csname rail@t@#1\endcsname}}
+
+% \rail@token{TEXT}[ANNOT] : format token TEXT with annotation
+% \rail@ltoken{TEXT}[ANNOT] : format token TEXT with annotation, arrow left
+% \rail@rtoken{TEXT}[ANNOT] : format token TEXT with annotation, arrow right
+%
+% \rail@ctoken{TEXT}[ANNOT] : format token TEXT centered with annotation
+% \rail@lctoken{TEXT}[ANNOT] : format token TEXT centered with annotation, arrow left
+% \rail@rctoken{TEXT}[ANNOT] : format token TEXT centered with annotation, arrow right
+%
+% \rail@nont{TEXT}[ANNOT] : format nonterminal TEXT with annotation
+% \rail@lnont{TEXT}[ANNOT] : format nonterminal TEXT with annotation, arrow left
+% \rail@rnont{TEXT}[ANNOT] : format nonterminal TEXT with annotation. arrow right
+%
+% \rail@cnont{TEXT}[ANNOT] : format nonterminal TEXT centered with annotation
+% \rail@lcnont{TEXT}[ANNOT] : format nonterminal TEXT centered with annotation,
+%                             arrow left
+% \rail@rcnont{TEXT}[ANNOT] : format nonterminal TEXT centered with annotation,
+%                             arrow right
+%
+% \rail@term{TEXT}[ANNOT] : format terminal TEXT with annotation
+% \rail@lterm{TEXT}[ANNOT] : format terminal TEXT with annotation, arrow left
+% \rail@rterm{TEXT}[ANNOT] : format terminal TEXT with annotation, arrow right
+%
+% \rail@cterm{TEXT}[ANNOT] : format terminal TEXT centered with annotation
+% \rail@lcterm{TEXT}[ANNOT] : format terminal TEXT centered with annotation, arrow left
+% \rail@rcterm{TEXT}[ANNOT] : format terminal TEXT centered with annotation,
+%                             arrow right
+%
+% \rail@annote[TEXT] : format TEXT as annotation
+
+\def\rail@token#1[#2]{
+\rail@setbox{%
+{\rail@termfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@oval
+}
+
+\def\rail@ltoken#1[#2]{
+\rail@setbox{%
+{\rail@termfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vloval
+}
+
+\def\rail@rtoken#1[#2]{
+\rail@setbox{%
+{\rail@termfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vroval
+}
+
+\def\rail@ctoken#1[#2]{
+\rail@setbox{%
+{\rail@termfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@coval
+}
+
+\def\rail@lctoken#1[#2]{
+\rail@setbox{%
+{\rail@termfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vlcoval
+}
+
+\def\rail@rctoken#1[#2]{
+\rail@setbox{%
+{\rail@termfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vrcoval
+}
+
+\def\rail@nont#1[#2]{
+\rail@setbox{%
+{\rail@nontfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@frame
+}
+
+\def\rail@lnont#1[#2]{
+\rail@setbox{%
+{\rail@nontfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vlframe
+}
+
+\def\rail@rnont#1[#2]{
+\rail@setbox{%
+{\rail@nontfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vrframe
+}
+
+\def\rail@cnont#1[#2]{
+\rail@setbox{%
+{\rail@nontfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@cframe
+}
+
+\def\rail@lcnont#1[#2]{
+\rail@setbox{%
+{\rail@nontfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vlcframe
+}
+
+\def\rail@rcnont#1[#2]{
+\rail@setbox{%
+{\rail@nontfont \rail@expand{#1}}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vrcframe
+}
+
+\def\rail@term#1[#2]{
+\rail@setbox{%
+{\rail@termfont #1}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@oval
+}
+
+\def\rail@lterm#1[#2]{
+\rail@setbox{%
+{\rail@termfont #1}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vloval
+}
+
+\def\rail@rterm#1[#2]{
+\rail@setbox{%
+{\rail@termfont #1}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vroval
+}
+
+\def\rail@cterm#1[#2]{
+\rail@setbox{%
+{\rail@termfont #1}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@coval
+}
+
+\def\rail@lcterm#1[#2]{
+\rail@setbox{%
+{\rail@termfont #1}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vlcoval
+}
+
+\def\rail@rcterm#1[#2]{
+\rail@setbox{%
+{\rail@termfont #1}\ifx\@empty#2\else\ {\rail@annofont #2}\fi
+}
+\rail@vrcoval
+}
+
+\def\rail@annote[#1]{
+\rail@setbox{\rail@annofont #1}
+\rail@text
+}
+
+% \rail@box : temporary box for \rail@oval and \rail@frame
+%
+% \rail@setbox{TEXT} : set \rail@box to TEXT, set \rail@tmpa to width
+%
+% \rail@oval : format \rail@box of width \rail@tmpa inside an oval
+% \rail@vloval : format \rail@box of width \rail@tmpa inside an oval, vector left
+% \rail@vroval : format \rail@box of width \rail@tmpa inside an oval, vector right
+%
+% \rail@coval : same as \rail@oval, but centered between \rail@x and
+%               \rail@mx
+% \rail@vlcoval : same as \rail@oval, but centered between \rail@x and
+%                 \rail@mx, vector left
+% \rail@vrcoval : same as \rail@oval, but centered between \rail@x and
+%                 \rail@mx, vector right
+%
+% \rail@frame : format \rail@box of width \rail@tmpa inside a frame
+% \rail@vlframe : format \rail@box of width \rail@tmpa inside a frame, vector left
+% \rail@vrframe : format \rail@box of width \rail@tmpa inside a frame, vector right
+%
+% \rail@cframe : same as \rail@frame, but centered between \rail@x and
+%                \rail@mx
+% \rail@vlcframe : same as \rail@frame, but centered between \rail@x and
+%                  \rail@mx, vector left
+% \rail@vrcframe : same as \rail@frame, but centered between \rail@x and
+%                  \rail@mx, vector right
+%
+% \rail@text : format \rail@box of width \rail@tmpa above the line
+
+\newbox\rail@box
+
+\def\rail@setbox#1{
+\setbox\rail@box\hbox{\strut#1}
+\rail@tmpa=\wd\rail@box
+\divide\rail@tmpa by \railunit
+}
+
+\def\rail@oval{
+\advance\rail@x by \rail@boxlf
+\rail@eline
+\advance\rail@tmpa by \rail@ovalsp
+\ifnum\rail@tmpa<\rail@boxht\rail@tmpa=\rail@boxht\fi
+\rail@tmpb=\rail@tmpa
+\divide\rail@tmpb by 2
+\advance\rail@y by -\rail@boxhht
+\rail@put{\makebox(\number\rail@tmpa,\number\rail@boxht){\box\rail@box}}
+\advance\rail@y by \rail@boxhht
+\advance\rail@x by \rail@tmpb
+\rail@put{\oval(\number\rail@tmpa,\number\rail@boxht)}
+\advance\rail@x by \rail@tmpb
+\rail@ex=\rail@x
+\advance\rail@x by \rail@boxrt
+}
+
+\def\rail@vloval{
+\advance\rail@x by \rail@boxlf
+\rail@eline
+\advance\rail@tmpa by \rail@ovalsp
+\ifnum\rail@tmpa<\rail@boxht\rail@tmpa=\rail@boxht\fi
+\rail@tmpb=\rail@tmpa
+\divide\rail@tmpb by 2
+\advance\rail@y by -\rail@boxhht
+\rail@put{\makebox(\number\rail@tmpa,\number\rail@boxht){\box\rail@box}}
+\advance\rail@y by \rail@boxhht
+\advance\rail@x by \rail@tmpb
+\rail@put{\oval(\number\rail@tmpa,\number\rail@boxht)}
+\advance\rail@x by \rail@tmpb
+\rail@ex=\rail@x
+\advance\rail@x by \rail@boxrt
+\rail@vleline
+}
+
+\def\rail@vroval{
+\advance\rail@x by \rail@boxlf
+\rail@vreline
+\advance\rail@tmpa by \rail@ovalsp
+\ifnum\rail@tmpa<\rail@boxht\rail@tmpa=\rail@boxht\fi
+\rail@tmpb=\rail@tmpa
+\divide\rail@tmpb by 2
+\advance\rail@y by -\rail@boxhht
+\rail@put{\makebox(\number\rail@tmpa,\number\rail@boxht){\box\rail@box}}
+\advance\rail@y by \rail@boxhht
+\advance\rail@x by \rail@tmpb
+\rail@put{\oval(\number\rail@tmpa,\number\rail@boxht)}
+\advance\rail@x by \rail@tmpb
+\rail@ex=\rail@x
+\advance\rail@x by \rail@boxrt
+}
+
+\def\rail@coval{
+\rail@tmpb=\rail@tmpa
+\advance\rail@tmpb by \rail@ovalsp
+\ifnum\rail@tmpb<\rail@boxht\rail@tmpb=\rail@boxht\fi
+\advance\rail@tmpb by \rail@boxlf
+\advance\rail@tmpb by \rail@boxrt
+\rail@tmpc=\rail@mx
+\advance\rail@tmpc by -\rail@x
+\advance\rail@tmpc by -\rail@tmpb
+\divide\rail@tmpc by 2
+\ifnum\rail@tmpc>0
+\advance\rail@x by \rail@tmpc
+\fi
+\rail@oval
+}
+
+\def\rail@vlcoval{
+\rail@tmpb=\rail@tmpa
+\advance\rail@tmpb by \rail@ovalsp
+\ifnum\rail@tmpb<\rail@boxht\rail@tmpb=\rail@boxht\fi
+\advance\rail@tmpb by \rail@boxlf
+\advance\rail@tmpb by \rail@boxrt
+\rail@tmpc=\rail@mx
+\advance\rail@tmpc by -\rail@x
+\advance\rail@tmpc by -\rail@tmpb
+\divide\rail@tmpc by 2
+\ifnum\rail@tmpc>0
+\advance\rail@x by \rail@tmpc
+\fi
+\rail@vloval
+}
+
+\def\rail@vrcoval{
+\rail@tmpb=\rail@tmpa
+\advance\rail@tmpb by \rail@ovalsp
+\ifnum\rail@tmpb<\rail@boxht\rail@tmpb=\rail@boxht\fi
+\advance\rail@tmpb by \rail@boxlf
+\advance\rail@tmpb by \rail@boxrt
+\rail@tmpc=\rail@mx
+\advance\rail@tmpc by -\rail@x
+\advance\rail@tmpc by -\rail@tmpb
+\divide\rail@tmpc by 2
+\ifnum\rail@tmpc>0
+\advance\rail@x by \rail@tmpc
+\fi
+\rail@vroval
+}
+
+\def\rail@frame{
+\advance\rail@x by \rail@boxlf
+\rail@eline
+\advance\rail@tmpa by \rail@framesp
+\ifnum\rail@tmpa<\rail@boxht\rail@tmpa=\rail@boxht\fi
+\advance\rail@y by -\rail@boxhht
+\rail@put{\framebox(\number\rail@tmpa,\number\rail@boxht){\box\rail@box}}
+\advance\rail@y by \rail@boxhht
+\advance\rail@x by \rail@tmpa
+\rail@ex=\rail@x
+\advance\rail@x by \rail@boxrt
+}
+
+\def\rail@vlframe{
+\advance\rail@x by \rail@boxlf
+\rail@eline
+\advance\rail@tmpa by \rail@framesp
+\ifnum\rail@tmpa<\rail@boxht\rail@tmpa=\rail@boxht\fi
+\advance\rail@y by -\rail@boxhht
+\rail@put{\framebox(\number\rail@tmpa,\number\rail@boxht){\box\rail@box}}
+\advance\rail@y by \rail@boxhht
+\advance\rail@x by \rail@tmpa
+\rail@ex=\rail@x
+\advance\rail@x by \rail@boxrt
+\rail@vleline
+}
+
+\def\rail@vrframe{
+\advance\rail@x by \rail@boxlf
+\rail@vreline
+\advance\rail@tmpa by \rail@framesp
+\ifnum\rail@tmpa<\rail@boxht\rail@tmpa=\rail@boxht\fi
+\advance\rail@y by -\rail@boxhht
+\rail@put{\framebox(\number\rail@tmpa,\number\rail@boxht){\box\rail@box}}
+\advance\rail@y by \rail@boxhht
+\advance\rail@x by \rail@tmpa
+\rail@ex=\rail@x
+\advance\rail@x by \rail@boxrt
+}
+
+\def\rail@cframe{
+\rail@tmpb=\rail@tmpa
+\advance\rail@tmpb by \rail@framesp
+\ifnum\rail@tmpb<\rail@boxht\rail@tmpb=\rail@boxht\fi
+\advance\rail@tmpb by \rail@boxlf
+\advance\rail@tmpb by \rail@boxrt
+\rail@tmpc=\rail@mx
+\advance\rail@tmpc by -\rail@x
+\advance\rail@tmpc by -\rail@tmpb
+\divide\rail@tmpc by 2
+\ifnum\rail@tmpc>0
+\advance\rail@x by \rail@tmpc
+\fi
+\rail@frame
+}
+
+\def\rail@vlcframe{
+\rail@tmpb=\rail@tmpa
+\advance\rail@tmpb by \rail@framesp
+\ifnum\rail@tmpb<\rail@boxht\rail@tmpb=\rail@boxht\fi
+\advance\rail@tmpb by \rail@boxlf
+\advance\rail@tmpb by \rail@boxrt
+\rail@tmpc=\rail@mx
+\advance\rail@tmpc by -\rail@x
+\advance\rail@tmpc by -\rail@tmpb
+\divide\rail@tmpc by 2
+\ifnum\rail@tmpc>0
+\advance\rail@x by \rail@tmpc
+\fi
+\rail@vlframe
+}
+
+\def\rail@vrcframe{
+\rail@tmpb=\rail@tmpa
+\advance\rail@tmpb by \rail@framesp
+\ifnum\rail@tmpb<\rail@boxht\rail@tmpb=\rail@boxht\fi
+\advance\rail@tmpb by \rail@boxlf
+\advance\rail@tmpb by \rail@boxrt
+\rail@tmpc=\rail@mx
+\advance\rail@tmpc by -\rail@x
+\advance\rail@tmpc by -\rail@tmpb
+\divide\rail@tmpc by 2
+\ifnum\rail@tmpc>0
+\advance\rail@x by \rail@tmpc
+\fi
+\rail@vrframe
+}
+
+\def\rail@text{
+\advance\rail@x by \rail@textlf
+\advance\rail@y by \rail@textup
+\rail@put{\box\rail@box}
+\advance\rail@y by -\rail@textup
+\advance\rail@x by \rail@tmpa
+\advance\rail@x by \rail@textrt
+}
+
+% alternatives
+%
+% \rail@jx \rail@jy : current join point
+%
+% \rail@gx \rail@gy \rail@gex \rail@grx : global versions of \rail@x etc,
+%                                         to pass values over group closings
+%
+% \rail@mx : maximum x so far
+%
+% \rail@sy : starting \rail@y for alternatives
+%
+% \rail@jput : put at (\rail@jx,\rail@jy)
+%
+% \rail@joval[PART] : put \oval[PART] with adjust
+
+\newcount\rail@jx
+\newcount\rail@jy
+
+\newcount\rail@gx
+\newcount\rail@gy
+\newcount\rail@gex
+\newcount\rail@grx
+
+\newcount\rail@sy
+\newcount\rail@mx
+
+\def\rail@jput{
+\put(\number\rail@jx,\number\rail@jy)
+}
+
+\def\rail@joval[#1]{
+\advance\rail@jx by \rail@joinadj
+\rail@jput{\oval(\number\rail@joinsz,\number\rail@joinsz)[#1]}
+\advance\rail@jx by -\rail@joinadj
+}
+
+% \rail@barsplit : incoming split for '|'
+%
+% \rail@plussplit : incoming split for '+'
+%
+
+\def\rail@barsplit{
+\advance\rail@jy by -\rail@joinhsz
+\rail@joval[tr]
+\advance\rail@jx by \rail@joinhsz
+}
+
+\def\rail@plussplit{
+\advance\rail@jy by -\rail@joinhsz
+\advance\rail@jx by \rail@joinsz
+\rail@joval[tl]
+\advance\rail@jx by -\rail@joinhsz
+}
+
+% \rail@alt{SPLIT} : start alternatives with incoming split SPLIT
+%
+
+\def\rail@alt#1{
+\rail@sy=\rail@y
+\rail@jx=\rail@x
+\rail@jy=\rail@y
+\advance\rail@x by \rail@joinsz
+\rail@mx=0
+\let\rail@list=\@empty
+\let\rail@comma=\@empty
+\let\rail@split=#1
+\begingroup
+\rail@sx=\rail@x
+\rail@rx=0
+}
+
+% \rail@nextalt{FIX}{Y} : start next alternative at vertical position Y
+%                         and fix-up FIX
+%
+
+\def\rail@nextalt#1#2{
+\global\rail@gx=\rail@x
+\global\rail@gy=\rail@y
+\global\rail@gex=\rail@ex
+\global\rail@grx=\rail@rx
+\endgroup
+#1
+\ifnum\rail@gx>\rail@mx\rail@mx=\rail@gx\fi
+\ifnum\rail@grx>\rail@mx\rail@mx=\rail@grx\fi
+\edef\rail@list{\rail@list\rail@comma\number\rail@gex:\number\rail@gy}
+\def\rail@comma{,}
+\rail@split
+\let\rail@split=\@empty
+\rail@sety{#2}
+\rail@tmpa=\rail@jy
+\advance\rail@tmpa by -\rail@y
+\advance\rail@tmpa by -\rail@joinhsz
+\rail@jput{\line(0,-1){\number\rail@tmpa}}
+\rail@jy=\rail@y
+\advance\rail@jy by \rail@joinhsz
+\advance\rail@jx by \rail@joinhsz
+\rail@joval[bl]
+\advance\rail@jx by -\rail@joinhsz
+\rail@ex=\rail@x
+\begingroup
+\rail@sx=\rail@x
+\rail@rx=0
+}
+
+% \rail@barjoin : outgoing join for first '|' alternative
+%
+% \rail@plusjoin : outgoing join for first '+' alternative
+%
+% \rail@altjoin : join for subsequent alternative
+%
+
+\def\rail@barjoin{
+\ifnum\rail@y<\rail@sy
+\global\rail@gex=\rail@jx
+\else
+\global\rail@gex=\rail@ex
+\fi
+\advance\rail@jy by -\rail@joinhsz
+\rail@joval[tl]
+\advance\rail@jx by -\rail@joinhsz
+\ifnum\rail@y<\rail@sy
+\rail@altjoin
+\fi
+}
+
+\def\rail@plusjoin{
+\global\rail@gex=\rail@ex
+\advance\rail@jy by -\rail@joinhsz
+\advance\rail@jx by -\rail@joinsz
+\rail@joval[tr]
+\advance\rail@jx by \rail@joinhsz
+}
+
+\def\rail@altjoin{
+\rail@eline
+\rail@tmpa=\rail@jy
+\advance\rail@tmpa by -\rail@y
+\advance\rail@tmpa by -\rail@joinhsz
+\rail@jput{\line(0,-1){\number\rail@tmpa}}
+\rail@jy=\rail@y
+\advance\rail@jy by \rail@joinhsz
+\advance\rail@jx by -\rail@joinhsz
+\rail@joval[br]
+\advance\rail@jx by \rail@joinhsz
+}
+
+% \rail@eltsplit EX:Y; : split EX:Y into \rail@ex \rail@y
+%
+% \rail@endalt{JOIN} : end alternatives with outgoing join JOIN
+
+\def\rail@eltsplit#1:#2;{\rail@ex=#1\rail@y=#2}
+
+\def\rail@endalt#1{
+\global\rail@gx=\rail@x
+\global\rail@gy=\rail@y
+\global\rail@gex=\rail@ex
+\global\rail@grx=\rail@rx
+\endgroup
+\ifnum\rail@gx>\rail@mx\rail@mx=\rail@gx\fi
+\ifnum\rail@grx>\rail@mx\rail@mx=\rail@grx\fi
+\edef\rail@list{\rail@list\rail@comma\number\rail@gex:\number\rail@gy}
+\rail@x=\rail@mx
+\rail@jx=\rail@x
+\rail@jy=\rail@sy
+\advance\rail@jx by \rail@joinsz
+\let\rail@join=#1
+\@for\rail@elt:=\rail@list\do{
+\expandafter\rail@eltsplit\rail@elt;
+\rail@join
+\let\rail@join=\rail@altjoin
+}
+\rail@x=\rail@mx
+\rail@y=\rail@sy
+\rail@ex=\rail@gex
+\advance\rail@x by \rail@joinsz
+}
+
+% \rail@bar : start '|' alternatives
+%
+% \rail@nextbar : next '|' alternative
+%
+% \rail@endbar : end '|' alternatives
+%
+
+\def\rail@bar{
+\rail@alt\rail@barsplit
+}
+
+\def\rail@nextbar{
+\rail@nextalt\relax
+}
+
+\def\rail@endbar{
+\rail@endalt\rail@barjoin
+}
+
+% \rail@plus : start '+' alternatives
+%
+% \rail@nextplus: next '+' alternative
+%
+% \rail@endplus : end '+' alternatives
+%
+
+\def\rail@plus{
+\rail@alt\rail@plussplit
+}
+
+\def\rail@nextplus{
+\rail@nextalt\rail@fixplus
+}
+
+\def\rail@fixplus{
+\ifnum\rail@gy<\rail@sy
+\begingroup
+\rail@x=\rail@gx
+\rail@y=\rail@gy
+\rail@ex=\rail@gex
+\rail@rx=\rail@grx
+\ifnum\rail@x<\rail@rx
+\rail@x=\rail@rx
+\fi
+\rail@eline
+\rail@jx=\rail@x
+\rail@jy=\rail@y
+\advance\rail@jy by \rail@joinhsz
+\rail@joval[br]
+\advance\rail@jx by \rail@joinhsz
+\rail@tmpa=\rail@sy
+\advance\rail@tmpa by -\rail@joinhsz
+\advance\rail@tmpa by -\rail@jy
+\rail@jput{\line(0,1){\number\rail@tmpa}}
+\rail@jy=\rail@sy
+\advance\rail@jy by -\rail@joinhsz
+\advance\rail@jx by \rail@joinhsz
+\rail@joval[tl]
+\advance\rail@jy by \rail@joinhsz
+\global\rail@gx=\rail@jx
+\global\rail@gy=\rail@jy
+\global\rail@gex=\rail@gx
+\global\rail@grx=\rail@rx
+\endgroup
+\fi
+}
+
+\def\rail@endplus{
+\rail@endalt\rail@plusjoin
+}
+
+% \rail@cr{Y} : carriage return to vertical position Y
+
+\def\rail@cr#1{
+\rail@tmpa=\rail@sx
+\advance\rail@tmpa by \rail@joinsz
+\ifnum\rail@x<\rail@tmpa\rail@x=\rail@tmpa\fi
+\rail@eline
+\rail@jx=\rail@x
+\rail@jy=\rail@y
+\advance\rail@x by \rail@joinsz
+\ifnum\rail@x>\rail@rx\rail@rx=\rail@x\fi
+\advance\rail@jy by -\rail@joinhsz
+\rail@joval[tr]
+\advance\rail@jx by \rail@joinhsz
+\rail@sety{#1}
+\rail@tmpa=\rail@jy
+\advance\rail@tmpa by -\rail@y
+\advance\rail@tmpa by -\rail@boxsp
+\advance\rail@tmpa by -\rail@joinhsz
+\rail@jput{\line(0,-1){\number\rail@tmpa}}
+\rail@jy=\rail@y
+\advance\rail@jy by \rail@boxsp
+\advance\rail@jy by \rail@joinhsz
+\advance\rail@jx by -\rail@joinhsz
+\rail@joval[br]
+\advance\rail@jy by -\rail@joinhsz
+\rail@tmpa=\rail@jx
+\advance\rail@tmpa by -\rail@sx
+\advance\rail@tmpa by -\rail@joinhsz
+\rail@jput{\line(-1,0){\number\rail@tmpa}}
+\rail@jx=\rail@sx
+\advance\rail@jx by \rail@joinhsz
+\advance\rail@jy by -\rail@joinhsz
+\rail@joval[tl]
+\advance\rail@jx by -\rail@joinhsz
+\rail@tmpa=\rail@boxsp
+\advance\rail@tmpa by -\rail@joinsz
+\rail@jput{\line(0,-1){\number\rail@tmpa}}
+\advance\rail@jy by -\rail@tmpa
+\advance\rail@jx by \rail@joinhsz
+\rail@joval[bl]
+\rail@x=\rail@jx
+\rail@ex=\rail@x
+}
+
+% default setup for Isabelle
+\newenvironment{railoutput}%
+{\begin{list}{}{\rail@param}\def\rail@expand{\relax}\makeatletter}{\makeatother\end{list}}
+
+\def\rail@termfont{\isabellestyle{tt}}
+\def\rail@nontfont{\isabellestyle{it}}
+\def\rail@namefont{\isabellestyle{it}}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys2/Journal/root.aux	Mon Nov 01 10:40:21 2021 +0000
@@ -0,0 +1,96 @@
+\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}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys2/Journal/root.bib	Mon Nov 01 10:40:21 2021 +0000
@@ -0,0 +1,340 @@
+@article{HosoyaVouillonPierce2005,
+  author = {H.~Hosoya and J.~Vouillon and B.~C.~Pierce},
+  title = {{R}egular {E}xpression {T}ypes for {XML}},
+  journal = {ACM Transactions on Programming Languages and Systems (TOPLAS)},
+  year = {2005},
+  volume = 27,
+  number = 1,
+  pages = {46--90}
+}
+
+
+@Misc{POSIX,
+  title =     {{T}he {O}pen {G}roup {B}ase {S}pecification {I}ssue 6 {IEEE} {S}td 1003.1 2004 {E}dition},
+  year =      {2004},
+  note =    {\url{http://pubs.opengroup.org/onlinepubs/009695399/basedefs/xbd_chap09.html}}
+}
+
+
+@InProceedings{AusafDyckhoffUrban2016,
+  author =   {F.~Ausaf and R.~Dyckhoff and C.~Urban},
+  title = 	 {{POSIX} {L}exing with {D}erivatives of {R}egular {E}xpressions ({P}roof {P}earl)},
+  year = 	 {2016},
+  booktitle = {Proc.~of the 7th International Conference on Interactive Theorem Proving (ITP)},
+  volume = 	 {9807},
+  series = 	 {LNCS},
+  pages =        {69--86}
+}
+
+@article{aduAFP16,
+  author =   {F.~Ausaf and R.~Dyckhoff and C.~Urban},
+  title =    {{POSIX} {L}exing with {D}erivatives of {R}egular {E}xpressions},
+  journal =  {Archive of Formal Proofs},
+  year =     2016,
+  note =     {\url{http://www.isa-afp.org/entries/Posix-Lexing.shtml}, Formal proof development},
+  ISSN =     {2150-914x}
+}
+
+
+@TechReport{CrashCourse2014,
+  author =       {N.~B.~B.~Grathwohl and F.~Henglein and U.~T.~Rasmussen},
+  title =        {{A} {C}rash-{C}ourse in {R}egular {E}xpression {P}arsing and {R}egular 
+                  {E}xpressions as {T}ypes},
+  institution =  {University of Copenhagen},
+  year =         {2014},
+  annote =       {draft report}
+}
+
+@inproceedings{Sulzmann2014,
+  author    = {M.~Sulzmann and K.~Lu},
+  title     = {{POSIX} {R}egular {E}xpression {P}arsing with {D}erivatives},
+  booktitle = {Proc.~of the 12th International Conference on Functional and Logic Programming (FLOPS)},
+  pages     = {203--220},
+  year      = {2014},
+  volume =    {8475},
+  series =    {LNCS}
+}
+
+@inproceedings{Sulzmann2014b,
+  author    = {M.~Sulzmann and P.~van Steenhoven},
+  title     = {{A} {F}lexible and {E}fficient {ML} {L}exer {T}ool {B}ased on {E}xtended {R}egular
+               {E}xpression {S}ubmatching},
+  booktitle = {Proc.~of the 23rd International Conference on Compiler Construction (CC)},
+  pages     = {174--191},
+  year      = {2014},
+  volume    = {8409},
+  series =    {LNCS}
+}
+
+@book{Pierce2015,
+  author = {B.~C.~Pierce and C.~Casinghino and M.~Gaboardi and
+            M.~Greenberg and C.~Hri\c{t}cu and 
+            V.~Sj\"{o}berg and B.~Yorgey},
+  title = {{S}oftware {F}oundations},
+  year = {2015},
+  publisher = {Electronic textbook},
+  note = {\url{http://www.cis.upenn.edu/~bcpierce/sf}}
+}
+
+@Misc{Kuklewicz,
+  author = 	 {C.~Kuklewicz},
+  title = 	 {{R}egex {P}osix},
+  howpublished = "\url{https://wiki.haskell.org/Regex_Posix}"
+}
+
+@article{Vansummeren2006,
+  author = {S.~Vansummeren},
+  title = {{T}ype {I}nference for {U}nique {P}attern {M}atching},
+  year = {2006},
+  journal = {ACM Transactions on Programming Languages and Systems},
+  volume = {28},
+  number = {3},
+  pages = {389--428}
+}
+
+@InProceedings{Asperti12,
+  author =       {A.~Asperti},
+  title =        {{A} {C}ompact {P}roof of {D}ecidability for {R}egular {E}xpression {E}quivalence},
+  booktitle = {Proc.~of the 3rd International Conference on Interactive Theorem Proving (ITP)},
+  pages =     {283--298},
+  year =      {2012},
+  volume =    {7406},
+  series =    {LNCS}
+}
+
+@inproceedings{Frisch2004,
+  author    = {A.~Frisch and L.~Cardelli},
+  title     = {{G}reedy {R}egular {E}xpression {M}atching},
+  booktitle = {Proc.~of the 31st International Conference on Automata, Languages and Programming (ICALP)},
+  pages     = {618--629},
+  year      = {2004},
+  volume =    {3142},
+  series =    {LNCS}
+}
+
+@ARTICLE{Antimirov95,
+    author = {V.~Antimirov},
+    title = {{P}artial {D}erivatives of {R}egular {E}xpressions and 
+     {F}inite {A}utomata {C}onstructions},
+    journal = {Theoretical Computer Science},
+    year = {1995},
+    volume = {155},
+    pages = {291--319}
+}
+
+@inproceedings{Nipkow98,
+ author={T.~Nipkow},
+ title={{V}erified {L}exical {A}nalysis},
+ booktitle={Proc.~of the 11th International Conference on Theorem Proving in Higher Order Logics (TPHOLs)},
+ series={LNCS},
+ volume=1479,
+ pages={1--15},
+ year=1998
+}
+
+@article{Brzozowski1964,
+  author    = {J.~A.~Brzozowski},
+  title     = {{D}erivatives of {R}egular {E}xpressions},
+  journal   = {Journal of the {ACM}},
+  volume    = {11},
+  number    = {4},
+  pages     = {481--494},
+  year      = {1964}
+}
+
+@article{Leroy2009,
+  author = {X.~Leroy},
+  title = {{F}ormal {V}erification of a {R}ealistic {C}ompiler},
+  journal = {Communications of the ACM},
+  year = 2009,
+  volume = 52,
+  number = 7,
+  pages = {107--115}
+}
+
+@InProceedings{Paulson2015,
+  author = 	 {L.~C.~Paulson},
+  title = 	 {{A} {F}ormalisation of {F}inite {A}utomata {U}sing {H}ereditarily {F}inite {S}ets},
+  booktitle = {Proc.~of the 25th International Conference on Automated Deduction (CADE)},
+  pages = 	 {231--245},
+  year = 	 {2015},
+  volume = 	 {9195},
+  series = 	 {LNAI}
+}
+
+@Article{Wu2014,
+  author = 	 {C.~Wu and X.~Zhang and C.~Urban},
+  title = 	 {{A} {F}ormalisation of the {M}yhill-{N}erode {T}heorem based on {R}egular {E}xpressions},
+  journal = 	 {Journal of Automatic Reasoning},
+  year = 	 {2014},
+  volume = 	 {52},
+  number = 	 {4},
+  pages = 	 {451--480}
+}
+
+@InProceedings{Regehr2011,
+  author = 	 {X.~Yang and Y.~Chen and E.~Eide and J.~Regehr},
+  title = 	 {{F}inding and {U}nderstanding {B}ugs in {C} {C}ompilers},
+  pages = 	 {283--294},
+  year = 	 {2011},
+  booktitle =    {Proc.~of the 32nd ACM SIGPLAN Conference on Programming Language Design and 
+                  Implementation (PLDI)}
+}
+
+@Article{Norrish2014,
+  author = 	 {A.~Barthwal and M.~Norrish},
+  title = 	 {{A} {M}echanisation of {S}ome {C}ontext-{F}ree {L}anguage {T}heory in {HOL4}},
+  journal          = {Journal of Computer and System Sciences},
+  year = 	 {2014},
+  volume = 	 {80},
+  number = 	 {2},
+  pages = 	 {346--362}
+}
+
+@Article{Thompson1968,
+ author = {K.~Thompson},
+ title = {{P}rogramming {T}echniques: {R}egular {E}xpression {S}earch {A}lgorithm},
+ journal = {Communications of the ACM},
+ issue_date = {June 1968},
+ volume = {11},
+ number = {6},
+ year = {1968},
+ pages = {419--422}
+}
+
+@article{Owens2009,
+  author    = {S.~Owens and J.~H.~Reppy and A.~Turon},
+  title     = {{R}egular-{E}xpression {D}erivatives {R}e-{E}xamined},
+  journal   = {Journal of Functinal Programming},
+  volume    = {19},
+  number    = {2},
+  pages     = {173--190},
+  year      = {2009}
+}
+
+@inproceedings{Sulzmann2015,
+  author    = {M.~Sulzmann and P.~Thiemann},
+  title     = {{D}erivatives for {R}egular {S}huffle {E}xpressions},
+  booktitle = {Proc.~of the 9th International Conference on Language and Automata Theory 
+               and Applications (LATA)},
+  pages     = {275--286},
+  year      = {2015},
+  volume = 	 {8977},
+  series = 	 {LNCS}
+}
+
+@inproceedings{Chen2012,
+  author    = {H.~Chen and S.~Yu},
+  title     = {{D}erivatives of {R}egular {E}xpressions and an {A}pplication},
+  booktitle = {Proc.~in the International Workshop on Theoretical
+               Computer Science (WTCS)},
+  pages     = {343--356},
+  year      = {2012},
+  volume = 	 {7160},
+  series = 	 {LNCS}
+}
+
+@article{Krauss2011,
+  author={A.~Krauss and T.~Nipkow},
+  title={{P}roof {P}earl: {R}egular {E}xpression {E}quivalence and {R}elation {A}lgebra},
+  journal={Journal of Automated Reasoning},
+  volume=49,
+  pages={95--106},
+  year=2012
+}
+
+@InProceedings{Traytel2015,
+  author =	{D.~Traytel},
+  title =	{{A} {C}oalgebraic {D}ecision {P}rocedure for {WS1S}},
+  booktitle =	{Proc.~of the 24th Annual Conference on Computer Science Logic (CSL)},
+  pages =	{487--503},
+  series =	{LIPIcs},
+  year =	{2015},
+  volume =	{41}
+}
+
+@inproceedings{Traytel2013,
+  author={D.~Traytel and T.~Nipkow},
+  title={{A} {V}erified {D}ecision {P}rocedure for {MSO} on 
+         {W}ords {B}ased on {D}erivatives of {R}egular {E}xpressions},
+  booktitle={Proc.~of the 18th ACM SIGPLAN International Conference on Functional Programming (ICFP)},
+  pages={3-12},
+  year=2013
+}
+
+@InProceedings{Coquand2012,
+  author =       {T.~Coquand and V.~Siles},
+  title =        {{A} {D}ecision {P}rocedure for {R}egular {E}xpression {E}quivalence in {T}ype {T}heory},
+  booktitle = {Proc.~of the 1st International Conference on Certified Programs and Proofs (CPP)},
+  pages =     {119--134},
+  year =      {2011},
+  volume =    {7086},
+  series =    {LNCS}
+}
+
+@InProceedings{Almeidaetal10,
+  author =       {J.~B.~Almeida and N.~Moriera and D.~Pereira and S.~M.~de Sousa},
+  title =        {{P}artial {D}erivative {A}utomata {F}ormalized in {C}oq},
+  booktitle =    {Proc.~of the 15th International Conference on Implementation
+                  and Application of Automata (CIAA)},
+  pages =        {59-68},
+  year =         {2010},
+  volume =       {6482},
+  series =       {LNCS}
+}
+
+@article{Owens2008,
+  author    = {S.~Owens and K.~Slind},
+  title     = {{A}dapting {F}unctional {P}rograms to {H}igher {O}rder {L}ogic},
+  journal   = {Higher-Order and Symbolic Computation},
+  volume    = {21},
+  number    = {4},
+  year      = {2008},
+  pages     = {377--409}
+}
+
+
+@article{Owens2,
+  author    = {S.~Owens and K.~Slind},
+  title     = {Adapting functional programs to higher order logic},
+  journal   = {Higher-Order and Symbolic Computation},
+  volume    = {21},
+  number    = {4},
+  pages     = {377--409},
+  year      = {2008},
+  url       = {http://dx.doi.org/10.1007/s10990-008-9038-0},
+  doi       = {10.1007/s10990-008-9038-0},
+  timestamp = {Wed, 16 Dec 2009 13:51:02 +0100},
+  biburl    = {http://dblp.uni-trier.de/rec/bib/journals/lisp/OwensS08},
+  bibsource = {dblp computer science bibliography, http://dblp.org}
+}
+
+@misc{PCRE,
+  title = "{PCRE - Perl Compatible Regular Expressions}",
+  url = {http://www.pcre.org},
+}
+
+
+
+@InProceedings{OkuiSuzuki2010,
+  author =       {S.~Okui and T.~Suzuki},
+  title =        {{D}isambiguation in {R}egular {E}xpression {M}atching via
+                  {P}osition {A}utomata with {A}ugmented {T}ransitions},
+  booktitle = {Proc.~of the 15th International Conference on Implementation
+                  and Application of Automata (CIAA)},
+  year =      {2010},
+  volume =    {6482},
+  series =    {LNCS},
+  pages =     {231--240}
+}
+
+
+
+@TechReport{OkuiSuzukiTech,
+  author =       {S.~Okui and T.~Suzuki},
+  title =        {{D}isambiguation in {R}egular {E}xpression {M}atching via
+                  {P}osition {A}utomata with {A}ugmented {T}ransitions},
+  institution =  {University of Aizu},
+  year =         {2013}
+}
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys2/Journal/root.log	Mon Nov 01 10:40:21 2021 +0000
@@ -0,0 +1,1995 @@
+This is LuaTeX, Version 1.07.0 (TeX Live 2018)  (format=lualatex 2019.2.7)  1 NOV 2021 09:38
+ 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.091 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 82.
+
+
+LaTeX Warning: Citation `Owens2008' on page 41 undefined on input line 96.
+
+
+LaTeX Warning: Citation `Krauss2011' on page 41 undefined on input line 97.
+
+
+LaTeX Warning: Citation `Coquand2012' on page 41 undefined on input line 98.
+
+[41]
+
+LaTeX Warning: Citation `Frisch2004' on page 42 undefined on input line 103.
+
+
+LaTeX Warning: Citation `POSIX' on page 42 undefined on input line 104.
+
+
+LaTeX Warning: Citation `Kuklewicz' on page 42 undefined on input line 104.
+
+
+LaTeX Warning: Citation `OkuiSuzuki2010' on page 42 undefined on input line 104.
+
+
+
+LaTeX Warning: Citation `Sulzmann2014' on page 42 undefined on input line 104.
+
+
+LaTeX Warning: Citation `Vansummeren2006' on page 42 undefined on input line 104
+.
+
+
+LaTeX Warning: Citation `POSIX' on page 42 undefined on input line 121.
+
+[42]
+
+LaTeX Warning: Citation `Sulzmann2014' on page 43 undefined on input line 160.
+
+
+LaTeX Warning: Citation `HosoyaVouillonPierce2005' on page 43 undefined on input
+ line 182.
+
+
+LaTeX Warning: Citation `Sulzmann2014' on page 43 undefined on input line 204.
+
+
+LaTeX Warning: Citation `Frisch2004' on page 43 undefined on input line 208.
+
+
+LaTeX Warning: Citation `Sulzmann2014' on page 43 undefined on input line 213.
+
+
+LaTeX Warning: Citation `Kuklewicz' on page 43 undefined on input line 221.
+
+
+LaTeX Warning: Citation `Sulzmann2014' on page 43 undefined on input line 222.
+
+
+LaTeX Warning: Citation `CrashCourse2014' on page 43 undefined on input line 223
+.
+
+[43]
+
+LaTeX Warning: Citation `Sulzmann2014' on page 44 undefined on input line 235.
+
+
+LaTeX Warning: Citation `Vansummeren2006' on page 44 undefined on input line 237
+.
+
+
+LaTeX Warning: Citation `Sulzmann2014' on page 44 undefined on input line 238.
+
+(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 242.
+
+
+LaTeX Warning: Citation `OkuiSuzuki2010' on page 44 undefined on input line 246.
+
+
+(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 325.
+
+
+LaTeX Warning: Citation `Krauss2011' on page 45 undefined on input line 330.
+
+(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 345.
+
+
+
+Overfull \hbox (1.4002pt too wide) in paragraph at lines 352--373
+ [] 
+ []
+
+[45]
+
+LaTeX Warning: Citation `Sulzmann2014' on page 46 undefined on input line 408.
+
+
+LaTeX Warning: Citation `Sulzmann2014' on page 46 undefined on input line 432.
+
+
+LaTeX Warning: Citation `Frisch2004' on page 46 undefined on input line 449.
+
+
+LaTeX Warning: Citation `Sulzmann2014' on page 46 undefined on input line 450.
+
+
+LaTeX Warning: Citation `AusafDyckhoffUrban2016' on page 46 undefined on input l
+ine 482.
+
+[46]
+
+LaTeX Warning: Citation `OkuiSuzuki2010' on page 47 undefined on input line 519.
+
+
+
+LaTeX Warning: Citation `Frisch2004' on page 47 undefined on input line 519.
+
+
+Underfull \hbox (badness 10000) in paragraph at lines 579--580
+
+ []
+
+
+LaTeX Warning: Citation `Sulzmann2014' on page 47 undefined on input line 590.
+
+
+LaTeX Warning: Citation `Sulzmann2014' on page 47 undefined on input line 590.
+
+[47]
+
+LaTeX Warning: Citation `Sulzmann2014' on page 48 undefined on input line 610.
+
+
+Underfull \vbox (badness 10000) has occurred while \output is active []
+
+ [48]
+[49]
+
+LaTeX Warning: Citation `Sulzmann2014' on page 50 undefined on input line 718.
+
+
+LaTeX Warning: Citation `Vansummeren2006' on page 50 undefined on input line 725
+.
+
+(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 938.
+
+
+LaTeX Warning: Citation `Sulzmann2014' on page 53 undefined on input line 940.
+
+
+LaTeX Warning: Citation `OkuiSuzuki2010' on page 53 undefined on input line 947.
+
+
+
+LaTeX Warning: Citation `OkuiSuzukiTech' on page 53 undefined on input line 947.
+
+
+[53] [54]
+
+LaTeX Warning: Citation `OkuiSuzuki2010' on page 55 undefined on input line 1069
+.
+
+
+LaTeX Warning: Citation `OkuiSuzuki2010' on page 55 undefined on input line 1090
+.
+
+
+LaTeX Warning: Citation `OkuiSuzuki2010' on page 55 undefined on input line 1130
+.
+
+[55]
+Overfull \hbox (47.18065pt too wide) in paragraph at lines 1154--1184
+[] [] 
+ []
+
+[56] [57]
+
+LaTeX Warning: Citation `Sulzmann2014' on page 58 undefined on input line 1351.
+
+(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 1370.
+
+
+Underfull \vbox (badness 10000) has occurred while \output is active []
+
+ [58]
+Overfull \hbox (4.61574pt too wide) in paragraph at lines 1434--1435
+\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 1434--1435
+\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 1434--1435
+\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 1434--1435
+\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 1434--1435
+\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 1434--1435
+\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 1434--1435
+\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 1434--1435
+\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 1438--1447
+\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 1438--1447
+\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 1438--1447
+\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 Warning: File `root.out' has changed.
+(rerunfilecheck)                Rerun to get outlines right
+(rerunfilecheck)                or use package `bookmark'.
+
+Package rerunfilecheck Info: Checksums for `root.out':
+(rerunfilecheck)             Before: 40AB4FCF7DFB133886AD9117C6B3D022;396
+(rerunfilecheck)             After:  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.
+
+
+LaTeX Warning: Label(s) may have changed. Rerun to get cross-references right.
+
+Package atveryend Info: Empty hook `AtVeryVeryEnd' on input line 96.
+)
+
+Here is how much of LuaTeX's memory you used:
+ 29104 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:290,4:52,5:2358,6:47,7:12582,8:52,9:593,10:23,11:398
+ 31952 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
+
+warning  (pdf backend): unreferenced destination with name 'equation.1.6.2'
+
+warning  (pdf backend): unreferenced destination with name 'theorem.1.4.3'
+
+warning  (pdf backend): unreferenced destination with name 'theorem.1.3.2'
+
+warning  (pdf backend): unreferenced destination with name 'theorem.1.3.1'
+
+warning  (pdf backend): unreferenced destination with name 'equation.1.2.1'
+</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, 375818 bytes).
+
+PDF statistics: 526 PDF objects out of 1000 (max. 8388607)
+ 411 compressed objects within 5 object streams
+ 112 named destinations out of 1000 (max. 131072)
+ 72 words of extra memory for PDF output out of 10000 (max. 100000000)
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys2/Journal/root.out	Mon Nov 01 10:40:21 2021 +0000
@@ -0,0 +1,9 @@
+\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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys2/Journal/session.tex	Mon Nov 01 10:40:21 2021 +0000
@@ -0,0 +1,8 @@
+\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