ChengsongTanPhdThesis/Chapters/Inj.tex
author Chengsong
Sun, 09 Jul 2023 00:29:02 +0100
changeset 657 00171b627b8d
parent 651 deb35fd780fe
child 666 6da4516ea87d
permissions -rwxr-xr-x
Fixed some annotated/unannotated a/r notation inconsistencies.

% Chapter Template

\chapter{Regular Expressions and POSIX Lexing} % Main chapter title

\label{Inj} % In chapter 2 \ref{Chapter2} we will introduce the concepts
%and notations we 
% used for describing the lexing algorithm by Sulzmann and Lu,
%and then give the algorithm and its variant and discuss
%why more aggressive simplifications are needed. 

In this chapter, we define the basic notions 
for regular languages and regular expressions.
This is essentially a description in ``English''
the functions and datatypes of our formalisation in Isabelle/HOL.
We also define what $\POSIX$ lexing means, 
followed by the first lexing algorithm by Sulzmanna and Lu \parencite{Sulzmann2014} 
that produces the output conforming
to the $\POSIX$ standard\footnote{In what follows 
we choose to use the Isabelle-style notation
for function applications, where
the parameters of a function are not enclosed
inside a pair of parentheses (e.g. $f \;x \;y$
instead of $f(x,\;y)$). This is mainly
to make the text visually more concise.}.


\section{Technical Overview}

Consider for example the regular expression $(a^*)^*\,b$ and 
strings of the form $aa..a$. These strings cannot be matched by this regular
expression: Obviously the expected $b$ in the last
position is missing. One would assume that modern regular expression
matching engines can find this out very quickly. Surprisingly, if one tries
this example in JavaScript, Python or Java 8, even with small strings, 
say of lenght of around 30 $a$'s,
the decision takes an absurd amount of time to finish (see graphs in figure \ref{fig:aStarStarb}).
The algorithms clearly show exponential behaviour, and as can be seen
is triggered by some relatively simple regular expressions.
Java 9 and newer
versions improve this behaviour somewhat, but are still slow compared 
with the approach we are going to use in this thesis.



