532
|
1 |
% Chapter Template
|
|
2 |
|
|
3 |
\chapter{A Better Bound and Other Extensions} % Main chapter title
|
|
4 |
|
|
5 |
\label{Cubic} %In Chapter 5\ref{Chapter5} we discuss stronger simplifications to improve the finite bound
|
|
6 |
%in Chapter 4 to a polynomial one, and demonstrate how one can extend the
|
|
7 |
%algorithm to include constructs such as bounded repetitions and negations.
|
590
|
8 |
\lstset{style=myScalastyle}
|
|
9 |
|
|
10 |
|
|
11 |
This chapter is a ``miscellaneous''
|
|
12 |
chapter which records various
|
|
13 |
extensions to our $\blexersimp$'s formalisations.\\
|
|
14 |
Firstly we present further improvements
|
|
15 |
made to our lexer algorithm $\blexersimp$.
|
|
16 |
We devise a stronger simplification algorithm,
|
|
17 |
called $\bsimpStrong$, which can prune away
|
|
18 |
similar components in two regular expressions at the same
|
|
19 |
alternative level,
|
|
20 |
even if these regular expressions are not exactly the same.
|
|
21 |
We call the lexer that uses this stronger simplification function
|
|
22 |
$\blexerStrong$.
|
|
23 |
We conjecture that both
|
|
24 |
\begin{center}
|
|
25 |
$\blexerStrong \;r \; s = \blexer\; r\;s$
|
|
26 |
\end{center}
|
|
27 |
and
|
|
28 |
\begin{center}
|
|
29 |
$\llbracket \bdersStrong{a}{s} \rrbracket = O(\llbracket a \rrbracket^3)$
|
|
30 |
\end{center}
|
|
31 |
hold, but formalising
|
|
32 |
them is still work in progress.
|
|
33 |
We give reasons why the correctness and cubic size bound proofs
|
|
34 |
can be achieved,
|
|
35 |
by exploring the connection between the internal
|
|
36 |
data structure of our $\blexerStrong$ and
|
|
37 |
Animirov's partial derivatives.\\
|
|
38 |
Secondly, we extend our $\blexersimp$
|
|
39 |
to support bounded repetitions ($r^{\{n\}}$).
|
|
40 |
We update our formalisation of
|
|
41 |
the correctness and finiteness properties to
|
|
42 |
include this new construct. With bounded repetitions
|
|
43 |
we are able to out-compete other verified lexers such as
|
|
44 |
Verbatim++ on regular expressions which involve a lot of
|
|
45 |
repetitions. We also present the idempotency property proof
|
|
46 |
of $\bsimp$, which leverages the idempotency proof of $\rsimp$.
|
|
47 |
This reinforces our claim that the fixpoint construction
|
|
48 |
originally required by Sulzmann and Lu can be removed in $\blexersimp$.
|
|
49 |
\\
|
|
50 |
Last but not least, we present our efforts and challenges we met
|
|
51 |
in further improving the algorithm by data
|
|
52 |
structures such as zippers.
|
|
53 |
|
|
54 |
|
|
55 |
|
532
|
56 |
%----------------------------------------------------------------------------------------
|
|
57 |
% SECTION strongsimp
|
|
58 |
%----------------------------------------------------------------------------------------
|
|
59 |
\section{A Stronger Version of Simplification}
|
|
60 |
%TODO: search for isabelle proofs of algorithms that check equivalence
|
590
|
61 |
In our bitcoded lexing algorithm, (sub)terms represent (sub)matches.
|
591
|
62 |
For example, the regular expression
|
|
63 |
\[
|
|
64 |
aa \cdot a^*+ a \cdot a^* + aa\cdot a^*
|
|
65 |
\]
|
|
66 |
contains three terms,
|
|
67 |
expressing three possibilities it will match future input.
|
|
68 |
The first and the third terms are identical, which means we can eliminate
|
|
69 |
the latter as we know it will not be picked up by $\bmkeps$.
|
|
70 |
In $\bsimps$, the $\distinctBy$ function takes care of this.
|
|
71 |
The criteria $\distinctBy$ uses for removing a duplicate
|
|
72 |
$a_2$ in the list
|
|
73 |
\begin{center}
|
|
74 |
$rs_a@[a_1]@rs_b@[a_2]@rs_c$
|
|
75 |
\end{center}
|
|
76 |
is that
|
533
|
77 |
\begin{center}
|
591
|
78 |
$\rerase{a_1} = \rerase{a_2}$.
|
|
79 |
\end{center}
|
|
80 |
It can be characterised as the $LD$
|
|
81 |
rewrite rule in \ref{rrewriteRules}.\\
|
|
82 |
The problem , however, is that identical components
|
|
83 |
in two slightly different regular expressions cannot be removed:
|
|
84 |
\begin{figure}[H]
|
|
85 |
\[
|
|
86 |
(a+b+d) \cdot r_1 + (a+c+e) \cdot r_1 \stackrel{?}{\rightsquigarrow} (a+b+d) \cdot r_1 + (c+e) \cdot r_1
|
|
87 |
\]
|
|
88 |
\caption{Desired simplification, but not done in $\blexersimp$}\label{partialDedup}
|
|
89 |
\end{figure}
|
|
90 |
\noindent
|
|
91 |
A simplification like this actually
|
|
92 |
cannot be omitted,
|
|
93 |
as without it the size could blow up even with our $\simp$ function:
|
|
94 |
\begin{figure}[H]
|
|
95 |
\centering
|
|
96 |
\begin{tikzpicture}
|
|
97 |
\begin{axis}[
|
|
98 |
%xlabel={$n$},
|
|
99 |
myplotstyle,
|
|
100 |
xlabel={input length},
|
|
101 |
ylabel={size},
|
|
102 |
]
|
|
103 |
\addplot[blue,mark=*, mark options={fill=white}] table {bsimpExponential.data};
|
|
104 |
\end{axis}
|
533
|
105 |
\end{tikzpicture}
|
591
|
106 |
\caption{Runtime of $\blexersimp$ for matching
|
|
107 |
$\protect((a^* + (aa)^* + \ldots + (aaaaa)^* )^*)^*$
|
|
108 |
with strings
|
|
109 |
of the form $\protect\underbrace{aa..a}_{n}$.}\label{blexerExp}
|
|
110 |
\end{figure}
|
|
111 |
\noindent
|
|
112 |
We would like to apply the rewriting at some stage
|
|
113 |
\begin{figure}[H]
|
|
114 |
\[
|
|
115 |
(a+b+d) \cdot r_1 \longrightarrow a \cdot r_1 + b \cdot r_1 + d \cdot r_1
|
|
116 |
\]
|
|
117 |
\caption{Desired simplification, but not done in $\blexersimp$}\label{desiredSimp}
|
533
|
118 |
\end{figure}
|
|
119 |
\noindent
|
591
|
120 |
in our $\simp$ function,
|
|
121 |
so that it makes the simplification in \ref{partialDedup} possible.
|
|
122 |
For example,
|
|
123 |
can a function like the below one be used?
|
|
124 |
%TODO: simp' that spills things
|
533
|
125 |
|
591
|
126 |
Unfortunately,
|
|
127 |
if we introduce them in our
|
|
128 |
setting we would lose the POSIX property of our calculated values.
|
|
129 |
For example given the regular expression
|
|
130 |
\begin{center}
|
|
131 |
$(a + ab)(bc + c)$
|
|
132 |
\end{center}
|
|
133 |
and the string
|
|
134 |
\begin{center}
|
|
135 |
$ab$,
|
533
|
136 |
\end{center}
|
591
|
137 |
then our algorithm generates the following
|
|
138 |
correct POSIX value
|
|
139 |
\begin{center}
|
|
140 |
$\Seq \; (\Right \; ab) \; (\Right \; c)$.
|
533
|
141 |
\end{center}
|
591
|
142 |
Essentially it matches the string with the longer Right-alternative
|
|
143 |
in the first sequence (and
|
|
144 |
then the 'rest' with the character regular expression $c$ from the second sequence).
|
|
145 |
If we add the simplification above, then we obtain the following value
|
|
146 |
\begin{center}
|
|
147 |
$\Left \; (\Seq \; a \; (\Left \; bc))$
|
|
148 |
\end{center}
|
|
149 |
where the $\Left$-alternatives get priority.
|
|
150 |
However this violates the POSIX rules.
|
|
151 |
The reason for getting this undesired value
|
|
152 |
is that the new rule splits this regular expression up into
|
|
153 |
\begin{center}
|
|
154 |
$a\cdot(b c + c) + ab \cdot (bc + c)$,
|
533
|
155 |
\end{center}
|
591
|
156 |
which becomes a regular expression with a
|
|
157 |
totally different structure--the original
|
|
158 |
was a sequence, and now it becomes an alternative.
|
|
159 |
With an alternative the maximum munch rule no longer works.\\
|
|
160 |
A method to reconcile this is to do the
|
|
161 |
transformation in \ref{desiredSimp} ``non-invasively'',
|
|
162 |
meaning that we traverse the list of regular expressions
|
|
163 |
\begin{center}
|
|
164 |
$rs_a@[a]@rs_c$
|
|
165 |
\end{center}
|
|
166 |
in the alternative
|
|
167 |
\begin{center}
|
|
168 |
$\sum ( rs_a@[a]@rs_c)$
|
533
|
169 |
\end{center}
|
591
|
170 |
using a function similar to $\distinctBy$,
|
|
171 |
but this time
|
|
172 |
we allow a more general list rewrite:
|
592
|
173 |
\begin{mathpar}\label{cubicRule}
|
591
|
174 |
\inferrule{\vspace{0mm} }{rs_a@[a]@rs_c
|
|
175 |
\stackrel{s}{\rightsquigarrow }
|
|
176 |
rs_a@[\textit{prune} \; a \; rs_a]@rs_c }
|
592
|
177 |
\end{mathpar}
|
591
|
178 |
%L \; a_1' = L \; a_1 \setminus (\cup_{a \in rs_a} L \; a)
|
|
179 |
where $\textit{prune} \;a \; acc$ traverses $a$
|
|
180 |
without altering the structure of $a$, removing components in $a$
|
|
181 |
that have appeared in the accumulator $acc$.
|
|
182 |
For example
|
|
183 |
\begin{center}
|
|
184 |
$\textit{prune} \;\;\; (r_a+r_f+r_g+r_h)r_d \;\; \; [(r_a+r_b+r_c)r_d, (r_e+r_f)r_d] $
|
|
185 |
\end{center}
|
|
186 |
should be equal to
|
|
187 |
\begin{center}
|
|
188 |
$(r_g+r_h)r_d$
|
|
189 |
\end{center}
|
|
190 |
because $r_gr_d$ and
|
|
191 |
$r_hr_d$ are the only terms
|
|
192 |
that have not appeared in the accumulator list
|
|
193 |
\begin{center}
|
|
194 |
$[(r_a+r_b+r_c)r_d, (r_e+r_f)r_d]$.
|
|
195 |
\end{center}
|
|
196 |
We implemented
|
|
197 |
function $\textit{prune}$ in Scala,
|
|
198 |
and incorporated into our lexer,
|
|
199 |
by replacing the $\simp$ function
|
|
200 |
with a stronger version called $\bsimpStrong$
|
|
201 |
that prunes regular expressions.
|
|
202 |
We call this lexer $\blexerStrong$.
|
|
203 |
$\blexerStrong$ is able to drastically reduce the
|
|
204 |
internal data structure size which could
|
|
205 |
trigger exponential behaviours in
|
|
206 |
$\blexersimp$.
|
590
|
207 |
\begin{figure}[H]
|
533
|
208 |
\centering
|
|
209 |
\begin{tabular}{@{}c@{\hspace{0mm}}c@{\hspace{0mm}}c@{}}
|
|
210 |
\begin{tikzpicture}
|
|
211 |
\begin{axis}[
|
535
|
212 |
%xlabel={$n$},
|
|
213 |
myplotstyle,
|
|
214 |
xlabel={input length},
|
533
|
215 |
ylabel={size},
|
535
|
216 |
]
|
533
|
217 |
\addplot[red,mark=*, mark options={fill=white}] table {strongSimpCurve.data};
|
|
218 |
\end{axis}
|
|
219 |
\end{tikzpicture}
|
|
220 |
&
|
|
221 |
\begin{tikzpicture}
|
|
222 |
\begin{axis}[
|
535
|
223 |
%xlabel={$n$},
|
|
224 |
myplotstyle,
|
|
225 |
xlabel={input length},
|
533
|
226 |
ylabel={size},
|
535
|
227 |
]
|
533
|
228 |
\addplot[blue,mark=*, mark options={fill=white}] table {bsimpExponential.data};
|
|
229 |
\end{axis}
|
|
230 |
\end{tikzpicture}\\
|
535
|
231 |
\multicolumn{2}{l}{Graphs: Runtime for matching $((a^* + (aa)^* + \ldots + (aaaaa)^* )^*)^*$ with strings
|
591
|
232 |
of the form $\underbrace{aa..a}_{n}$.}
|
533
|
233 |
\end{tabular}
|
|
234 |
\caption{aaaaaStarStar} \label{fig:aaaaaStarStar}
|
|
235 |
\end{figure}
|
591
|
236 |
\begin{figure}[H]
|
|
237 |
\begin{lstlisting}
|
|
238 |
def atMostEmpty(r: Rexp) : Boolean = r match {
|
|
239 |
case ZERO => true
|
|
240 |
case ONE => true
|
|
241 |
case STAR(r) => atMostEmpty(r)
|
|
242 |
case SEQ(r1, r2) => atMostEmpty(r1) && atMostEmpty(r2)
|
|
243 |
case ALTS(r1, r2) => atMostEmpty(r1) && atMostEmpty(r2)
|
|
244 |
case CHAR(_) => false
|
|
245 |
}
|
|
246 |
|
|
247 |
|
|
248 |
def isOne(r: Rexp) : Boolean = r match {
|
|
249 |
case ONE => true
|
|
250 |
case SEQ(r1, r2) => isOne(r1) && isOne(r2)
|
|
251 |
case ALTS(r1, r2) => (isOne(r1) || isOne(r2)) && (atMostEmpty(r1) && atMostEmpty(r2))//rs.forall(atMostEmpty) && rs.exists(isOne)
|
|
252 |
case STAR(r0) => atMostEmpty(r0)
|
|
253 |
case CHAR(c) => false
|
|
254 |
case ZERO => false
|
|
255 |
}
|
|
256 |
|
|
257 |
//r = r' ~ tail' : If tail' matches tail => returns r'
|
|
258 |
def removeSeqTail(r: Rexp, tail: Rexp) : Rexp = r match {
|
|
259 |
case SEQ(r1, r2) =>
|
|
260 |
if(r2 == tail)
|
|
261 |
r1
|
|
262 |
else
|
|
263 |
ZERO
|
|
264 |
case r => ZERO
|
|
265 |
}
|
|
266 |
|
|
267 |
def prune(r: ARexp, acc: Set[Rexp]) : ARexp = r match{
|
|
268 |
case AALTS(bs, rs) => rs.map(r => prune(r, acc)).filter(_ != ZERO) match
|
|
269 |
{
|
|
270 |
//all components have been removed, meaning this is effectively a duplicate
|
|
271 |
//flats will take care of removing this AZERO
|
|
272 |
case Nil => AZERO
|
|
273 |
case r::Nil => fuse(bs, r)
|
|
274 |
case rs1 => AALTS(bs, rs1)
|
|
275 |
}
|
|
276 |
case ASEQ(bs, r1, r2) =>
|
|
277 |
//remove the r2 in (ra + rb)r2 to identify the duplicate contents of r1
|
|
278 |
prune(r1, acc.map(r => removeSeqTail(r, erase(r2)))) match {
|
|
279 |
//after pruning, returns 0
|
|
280 |
case AZERO => AZERO
|
|
281 |
//after pruning, got r1'.r2, where r1' is equal to 1
|
|
282 |
case r1p if(isOne(erase(r1p))) => fuse(bs ++ mkepsBC(r1p), r2)
|
|
283 |
//assemble the pruned head r1p with r2
|
|
284 |
case r1p => ASEQ(bs, r1p, r2)
|
|
285 |
}
|
|
286 |
//this does the duplicate component removal task
|
|
287 |
case r => if(acc(erase(r))) AZERO else r
|
|
288 |
}
|
|
289 |
\end{lstlisting}
|
|
290 |
\caption{pruning function together with its helper functions}
|
|
291 |
\end{figure}
|
|
292 |
\noindent
|
|
293 |
The benefits of using
|
|
294 |
$\textit{prune}$ such as refining the finiteness bound
|
|
295 |
to a cubic bound has not been formalised yet.
|
|
296 |
Therefore we choose to use Scala code rather than an Isabelle-style formal
|
|
297 |
definition like we did for $\simp$, as the definitions might change
|
|
298 |
to suit proof needs.
|
|
299 |
In the rest of the chapter we will use this convention consistently.
|
|
300 |
\begin{figure}[H]
|
|
301 |
\begin{lstlisting}
|
|
302 |
def distinctWith(rs: List[ARexp],
|
|
303 |
pruneFunction: (ARexp, Set[Rexp]) => ARexp,
|
|
304 |
acc: Set[Rexp] = Set()) : List[ARexp] =
|
|
305 |
rs match{
|
|
306 |
case Nil => Nil
|
|
307 |
case r :: rs =>
|
|
308 |
if(acc(erase(r)))
|
|
309 |
distinctWith(rs, pruneFunction, acc)
|
|
310 |
else {
|
|
311 |
val pruned_r = pruneFunction(r, acc)
|
|
312 |
pruned_r ::
|
|
313 |
distinctWith(rs,
|
|
314 |
pruneFunction,
|
|
315 |
turnIntoTerms(erase(pruned_r)) ++: acc
|
|
316 |
)
|
|
317 |
}
|
|
318 |
}
|
|
319 |
\end{lstlisting}
|
|
320 |
\caption{A Stronger Version of $\textit{distinctBy}$}
|
|
321 |
\end{figure}
|
|
322 |
\noindent
|
|
323 |
The function $\textit{prune}$ is used in $\distinctWith$.
|
|
324 |
$\distinctWith$ is a stronger version of $\distinctBy$
|
|
325 |
which not only removes duplicates as $\distinctBy$ would
|
|
326 |
do, but also uses the $\textit{pruneFunction}$
|
|
327 |
argument to prune away verbose components in a regular expression.\\
|
|
328 |
\begin{figure}[H]
|
|
329 |
\begin{lstlisting}
|
|
330 |
//a stronger version of simp
|
|
331 |
def bsimpStrong(r: ARexp): ARexp =
|
|
332 |
{
|
|
333 |
r match {
|
|
334 |
case ASEQ(bs1, r1, r2) => (bsimpStrong(r1), bsimpStrong(r2)) match {
|
|
335 |
//normal clauses same as simp
|
|
336 |
case (AZERO, _) => AZERO
|
|
337 |
case (_, AZERO) => AZERO
|
|
338 |
case (AONE(bs2), r2s) => fuse(bs1 ++ bs2, r2s)
|
|
339 |
//bs2 can be discarded
|
|
340 |
case (r1s, AONE(bs2)) => fuse(bs1, r1s) //assert bs2 == Nil
|
|
341 |
case (r1s, r2s) => ASEQ(bs1, r1s, r2s)
|
|
342 |
}
|
|
343 |
case AALTS(bs1, rs) => {
|
|
344 |
//distinctBy(flat_res, erase)
|
|
345 |
distinctWith(flats(rs.map(bsimpStrong(_))), prune) match {
|
|
346 |
case Nil => AZERO
|
|
347 |
case s :: Nil => fuse(bs1, s)
|
|
348 |
case rs => AALTS(bs1, rs)
|
|
349 |
}
|
|
350 |
}
|
|
351 |
//stars that can be treated as 1
|
|
352 |
case ASTAR(bs, r0) if(atMostEmpty(erase(r0))) => AONE(bs)
|
|
353 |
case r => r
|
|
354 |
}
|
|
355 |
}
|
|
356 |
\end{lstlisting}
|
|
357 |
\caption{The function $\bsimpStrong$ and $\bdersStrongs$}
|
|
358 |
\end{figure}
|
|
359 |
\noindent
|
|
360 |
$\distinctWith$, is in turn used in $\bsimpStrong$:
|
|
361 |
\begin{figure}[H]
|
|
362 |
\begin{lstlisting}
|
|
363 |
//Conjecture: [| bdersStrong(s, r) |] = O([| r |]^3)
|
|
364 |
def bdersStrong(s: List[Char], r: ARexp) : ARexp = s match {
|
|
365 |
case Nil => r
|
|
366 |
case c::s => bdersStrong(s, bsimpStrong(bder(c, r)))
|
|
367 |
}
|
|
368 |
\end{lstlisting}
|
|
369 |
\caption{The function $\bsimpStrong$ and $\bdersStrongs$}
|
|
370 |
\end{figure}
|
|
371 |
\noindent
|
|
372 |
We conjecture that the above Scala function $\bdersStrongs$,
|
|
373 |
written $\bdersStrong{\_}{\_}$ as an infix notation,
|
|
374 |
satisfies the following property:
|
|
375 |
\begin{conjecture}
|
|
376 |
$\llbracket \bdersStrong{a}{s} \rrbracket = O(\llbracket a \rrbracket^3)$
|
|
377 |
\end{conjecture}
|
|
378 |
The stronger version of $\blexersimp$'s
|
|
379 |
code in Scala looks like:
|
|
380 |
\begin{figure}[H]
|
|
381 |
\begin{lstlisting}
|
|
382 |
def strongBlexer(r: Rexp, s: String) : Option[Val] = {
|
|
383 |
Try(Some(decode(r, strong_blex_simp(internalise(r), s.toList)))).getOrElse(None)
|
|
384 |
}
|
|
385 |
def strong_blex_simp(r: ARexp, s: List[Char]) : Bits = s match {
|
|
386 |
case Nil => {
|
|
387 |
if (bnullable(r)) {
|
|
388 |
mkepsBC(r)
|
|
389 |
}
|
|
390 |
else
|
|
391 |
throw new Exception("Not matched")
|
|
392 |
}
|
|
393 |
case c::cs => {
|
|
394 |
strong_blex_simp(strongBsimp(bder(c, r)), cs)
|
|
395 |
}
|
|
396 |
}
|
|
397 |
\end{lstlisting}
|
|
398 |
\end{figure}
|
|
399 |
\noindent
|
|
400 |
We would like to preserve the correctness like the one
|
|
401 |
we had for $\blexersimp$:
|
|
402 |
\begin{conjecture}\label{cubicConjecture}
|
|
403 |
$\blexerStrong \;r \; s = \blexer\; r\;s$
|
|
404 |
\end{conjecture}
|
592
|
405 |
\noindent
|
|
406 |
We introduce the new rule \ref{cubicRule}
|
|
407 |
and augment our proofs in chapter \ref{Bitcoded2}.
|
|
408 |
The idea is to maintain the properties like
|
|
409 |
$r \stackrel{*}{\rightsquigarrow} \textit{bsimp} \; r$ etc.
|
591
|
410 |
|
592
|
411 |
In the next section,
|
|
412 |
we will describe why we
|
|
413 |
believe a cubic bound can be achieved.
|
|
414 |
We give an introduction to the
|
|
415 |
partial derivatives,
|
|
416 |
which was invented by Antimirov \cite{Antimirov95},
|
|
417 |
and then link it with the result of the function
|
|
418 |
$\bdersStrongs$.
|
|
419 |
|
|
420 |
\section{Antimirov's partial derivatives}
|
591
|
421 |
This suggests a link towrads "partial derivatives"
|
592
|
422 |
introduced The idea behind Antimirov's partial derivatives
|
591
|
423 |
is to do derivatives in a similar way as suggested by Brzozowski,
|
|
424 |
but maintain a set of regular expressions instead of a single one:
|
|
425 |
|
|
426 |
%TODO: antimirov proposition 3.1, needs completion
|
|
427 |
\begin{center}
|
|
428 |
\begin{tabular}{ccc}
|
|
429 |
$\partial_x(a+b)$ & $=$ & $\partial_x(a) \cup \partial_x(b)$\\
|
|
430 |
$\partial_x(\ONE)$ & $=$ & $\phi$
|
|
431 |
\end{tabular}
|
|
432 |
\end{center}
|
|
433 |
|
|
434 |
Rather than joining the calculated derivatives $\partial_x a$ and $\partial_x b$ together
|
|
435 |
using the alternatives constructor, Antimirov cleverly chose to put them into
|
|
436 |
a set instead. This breaks the terms in a derivative regular expression up,
|
|
437 |
allowing us to understand what are the "atomic" components of it.
|
|
438 |
For example, To compute what regular expression $x^*(xx + y)^*$'s
|
|
439 |
derivative against $x$ is made of, one can do a partial derivative
|
|
440 |
of it and get two singleton sets $\{x^* \cdot (xx + y)^*\}$ and $\{x \cdot (xx + y) ^* \}$
|
|
441 |
from $\partial_x(x^*) \cdot (xx + y) ^*$ and $\partial_x((xx + y)^*)$.
|
|
442 |
To get all the "atomic" components of a regular expression's possible derivatives,
|
|
443 |
there is a procedure Antimirov called $\textit{lf}$, short for "linear forms", that takes
|
|
444 |
whatever character is available at the head of the string inside the language of a
|
|
445 |
regular expression, and gives back the character and the derivative regular expression
|
|
446 |
as a pair (which he called "monomial"):
|
|
447 |
\begin{center}
|
|
448 |
\begin{tabular}{ccc}
|
|
449 |
$\lf(\ONE)$ & $=$ & $\phi$\\
|
|
450 |
$\lf(c)$ & $=$ & $\{(c, \ONE) \}$\\
|
|
451 |
$\lf(a+b)$ & $=$ & $\lf(a) \cup \lf(b)$\\
|
|
452 |
$\lf(r^*)$ & $=$ & $\lf(r) \bigodot \lf(r^*)$\\
|
|
453 |
\end{tabular}
|
|
454 |
\end{center}
|
|
455 |
%TODO: completion
|
|
456 |
|
|
457 |
There is a slight difference in the last three clauses compared
|
|
458 |
with $\partial$: instead of a dot operator $ \textit{rset} \cdot r$ that attaches the regular
|
|
459 |
expression $r$ with every element inside $\textit{rset}$ to create a set of
|
|
460 |
sequence derivatives, it uses the "circle dot" operator $\bigodot$ which operates
|
|
461 |
on a set of monomials (which Antimirov called "linear form") and a regular
|
|
462 |
expression, and returns a linear form:
|
|
463 |
\begin{center}
|
|
464 |
\begin{tabular}{ccc}
|
|
465 |
$l \bigodot (\ZERO)$ & $=$ & $\phi$\\
|
|
466 |
$l \bigodot (\ONE)$ & $=$ & $l$\\
|
|
467 |
$\phi \bigodot t$ & $=$ & $\phi$\\
|
|
468 |
$\{ (x, \ZERO) \} \bigodot t$ & $=$ & $\{(x,\ZERO) \}$\\
|
|
469 |
$\{ (x, \ONE) \} \bigodot t$ & $=$ & $\{(x,t) \}$\\
|
|
470 |
$\{ (x, p) \} \bigodot t$ & $=$ & $\{(x,p\cdot t) \}$\\
|
|
471 |
$\lf(a+b)$ & $=$ & $\lf(a) \cup \lf(b)$\\
|
|
472 |
$\lf(r^*)$ & $=$ & $\lf(r) \cdot \lf(r^*)$\\
|
|
473 |
\end{tabular}
|
|
474 |
\end{center}
|
|
475 |
%TODO: completion
|
|
476 |
|
|
477 |
Some degree of simplification is applied when doing $\bigodot$, for example,
|
|
478 |
$l \bigodot (\ZERO) = \phi$ corresponds to $r \cdot \ZERO \rightsquigarrow \ZERO$,
|
|
479 |
and $l \bigodot (\ONE) = l$ to $l \cdot \ONE \rightsquigarrow l$, and
|
|
480 |
$\{ (x, \ZERO) \} \bigodot t = \{(x,\ZERO) \}$ to $\ZERO \cdot x \rightsquigarrow \ZERO$,
|
|
481 |
and so on.
|
|
482 |
|
|
483 |
With the function $\lf$ one can compute all possible partial derivatives $\partial_{UNIV}(r)$ of a regular expression $r$ with
|
|
484 |
an iterative procedure:
|
|
485 |
\begin{center}
|
|
486 |
\begin{tabular}{llll}
|
|
487 |
$\textit{while}$ & $(\Delta_i \neq \phi)$ & & \\
|
|
488 |
& $\Delta_{i+1}$ & $ =$ & $\lf(\Delta_i) - \PD_i$ \\
|
|
489 |
& $\PD_{i+1}$ & $ =$ & $\Delta_{i+1} \cup \PD_i$ \\
|
|
490 |
$\partial_{UNIV}(r)$ & $=$ & $\PD$ &
|
|
491 |
\end{tabular}
|
|
492 |
\end{center}
|
|
493 |
|
|
494 |
|
|
495 |
$(r_1 + r_2) \cdot r_3 \longrightarrow (r_1 \cdot r_3) + (r_2 \cdot r_3)$,
|
|
496 |
|
|
497 |
|
532
|
498 |
|
|
499 |
|
594
|
500 |
\section{The NTIMES Constructor, and the Size Bound and Correctness Proof for it}
|
592
|
501 |
The NTIMES construct has the following closed form:
|
|
502 |
\begin{verbatim}
|
|
503 |
"rders_simp (RNTIMES r0 (Suc n)) (c#s) =
|
|
504 |
rsimp ( RALTS ( (map (optermsimp r0 ) (nupdates s r0 [Some ([c], n)]) ) ))"
|
|
505 |
\end{verbatim}
|
532
|
506 |
%----------------------------------------------------------------------------------------
|
|
507 |
% SECTION 1
|
|
508 |
%----------------------------------------------------------------------------------------
|
|
509 |
|
|
510 |
\section{Adding Support for the Negation Construct, and its Correctness Proof}
|
|
511 |
We now add support for the negation regular expression:
|
|
512 |
\[ r ::= \ZERO \mid \ONE
|
|
513 |
\mid c
|
|
514 |
\mid r_1 \cdot r_2
|
|
515 |
\mid r_1 + r_2
|
|
516 |
\mid r^*
|
|
517 |
\mid \sim r
|
|
518 |
\]
|
|
519 |
The $\textit{nullable}$ function's clause for it would be
|
|
520 |
\[
|
|
521 |
\textit{nullable}(~r) = \neg \nullable(r)
|
|
522 |
\]
|
|
523 |
The derivative would be
|
|
524 |
\[
|
|
525 |
~r \backslash c = ~ (r \backslash c)
|
|
526 |
\]
|
|
527 |
|
591
|
528 |
The most tricky part of lexing for the $~r$ regular expression
|
532
|
529 |
is creating a value for it.
|
|
530 |
For other regular expressions, the value aligns with the
|
591
|
531 |
structure of the regular expression:
|
532
|
532 |
\[
|
|
533 |
\vdash \Seq(\Char(a), \Char(b)) : a \cdot b
|
|
534 |
\]
|
591
|
535 |
But for the $~r$ regular expression, $s$ is a member of it if and only if
|
532
|
536 |
$s$ does not belong to $L(r)$.
|
|
537 |
That means when there
|
591
|
538 |
is a match for the not regular expression, it is not possible to generate how the string $s$ matched
|
532
|
539 |
with $r$.
|
|
540 |
What we can do is preserve the information of how $s$ was not matched by $r$,
|
|
541 |
and there are a number of options to do this.
|
|
542 |
|
591
|
543 |
We could give a partial value when there is a partial match for the regular expression inside
|
532
|
544 |
the $\mathbf{not}$ construct.
|
|
545 |
For example, the string $ab$ is not in the language of $(a\cdot b) \cdot c$,
|
|
546 |
A value for it could be
|
|
547 |
\[
|
|
548 |
\vdash \textit{Not}(\Seq(\Char(a), \Char(b))) : ~((a \cdot b ) \cdot c)
|
|
549 |
\]
|
|
550 |
The above example demonstrates what value to construct
|
|
551 |
when the string $s$ is at most a real prefix
|
|
552 |
of the strings in $L(r)$. When $s$ instead is not a prefix of any strings
|
|
553 |
in $L(r)$, it becomes unclear what to return as a value inside the $\textit{Not}$
|
|
554 |
constructor.
|
|
555 |
|
|
556 |
Another option would be to either store the string $s$ that resulted in
|
|
557 |
a mis-match for $r$ or a dummy value as a placeholder:
|
|
558 |
\[
|
533
|
559 |
\vdash \textit{Not}(abcd) : ~( r_1 )
|
532
|
560 |
\]
|
|
561 |
or
|
|
562 |
\[
|
533
|
563 |
\vdash \textit{Not}(\textit{Dummy}) : ~( r_1 )
|
532
|
564 |
\]
|
|
565 |
We choose to implement this as it is most straightforward:
|
|
566 |
\[
|
|
567 |
\mkeps(~(r)) = \textit{if}(\nullable(r)) \; \textit{Error} \; \textit{else} \; \textit{Not}(\textit{Dummy})
|
|
568 |
\]
|
|
569 |
|
|
570 |
%----------------------------------------------------------------------------------------
|
|
571 |
% SECTION 2
|
|
572 |
%----------------------------------------------------------------------------------------
|
|
573 |
|
|
574 |
\section{Bounded Repetitions}
|
|
575 |
|
|
576 |
|