author | Christian Urban <christian.urban@kcl.ac.uk> |
Fri, 28 Jan 2022 12:02:25 +0000 | |
changeset 398 | dac6d27c99c6 |
parent 397 | e1b74d618f1b |
child 400 | 46e5566ad4ba |
permissions | -rw-r--r-- |
396 | 1 |
(*<*) |
2 |
theory Paper |
|
3 |
imports |
|
4 |
"../Lexer" |
|
5 |
"../Simplifying" |
|
6 |
"../Positions" |
|
7 |
"../SizeBound4" |
|
8 |
"HOL-Library.LaTeXsugar" |
|
9 |
begin |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
10 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
11 |
declare [[show_question_marks = false]] |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
12 |
|
398 | 13 |
notation (latex output) |
14 |
If ("(\<^latex>\<open>\\textrm{\<close>if\<^latex>\<open>}\<close> (_)/ \<^latex>\<open>\\textrm{\<close>then\<^latex>\<open>}\<close> (_)/ \<^latex>\<open>\\textrm{\<close>else\<^latex>\<open>}\<close> (_))" 10) and |
|
15 |
Cons ("_\<^latex>\<open>\\mbox{$\\,$}\<close>::\<^latex>\<open>\\mbox{$\\,$}\<close>_" [75,73] 73) |
|
16 |
||
17 |
||
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
18 |
abbreviation |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
19 |
"der_syn r c \<equiv> der c r" |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
20 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
21 |
notation (latex output) |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
22 |
der_syn ("_\\_" [79, 1000] 76) and |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
23 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
24 |
ZERO ("\<^bold>0" 81) and |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
25 |
ONE ("\<^bold>1" 81) and |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
26 |
CH ("_" [1000] 80) and |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
27 |
ALT ("_ + _" [77,77] 78) and |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
28 |
SEQ ("_ \<cdot> _" [77,77] 78) and |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
29 |
STAR ("_\<^sup>\<star>" [79] 78) |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
30 |
|
396 | 31 |
(*>*) |
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
32 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
33 |
section {* Introduction *} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
34 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
35 |
text {* |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
36 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
37 |
Brzozowski \cite{Brzozowski1964} introduced the notion of the {\em |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
38 |
derivative} @{term "der c r"} of a regular expression \<open>r\<close> w.r.t.\ |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
39 |
a character~\<open>c\<close>, and showed that it gave a simple solution to the |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
40 |
problem of matching a string @{term s} with a regular expression @{term |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
41 |
r}: if the derivative of @{term r} w.r.t.\ (in succession) all the |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
42 |
characters of the string matches the empty string, then @{term r} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
43 |
matches @{term s} (and {\em vice versa}). The derivative has the |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
44 |
property (which may almost be regarded as its specification) that, for |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
45 |
every string @{term s} and regular expression @{term r} and character |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
46 |
@{term c}, one has @{term "cs \<in> L(r)"} if and only if \mbox{@{term "s \<in> L(der c r)"}}. |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
47 |
The beauty of Brzozowski's derivatives is that |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
48 |
they are neatly expressible in any functional language, and easily |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
49 |
definable and reasoned about in theorem provers---the definitions just |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
50 |
consist of inductive datatypes and simple recursive functions. A |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
51 |
mechanised correctness proof of Brzozowski's matcher in for example HOL4 |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
52 |
has been mentioned by Owens and Slind~\cite{Owens2008}. Another one in |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
53 |
Isabelle/HOL is part of the work by Krauss and Nipkow \cite{Krauss2011}. |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
54 |
And another one in Coq is given by Coquand and Siles \cite{Coquand2012}. |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
55 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
56 |
If a regular expression matches a string, then in general there is more |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
57 |
than one way of how the string is matched. There are two commonly used |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
58 |
disambiguation strategies to generate a unique answer: one is called |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
59 |
GREEDY matching \cite{Frisch2004} and the other is POSIX |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
60 |
matching~\cite{POSIX,Kuklewicz,OkuiSuzuki2010,Sulzmann2014,Vansummeren2006}. |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
61 |
For example consider the string @{term xy} and the regular expression |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
62 |
\mbox{@{term "STAR (ALT (ALT x y) xy)"}}. Either the string can be |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
63 |
matched in two `iterations' by the single letter-regular expressions |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
64 |
@{term x} and @{term y}, or directly in one iteration by @{term xy}. The |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
65 |
first case corresponds to GREEDY matching, which first matches with the |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
66 |
left-most symbol and only matches the next symbol in case of a mismatch |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
67 |
(this is greedy in the sense of preferring instant gratification to |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
68 |
delayed repletion). The second case is POSIX matching, which prefers the |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
69 |
longest match. |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
70 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
71 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
72 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
73 |
\begin{figure}[t] |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
74 |
\begin{center} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
75 |
\begin{tikzpicture}[scale=2,node distance=1.3cm, |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
76 |
every node/.style={minimum size=6mm}] |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
77 |
\node (r1) {@{term "r\<^sub>1"}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
78 |
\node (r2) [right=of r1]{@{term "r\<^sub>2"}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
79 |
\draw[->,line width=1mm](r1)--(r2) node[above,midway] {@{term "der a DUMMY"}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
80 |
\node (r3) [right=of r2]{@{term "r\<^sub>3"}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
81 |
\draw[->,line width=1mm](r2)--(r3) node[above,midway] {@{term "der b DUMMY"}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
82 |
\node (r4) [right=of r3]{@{term "r\<^sub>4"}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
83 |
\draw[->,line width=1mm](r3)--(r4) node[above,midway] {@{term "der c DUMMY"}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
84 |
\draw (r4) node[anchor=west] {\;\raisebox{3mm}{@{term nullable}}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
85 |
\node (v4) [below=of r4]{@{term "v\<^sub>4"}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
86 |
\draw[->,line width=1mm](r4) -- (v4); |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
87 |
\node (v3) [left=of v4] {@{term "v\<^sub>3"}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
88 |
\draw[->,line width=1mm](v4)--(v3) node[below,midway] {\<open>inj r\<^sub>3 c\<close>}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
89 |
\node (v2) [left=of v3]{@{term "v\<^sub>2"}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
90 |
\draw[->,line width=1mm](v3)--(v2) node[below,midway] {\<open>inj r\<^sub>2 b\<close>}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
91 |
\node (v1) [left=of v2] {@{term "v\<^sub>1"}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
92 |
\draw[->,line width=1mm](v2)--(v1) node[below,midway] {\<open>inj r\<^sub>1 a\<close>}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
93 |
\draw (r4) node[anchor=north west] {\;\raisebox{-8mm}{@{term "mkeps"}}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
94 |
\end{tikzpicture} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
95 |
\end{center} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
96 |
\mbox{}\\[-13mm] |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
97 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
98 |
\caption{The two phases of the algorithm by Sulzmann \& Lu \cite{Sulzmann2014}, |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
99 |
matching the string @{term "[a,b,c]"}. The first phase (the arrows from |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
100 |
left to right) is \Brz's matcher building successive derivatives. If the |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
101 |
last regular expression is @{term nullable}, then the functions of the |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
102 |
second phase are called (the top-down and right-to-left arrows): first |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
103 |
@{term mkeps} calculates a value @{term "v\<^sub>4"} witnessing |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
104 |
how the empty string has been recognised by @{term "r\<^sub>4"}. After |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
105 |
that the function @{term inj} ``injects back'' the characters of the string into |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
106 |
the values. |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
107 |
\label{Sulz}} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
108 |
\end{figure} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
109 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
110 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
111 |
*} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
112 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
113 |
section {* Background *} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
114 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
115 |
text {* |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
116 |
Sulzmann-Lu algorithm with inj. State that POSIX rules. |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
117 |
metion slg is correct. |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
118 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
119 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
120 |
\begin{figure}[t] |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
121 |
\begin{center} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
122 |
\begin{tabular}{c} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
123 |
@{thm[mode=Axiom] Posix.intros(1)}\<open>P\<close>@{term "ONE"} \qquad |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
124 |
@{thm[mode=Axiom] Posix.intros(2)}\<open>P\<close>@{term "c"}\medskip\\ |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
125 |
@{thm[mode=Rule] Posix.intros(3)[of "s" "r\<^sub>1" "v" "r\<^sub>2"]}\<open>P+L\<close>\qquad |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
126 |
@{thm[mode=Rule] Posix.intros(4)[of "s" "r\<^sub>2" "v" "r\<^sub>1"]}\<open>P+R\<close>\medskip\\ |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
127 |
$\mprset{flushleft} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
128 |
\inferrule |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
129 |
{@{thm (prem 1) Posix.intros(5)[of "s\<^sub>1" "r\<^sub>1" "v\<^sub>1" "s\<^sub>2" "r\<^sub>2" "v\<^sub>2"]} \qquad |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
130 |
@{thm (prem 2) Posix.intros(5)[of "s\<^sub>1" "r\<^sub>1" "v\<^sub>1" "s\<^sub>2" "r\<^sub>2" "v\<^sub>2"]} \\\\ |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
131 |
@{thm (prem 3) Posix.intros(5)[of "s\<^sub>1" "r\<^sub>1" "v\<^sub>1" "s\<^sub>2" "r\<^sub>2" "v\<^sub>2"]}} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
132 |
{@{thm (concl) Posix.intros(5)[of "s\<^sub>1" "r\<^sub>1" "v\<^sub>1" "s\<^sub>2" "r\<^sub>2" "v\<^sub>2"]}}$\<open>PS\<close>\\ |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
133 |
@{thm[mode=Axiom] Posix.intros(7)}\<open>P[]\<close>\medskip\\ |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
134 |
$\mprset{flushleft} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
135 |
\inferrule |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
136 |
{@{thm (prem 1) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]} \qquad |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
137 |
@{thm (prem 2) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]} \qquad |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
138 |
@{thm (prem 3) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]} \\\\ |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
139 |
@{thm (prem 4) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]}} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
140 |
{@{thm (concl) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]}}$\<open>P\<star>\<close> |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
141 |
\end{tabular} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
142 |
\end{center} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
143 |
\caption{Our inductive definition of POSIX values.}\label{POSIXrules} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
144 |
\end{figure} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
145 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
146 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
147 |
*} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
148 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
149 |
section {* Bitcoded Derivatives *} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
150 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
151 |
text {* |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
152 |
bitcoded regexes / decoding / bmkeps |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
153 |
gets rid of the second phase (only single phase) |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
154 |
correctness |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
155 |
*} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
156 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
157 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
158 |
section {* Simplification *} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
159 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
160 |
text {* |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
161 |
not direct correspondence with PDERs, because of example |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
162 |
problem with retrieve |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
163 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
164 |
correctness |
398 | 165 |
|
166 |
||
167 |
||
168 |
\begin{figure}[t] |
|
169 |
\begin{center} |
|
170 |
\begin{tabular}{c} |
|
171 |
@{thm[mode=Axiom] bs1}\qquad |
|
172 |
@{thm[mode=Axiom] bs2}\qquad |
|
173 |
@{thm[mode=Axiom] bs3}\\ |
|
174 |
@{thm[mode=Rule] bs4}\qquad |
|
175 |
@{thm[mode=Rule] bs5}\\ |
|
176 |
@{thm[mode=Rule] bs6}\qquad |
|
177 |
@{thm[mode=Rule] bs7}\\ |
|
178 |
@{thm[mode=Rule] bs8}\\ |
|
179 |
@{thm[mode=Axiom] ss1}\qquad |
|
180 |
@{thm[mode=Rule] ss2}\qquad |
|
181 |
@{thm[mode=Rule] ss3}\\ |
|
182 |
@{thm[mode=Axiom] ss4}\qquad |
|
183 |
@{thm[mode=Axiom] ss5}\qquad |
|
184 |
@{thm[mode=Rule] ss6}\\ |
|
185 |
\end{tabular} |
|
186 |
\end{center} |
|
187 |
\caption{???}\label{SimpRewrites} |
|
188 |
\end{figure} |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
189 |
*} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
190 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
191 |
section {* Bound - NO *} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
192 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
193 |
section {* Bounded Regex / Not *} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
194 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
195 |
section {* Conclusion *} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
196 |
|
396 | 197 |
text {* |
198 |
||
199 |
\cite{AusafDyckhoffUrban2016} |
|
200 |
||
201 |
%%\bibliographystyle{plain} |
|
202 |
\bibliography{root} |
|
203 |
*} |
|
204 |
||
205 |
(*<*) |
|
206 |
end |
|
207 |
(*>*) |