This superlinear blowup in regular expression engines
has caused grief in ``real life'' where it is 
given the name ``catastrophic backtracking'' or ``evil'' regular expressions.
For example, on 20 July 2016 one evil
regular expression brought the webpage
\href{http://stackexchange.com}{Stack Exchange} to its
knees.\footnote{\url{https://stackstatus.net/post/147710624694/outage-postmortem-july-20-2016}(Last accessed in 2019)}
In this instance, a regular expression intended to just trim white
spaces from the beginning and the end of a line actually consumed
massive amounts of CPU resources---causing the web servers to grind to a
halt. In this example, the time needed to process
the string was 
$O(n^2)$ with respect to the string length $n$. This
quadratic overhead was enough for the homepage of Stack Exchange to
respond so slowly that the load balancer assumed a $\mathit{DoS}$ 
attack and therefore stopped the servers from responding to any
requests. This made the whole site become unavailable. 

\begin{figure}[p]
\begin{center}
\begin{tabular}{@{}c@{\hspace{0mm}}c@{}}
\begin{tikzpicture}
\begin{axis}[
    xlabel={$n$},
    x label style={at={(1.05,-0.05)}},
    ylabel={time in secs},
    enlargelimits=false,
    xtick={0,5,...,30},
    xmax=33,
    ymax=35,
    ytick={0,5,...,30},
    scaled ticks=false,
    axis lines=left,
    width=5cm,
    height=4cm, 
    legend entries={JavaScript},  
    legend pos=north west,
    legend cell align=left]
\addplot[red,mark=*, mark options={fill=white}] table {re-js.data};
\end{axis}
\end{tikzpicture}
  &
\begin{tikzpicture}
\begin{axis}[
    xlabel={$n$},
    x label style={at={(1.05,-0.05)}},
    %ylabel={time in secs},
    enlargelimits=false,
    xtick={0,5,...,30},
    xmax=33,
    ymax=35,
    ytick={0,5,...,30},
    scaled ticks=false,
    axis lines=left,
    width=5cm,
    height=4cm, 
    legend entries={Python},  
    legend pos=north west,
    legend cell align=left]
\addplot[blue,mark=*, mark options={fill=white}] table {re-python2.data};
\end{axis}
\end{tikzpicture}\\ 
\begin{tikzpicture}
\begin{axis}[
    xlabel={$n$},
    x label style={at={(1.05,-0.05)}},
    ylabel={time in secs},
    enlargelimits=false,
    xtick={0,5,...,30},
    xmax=33,
    ymax=35,
    ytick={0,5,...,30},
    scaled ticks=false,
    axis lines=left,
    width=5cm,
    height=4cm, 
    legend entries={Java 8},  
    legend pos=north west,
    legend cell align=left]
\addplot[cyan,mark=*, mark options={fill=white}] table {re-java.data};
\end{axis}
\end{tikzpicture}
  &
\begin{tikzpicture}
\begin{axis}[
    xlabel={$n$},
    x label style={at={(1.05,-0.05)}},
    %ylabel={time in secs},
    enlargelimits=false,
    xtick={0,5,...,30},
    xmax=33,
    ymax=35,
    ytick={0,5,...,30},
    scaled ticks=false,
    axis lines=left,
    width=5cm,
    height=4cm, 
    legend entries={Dart},  
    legend pos=north west,
    legend cell align=left]
\addplot[green,mark=*, mark options={fill=white}] table {re-dart.data};
\end{axis}
\end{tikzpicture}\\ 
\begin{tikzpicture}
\begin{axis}[
    xlabel={$n$},
    x label style={at={(1.05,-0.05)}},
    ylabel={time in secs},
    enlargelimits=false,
    xtick={0,5,...,30},
    xmax=33,
    ymax=35,
    ytick={0,5,...,30},
    scaled ticks=false,
    axis lines=left,
    width=5cm,
    height=4cm, 
    legend entries={Swift},  
    legend pos=north west,
    legend cell align=left]
\addplot[purple,mark=*, mark options={fill=white}] table {re-swift.data};
\end{axis}
\end{tikzpicture}
  & 
\begin{tikzpicture}
\begin{axis}[
    xlabel={$n$},
    x label style={at={(1.05,-0.05)}},
    %ylabel={time in secs},
    enlargelimits=true,
    %xtick={0,5000,...,40000},
    %xmax=40000,
    %ymax=35,
    restrict x to domain*=0:40000,
    restrict y to domain*=0:35,
    %ytick={0,5,...,30},
    %scaled ticks=false,
    axis lines=left,
    width=5cm,
    height=4cm, 
    legend entries={Java9+},  
    legend pos=north west,
    legend cell align=left]
\addplot[orange,mark=*, mark options={fill=white}] table {re-java9.data};
\end{axis}
\end{tikzpicture}\\ 
\multicolumn{2}{c}{Graphs}
\end{tabular}    
\end{center}
\caption{Graphs showing runtime for matching $(a^*)^*\,b$ with strings 
           of the form $\protect\underbrace{aa..a}_{n}$ in various existing regular expression libraries.
   The reason for their superlinear behaviour is that they do a depth-first-search
   using NFAs.
   If the string does not match, the regular expression matching
   engine starts to explore all possibilities. 
}\label{fig:aStarStarb}
\end{figure}\afterpage{\clearpage}

A more recent example is a global outage of all Cloudflare servers on 2 July
2019. A poorly written regular expression exhibited catastrophic backtracking
and exhausted CPUs that serve HTTP traffic. Although the outage
had several causes, at the heart was a regular expression that
was used to monitor network
traffic.\footnote{\url{https://blog.cloudflare.com/details-of-the-cloudflare-outage-on-july-2-2019/}(Last accessed in 2022)}
These problems with regular expressions 
are not isolated events that happen
very rarely, 
%but actually widespread.
%They occur so often that they have a 
but they occur actually often enough that they have a
name: Regular-Expression-Denial-Of-Service (ReDoS)
attacks.
Davis et al. \cite{Davis18} detected more
than 1000 evil regular expressions
in Node.js, Python core libraries, npm and pypi. 
They therefore concluded that evil regular expressions
are a real problem rather than just "a parlour trick".

The work in this thesis aims to address this issue
with the help of formal proofs.
We describe a lexing algorithm based
on Brzozowski derivatives with verified correctness 
and a finiteness property for the size of derivatives
(which are all done in Isabelle/HOL).
Such properties %guarantee the absence of 
are an important step in preventing
catastrophic backtracking once and for all.
We will give more details in the next sections
on (i) why the slow cases in graph \ref{fig:aStarStarb}
can occur in traditional regular expression engines
and (ii) why we choose our 
approach based on Brzozowski derivatives and formal proofs.


\section{Preliminaries}%Regex, and the Problems with Regex Matchers}
Regular expressions and regular expression matchers 
have clearly been studied for many, many years.
Theoretical results in automata theory state 
that basic regular expression matching should be linear
w.r.t the input.
This assumes that the regular expression
$r$ was pre-processed and turned into a
deterministic finite automaton (DFA) before matching \cite{Sakarovitch2009}.
By basic we mean textbook definitions such as the one
below, involving only regular expressions for characters, alternatives,
sequences, and Kleene stars:
\[
	r ::= c | r_1 + r_2 | r_1 \cdot r_2 | r^*
\]
Modern regular expression matchers used by programmers,
however,
support much richer constructs, such as bounded repetitions,
negations,
and back-references.
To differentiate, we use the word \emph{regex} to refer
to those expressions with richer constructs while reserving the
term \emph{regular expression}
for the more traditional meaning in formal languages theory.
We follow this convention 
in this thesis.
In the future, we aim to support all the popular features of regexes, 
but for this work we mainly look at basic regular expressions
and bounded repetitions.



%Most modern regex libraries
%the so-called PCRE standard (Peral Compatible Regular Expressions)
%has the back-references
Regexes come with a number of constructs
that make it more convenient for 
programmers to write regular expressions.
Depending on the types of constructs
the task of matching and lexing with them
will have different levels of complexity.
Some of those constructs are syntactic sugars that are
simply short hand notations
that save the programmers a few keystrokes.
These will not cause problems for regex libraries.
For example the
non-binary alternative involving three or more choices just means:
\[
	(a | b | c) \stackrel{means}{=} ((a + b)+ c)
\]
Similarly, the range operator
%used to express the alternative
%of all characters between its operands, 
is just a concise way
of expressing an alternative of consecutive characters:
\[
	[0~-9]\stackrel{means}{=} (0 | 1 | \ldots | 9 )  
\]
for an alternative. The
wildcard character '$.$' is used to refer to any single character,
\[
	. \stackrel{means}{=} [0-9a-zA-Z+-()*\&\ldots]
\]
except the newline.

\subsection{Bounded Repetitions}
More interesting are bounded repetitions, which can 
make the regular expressions much
more compact.
Normally there are four kinds of bounded repetitions:
$r^{\{n\}}$, $r^{\{\ldots m\}}$, $r^{\{n\ldots \}}$ and $r^{\{n\ldots m\}}$
(where $n$ and $m$ are constant natural numbers).
Like the star regular expressions, the set of strings or language
a bounded regular expression can match
is defined using the power operation on sets:
\begin{center}
	\begin{tabular}{lcl}
		$L \; r^{\{n\}}$ & $\dn$ & $(L \; r)^n$\\
		$L \; r^{\{\ldots m\}}$ & $\dn$ & $\bigcup_{0 \leq i \leq m}. (L \; r)^i$\\
		$L \; r^{\{n\ldots \}}$ & $\dn$ & $\bigcup_{n \leq i}. (L \; r)^i$\\
		$L \; r^{\{n \ldots m\}}$ & $\dn$ & $\bigcup_{n \leq i \leq m}. (L \; r)^i$
	\end{tabular}
\end{center}
The attraction of bounded repetitions is that they can be
used to avoid a size blow up: for example $r^{\{n\}}$
is a shorthand for
the much longer regular expression:
\[
	\underbrace{r\ldots r}_\text{n copies of r}.
\]
%Therefore, a naive algorithm that simply unfolds
%them into their desugared forms
%will suffer from at least an exponential runtime increase.


The problem with matching 
such bounded repetitions
is that tools based on the classic notion of
automata need to expand $r^{\{n\}}$ into $n$ connected 
copies of the automaton for $r$. This leads to very inefficient matching
algorithms  or algorithms that consume large amounts of memory.
Implementations using $\DFA$s will
in such situations
either become excruciatingly slow 
(for example Verbatim++ \cite{Verbatimpp}) or run
out of memory (for example $\mathit{LEX}$ and 
$\mathit{JFLEX}$\footnote{LEX and JFLEX are lexer generators
in C and JAVA that generate $\mathit{DFA}$-based
lexers. The user provides a set of regular expressions
and configurations, and then 
gets an output program encoding a minimized $\mathit{DFA}$
that can be compiled and run. 
When given the above countdown regular expression,
a small $n$ (say 20) would result in a program representing a
DFA
with millions of states.}) for large counters.
A classic example for this phenomenon is the regular expression $(a+b)^*  a (a+b)^{n}$
where the minimal DFA requires at least $2^{n+1}$ states.
For example, when $n$ is equal to 2,
the corresponding $\mathit{NFA}$ looks like:
\vspace{6mm}
\begin{center}
\begin{tikzpicture}[shorten >=1pt,node distance=2cm,on grid,auto] 
   \node[state,initial] (q_0)   {$q_0$}; 
   \node[state, red] (q_1) [right=of q_0] {$q_1$}; 
   \node[state, red] (q_2) [right=of q_1] {$q_2$}; 
   \node[state, accepting, red](q_3) [right=of q_2] {$q_3$};
    \path[->] 
    (q_0) edge  node {a} (q_1)
    	  edge [loop below] node {a,b} ()
    (q_1) edge  node  {a,b} (q_2)
    (q_2) edge  node  {a,b} (q_3);
\end{tikzpicture}
\end{center}
and when turned into a DFA by the subset construction
requires at least $2^3$ states.\footnote{The 
red states are "countdown states" which count down 
the number of characters needed in addition to the current
string to make a successful match.
For example, state $q_1$ indicates a match that has
gone past the $(a|b)^*$ part of $(a|b)^*a(a|b)^{\{2\}}$,
and just consumed the "delimiter" $a$ in the middle, and 
needs to match 2 more iterations of $(a|b)$ to complete.
State $q_2$ on the other hand, can be viewed as a state
after $q_1$ has consumed 1 character, and just waits
for 1 more character to complete.
The state $q_3$ is the last (accepting) state, requiring 0 
more characters.
Depending on the suffix of the
input string up to the current read location,
the states $q_1$ and $q_2$, $q_3$
may or may
not be active.
A $\mathit{DFA}$ for such an $\mathit{NFA}$ would
contain at least $2^3$ non-equivalent states that cannot be merged, 
because the subset construction during determinisation will generate
all the elements in the power set $\mathit{Pow}\{q_1, q_2, q_3\}$.
Generalizing this to regular expressions with larger
bounded repetitions number, we have that
regexes shaped like $r^*ar^{\{n\}}$ when converted to $\mathit{DFA}$s
would require at least $2^{n+1}$ states, if $r$ itself contains
more than 1 string.
This is to represent all different 
scenarios in which "countdown" states are active.}


Bounded repetitions are important because they
tend to occur frequently in practical use,
for example in the regex library RegExLib, in
the rules library of Snort \cite{Snort1999}\footnote{
Snort is a network intrusion detection (NID) tool
for monitoring network traffic.
The network security community curates a list
of malicious patterns written as regexes,
which is used by Snort's detection engine
to match against network traffic for any hostile
activities such as buffer overflow attacks.}, 
as well as in XML Schema definitions (XSDs).
According to Bj\"{o}rklund et al \cite{xml2015},
more than half of the 
XSDs they found on the Maven.org central repository
have bounded regular expressions in them.
Often the counters are quite large, with the largest being
close to ten million. 
A smaller sample XSD they gave
is:
\lstset{
	basicstyle=\fontsize{8.5}{9}\ttfamily,
  language=XML,
  morekeywords={encoding,
    xs:schema,xs:element,xs:complexType,xs:sequence,xs:attribute}
}
\begin{lstlisting}
<sequence minOccurs="0" maxOccurs="65535">
	<element name="TimeIncr" type="mpeg7:MediaIncrDurationType"/>
 	<element name="MotionParams" type="float" minOccurs="2" maxOccurs="12"/>
</sequence>
\end{lstlisting}
This can be seen as the regex
$(ab^{2\ldots 12})^{0 \ldots 65535}$, where $a$ and $b$ are themselves
regular expressions 
satisfying certain constraints (such as 
satisfying the floating point number format).
It is therefore quite unsatisfying that 
some regular expressions matching libraries
impose adhoc limits
for bounded regular expressions:
For example, in the regular expression matching library in the Go
language the regular expression $a^{1001}$ is not permitted, because no counter
can be above 1000, and in the built-in Rust regular expression library
expressions such as $a^{\{1000\}\{100\}\{5\}}$ give an error message
for being too big. 
As Becchi and Crawley \cite{Becchi08}  have pointed out,
the reason for these restrictions
is that they simulate a non-deterministic finite
automata (NFA) with a breadth-first search.
This way the number of active states could
be equal to the counter number.
When the counters are large, 
the memory requirement could become
infeasible, and a regex engine
like in Go will reject this pattern straight away.
\begin{figure}[H]
\begin{center}
\begin{tikzpicture} [node distance = 2cm, on grid, auto]
 
    	\node (q0) [state, initial] {$0$};
	\node (q1) [state, right = of q0] {$1$};
	%\node (q2) [state, right = of q1] {$2$};
	\node (qdots) [right = of q1] {$\ldots$};
	\node (qn) [state, right = of qdots] {$n$};
	\node (qn1) [state, right = of qn] {$n+1$};
	\node (qn2) [state, right = of qn1] {$n+2$};
	\node (qn3) [state, accepting, right = of qn2] {$n+3$}; 
 
\path [-stealth, thick]
	(q0) edge [loop above] node {a} ()
    (q0) edge node {a}   (q1) 
    %(q1) edge node {.}   (q2)
    (q1) edge node {.}   (qdots)
    (qdots) edge node {.} (qn)
    (qn) edge node {.} (qn1)
    (qn1) edge node {b} (qn2)
    (qn2) edge node {$c$} (qn3);
\end{tikzpicture}
%\begin{tikzpicture}[shorten >=1pt,node distance=2cm,on grid,auto] 
%   \node[state,initial] (q_0)   {$0$}; 
%   \node[state, ] (q_1) [right=of q_0] {$1$}; 
%   \node[state, ] (q_2) [right=of q_1] {$2$}; 
%   \node[state,
%   \node[state, accepting, ](q_3) [right=of q_2] {$3$};
%    \path[->] 
%    (q_0) edge  node {a} (q_1)
%    	  edge [loop below] node {a,b} ()
%    (q_1) edge  node  {a,b} (q_2)
%    (q_2) edge  node  {a,b} (q_3);
%\end{tikzpicture}
\end{center}
\caption{The example given by Becchi and Crawley
	that NFA simulation can consume large
	amounts of memory: $.^*a.^{\{n\}}bc$ matching
	strings of the form $aaa\ldots aaaabc$.
	When traversing in a breadth-first manner,
all states from 0 till $n+1$ will become active.}\label{fig:inj}

\end{figure}
%Languages like $\mathit{Go}$ and $\mathit{Rust}$ use this
%type of $\mathit{NFA}$ simulation and guarantees a linear runtime
%in terms of input string length.
%TODO:try out these lexers
These problems can of course be solved in matching algorithms where 
automata go beyond the classic notion and for instance include explicit
counters \cite{Turo_ov__2020}.
These solutions can be quite efficient,
with the ability to process
gigabits of strings input per second
even with large counters \cite{Becchi08}.
These practical solutions do not come with
formal guarantees, and as pointed out by
Kuklewicz \cite{KuklewiczHaskell}, can be error-prone.
%But formal reasoning about these automata especially in Isabelle 
%can be challenging
%and un-intuitive. 
%Therefore, we take correctness and runtime claims made about these solutions
%with a grain of salt.

In the work reported in \cite{FoSSaCS2023} and here, 
we add better support using derivatives
for bounded regular expression $r^{\{n\}}$.
Our results
extend straightforwardly to
repetitions with intervals such as 
$r^{\{n\ldots m\}}$.
The merit of Brzozowski derivatives (more on this later)
on this problem is that
it can be naturally extended to support bounded repetitions.
Moreover these extensions are still made up of only small
inductive datatypes and recursive functions,
making it handy to deal with them in theorem provers.
%The point here is that Brzozowski derivatives and the algorithms by Sulzmann and Lu can be
%straightforwardly extended to deal with bounded regular expressions
%and moreover the resulting code still consists of only simple
%recursive functions and inductive datatypes.
Finally, bounded regular expressions do not destroy our finite
boundedness property, which we shall prove later on.





\subsection{Back-References}
The other way to simulate an $\mathit{NFA}$ for matching is choosing  
a single transition each time, keeping all the other options in 
a queue or stack, and backtracking if that choice eventually 
fails. 
This method, often called a  "depth-first-search", 
is efficient in many cases, but could end up
with exponential run time.
The backtracking method is employed in regex libraries
that support \emph{back-references}, for example
in Java and Python.
%\section{Back-references and The Terminology Regex}

%When one constructs an $\NFA$ out of a regular expression
%there is often very little to be done in the first phase, one simply 
%construct the $\NFA$ states based on the structure of the input regular expression.

%In the lexing phase, one can simulate the $\mathit{NFA}$ running in two ways:
%one by keeping track of all active states after consuming 
%a character, and update that set of states iteratively.
%This can be viewed as a breadth-first-search of the $\mathit{NFA}$
%for a path terminating
%at an accepting state.



Consider the following regular expression where the sequence
operator is omitted for brevity:
\begin{center}
	$r_1r_2r_3r_4$
\end{center}
In this example,
one could label sub-expressions of interest 
by parenthesizing them and giving 
them a number in the order in which their opening parentheses appear.
One possible way of parenthesizing and labelling is: 
\begin{center}
	$\underset{1}{(}r_1\underset{2}{(}r_2\underset{3}{(}r_3)\underset{4}{(}r_4)))$
\end{center}
The sub-expressions
$r_1r_2r_3r_4$, $r_1r_2r_3$, $r_3$ and $r_4$ are labelled
by 1 to 4, and can be ``referred back'' by their respective numbers. 
%These sub-expressions are called "capturing groups".
To do so, one uses the syntax $\backslash i$ 
to denote that we want the sub-string 
of the input matched by the i-th
sub-expression to appear again, 
exactly the same as it first appeared: 
\begin{center}
$\ldots\underset{\text{i-th lparen}}{(}{r_i})\ldots 
\underset{s_i \text{ which just matched} \;r_i}{\backslash i} \ldots$
\end{center}
Once the sub-string $s_i$ for the sub-expression $r_i$
has been fixed, there is no variability on what the back-reference
label $\backslash i$ can be---it is tied to $s_i$.
The overall string will look like $\ldots s_i \ldots s_i \ldots $
%The backslash and number $i$ are the
%so-called "back-references".
%Let $e$ be an expression made of regular expressions 
%and back-references. $e$ contains the expression $e_i$
%as its $i$-th capturing group.
%The semantics of back-reference can be recursively
%written as:
%\begin{center}
%	\begin{tabular}{c}
%		$L ( e \cdot \backslash i) = \{s @ s_i \mid s \in L (e)\quad s_i \in L(r_i)$\\
%		$s_i\; \text{match of ($e$, $s$)'s $i$-th capturing group string}\}$
%	\end{tabular}
%\end{center}
A concrete example
for back-references is
\begin{center}
$(.^*)\backslash 1$,
\end{center}
which matches
strings that can be split into two identical halves,
for example $\mathit{foofoo}$, $\mathit{ww}$ and so on.
Note that this is different from 
repeating the  sub-expression verbatim like
\begin{center}
	$(.^*)(.^*)$,
\end{center}
which does not impose any restrictions on what strings the second 
sub-expression $.^*$
might match.
Another example for back-references is
\begin{center}
$(.)(.)\backslash 2\backslash 1$
\end{center}
which matches four-character palindromes
like $abba$, $x??x$ and so on.

Back-references are a regex construct 
that programmers find quite useful.
According to Becchi and Crawley \cite{Becchi08},
6\% of Snort rules (up until 2008) use them.
The most common use of back-references
is to express well-formed html files,
where back-references are convenient for matching
opening and closing tags like 
\begin{center}
	$\langle html \rangle \ldots \langle / html \rangle$
\end{center}
A regex describing such a format
is
\begin{center}
	$\langle (.^+) \rangle \ldots \langle / \backslash 1 \rangle$
\end{center}
Despite being useful, the expressive power of regexes 
go beyond regular languages 
once back-references are included.
In fact, they allow the regex construct to express 
languages that cannot be contained in context-free
languages either.
For example, the back-reference $(a^*)b\backslash1 b \backslash 1$
expresses the language $\{a^n b a^n b a^n\mid n \in \mathbb{N}\}$,
which cannot be expressed by 
context-free grammars \cite{campeanu2003formal}.
Such a language is contained in the context-sensitive hierarchy
of formal languages. 
Also solving the matching problem involving back-references
is known to be NP-complete \parencite{alfred2014algorithms}.
Regex libraries supporting back-references such as 
PCRE \cite{pcre} therefore have to
revert to a depth-first search algorithm including backtracking.
What is unexpected is that even in the cases 
not involving back-references, there is still
a (non-negligible) chance they might backtrack super-linearly,
as shown in the graphs in figure \ref{fig:aStarStarb}.

Summing up, we can categorise existing 
practical regex libraries into two kinds:
(i) The ones  with  linear
time guarantees like Go and Rust. The downside with them is that
they impose restrictions
on the regular expressions (not allowing back-references, 
bounded repetitions cannot exceed an ad hoc limit etc.).
And (ii) those 
that allow large bounded regular expressions and back-references
at the expense of using backtracking algorithms.
They can potentially ``grind to a halt''
on some very simple cases, resulting 
ReDoS attacks if exposed to the internet.

The problems with both approaches are the motivation for us 
to look again at the regular expression matching problem. 
Another motivation is that regular expression matching algorithms
that follow the POSIX standard often contain errors and bugs 
as we shall explain next.

%We would like to have regex engines that can 
%deal with the regular part (e.g.
%bounded repetitions) of regexes more
%efficiently.
%Also we want to make sure that they do it correctly.
%It turns out that such aim is not so easy to achieve.
 %TODO: give examples such as RE2 GOLANG 1000 restriction, rust no repetitions 
% For example, the Rust regex engine claims to be linear, 
% but does not support lookarounds and back-references.
% The GoLang regex library does not support over 1000 repetitions.  
% Java and Python both support back-references, but shows
%catastrophic backtracking behaviours on inputs without back-references(
%when the language is still regular).
 %TODO: test performance of Rust on (((((a*a*)b*)b){20})*)c  baabaabababaabaaaaaaaaababaaaababababaaaabaaabaaaaaabaabaabababaababaaaaaaaaababaaaababababaaaaaaaaaaaaac
 %TODO: verify the fact Rust does not allow 1000+ reps




%The time cost of regex matching algorithms in general
%involve two different phases, and different things can go differently wrong on 
%these phases.
%$\DFA$s usually have problems in the first (construction) phase
%, whereas $\NFA$s usually run into trouble
%on the second phase.


\section{Error-prone POSIX Implementations}
Very often there are multiple ways of matching a string
with a regular expression.
In such cases the regular expressions matcher needs to
disambiguate.
The more widely used strategy is called POSIX,
which roughly speaking always chooses the longest initial match.
The POSIX strategy is widely adopted in many regular expression matchers
because it is a natural strategy for lexers.
However, many implementations (including the C libraries
used by Linux and OS X distributions) contain bugs
or do not meet the specification they claim to adhere to.
Kuklewicz maintains a unit test repository which lists some
problems with existing regular expression engines \cite{KuklewiczHaskell}.
In some cases, they either fail to generate a
result when there exists a match,
or give results that are inconsistent with the POSIX standard.
A concrete example is the regex:
\begin{center}
	$(aba + ab + a)^* \text{and the string} \; ababa$
\end{center}
The correct POSIX match for the above
involves the entire string $ababa$, 
split into two Kleene star iterations, namely $[ab], [aba]$ at positions
$[0, 2), [2, 5)$
respectively.
But feeding this example to the different engines
listed at regex101 \footnote{
	regex101 is an online regular expression matcher which
	provides API for trying out regular expression
	engines of multiple popular programming languages like
Java, Python, Go, etc.} \parencite{regex101}. 
yields
only two incomplete matches: $[aba]$ at $[0, 3)$
and $a$ at $[4, 5)$.
Fowler \cite{fowler2003} and Kuklewicz \cite{KuklewiczHaskell} 
commented that most regex libraries are not
correctly implementing the central POSIX
rule, called the maximum munch rule.
Grathwohl \parencite{grathwohl2014crash} 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
People have recognised that the implementation complexity of POSIX rules also come from
the specification being not very precise.
The developers of the regexp package of Go 
\footnote{\url{https://pkg.go.dev/regexp\#pkg-overview}}
commented that
\begin{quote}\it
``
The POSIX rule is computationally prohibitive
and not even well-defined.
``
\end{quote}
There are many informal summaries of this disambiguation
strategy, which are often quite long and delicate.
For example Kuklewicz \cite{KuklewiczHaskell} 
described the POSIX rule as (section 1, last paragraph):
\begin{quote}
	\begin{itemize}
		\item
regular expressions (REs) take the leftmost starting match, and the longest match starting there
earlier subpatterns have leftmost-longest priority over later subpatterns\\
\item
higher-level subpatterns have leftmost-longest priority over their component subpatterns\\
\item
REs have right associative concatenation which can be changed with parenthesis\\
\item
parenthesized subexpressions return the match from their last usage\\
\item
text of component subexpressions must be contained in the text of the 
higher-level subexpressions\\
\item
if "p" and "q" can never match the same text then "p|q" and "q|p" are equivalent, up to trivial renumbering of captured subexpressions\\
\item
if "p" in "p*" is used to capture non-empty text then additional repetitions of "p" will not capture an empty string\\
\end{itemize}
\end{quote}
%The text above 
%is trying to capture something very precise,
%and is crying out for formalising.
Ribeiro and Du Bois \cite{RibeiroAgda2017} have 
formalised the notion of bit-coded regular expressions
and proved their relations with simple regular expressions in
the dependently-typed proof assistant Agda.
They also proved the soundness and completeness of a matching algorithm
based on the bit-coded regular expressions.
Ausaf et al. \cite{AusafDyckhoffUrban2016}
are the first to
give a quite simple formalised POSIX
specification in Isabelle/HOL, and also prove 
that their specification coincides with an earlier (unformalised) 
POSIX specification given by Okui and Suzuki \cite{Okui10}.
They then formally proved the correctness of
a lexing algorithm by Sulzmann and Lu \cite{Sulzmann2014}
with regards to that specification.
They also found that the informal POSIX
specification by Sulzmann and Lu needs to be substantially 
revised in order for the correctness proof to go through.
Their original specification and proof were unfixable
according to Ausaf et al.


In the next section, we will briefly
introduce Brzozowski derivatives and Sulzmann
and Lu's algorithm, which the main point of this thesis builds on.
%We give a taste of what they 
%are like and why they are suitable for regular expression
%matching and lexing.
\section{Formal Specification of POSIX Matching 
and Brzozowski Derivatives}
%Now we start with the central topic of the thesis: Brzozowski derivatives.
Brzozowski \cite{Brzozowski1964} first introduced the 
concept of a \emph{derivative} of regular expression in 1964.
The derivative of a regular expression $r$
with respect to a character $c$, is written as $r \backslash c$.
This operation tells us what $r$ transforms into
if we ``chop'' off a particular character $c$ 
from all strings in the language of $r$ (defined
later as $L \; r$).
%To give a flavour of Brzozowski derivatives, we present
%two straightforward clauses from it:
%\begin{center}
%	\begin{tabular}{lcl}
%		$d \backslash c$     & $\dn$ & 
%		$\mathit{if} \;c = d\;\mathit{then}\;\ONE\;\mathit{else}\;\ZERO$\\
%$(r_1 + r_2)\backslash c$     & $\dn$ & $r_1 \backslash c \,+\, r_2 \backslash c$\\
%	\end{tabular}
%\end{center}
%\noindent
%The first clause says that for the regular expression
%denoting a singleton set consisting of a single-character string $\{ d \}$,
%we check the derivative character $c$ against $d$,
%returning a set containing only the empty string $\{ [] \}$
%if $c$ and $d$ are equal, and the empty set $\varnothing$ otherwise.
%The second clause states that to obtain the regular expression
%representing all strings' head character $c$ being chopped off
%from $r_1 + r_2$, one simply needs to recursively take derivative
%of $r_1$ and $r_2$ and then put them together.
Derivatives have the property
that $s \in L \; (r\backslash c)$ if and only if 
$c::s \in L \; r$ where $::$ stands for list prepending.
%This property can be used on regular expressions
%matching and lexing--to test whether a string $s$ is in $L \; r$,
%one simply takes derivatives of $r$ successively with
%respect to the characters (in the correct order) in $s$,
%and then test whether the empty string is in the last regular expression.
With this property, derivatives can give a simple solution
to the problem of matching a string $s$ with a regular
expression $r$: if the derivative of $r$ w.r.t.\ (in
succession) all the characters of the string matches the empty string,
then $r$ matches $s$ (and {\em vice versa}).  
%This makes formally reasoning about these properties such
%as correctness and complexity smooth and intuitive.
There are several mechanised proofs of this property in various theorem
provers,
for example one by Owens and Slind \cite{Owens2008} in HOL4,
another one by Krauss and Nipkow \cite{Nipkow98} in Isabelle/HOL, and
yet another in Coq by Coquand and Siles \cite{Coquand2012}.

In addition, one can extend derivatives to bounded repetitions
relatively straightforwardly. For example, the derivative for 
this can be defined as:
\begin{center}
	\begin{tabular}{lcl}
		$r^{\{n\}} \backslash c$     & $\dn$ & $r \backslash c \cdot
		r^{\{n-1\}} (\text{when} n > 0)$\\
	\end{tabular}
\end{center}
\noindent
Experimental results suggest that  unlike DFA-based solutions
for bounded regular expressions,
derivatives can cope
large counters
quite well.

There have also been 
extensions of derivatives to other regex constructs.
For example, Owens et al include the derivatives
for the \emph{NOT} regular expression, which is
able to concisely express C-style comments of the form
$/* \ldots */$ (see \cite{Owens2008}).
Another extension for derivatives is
regular expressions with look-aheads, done
by Miyazaki and Minamide
\cite{Takayuki2019}.
%We therefore use Brzozowski derivatives on regular expressions 
%lexing 



Given the above definitions and properties of
Brzozowski derivatives, one quickly realises their potential
in generating a formally verified algorithm for lexing: the clauses and property
can be easily expressed in a functional programming language 
or converted to theorem prover
code, with great ease.
Perhaps this is the reason why derivatives have sparked quite a bit of interest
in the functional programming and theorem prover communities in the last
fifteen or so years (
\cite{Almeidaetal10}, \cite{Berglund14}, \cite{Berglund18},
\cite{Chen12} and \cite{Coquand2012}
to name a few), despite being buried in the ``sands of time'' \cite{Owens2008}
after they were first published by Brzozowski.


However, there are two difficulties with derivative-based matchers:
First, Brzozowski's original matcher only generates a yes/no answer
for whether a regular expression matches a string or not.  This is too
little information in the context of lexing where separate tokens must
be identified and also classified (for example as keywords
or identifiers). 
Second, derivative-based matchers need to be more efficient in terms
of the sizes of derivatives.
Elegant and beautiful
as many implementations are,
they can be excruciatingly slow. 
For example, Sulzmann and Lu
claim a linear running time of their proposed algorithm,
but that was falsified by our experiments. The running time 
is actually $\Omega(2^n)$ in the worst case.
A similar claim about a theoretical runtime of $O(n^2)$ 
is made for the Verbatim \cite{Verbatim}
%TODO: give references
lexer, which calculates POSIX matches and is based on derivatives.
They formalized the correctness of the lexer, but not their complexity result.
In the performance evaluation section, they analyzed the run time
of matching $a$ with the string 
\begin{center}
	$\underbrace{a \ldots a}_{\text{n a's}}$.
\end{center}
\noindent
They concluded that the algorithm is quadratic in terms of 
the length of the input string.
When we tried out their extracted OCaml code with the example $(a+aa)^*$,
the time it took to match a string of 40 $a$'s was approximately 5 minutes.


\subsection{Sulzmann and Lu's Algorithm}
Sulzmann and Lu~\cite{Sulzmann2014} overcame the first 
problem with the yes/no answer 
by cleverly extending Brzozowski's matching
algorithm. Their extended version generates additional information on
\emph{how} a regular expression matches a string following the POSIX
rules for regular expression matching. They achieve this by adding a
second ``phase'' to Brzozowski's algorithm involving an injection
function. This injection function in a sense undoes the ``damage''
of the derivatives chopping off characters.
In earlier work, Ausaf et al provided the formal
specification of what POSIX matching means and proved in Isabelle/HOL
the correctness
of this extended algorithm accordingly
\cite{AusafDyckhoffUrban2016}.

The version of the algorithm proven correct
suffers however heavily from a 
second difficulty, where derivatives can
grow to arbitrarily big sizes. 
For example if we start with the
regular expression $(a+aa)^*$ and take
successive derivatives according to the character $a$, we end up with
a sequence of ever-growing derivatives like 

\def\ll{\stackrel{\_\backslash{} a}{\longrightarrow}}
\begin{center}
\begin{tabular}{rll}
$(a + aa)^*$ & $\ll$ & $(\ONE + \ONE{}a) \cdot (a + aa)^*$\\
& $\ll$ & $(\ZERO + \ZERO{}a + \ONE) \cdot (a + aa)^* \;+\; (\ONE + \ONE{}a) \cdot (a + aa)^*$\\
& $\ll$ & $(\ZERO + \ZERO{}a + \ZERO) \cdot (a + aa)^* + (\ONE + \ONE{}a) \cdot (a + aa)^* \;+\; $\\
& & $\qquad(\ZERO + \ZERO{}a + \ONE) \cdot (a + aa)^* + (\ONE + \ONE{}a) \cdot (a + aa)^*$\\
& $\ll$ & \ldots \hspace{15mm}(regular expressions of sizes 98, 169, 283, 468, 767, \ldots)
\end{tabular}
\end{center}
 
\noindent where after around 35 steps we usually run out of memory on a
typical computer.  Clearly, the
notation involving $\ZERO$s and $\ONE$s already suggests
simplification rules that can be applied to regular regular
expressions, for example $\ZERO{}\,r \Rightarrow \ZERO$, $\ONE{}\,r
\Rightarrow r$, $\ZERO{} + r \Rightarrow r$ and $r + r \Rightarrow
r$. While such simple-minded simplifications have been proved in 
the work by Ausaf et al. to preserve the correctness of Sulzmann and Lu's
algorithm \cite{AusafDyckhoffUrban2016}, they unfortunately do
\emph{not} help with limiting the growth of the derivatives shown
above: the growth is slowed, but the derivatives can still grow rather
quickly beyond any finite bound.

Therefore we want to look in this thesis at a second
algorithm by Sulzmann and Lu where they
overcame this ``growth problem'' \cite{Sulzmann2014}.
In this version, POSIX values are 
represented as bit sequences and such sequences are incrementally generated
when derivatives are calculated. The compact representation
of bit sequences and regular expressions allows them to define a more
``aggressive'' simplification method that keeps the size of the
derivatives finite no matter what the length of the string is.
They make some informal claims about the correctness and linear behaviour
of this version, but do not provide any supporting proof arguments, not
even ``pencil-and-paper'' arguments. They write about their bit-coded
\emph{incremental parsing method} (that is the algorithm to be formalised
in this dissertation)


  
  \begin{quote}\it
  ``Correctness Claim: We further claim that the incremental parsing
  method [..] in combination with the simplification steps [..]
  yields POSIX parse trees. We have tested this claim
  extensively [..] but yet
  have to work out all proof details.'' \cite[Page 14]{Sulzmann2014}
\end{quote}  
Ausaf and Urban made some initial progress towards the 
full correctness proof but still had to leave out the optimisation
Sulzmann and Lu proposed.
Ausaf  wrote \cite{Ausaf},
  \begin{quote}\it
``The next step would be to implement a more aggressive simplification procedure on annotated regular expressions and then prove the corresponding algorithm generates the same values as blexer. Alas due to time constraints we are unable to do so here.''
\end{quote}  
This thesis implements the aggressive simplifications envisioned
by Ausaf and Urban,
together with a formal proof of the correctness of those simplifications.


One of the most recent work in the context of lexing
%with this issue
is the Verbatim lexer by Egolf, Lasser and Fisher~\cite{Verbatim}.
This is relevant work for us and we will compare it later with 
our derivative-based matcher we are going to present.
There is also some newer work called
Verbatim++~\cite{Verbatimpp}, which does not use derivatives, 
but deterministic finite automaton instead.
We will also study this work in a later section.
%An example that gives problem to automaton approaches would be
%the regular expression $(a|b)^*a(a|b)^{\{n\}}$.
%It requires at least $2^{n+1}$ states to represent
%as a DFA.


\section{Basic Concepts}
Formal language theory usually starts with an alphabet 
denoting a set of characters.
Here we use the datatype of characters from Isabelle,
which roughly corresponds to the ASCII characters.
In what follows, we shall leave the information about the alphabet
implicit.
Then using the usual bracket notation for lists,
we can define strings made up of characters as: 
\begin{center}
\begin{tabular}{lcl}
$\textit{s}$ & $\dn$ & $[] \; |\; c  :: s$
\end{tabular}
\end{center}
where $c$ is a variable ranging over characters.
The $::$ stands for list cons and $[]$ for the empty
list.
For brevity, a singleton list is sometimes written as $[c]$.
Strings can be concatenated to form longer strings in the same
way we concatenate two lists, which we shall write as $s_1 @ s_2$.
We omit the precise 
recursive definition here.
We overload this concatenation operator for two sets of strings:
\begin{center}
\begin{tabular}{lcl}
$A @ B $ & $\dn$ & $\{s_A @ s_B \mid s_A \in A \land s_B \in B \}$\\
\end{tabular}
\end{center}
We also call the above \emph{language concatenation}.
The power of a language is defined recursively, using the 
concatenation operator $@$:
\begin{center}
\begin{tabular}{lcl}
$A^0 $ & $\dn$ & $\{ [] \}$\\
$A^{n+1}$ & $\dn$ & $A @ A^n$
\end{tabular}
\end{center}
The union of all powers of a language   
can be used to define the Kleene star operator:
\begin{center}
\begin{tabular}{lcl}
 $A*$ & $\dn$ & $\bigcup_{i \geq 0} A^i$ \\
\end{tabular}
\end{center}

\noindent
However, to obtain a more convenient induction principle 
in Isabelle/HOL, 
we instead define the Kleene star
as an inductive set: 

\begin{center}
\begin{mathpar}
	\inferrule{\mbox{}}{[] \in A*\\}

\inferrule{s_1 \in A \;\; s_2 \in A*}{s_1 @ s_2 \in A*}
\end{mathpar}
\end{center}
\noindent
We also define an operation of "chopping off" a character from
a language, which we call $\Der$, meaning \emph{Derivative} (for a language):
\begin{center}
\begin{tabular}{lcl}
$\textit{Der} \;c \;A$ & $\dn$ & $\{ s \mid c :: s \in A \}$\\
\end{tabular}
\end{center}
\noindent
This can be generalised to ``chopping off'' a string 
from all strings within a set $A$, 
namely:
\begin{center}
\begin{tabular}{lcl}
$\textit{Ders} \;s \;A$ & $\dn$ & $\{ s' \mid s@s' \in A \}$\\
\end{tabular}
\end{center}
\noindent
which is essentially the left quotient $A \backslash L$ of $A$ against 
the singleton language with $L = \{s\}$.
However, for our purposes here, the $\textit{Ders}$ definition with 
a single string is sufficient.

The reason for defining derivatives
is that they provide another approach
to test membership of a string in 
a set of strings. 
For example, to test whether the string
$bar$ is contained in the set $\{foo, bar, brak\}$, one can take derivative of the set with
respect to the string $bar$:
\begin{center}
\begin{tabular}{lll}
	$S = \{foo, bar, brak\}$ & $ \stackrel{\backslash b}{\rightarrow }$ & 
	$\{ar, rak\}$ \\
				 & $\stackrel{\backslash a}{\rightarrow}$ & $\{r \}$\\
				 & $\stackrel{\backslash r}{\rightarrow}$ & $\{[]\}$\\
				 %& $\stackrel{[] \in S \backslash bar}{\longrightarrow}$ & $bar \in S$\\
\end{tabular}	
\end{center}
\noindent
and in the end, test whether the set
contains the empty string.\footnote{We use the infix notation $A\backslash c$
	instead of $\Der \; c \; A$ for brevity, as it will always be
	clear from the context that we are operating
on languages rather than regular expressions.}

In general, if we have a language $S$,
then we can test whether $s$ is in $S$
by testing whether $[] \in S \backslash s$.
With the sequencing, Kleene star, and $\textit{Der}$ operator on languages,
we have a  few properties of how the language derivative can be defined using 
sub-languages.
For example, for the sequence operator, we have
something similar to a ``chain rule'':
\begin{lemma}
\[
	\Der \; c \; (A @ B) =
	\begin{cases}
	((\Der \; c \; A) \, @ \, B ) \cup (\Der \; c\; B) , &  \text{if} \;  [] \in A  \\
	 (\Der \; c \; A) \,  @ \, B, & \text{otherwise}
	 \end{cases}	
\]
\end{lemma}
\noindent
This lemma states that if $A$ contains the empty string, $\Der$ can "pierce" through it
and get to $B$.
The language derivative for $A*$ can be described using the language derivative
of $A$:
\begin{lemma}
$\textit{Der} \;c \;(A*) = (\textit{Der}\; c A) @ (A*)$\\
\end{lemma}
\begin{proof}
There are two inclusions to prove:
\begin{itemize}
\item{$\subseteq$}:\\
The set 
\[ S_1 = \{s \mid c :: s \in A*\} \]
is enclosed in the set
\[ S_2 = \{s_1 @ s_2 \mid s_1 \, s_2.\;  s_1 \in \{s \mid c :: s \in A\} \land s_2 \in A* \}. \]
This is because for any string $c::s$ satisfying $c::s \in A*$,
%whenever you have a string starting with a character 
%in the language of a Kleene star $A*$, 
%then that
the character $c$, together with a prefix of $s$
%immediately after $c$ 
forms the first iteration of $A*$, 
and the rest of the $s$ is also $A*$.
This coincides with the definition of $S_2$.
\item{$\supseteq$}:\\
Note that
\[ \Der \; c \; (A*) = \Der \; c \;  (\{ [] \} \cup (A @ A*) ) \]
holds.
Also the following holds:
\[ \Der \; c \;  (\{ [] \} \cup (A @ A*) ) = \Der\; c \; (A @ A*) \]
where the $\textit{RHS}$ can be rewritten
as \[ (\Der \; c\; A) @ A* \cup (\Der \; c \; (A*)) \]
which of course contains $\Der \; c \; A @ A*$.
\end{itemize}
\end{proof}

\noindent
The clever idea of Brzozowski was to find counterparts of $\Der$ and $\Ders$
for regular expressions.
To introduce them, we need to first give definitions for regular expressions,
which we shall do next.

\subsection{Regular Expressions and Their Meaning}
The \emph{basic regular expressions} are defined inductively
 by the following grammar:
\[			r ::=   \ZERO \mid  \ONE
			 \mid  c  
			 \mid  r_1 \cdot r_2
			 \mid  r_1 + r_2   
			 \mid r^*         
\]
\noindent
We call them basic because we will introduce
additional constructors in later chapters, such as negation
and bounded repetitions.
We use $\ZERO$ for the regular expression that
matches no string, and $\ONE$ for the regular
expression that matches only the empty string.\footnote{
Some authors
also use $\phi$ and $\epsilon$ for $\ZERO$ and $\ONE$
but we prefer this notation.} 
The sequence regular expression is written $r_1\cdot r_2$
and sometimes we omit the dot if it is clear which
regular expression is meant; the alternative
is written $r_1 + r_2$.
The \emph{language} or meaning of 
a regular expression is defined recursively as
a set of strings:
%TODO: FILL in the other defs
\begin{center}
\begin{tabular}{lcl}
$L \; \ZERO$ & $\dn$ & $\phi$\\
$L \; \ONE$ & $\dn$ & $\{[]\}$\\
$L \; c$ & $\dn$ & $\{[c]\}$\\
$L \; (r_1 + r_2)$ & $\dn$ & $ L \; r_1 \cup L \; r_2$\\
$L \; (r_1 \cdot r_2)$ & $\dn$ & $ L \; r_1 @ L \; r_2$\\
$L \; (r^*)$ & $\dn$ & $ (L\;r)*$
\end{tabular}
\end{center}
\noindent
%Now with language derivatives of a language and regular expressions and
%their language interpretations in place, we are ready to define derivatives on regular expressions.
With $L$, we are ready to introduce Brzozowski derivatives on regular expressions.
We do so by first introducing what properties they should satisfy.

\subsection{Brzozowski Derivatives and a Regular Expression Matcher}
%Recall, the language derivative acts on a set of strings
%and essentially chops off a particular character from
%all strings in that set, Brzozowski defined a derivative operation on regular expressions
%so that after derivative $L(r\backslash c)$ 
%will look as if it was obtained by doing a language derivative on $L(r)$:
%Recall that the language derivative acts on a 
%language (set of strings).
%One can decide whether a string $s$ belongs
%to a language $S$ by taking derivative with respect to
%that string and then checking whether the empty 
%string is in the derivative:
%\begin{center}
%\parskip \baselineskip
%\def\myupbracefill#1{\rotatebox{90}{\stretchto{\{}{#1}}}
%\def\rlwd{.5pt}
%\newcommand\notate[3]{%
%  \unskip\def\useanchorwidth{T}%
%  \setbox0=\hbox{#1}%
%  \def\stackalignment{c}\stackunder[-6pt]{%
%    \def\stackalignment{c}\stackunder[-1.5pt]{%
%      \stackunder[-2pt]{\strut #1}{\myupbracefill{\wd0}}}{%
%    \rule{\rlwd}{#2\baselineskip}}}{%
%  \strut\kern7pt$\hookrightarrow$\rlap{ \footnotesize#3}}\ignorespaces%
%}
%\Longstack{
%\notate{$\{ \ldots ,\;$ 
%	\notate{s}{1}{$(c_1 :: s_1)$} 
%	$, \; \ldots \}$ 
%}{1}{$S_{start}$} 
%}
%\Longstack{
%	$\stackrel{\backslash c_1}{\longrightarrow}$
%}
%\Longstack{
%	$\{ \ldots,\;$  \notate{$s_1$}{1}{$(c_2::s_2)$} 
%	$,\; \ldots \}$
%}
%\Longstack{
%	$\stackrel{\backslash c_2}{\longrightarrow}$ 
%}
%\Longstack{
%	$\{ \ldots,\;  s_2
%	,\; \ldots \}$
%}
%\Longstack{
%	$ \xdashrightarrow{\backslash c_3\ldots\ldots} $	
%}
%\Longstack{
%	\notate{$\{\ldots, [], \ldots\}$}{1}{$S_{end} = 
%	S_{start}\backslash s$}
%}
%\end{center}
%\begin{center}
%	$s \in S_{start} \iff [] \in S_{end}$
%\end{center}
%\noindent
Brzozowski noticed that $\Der$
can be ``mirrored'' on regular expressions which
he calls the derivative of a regular expression $r$
with respect to a character $c$, written
$r \backslash c$. This infix operator
takes regular expression $r$ as input
and a character as a right operand.
The derivative operation on regular expression
is defined such that the language of the derivative result 
coincides with the language of the original 
regular expression being taken 
derivative with respect to the same characters, namely
\begin{property}

\[
	L \; (r \backslash c) = \Der \; c \; (L \; r)
\]
\end{property}
\noindent
Next, we give the recursive definition of derivative on
regular expressions so that it satisfies the properties above.
%The derivative function, written $r\backslash c$, 
%takes a regular expression $r$ and character $c$, and
%returns a new regular expression representing
%the original regular expression's language $L \; r$
%being taken the language derivative with respect to $c$.
\begin{table}
	\begin{center}
\begin{tabular}{lcl}
		$\ZERO \backslash c$ & $\dn$ & $\ZERO$\\  
		$\ONE \backslash c$  & $\dn$ & $\ZERO$\\
		$d \backslash c$     & $\dn$ & 
		$\mathit{if} \;c = d\;\mathit{then}\;\ONE\;\mathit{else}\;\ZERO$\\
$(r_1 + r_2)\backslash c$     & $\dn$ & $r_1 \backslash c \,+\, r_2 \backslash c$\\
$(r_1 \cdot r_2)\backslash c$ & $\dn$ & $\mathit{if} \, [] \in L(r_1)$\\
	&   & $\mathit{then}\;(r_1\backslash c) \cdot r_2 \,+\, r_2\backslash c$\\
	&   & $\mathit{else}\;(r_1\backslash c) \cdot r_2$\\
	$(r^*)\backslash c$           & $\dn$ & $(r\backslash c) \cdot r^*$\\
\end{tabular}
	\end{center}
\caption{Derivative on Regular Expressions}
\label{table:der}
\end{table}
\noindent
The most involved cases are the sequence case
and the star case.
The sequence case says that if the first regular expression
contains an empty string, then the second component of the sequence
needs to be considered, as its derivative will contribute to the
result of this derivative:
\begin{center}
	\begin{tabular}{lcl}
		$(r_1 \cdot r_2 ) \backslash c$ & $\dn$ & 
		$\textit{if}\;\,([] \in L(r_1))\; 
		\textit{then} \; (r_1 \backslash c) \cdot r_2 + r_2 \backslash c$ \\
		& & $\textit{else} \; (r_1 \backslash c) \cdot r_2$
	\end{tabular}
\end{center}
\noindent
Notice how this closely resembles
the language derivative operation $\Der$:
\begin{center}
	\begin{tabular}{lcl}
		$\Der \; c \; (A @ B)$ & $\dn$ & 
		$ \textit{if} \;\,  [] \in A \; 
		\textit{then} \;\, ((\Der \; c \; A) @ B ) \cup 
		\Der \; c\; B$\\
		& & $\textit{else}\; (\Der \; c \; A) @ B$\\
	\end{tabular}
\end{center}
\noindent
The derivative of the star regular expression $r^*$ 
unwraps one iteration of $r$, turns it into $r\backslash c$,
and attaches the original $r^*$
after $r\backslash c$, so that 
we can further unfold it as many times as needed:
\[
	(r^*) \backslash c \dn (r \backslash c)\cdot r^*.
\]
Again,
the structure is the same as the language derivative of the Kleene star: 
\[
	\textit{Der} \;c \;(A*) \dn (\textit{Der}\; c A) @ (A*)
\]
In the above definition of $(r_1\cdot r_2) \backslash c$,
the $\textit{if}$ clause's
boolean condition 
$[] \in L(r_1)$ needs to be 
somehow recursively computed.
We call such a function that checks
whether the empty string $[]$ is 
in the language of a regular expression $\nullable$:
\begin{center}
		\begin{tabular}{lcl}
			$\nullable(\ZERO)$     & $\dn$ & $\mathit{false}$ \\  
			$\nullable(\ONE)$      & $\dn$ & $\mathit{true}$ \\
			$\nullable(c)$ 	       & $\dn$ & $\mathit{false}$ \\
			$\nullable(r_1 + r_2)$ & $\dn$ & $\nullable(r_1) \vee \nullable(r_2)$ \\
			$\nullable(r_1\cdot r_2)$  & $\dn$ & $\nullable(r_1) \wedge \nullable(r_2)$ \\
			$\nullable(r^*)$       & $\dn$ & $\mathit{true}$ \\
		\end{tabular}
\end{center}
\noindent
The $\ZERO$ regular expression
does not contain any string and
therefore is not \emph{nullable}.
$\ONE$ is \emph{nullable} 
by definition. 
The character regular expression $c$
corresponds to the singleton set $\{c\}$, 
and therefore does not contain the empty string.
The alternative regular expression is nullable
if at least one of its children is nullable.
The sequence regular expression
would require both children to have the empty string
to compose an empty string, and the Kleene star
is always nullable because it naturally
contains the empty string.  
\noindent
We have the following two correspondences between 
derivatives on regular expressions and
derivatives on a set of strings:
\begin{lemma}\label{derDer}
	\mbox{}
	\begin{itemize}
		\item
$\textit{Der} \; c \; L(r) = L (r\backslash c)$
\item
	$c\!::\!s \in L(r)$ \textit{iff} $s \in L(r\backslash c)$.
	\end{itemize}
\end{lemma}
\begin{proof}
	By induction on $r$.
\end{proof}
\noindent
which are the main properties of derivatives
that enables us later to reason about the correctness of
derivative-based matching.
We can generalise the derivative operation shown above for single characters
to strings as follows:

\begin{center}
\begin{tabular}{lcl}
$r \backslash_s (c\!::\!s) $ & $\dn$ & $(r \backslash c) \backslash_s s$ \\
$r \backslash_s [\,] $ & $\dn$ & $r$
\end{tabular}
\end{center}

\noindent
When there is no ambiguity, we will 
omit the subscript and use $\backslash$ instead
of $\backslash_s$ to denote
string derivatives for brevity.
Brzozowski's derivative-based
regular-expression matching algorithm can then be described as:

\begin{definition}
$\textit{match}\;s\;r \;\dn\; \nullable \; (r\backslash s)$
\end{definition}

\noindent
Assuming the string is given as a sequence of characters, say $c_0c_1 \ldots c_n$, 
this algorithm, presented graphically, is as follows:

\begin{equation}\label{matcher}
\begin{tikzcd}
r_0 \arrow[r, "\backslash c_0"]  & r_1 \arrow[r, "\backslash c_1"] & 
r_2 \arrow[r, dashed]  & r_n  \arrow[r,"\textit{nullable}?"] & 
\;\textrm{true}/\textrm{false}
\end{tikzcd}
\end{equation}

\noindent
 It is relatively
easy to show that this matcher is correct, namely
\begin{lemma}
	$\textit{match} \; s\; r  = \textit{true} \; \textit{iff} \; s \in L(r)$
\end{lemma}
\begin{proof}
	By induction on $s$ using the property of derivatives:
	lemma \ref{derDer}.
\end{proof}
\begin{figure}
\begin{center}
\begin{tikzpicture}
\begin{axis}[
    xlabel={$n$},
    ylabel={time in secs},
    %ymode = log,
    legend entries={Naive Matcher},  
    legend pos=north west,
    legend cell align=left]
\addplot[red,mark=*, mark options={fill=white}] table {NaiveMatcher.data};
\end{axis}
\end{tikzpicture} 
\caption{Matching the regular expression $(a^*)^*b$ against strings of the form
$\protect\underbrace{aa\ldots a}_\text{n \textit{a}s}
$ using Brzozowski's original algorithm}\label{NaiveMatcher}
\end{center}
\end{figure}
\noindent
If we implement the above algorithm naively, however,
the algorithm can be excruciatingly slow, as shown in 
\ref{NaiveMatcher}.
Note that both axes are in logarithmic scale.
Around two dozen characters
this algorithm already ``explodes'' with the regular expression 
$(a^*)^*b$.
To improve this situation, we need to introduce simplification
rules for the intermediate results,
such as $r + r \rightarrow r$ or $\ONE \cdot r \rightarrow r$,
and make sure those rules do not change the 
language of the regular expression.
One simple-minded simplification function
that achieves these requirements 
is given below (see Ausaf et al. \cite{AusafDyckhoffUrban2016}):
\begin{center}
	\begin{tabular}{lcl}
		$\simp \; r_1 \cdot r_2 $ & $ \dn$ & 
		$(\simp \; r_1,  \simp \; r_2) \; \textit{match}$\\
					  & & $\quad \case \; (\ZERO, \_) \Rightarrow \ZERO$\\
					  & & $\quad \case \; (\_, \ZERO) \Rightarrow \ZERO$\\
					  & & $\quad \case \; (\ONE, r_2') \Rightarrow r_2'$\\
					  & & $\quad \case \; (r_1', \ONE) \Rightarrow r_1'$\\
					  & & $\quad \case \; (r_1', r_2') \Rightarrow r_1'\cdot r_2'$\\
		$\simp \; r_1 + r_2$ & $\dn$ & $(\simp \; r_1, \simp \; r_2) \textit{match}$\\
				     & & $\quad \; \case \; (\ZERO, r_2') \Rightarrow r_2'$\\
				     & & $\quad \; \case \; (r_1', \ZERO) \Rightarrow r_1'$\\
				     & & $\quad \; \case \; (r_1', r_2') \Rightarrow r_1' + r_2'$\\
		$\simp \; r$ & $\dn$ & $r\quad\quad (otherwise)$
				   
	\end{tabular}
\end{center}
If we repeatedly apply this simplification  
function during the matching algorithm, 
we have a matcher with simplification:
\begin{center}
	\begin{tabular}{lcl}
		$\derssimp \; [] \; r$ & $\dn$ & $r$\\
		$\derssimp \; c :: cs \; r$ & $\dn$ & $\derssimp \; cs \; (\simp \; (r \backslash c))$\\
		$\textit{matcher}_{simp}\; s \; r $ & $\dn$ & $\nullable \; (\derssimp \; s\;r)$
	\end{tabular}
\end{center}
\begin{figure}
\begin{tikzpicture}
\begin{axis}[
    xlabel={$n$},
    ylabel={time in secs},
    %ymode = log,
    %xmode = log,
    grid = both,
    legend entries={Matcher With Simp},  
    legend pos=north west,
    legend cell align=left]
\addplot[red,mark=*, mark options={fill=white}] table {BetterMatcher.data};
\end{axis}
\end{tikzpicture} 
\caption{$(a^*)^*b$ 
against 
$\protect\underbrace{aa\ldots a}_\text{n \textit{a}s}$ Using $\textit{matcher}_{simp}$}\label{BetterMatcher}
\end{figure}
\noindent
The running time of $\textit{ders}\_\textit{simp}$
on the same example of Figure \ref{NaiveMatcher}
is now ``tame''  in terms of the length of inputs,
as shown in Figure \ref{BetterMatcher}.

So far, the story is use Brzozowski derivatives and
simplify as much as possible, and at the end test
whether the empty string is recognised 
by the final derivative.
But what if we want to 
do lexing instead of just getting a true/false answer?
Sulzmann and Lu \cite{Sulzmann2014} first came up with a nice and 
elegant (arguably as beautiful as the definition of the 
Brzozowski derivative) solution for this.

\section{Values and the Lexing Algorithm by Sulzmann and Lu}
In this section, we present a two-phase regular expression lexing 
algorithm.
The first phase takes successive derivatives with 
respect to the input string,
and the second phase does the reverse, \emph{injecting} back
characters, in the meantime constructing a lexing result.
We will introduce the injection phase in detail slightly
later, but as a preliminary we have to first define 
the datatype for lexing results, 
called \emph{value} or
sometimes also \emph{lexical value}.  
Values and regular
expressions correspond to each other 
as illustrated in the following
table:

\begin{center}
	\begin{tabular}{c@{\hspace{20mm}}c}
		\begin{tabular}{@{}rrl@{}}
			\multicolumn{3}{@{}l}{\textbf{Regular Expressions}}\medskip\\
			$r$ & $::=$  & $\ZERO$\\
			& $\mid$ & $\ONE$   \\
			& $\mid$ & $c$          \\
			& $\mid$ & $r_1 \cdot r_2$\\
			& $\mid$ & $r_1 + r_2$   \\
			\\
			& $\mid$ & $r^*$         \\
		\end{tabular}
		&
		\begin{tabular}{@{\hspace{0mm}}rrl@{}}
			\multicolumn{3}{@{}l}{\textbf{Values}}\medskip\\
			$v$ & $::=$  & \\
			&        & $\Empty$   \\
			& $\mid$ & $\Char \; c$          \\
			& $\mid$ & $\Seq\,v_1\, v_2$\\
			& $\mid$ & $\Left \;v$   \\
			& $\mid$ & $\Right\;v$  \\
			& $\mid$ & $\Stars\,[v_1,\ldots\,v_n]$ \\
		\end{tabular}
	\end{tabular}
\end{center}
\noindent
A value has an underlying string, which 
can be calculated by the ``flatten" function $|\_|$:
\begin{center}
	\begin{tabular}{lcl}
		$|\Empty|$ & $\dn$ &  $[]$\\
		$|\Char \; c|$ & $ \dn$ & $ [c]$\\
		$|\Seq \; v_1, \;v_2|$ & $ \dn$ & $ v_1| @ |v_2|$\\
		$|\Left \; v|$ & $ \dn$ & $ |v|$\\
		$|\Right \; v|$ & $ \dn$ & $ |v|$\\
		$|\Stars \; []|$ & $\dn$ & $[]$\\
		$|\Stars \; v::vs|$ &  $\dn$ & $ |v| @ |\Stars(vs)|$
	\end{tabular}
\end{center}
Sulzmann and Lu used a binary predicate, written $\vdash v:r $,
to indicate that a value $v$ could be generated from a lexing algorithm
with input $r$. They call it the value inhabitation relation,
defined by the rules.
\begin{figure}[H]
\begin{mathpar}
	\inferrule{\mbox{}}{\vdash \Char \; c : \mathbf{c}} \hspace{2em}

	\inferrule{\mbox{}}{\vdash \Empty :  \ONE} \hspace{2em}

\inferrule{\vdash v_1 : r_1 \;\; \vdash v_2 : r_2 }{\vdash \Seq \; v_1,\; v_2 : (r_1 \cdot r_2)}

\inferrule{\vdash v_1 : r_1}{\vdash \Left \; v_1 : r_1+r_2}

\inferrule{\vdash v_2 : r_2}{\vdash \Right \; v_2:r_1 + r_2}

\inferrule{\forall v \in vs. \vdash v:r \land  |v| \neq []}{\vdash \Stars \; vs : r^*}
\end{mathpar}
\caption{The inhabitation relation for values and regular expressions}\label{fig:inhab}
\end{figure}
\noindent
The condition $|v| \neq []$ in the premise of star's rule
is to make sure that for a given pair of regular 
expression $r$ and string $s$, the number of values 
satisfying $|v| = s$ and $\vdash v:r$ is finite.
This additional condition was
imposed by Ausaf and Urban to make their proofs easier.
Given a string and a regular expression, there can be
multiple values for it. For example, both
$\vdash \Seq(\Left \; ab)(\Right \; c):(ab+a)(bc+c)$ and
$\vdash \Seq(\Right\; a)(\Left \; bc ):(ab+a)(bc+c)$ hold
and the values both flatten to $abc$.
Lexers therefore have to disambiguate and choose only
one of the values to be generated. $\POSIX$ is one of the
disambiguation strategies that is widely adopted.

Ausaf et al. \cite{AusafDyckhoffUrban2016} 
formalised the property 
as a ternary relation.
The $\POSIX$ value $v$ for a regular expression
$r$ and string $s$, denoted as $(s, r) \rightarrow v$, can be specified 
in the following rules\footnote{The names of the rules are used
as they were originally given in \cite{AusafDyckhoffUrban2016}.}:
\begin{figure}[p]
\begin{mathpar}
	\inferrule[P1]{\mbox{}}{([], \ONE) \rightarrow \Empty}
		
	\inferrule[PC]{\mbox{}}{([c], c) \rightarrow \Char \; c}

	\inferrule[P+L]{(s,r_1)\rightarrow v_1}{(s, r_1+r_2)\rightarrow \Left \; v_1}

	\inferrule[P+R]{(s,r_2)\rightarrow v_2\\ s \notin L \; r_1}{(s, r_1+r_2)\rightarrow \Right \; v_2}

	\inferrule[PS]{(s_1, v_1) \rightarrow r_1 \\ (s_2, v_2)\rightarrow r_2\\
		\nexists s_3 \; s_4. s_3 \neq [] \land s_3 @ s_4 = s_2 \land 
		s_1@ s_3 \in L \; r_1 \land s_4 \in L \; r_2}{(s_1 @ s_2, r_1\cdot r_2) \rightarrow
	\Seq \; v_1 \; v_2}

	\inferrule[P{[]}]{\mbox{}}{([], r^*) \rightarrow \Stars([])}

	\inferrule[P*]{(s_1, v) \rightarrow v \\ (s_2, r^*) \rightarrow \Stars \; vs \\
		|v| \neq []\\ \nexists s_3 \; s_4. s_3 \neq [] \land s_3@s_4 = s_2 \land
		s_1@s_3 \in L \; r \land s_4 \in L \; r^*}{(s_1@s_2, r^*)\rightarrow \Stars \;
	(v::vs)}
\end{mathpar}
\caption{The inductive POSIX rules given by Ausaf et al.
	\cite{AusafDyckhoffUrban2016}.
This ternary relation, written $(s, r) \rightarrow v$, 
formalises the POSIX constraints on the
value $v$ given a string $s$ and 
regular expression $r$.
}
\label{fig:POSIXDef}
\end{figure}\afterpage{\clearpage}
\noindent

%\begin{figure}
%\begin{tikzpicture}[]
%    \node [minimum width = 6cm, rectangle split, rectangle split horizontal, 
%	    rectangle split parts=2, rectangle split part fill={red!30,blue!20}, style={draw, rounded corners, inner sep=10pt}]
%	    (node1)
%	    {$r_{token1}$
%	    \nodepart{two}  $\;\;\; \quad r_{token2}\;\;\;\quad$ };
%	    %\node [left = 6.0cm of node1] (start1) {hi};
%	    \node [left = 0.2cm of node1] (middle) {$v.s.$};
%    \node [minimum width = 6cm, left = 0.2cm of middle, rectangle split, rectangle split horizontal, 
%	    rectangle split parts=2, rectangle split part fill={red!30,blue!20}, style={draw, rounded corners, inner sep=10pt}]
%	    (node2)
%	    {$\quad\;\;\;r_{token1}\quad\;\;\;$
%	    \nodepart{two}  $r_{token2}$ };
%	    \node [below = 0.1cm of node2] (text1) {\checkmark preferred by POSIX};
%		\node [above = 1.5cm of middle, minimum width = 6cm, 
%			rectangle, style={draw, rounded corners, inner sep=10pt}] 
%			(topNode) {$s$};
%	    \path[->,draw]
%	        (topNode) edge node {split $A$} (node2)
%	        (topNode) edge node {split $B$} (node1)
%		;
%			
%
%\end{tikzpicture}
%\caption{Maximum munch example: $s$ matches $r_{token1} \cdot r_{token2}$}\label{munch}
%\end{figure}
The above $\POSIX$ rules follow the intuition described below: 
\begin{itemize}
	\item (Left Priority)\\
		Match the leftmost regular expression when multiple options for matching
		are available. See P+L and P+R where in P+R $s$ cannot
		be in the language of $L \; r_1$.
	\item (Maximum munch)\\
		Always match a subpart as much as possible before proceeding
		to the next part of the string.
		For example, when the string $s$ matches 
		$r_{part1}\cdot r_{part2}$, and we have two ways $s$ can be split:
		Then the split that matches a longer string for the first part
		$r_{part1}$ is preferred by this maximum munch rule.
		The side-condition 
		\begin{center}
		$\nexists s_3 \; s_4. s_3 \neq [] \land s_3 @ s_4 = s_2 \land 
		s_1@ s_3 \in L \; r_1 \land s_4 \in L \; r_2$
		\end{center}
		in PS causes this.
		%(See
		%\ref{munch} for an illustration).
\end{itemize}
\noindent
These disambiguation strategies can be 
quite practical.
For instance, when lexing a code snippet 
\[ 
	\textit{iffoo} = 3
\]
using a regular expression 
for keywords and 
identifiers:
%(for example, keyword is a nonempty string starting with letters 
%followed by alphanumeric characters or underscores):
\[
	r_{keyword} + r_{identifier}.
\]
If we want $\textit{iffoo}$ to be recognized
as an identifier
where identifiers are defined as usual (letters
followed by letters, numbers or underscores),
then a match with a keyword (if)
followed by
an identifier (foo) would be incorrect.
POSIX lexing generates what is included by lexing.

\noindent
We know that a POSIX 
value for regular expression $r$ is inhabited by $r$.
\begin{lemma}
$(r, s) \rightarrow v \implies \vdash v: r$
\end{lemma}
\noindent
The main property about a $\POSIX$ value is that 
given the same regular expression $r$ and string $s$,
one can always uniquely determine the $\POSIX$ value for it:
\begin{lemma}
$\textit{if} \,(s, r) \rightarrow v_1 \land (s, r) \rightarrow v_2\quad  \textit{then} \; v_1 = v_2$
\end{lemma}
\begin{proof}
By induction on $s$, $r$ and $v_1$. The inductive cases
are all the POSIX rules. 
Probably the most cumbersome cases are 
the sequence and star with non-empty iterations.
We shall give the details for proving the sequence case here.

When we have 
\[
	(s_1, r_1) \rightarrow v_1 \;\, and \;\, 
	(s_2, r_2) \rightarrow v_2  \;\, and \;\, 
	\nexists s_3 \; s_4. s_3 \neq [] \land s_3 @ s_4 = s_2 \land 
		s_1@ s_3 \in L \; r_1 \land s_4 \in L \; r_2
\]
we know that the last condition 
excludes the possibility of a 
string $s_1'$ longer than $s_1$ such that 
\[
(s_1', r_1) \rightarrow v_1'   \;\; 
and\;\; (s_2', r_2) \rightarrow v_2'\;\; and \;\;s_1' @s_2' = s 
\]
hold.
A shorter string $s_1''$ with $s_2''$ satisfying
\[
(s_1'', r_1) \rightarrow v_1''
\;\;and\;\; (s_2'', r_2) \rightarrow v_2'' \;\;and \;\;s_1'' @s_2'' = s 
\]
cannot possibly form a $\POSIX$ value either, because
by definition, there is a candidate
with a longer initial string
$s_1$. Therefore, we know that the POSIX
value $\Seq \; a \; b$ for $r_1 \cdot r_2$ matching
$s$ must have the 
property that 
\[
	|a| = s_1 \;\; and \;\; |b| = s_2.
\]
The goal is to prove that $a = v_1 $ and $b = v_2$.
If we have some other POSIX values $v_{10}$ and $v_{20}$ such that 
$(s_1, r_1) \rightarrow v_{10}$ and $(s_2, r_2) \rightarrow v_{20}$ hold,
then by induction hypothesis $v_{10} = v_1$ and $v_{20}= v_2$, 
which means this "other" $\POSIX$ value $\Seq(v_{10}, v_{20})$
is the same as $\Seq(v_1, v_2)$. 
\end{proof}
\noindent
We have now defined what a POSIX value is and shown that it is unique.
The problem is to generate
such a value in a lexing algorithm using derivatives.

\subsection{Sulzmann and Lu's Injection-based Lexing Algorithm}

Sulzmann and Lu extended Brzozowski's 
derivative-based matching
to a lexing algorithm by a second phase 
after the initial phase of successive derivatives.
This second phase generates a POSIX value 
if the regular expression matches the string.
The algorithm uses two functions called $\inj$ and $\mkeps$.
The function $\mkeps$ constructs a POSIX value from the last
derivative $r_n$:
\begin{ceqn}
\begin{equation}\label{graph:mkeps}
\begin{tikzcd}
r_0 \arrow[r, "\backslash c_0"]  & r_1 \arrow[r, "\backslash c_1"] & r_2 \arrow[r, dashed, "\ldots"] & r_n \arrow[d, "mkeps" description] \\
	        & 	              & 	            & v_n       
\end{tikzcd}
\end{equation}
\end{ceqn}
\noindent
In the above diagram, again we assume that
the input string $s$ is made of $n$ characters
$c_0c_1 \ldots c_{n-1}$ 
The last derivative operation 
$\backslash c_{n-1}$ generates the derivative $r_n$, for which
$\mkeps$ produces the value $v_n$. This value
tells us how the empty string is matched by the (nullable)
regular expression $r_n$, in a POSIX way.
The definition of $\mkeps$ is
	\begin{center}
		\begin{tabular}{lcl}
			$\mkeps \; \ONE$ 		& $\dn$ & $\Empty$ \\
			$\mkeps \; (r_{1}+r_{2})$	& $\dn$ 
						& $\textit{if}\; (\nullable \; r_{1}) \;\,
							\textit{then}\;\, \Left \; (\mkeps \; r_{1})$\\ 
						& & $\phantom{if}\; \textit{else}\;\, \Right \;(\mkeps \; r_{2})$\\
			$\mkeps \; (r_1 \cdot r_2)$ 	& $\dn$ & $\Seq\;(\mkeps\;r_1)\;(\mkeps \; r_2)$\\
			$\mkeps \; r^* $	        & $\dn$ & $\Stars\;[]$
		\end{tabular}
	\end{center}


\noindent 
The function prefers the left child $r_1$ of $r_1 + r_2$ 
to match an empty string if there is a choice.
When there is a star to match the empty string,
we give the $\Stars$ constructor an empty list, meaning
no iteration is taken.
The result of $\mkeps$ on a $\nullable$ $r$ 
is a POSIX value for $r$ and the empty string:
\begin{lemma}\label{mePosix}
$\nullable\; r \implies (r, []) \rightarrow (\mkeps\; r)$
\end{lemma}
\begin{proof}
	By induction on $r$.
\end{proof}
\noindent
After the $\mkeps$-call, Sulzmann and Lu inject back the characters one by one
in reverse order as they were chopped off in the derivative phase.
The function for this is called $\inj$. This function 
operates on values, unlike $\backslash$ which operates on regular expressions.
In the diagram below, $v_i$ stands for the (POSIX) value 
for how the regular expression 
$r_i$ matches the string $s_i$ consisting of the last $n-i$ characters
of $s$ (i.e. $s_i = c_i \ldots c_{n-1}$ ) from the previous lexical value $v_{i+1}$.
After injecting back $n$ characters, we get the lexical value for how $r_0$
matches $s$. 
\begin{figure}[H]
\begin{center}	
\begin{ceqn}
\begin{tikzcd}
r_0 \arrow[r, dashed] \arrow[d]& r_i \arrow[r, "\backslash c_i"]  \arrow[d]  & r_{i+1}  \arrow[r, dashed] \arrow[d]        & r_n \arrow[d, "mkeps" description] \\
v_0           \arrow[u]                 & v_i  \arrow[l, dashed]                              & v_{i+1} \arrow[l,"inj_{r_i} c_i"]                 & v_n \arrow[l, dashed]         
\end{tikzcd}
\end{ceqn}
\end{center}
\caption{The two-phase lexing algorithm by Sulzmann and Lu \cite{AusafDyckhoffUrban2016},
	matching the regular expression $r_0$ and string of the form $[c_0, c_1, \ldots, c_{n-1}]$.
	The first phase involves taking successive derivatives w.r.t the characters $c_0$,
	$c_1$, and so on. These are the same operations as they have appeared in the matcher
	\ref{matcher}. When the final derivative regular expression is nullable (contains the empty string),
	then the second phase starts. First, $\mkeps$ generates a POSIX value which tells us how $r_n$ matches
	the empty string, by always selecting the leftmost 
	nullable regular expression. After that, $\inj$ ``injects'' back the character in reverse order as they 
	appeared in the string, always preserving POSIXness.}\label{graph:inj}
\end{figure}
\noindent
The function $\textit{inj}$ as defined by Sulzmann and Lu
takes three arguments: a regular
expression ${r_{i}}$, before the character is chopped off, 
a character ${c_{i}}$ (the character we want to inject back) and 
the third argument $v_{i+1}$ the value we want to inject into. 
The result of an application 
$\inj \; r_i \; c_i \; v_{i+1}$ is a new value $v_i$ such that
\[
	(s_i, r_i) \rightarrow v_i
\]
holds.
The definition of $\textit{inj}$ is as follows: 
\begin{center}
\begin{tabular}{l@{\hspace{1mm}}c@{\hspace{5mm}}l}
  $\textit{inj}\;(c)\;c\,Empty$            & $\dn$ & $\Char\,c$\\
  $\textit{inj}\;(r_1 + r_2)\;c\; (\Left\; v)$ & $\dn$ & $\Left  \; (\textit{inj}\; r_1 \; c\,v)$\\
  $\textit{inj}\;(r_1 + r_2)\,c\; (\Right\;v)$ & $\dn$ & $\Right \; (\textit{inj}\;r_2\;c  \; v)$\\
  $\textit{inj}\;(r_1 \cdot r_2)\; c\;(\Seq \; v_1 \; v_2)$ & $\dn$  & 
  $\Seq \; (\textit{inj}\;r_1\;c\;v_1) \; v_2$\\
  $\textit{inj}\;(r_1 \cdot r_2)\; c\;(\Left \; (\Seq \; v_1\;v_2) )$ & 
  $\dn$  & $\Seq \; (\textit{inj}\,r_1\,c\,v_1)\; v_2$\\
  $\textit{inj}\;(r_1 \cdot r_2)\; c\; (\Right\; v)$ & $\dn$  & $\Seq\; (\textit{mkeps}\; r_1) \; (\textit{inj} \; r_2\;c\;v)$\\
  $\textit{inj}\;(r^*)\; c \; (\Seq \; v\; (\Stars\;vs))$         & $\dn$  & $\Stars\;\,((\textit{inj}\;r\;c\;v)\,::\,vs)$\\
\end{tabular}
\end{center}

\noindent 
The function recurses on 
the shape of regular
expressions and values.
Intuitively, each clause analyses 
how $r_i$ could have transformed when being 
derived by $c$, identifying which subpart
of $v_{i+1}$ has the ``hole'' 
to inject the character back into.
Once the character is
injected back to that sub-value; 
$\inj$ assembles all parts
to form a new value.

For instance, the last clause is an
injection into a sequence value $v_{i+1}$
whose second child
value is a star and the shape of the 
regular expression $r_i$ before injection 
is a star.
We therefore know 
the derivative 
starts on a star and ends as a sequence:
\[
	(r^*) \backslash c \longrightarrow r\backslash c \cdot r^*
\]
during which an iteration of the star
had just been unfolded, giving the below
value inhabitation relation:
\[
	\vdash \Seq \; v \; (\Stars \; vs) : (r\backslash c) \cdot r^*.
\]
The value list $vs$ corresponds to
matched star iterations,
and the ``hole'' lies in $v$ because
\[
	\vdash v: r\backslash c.
\]
Finally, 
$\inj \; r \;c \; v$ is prepended
to the previous list of iterations and then
wrapped under the $\Stars$
constructor, giving us $\Stars \; ((\inj \; r \; c \; v) ::vs)$.

Recall that lemma 
\ref{mePosix} tells us that
$\mkeps$ always generates the POSIX value.
The function $\inj$ preserves the POSIXness, provided
the value before injection is POSIX, namely
\begin{lemma}\label{injPosix}
	If$(r \backslash c, s) \rightarrow v $,
	then $(r, c :: s) \rightarrow (\inj r \; c\; v)$.
\end{lemma}
\begin{proof}
	By induction on $r$.
	The non-trivial cases are sequence and star.
	When $r = a \cdot b$, there can be
	three cases for the value $v$ satisfying $\vdash v:a\backslash c$.
	We give the reasoning why $\inj \; r \; c \; v$ is POSIX in each
	case.
	\begin{itemize}
		\item
			$v = \Seq \; v_a \; v_b$.\\
			The ``not nullable'' clause of the $\inj$ function is taken:
			\begin{center}
				\begin{tabular}{lcl}
					$\inj \; r \; c \; v$ &   $=$ & $ \inj \;\; (a \cdot b) \;\; c \;\; (\Seq \; v_a \; v_b) $\\
					& $=$ & $\Seq \; (\inj \;a \; c \; v_a) \; v_b$
				\end{tabular}
			\end{center}
			We know that there exists a unique pair of
			$s_a$ and $s_b$ satisfying	
				$(a \backslash c, s_a) \rightarrow v_a$,
				$(b , s_b) \rightarrow v_b$, and
				$\nexists s_3 \; s_4. s_3 \neq [] \land s_a @ s_3 \in 
				L \; (a\backslash c) \land
				s_4 \in L \; b$.
			The last condition gives us
			$\nexists s_3 \; s_4. s_3 \neq [] \land (c :: s_a )@ s_3 \in 
				L \; a \land
				s_4 \in L \; b$.
			By induction hypothesis, $(a, c::s_a) \rightarrow \inj \; a \; c \; v_a $ holds,
			and this gives us
			\[
				(a\cdot b, (c::s_a)@s_b) \rightarrow \Seq \; (\inj \; a\;c \;v_a) \; v_b.
			\]
		\item
			$v = \Left \; (\Seq \; v_a \; v_b)$\\
			The argument is almost identical to the above case,	
			except that a different clause of $\inj$ is taken:
			\begin{center}
				\begin{tabular}{lcl}
					$\inj \; r \; c \; v$ &   $=$ & $ \inj \;\; (a \cdot b) \;\; c \;\; (\Left \; (\Seq \; v_a \; v_b)) $\\
					& $=$ & $\Seq \; (\inj \;a \; c \; v_a) \; v_b$
				\end{tabular}
			\end{center}
			With similar reasoning, 

			\[
				(a\cdot b, (c::s_a)@s_b) \rightarrow \Seq \; (\inj \; a\;c \;v_a) \; v_b.
			\]
			again holds.
		\item 
			$v = \Right \; v_b$\\
			Again the injection result would be 
			\begin{center}
				\begin{tabular}{lcl}
					$\inj \; r \; c \; v$ &   $=$ & $ \inj \;\; (a \cdot b) \;\; c \;\; \Right \; (v_b) $\\
					& $=$ & $\Seq \; (\mkeps \; a) \; (\inj \;b \; c\; v_b)$
				\end{tabular}
			\end{center}
			We know that $a$ must be nullable,
			allowing us to call $\mkeps$ and get
			\[
				(a, []) \rightarrow \mkeps \; a.
			\]
			Also, by inductive hypothesis
			\[
				(b, c::s) \rightarrow \inj\; b \; c \; v_b
			\]
			holds.
			In addition, as
			$\Right \;v_b$  instead of $\Left \ldots$ is 
			the POSIX value for $v$, it must be the case
			that $s \notin L \;( (a\backslash c)\cdot b)$.
			This tells us that 
			\[
				\nexists s_3 \; s_4.
				s_3 @s_4 = s  \land s_3 \in L \; (a\backslash c) 
				\land s_4 \in L \; b
			\]
			which translates to
			\[
				\nexists s_3 \; s_4. \; s_3 \neq [] \land
				s_3 @s_4 = c::s  \land s_3 \in L \; a 
				\land s_4 \in L \; b.
			\]
			(Which says there cannot be a longer 
			initial split for $s$ other than the empty string.)
			Therefore we have $\Seq \; (\mkeps \; a) \;(\inj \;b \; c\; v_b)$
			as the POSIX value for $a\cdot b$.
	\end{itemize}
	The star case can be proven similarly.
\end{proof}
\noindent
Putting all together, Sulzmann and Lu obtained the following algorithm
outlined in the diagram \ref{graph:inj}:
\begin{center}
\begin{tabular}{lcl}
	$\lexer \; r \; [] $ & $=$ & $\textit{if} \; (\nullable \; r)\; \textit{then}\;  \Some(\mkeps \; r) \; \textit{else} \; \None$\\
	$\lexer \; r \;c::s$ & $=$ & $\textit{case}\; (\lexer \; (r\backslash c) \; s) \;\textit{of}\; $\\
	& & $\quad \phantom{\mid}\; \None \implies \None$\\
	& & $\quad \mid           \Some(v) \implies \Some(\inj \; r\; c\; v)$
\end{tabular}
\end{center}
\noindent
The central property of the $\lexer$ is that it gives the correct result
according to
POSIX rules. 
\begin{theorem}\label{lexerCorrectness}
	The $\lexer$ based on derivatives and injections is correct: 
	\begin{center}
		\begin{tabular}{lcl}
			$\lexer \; r \; s = \Some(v)$ & $ \Longleftrightarrow$ & $ (r, \; s) \rightarrow v$\\
			$\lexer \;r \; s = \None $ & $\Longleftrightarrow$ & $ \neg(\exists v. (r, s) \rightarrow v)$
		\end{tabular}
	\end{center}
\end{theorem} 
\begin{proof}
By induction on $s$. $r$ generalising over an arbitrary regular expression.
The $[]$ case is proven by an application of lemma \ref{mePosix}, and the inductive case
by lemma \ref{injPosix}.
\end{proof}
\noindent
As we did earlier in this chapter with the matcher, one can 
introduce simplification on the regular expression in each derivative step.
However, due to lexing, one needs to do a backward phase (w.r.t the forward derivative phase)
and ensure that
the values align with the regular expression at each step.
Therefore one has to
be careful not to break the correctness, as the injection 
function heavily relies on the structure of 
the regular expressions and values being aligned.
This can be achieved by recording some extra rectification functions
during the derivatives step and applying these rectifications in 
each run during the injection phase.
With extra care
one can show that POSIXness will not be affected
by the simplifications listed here \cite{AusafDyckhoffUrban2016}. 
\begin{center}
	\begin{tabular}{lcl}
		$\simp \; r_1 \cdot r_2 $ & $ \dn$ & 
		$(\simp \; r_1,  \simp \; r_2) \; \textit{match}$\\
					  & & $\quad \case \; (\ZERO, \_) \Rightarrow \ZERO$\\
					  & & $\quad \case \; (\_, \ZERO) \Rightarrow \ZERO$\\
					  & & $\quad \case \; (\ONE, r_2') \Rightarrow r_2'$\\
					  & & $\quad \case \; (r_1', \ONE) \Rightarrow r_1'$\\
					  & & $\quad \case \; (r_1', r_2') \Rightarrow r_1'\cdot r_2'$\\
		$\simp \; r_1 + r_2$ & $\dn$ & $(\simp \; r_1, \simp \; r_2) \textit{match}$\\
				     & & $\quad \; \case \; (\ZERO, r_2') \Rightarrow r_2'$\\
				     & & $\quad \; \case \; (r_1', \ZERO) \Rightarrow r_1'$\\
				     & & $\quad \; \case \; (r_1', r_2') \Rightarrow r_1' + r_2'$\\
		$\simp \; r$ & $\dn$ & $r\quad\quad (otherwise)$
				   
	\end{tabular}
\end{center}

However, one can still end up
with exploding derivatives,
even with the simple-minded simplification rules allowed
in an injection-based lexer.
\section{A Case Requiring More Aggressive Simplifications}
For example, when starting with the regular
expression $(a^* \cdot a^*)^*$ and building just over
a dozen successive derivatives 
w.r.t.~the character $a$, one obtains a derivative regular expression
with millions of nodes (when viewed as a tree)
even with the mentioned simplifications.
\begin{figure}[H]
\begin{center}
\begin{tikzpicture}
\begin{axis}[
    xlabel={$n$},
    ylabel={size},
    legend entries={Simple-Minded Simp, Naive Matcher},  
    legend pos=north west,
    legend cell align=left]
\addplot[red,mark=*, mark options={fill=white}] table {BetterWaterloo.data};
\addplot[blue,mark=*, mark options={fill=white}] table {BetterWaterloo1.data};
\end{axis}
\end{tikzpicture} 
\end{center}
\caption{Size of $(a^*\cdot a^*)^*$ against $\protect\underbrace{aa\ldots a}_\text{n \textit{a}s}$}
\end{figure}\label{fig:BetterWaterloo}
   
That is because Sulzmann and Lu's 
injection-based lexing algorithm keeps a lot of 
"useless" values that will not be used. 
These different ways of matching will grow exponentially with the string length.
Consider the case
\[
	r= (a^*\cdot a^*)^* \quad and \quad
	s=\underbrace{aa\ldots a}_\text{n \textit{a}s}
\]
as an example.
This is a highly ambiguous regular expression, with
many ways to split up the string into multiple segments for
different star iterations,
and for each segment 
multiple ways of splitting between 
the two $a^*$ sub-expressions.
When $n$ is equal to $1$, there are two lexical values for
the match:
\[
	\Stars \; [\Seq \; (\Stars \; [\Char \; a])\; (\Stars \; [])] \quad (v1)
\]
and
\[
	\Stars \; [\Seq \; (\Stars \; [])\; (\Stars \; [\Char \; a])] \quad (v2)
\]
The derivative of $\derssimp \;s \; r$ is
\[
	(a^*a^* + a^*)\cdot(a^*a^*)^*.
\]
The $a^*a^*$ and $a^*$ in the first child of the above sequence
correspond to value 1 and value 2, respectively.
When $n=2$, the number goes up to 7:
\[
	\Stars \; [\Seq \; (\Stars \; [\Char \; a, \Char \; a])\; (\Stars \; [])]
\]

\[
	\Stars \; [\Seq \; (\Stars \; [\Char \; a])\; (\Stars \; [\Char \; a])]
\]

\[
	\Stars \; [\Seq \; (\Stars \; [])\; (\Stars \; [\Char \; a, \Char \; a])]
\]

\[
	\Stars \; [\Seq \; (\Stars \; [\Char \; a])\; (\Stars \; []), \Seq \; (\Stars \; [\Char\;a])\; (\Stars\; []) ]
\]

\[
	\Stars \; [\Seq \; (\Stars \; [\Char \; a])\; (\Stars \; []), 
	 	   \Seq \; (\Stars \; [])\; (\Stars \; [\Char \; a])
		  ] 
\]

\[
	\Stars \; [
	 	   \Seq \; (\Stars \; [])\; (\Stars \; [\Char \; a]),
	 	   \Seq \; (\Stars \; [])\; (\Stars \; [\Char \; a])
		  ] 
\]
and
\[
	\Stars \; [
	 	   \Seq \; (\Stars \; [])\; (\Stars \; [\Char \; a]),
		   \Seq \; (\Stars \; [\Char \; a])\; (\Stars \; [])
		  ] 
\]
And $\derssimp \; aa \; (a^*a^*)^*$ is
\[
	((a^*a^* + a^*)+a^*)\cdot(a^*a^*)^* + 
	(a^*a^* + a^*)\cdot(a^*a^*)^*.
\]
which removes two out of the seven terms corresponding to the
seven distinct lexical values.

It is not surprising that there are exponentially many 
distinct lexical values that cannot be eliminated by 
the simple-minded simplification of $\derssimp$. 
A lexer without a good enough strategy to 
deduplicate will naturally
have an exponential runtime on highly
ambiguous regular expressions because there
are exponentially many matches.
For this particular example, it seems
that the number of distinct matches growth
speed is proportional to $(2n)!/(n!(n+1)!)$ ($n$ being the input length).

On the other hand, the
 $\POSIX$ value for $r= (a^*\cdot a^*)^*$  and 
$s=\underbrace{aa\ldots a}_\text{n \textit{a}s}$ is  
\[
	\Stars\,
	[\Seq \; (\Stars\,[\underbrace{\Char(a),\ldots,\Char(a)}_\text{n iterations}]), \Stars\,[]].
\]
At any moment, the  subterms in a regular expression
that will potentially result in a POSIX value is only
a minority among the many other terms,
and one can remove the ones that are not possible to 
be POSIX.
In the above example,
\begin{equation}\label{eqn:growth2}
	((a^*a^* + \underbrace{a^*}_\text{A})+\underbrace{a^*}_\text{duplicate of A})\cdot(a^*a^*)^* + 
	\underbrace{(a^*a^* + a^*)\cdot(a^*a^*)^*}_\text{further simp removes this}.
\end{equation}
can be further simplified by 
removing the underlined term first,
which would open up possibilities
of further simplification that removes the
underbraced part.
The result would be 
\[
	(\underbrace{a^*a^*}_\text{term 1} + \underbrace{a^*}_\text{term 2})\cdot(a^*a^*)^*.
\]
with corresponding values
\begin{center}
	\begin{tabular}{lr}
		$\Stars \; [\Seq \; (\Stars \; [\Char \; a, \Char \; a])\; (\Stars \; [])]$  & $(\text{term 1})$\\
		$\Stars \; [\Seq \; (\Stars \; [\Char \; a])\; (\Stars \; [\Char \; a])]  $ &  $(\text{term 2})$
	\end{tabular}
\end{center}
Other terms with an underlying value, such as
\[
	\Stars \; [\Seq \; (\Stars \; [])\; (\Stars \; [\Char \; a, \Char \; a])]
\]
do not to contribute a POSIX lexical value,
and therefore can be thrown away.

Ausaf et al. \cite{AusafDyckhoffUrban2016} 
have come up with some simplification steps, 
and incorporated the simplification into $\lexer$.
They call this lexer $\textit{lexer}_{simp}$ and proved that
\[
	\lexer \; r\; s = \textit{lexer}_{simp} \; r \; s
\]
The function $\textit{lexer}_{simp}$
involves some fiddly manipulation of value rectification,
which we omit here.
however those steps
are not yet sufficiently strong, to achieve the above effects.
And even with these relatively mild simplifications, the proof
is already quite a bit more complicated than the theorem \ref{lexerCorrectness}.
One would need to prove something like this: 
\[
	\textit{If}\; (\textit{snd} \; (\textit{simp} \; r\backslash c), s) \rightarrow  v  \;\;
	\textit{then}\;\; (r, c::s) \rightarrow 
	\inj\;\, r\,  \;c \;\, ((\textit{fst} \; (\textit{simp} \; r \backslash c))\; v).
\]
instead of the simple lemma \ref{injPosix}, where now $\textit{simp}$
not only has to return a simplified regular expression,
but also what specific simplifications 
have been done as a function on values
showing how one can transform the value
underlying the simplified regular expression
to the unsimplified one.

We therefore choose a slightly different approach
also described by Sulzmann and Lu to
get better simplifications, which uses
some augmented data structures compared to 
plain regular expressions.
We call them \emph{annotated}
regular expressions.
With annotated regular expressions,
we can avoid creating the intermediate values $v_1,\ldots v_n$ and a 
second phase altogether.
We introduce this new datatype and the 
corresponding algorithm in the next chapter.