author | Christian Urban <christian.urban@kcl.ac.uk> |
Sun, 06 Feb 2022 00:02:04 +0000 | |
changeset 416 | 57182b36ec01 |
parent 410 | 9261d980225d |
child 418 | 41a2a3b63853 |
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 |
402
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
29 |
STAR ("_\<^sup>\<star>" [79] 78) and |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
30 |
|
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
31 |
val.Void ("Empty" 78) and |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
32 |
val.Char ("Char _" [1000] 78) and |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
33 |
val.Left ("Left _" [79] 78) and |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
34 |
val.Right ("Right _" [1000] 78) and |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
35 |
val.Seq ("Seq _ _" [79,79] 78) and |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
36 |
val.Stars ("Stars _" [79] 78) and |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
37 |
|
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
38 |
Posix ("'(_, _') \<rightarrow> _" [63,75,75] 75) and |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
39 |
|
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
40 |
flat ("|_|" [75] 74) and |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
41 |
flats ("|_|" [72] 74) and |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
42 |
injval ("inj _ _ _" [79,77,79] 76) and |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
43 |
mkeps ("mkeps _" [79] 76) and |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
44 |
length ("len _" [73] 73) and |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
45 |
set ("_" [73] 73) and |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
46 |
|
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
47 |
AZERO ("ZERO" 81) and |
405
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
48 |
AONE ("ONE _" [79] 78) and |
402
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
49 |
ACHAR ("CHAR _ _" [79, 79] 80) and |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
50 |
AALTs ("ALTs _ _" [77,77] 78) and |
405
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
51 |
ASEQ ("SEQ _ _ _" [79, 79,79] 78) and |
402
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
52 |
ASTAR ("STAR _ _" [79, 79] 78) and |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
53 |
|
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
54 |
code ("code _" [79] 74) and |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
55 |
intern ("_\<^latex>\<open>\\mbox{$^\\uparrow$}\<close>" [900] 80) and |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
56 |
erase ("_\<^latex>\<open>\\mbox{$^\\downarrow$}\<close>" [1000] 74) and |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
57 |
bnullable ("nullable\<^latex>\<open>\\mbox{$_b$}\<close> _" [1000] 80) and |
405
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
58 |
bmkeps ("mkeps\<^latex>\<open>\\mbox{$_b$}\<close> _" [1000] 80) and |
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
59 |
|
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
60 |
srewrite ("_\<^latex>\<open>\\mbox{$\\,\\stackrel{s}{\\leadsto}$}\<close> _" [71, 71] 80) |
402
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
61 |
|
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
62 |
|
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
63 |
lemma better_retrieve: |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
64 |
shows "rs \<noteq> Nil ==> retrieve (AALTs bs (r#rs)) (Left v) = bs @ retrieve r v" |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
65 |
and "rs \<noteq> Nil ==> retrieve (AALTs bs (r#rs)) (Right v) = bs @ retrieve (AALTs [] rs) v" |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
66 |
apply (metis list.exhaust retrieve.simps(4)) |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
67 |
by (metis list.exhaust retrieve.simps(5)) |
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
68 |
|
396 | 69 |
(*>*) |
397
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 |
section {* Introduction *} |
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 |
text {* |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
74 |
|
400 | 75 |
In the last fifteen or so years, Brzozowski's derivatives of regular |
76 |
expressions have sparked quite a bit of interest in the functional |
|
77 |
programming and theorem prover communities. The beauty of |
|
78 |
Brzozowski's derivatives \cite{Brzozowski1964} is that they are neatly |
|
79 |
expressible in any functional language, and easily definable and |
|
80 |
reasoned about in theorem provers---the definitions just consist of |
|
81 |
inductive datatypes and simple recursive functions. A mechanised |
|
82 |
correctness proof of Brzozowski's matcher in for example HOL4 has been |
|
83 |
mentioned by Owens and Slind~\cite{Owens2008}. Another one in |
|
84 |
Isabelle/HOL is part of the work by Krauss and Nipkow |
|
85 |
\cite{Krauss2011}. And another one in Coq is given by Coquand and |
|
86 |
Siles \cite{Coquand2012}. |
|
87 |
||
88 |
||
89 |
The notion of derivatives |
|
90 |
\cite{Brzozowski1964}, written @{term "der c r"}, of a regular |
|
91 |
expression give a simple solution to the problem of matching a string |
|
92 |
@{term s} with a regular expression @{term r}: if the derivative of |
|
93 |
@{term r} w.r.t.\ (in succession) all the characters of the string |
|
94 |
matches the empty string, then @{term r} matches @{term s} (and {\em |
|
95 |
vice versa}). The derivative has the property (which may almost be |
|
96 |
regarded as its specification) that, for every string @{term s} and |
|
97 |
regular expression @{term r} and character @{term c}, one has @{term |
|
98 |
"cs \<in> L(r)"} if and only if \mbox{@{term "s \<in> L(der c r)"}}. |
|
99 |
||
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
100 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
101 |
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
|
102 |
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
|
103 |
disambiguation strategies to generate a unique answer: one is called |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
104 |
GREEDY matching \cite{Frisch2004} and the other is POSIX |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
105 |
matching~\cite{POSIX,Kuklewicz,OkuiSuzuki2010,Sulzmann2014,Vansummeren2006}. |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
106 |
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
|
107 |
\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
|
108 |
matched in two `iterations' by the single letter-regular expressions |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
109 |
@{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
|
110 |
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
|
111 |
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
|
112 |
(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
|
113 |
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
|
114 |
longest match. |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
115 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
116 |
|
402
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
117 |
\begin{center} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
118 |
\begin{tabular}{cc} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
119 |
\begin{tabular}{r@ {\hspace{2mm}}c@ {\hspace{2mm}}l} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
120 |
@{thm (lhs) der.simps(1)} & $\dn$ & @{thm (rhs) der.simps(1)}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
121 |
@{thm (lhs) der.simps(2)} & $\dn$ & @{thm (rhs) der.simps(2)}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
122 |
@{thm (lhs) der.simps(3)} & $\dn$ & @{thm (rhs) der.simps(3)}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
123 |
@{thm (lhs) der.simps(4)[of c "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) der.simps(4)[of c "r\<^sub>1" "r\<^sub>2"]}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
124 |
@{thm (lhs) der.simps(5)[of c "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{text "if"} @{term "nullable(r\<^sub>1)"}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
125 |
& & @{text "then"} @{term "ALT (SEQ (der c r\<^sub>1) r\<^sub>2) (der c r\<^sub>2)"}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
126 |
& & @{text "else"} @{term "SEQ (der c r\<^sub>1) r\<^sub>2"}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
127 |
% & & @{thm (rhs) der.simps(5)[of c "r\<^sub>1" "r\<^sub>2"]}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
128 |
@{thm (lhs) der.simps(6)} & $\dn$ & @{thm (rhs) der.simps(6)} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
129 |
\end{tabular} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
130 |
& |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
131 |
\begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
132 |
@{thm (lhs) nullable.simps(1)} & $\dn$ & @{thm (rhs) nullable.simps(1)}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
133 |
@{thm (lhs) nullable.simps(2)} & $\dn$ & @{thm (rhs) nullable.simps(2)}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
134 |
@{thm (lhs) nullable.simps(3)} & $\dn$ & @{thm (rhs) nullable.simps(3)}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
135 |
@{thm (lhs) nullable.simps(4)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) nullable.simps(4)[of "r\<^sub>1" "r\<^sub>2"]}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
136 |
@{thm (lhs) nullable.simps(5)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) nullable.simps(5)[of "r\<^sub>1" "r\<^sub>2"]}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
137 |
@{thm (lhs) nullable.simps(6)} & $\dn$ & @{thm (rhs) nullable.simps(6)}\medskip\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
138 |
\end{tabular} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
139 |
\end{tabular} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
140 |
\end{center} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
141 |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
142 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
143 |
\begin{figure}[t] |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
144 |
\begin{center} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
145 |
\begin{tikzpicture}[scale=2,node distance=1.3cm, |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
146 |
every node/.style={minimum size=6mm}] |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
147 |
\node (r1) {@{term "r\<^sub>1"}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
148 |
\node (r2) [right=of r1]{@{term "r\<^sub>2"}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
149 |
\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
|
150 |
\node (r3) [right=of r2]{@{term "r\<^sub>3"}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
151 |
\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
|
152 |
\node (r4) [right=of r3]{@{term "r\<^sub>4"}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
153 |
\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
|
154 |
\draw (r4) node[anchor=west] {\;\raisebox{3mm}{@{term nullable}}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
155 |
\node (v4) [below=of r4]{@{term "v\<^sub>4"}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
156 |
\draw[->,line width=1mm](r4) -- (v4); |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
157 |
\node (v3) [left=of v4] {@{term "v\<^sub>3"}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
158 |
\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
|
159 |
\node (v2) [left=of v3]{@{term "v\<^sub>2"}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
160 |
\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
|
161 |
\node (v1) [left=of v2] {@{term "v\<^sub>1"}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
162 |
\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
|
163 |
\draw (r4) node[anchor=north west] {\;\raisebox{-8mm}{@{term "mkeps"}}}; |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
164 |
\end{tikzpicture} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
165 |
\end{center} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
166 |
\mbox{}\\[-13mm] |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
167 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
168 |
\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
|
169 |
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
|
170 |
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
|
171 |
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
|
172 |
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
|
173 |
@{term mkeps} calculates a value @{term "v\<^sub>4"} witnessing |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
174 |
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
|
175 |
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
|
176 |
the values. |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
177 |
\label{Sulz}} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
178 |
\end{figure} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
179 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
180 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
181 |
*} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
182 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
183 |
section {* Background *} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
184 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
185 |
text {* |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
186 |
Sulzmann-Lu algorithm with inj. State that POSIX rules. |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
187 |
metion slg is correct. |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
188 |
|
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 |
\begin{figure}[t] |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
191 |
\begin{center} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
192 |
\begin{tabular}{c} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
193 |
@{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
|
194 |
@{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
|
195 |
@{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
|
196 |
@{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
|
197 |
$\mprset{flushleft} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
198 |
\inferrule |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
199 |
{@{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
|
200 |
@{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
|
201 |
@{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
|
202 |
{@{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
|
203 |
@{thm[mode=Axiom] Posix.intros(7)}\<open>P[]\<close>\medskip\\ |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
204 |
$\mprset{flushleft} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
205 |
\inferrule |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
206 |
{@{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
|
207 |
@{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
|
208 |
@{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
|
209 |
@{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
|
210 |
{@{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
|
211 |
\end{tabular} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
212 |
\end{center} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
213 |
\caption{Our inductive definition of POSIX values.}\label{POSIXrules} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
214 |
\end{figure} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
215 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
216 |
|
402
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
217 |
\begin{center} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
218 |
\begin{tabular}{lcl} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
219 |
@{thm (lhs) mkeps.simps(1)} & $\dn$ & @{thm (rhs) mkeps.simps(1)}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
220 |
@{thm (lhs) mkeps.simps(2)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) mkeps.simps(2)[of "r\<^sub>1" "r\<^sub>2"]}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
221 |
@{thm (lhs) mkeps.simps(3)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) mkeps.simps(3)[of "r\<^sub>1" "r\<^sub>2"]}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
222 |
@{thm (lhs) mkeps.simps(4)} & $\dn$ & @{thm (rhs) mkeps.simps(4)}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
223 |
\end{tabular} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
224 |
\end{center} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
225 |
|
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
226 |
\begin{center} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
227 |
\begin{tabular}{l@ {\hspace{5mm}}lcl} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
228 |
\textit{(1)} & @{thm (lhs) injval.simps(1)} & $\dn$ & @{thm (rhs) injval.simps(1)}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
229 |
\textit{(2)} & @{thm (lhs) injval.simps(2)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1"]} & $\dn$ & |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
230 |
@{thm (rhs) injval.simps(2)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1"]}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
231 |
\textit{(3)} & @{thm (lhs) injval.simps(3)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]} & $\dn$ & |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
232 |
@{thm (rhs) injval.simps(3)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
233 |
\textit{(4)} & @{thm (lhs) injval.simps(4)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]} & $\dn$ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
234 |
& @{thm (rhs) injval.simps(4)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
235 |
\textit{(5)} & @{thm (lhs) injval.simps(5)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]} & $\dn$ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
236 |
& @{thm (rhs) injval.simps(5)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
237 |
\textit{(6)} & @{thm (lhs) injval.simps(6)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]} & $\dn$ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
238 |
& @{thm (rhs) injval.simps(6)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
239 |
\textit{(7)} & @{thm (lhs) injval.simps(7)[of "r" "c" "v" "vs"]} & $\dn$ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
240 |
& @{thm (rhs) injval.simps(7)[of "r" "c" "v" "vs"]}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
241 |
\end{tabular} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
242 |
\end{center} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
243 |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
244 |
*} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
245 |
|
402
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
246 |
section {* Bitcoded Regular Expressions and Derivatives *} |
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
247 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
248 |
text {* |
405
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
249 |
|
416
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
250 |
In the second part of their paper \cite{Sulzmann2014}, |
405
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
251 |
Sulzmann and Lu describe another algorithm that generates POSIX |
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
252 |
values but dispences with the second phase where characters are |
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
253 |
injected ``back'' into values. For this they annotate bitcodes to |
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
254 |
regular expressions, which we define in Isabelle/HOL as the datatype |
402
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
255 |
|
405
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
256 |
\begin{center} |
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
257 |
\begin{tabular}{lcl} |
416
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
258 |
@{term breg} & $::=$ & @{term "AZERO"} $\quad\mid\quad$ @{term "AONE bs"}\\ |
405
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
259 |
& $\mid$ & @{term "ACHAR bs c"}\\ |
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
260 |
& $\mid$ & @{term "AALTs bs rs"}\\ |
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
261 |
& $\mid$ & @{term "ASEQ bs r\<^sub>1 r\<^sub>2"}\\ |
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
262 |
& $\mid$ & @{term "ASTAR bs r"} |
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
263 |
\end{tabular} |
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
264 |
\end{center} |
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
265 |
|
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
266 |
\noindent where @{text bs} stands for a bitsequences; @{text r}, |
416
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
267 |
@{text "r\<^sub>1"} and @{text "r\<^sub>2"} for bitcoded regular |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
268 |
expressions; and @{text rs} for lists of bitcoded regular |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
269 |
expressions. The binary alternative @{text "ALT bs r\<^sub>1 r\<^sub>2"} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
270 |
is just an abbreviation for @{text "ALTs bs [r\<^sub>1, r\<^sub>2]"}. |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
271 |
For bitsequences we just use lists made up of the |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
272 |
constants @{text Z} and @{text S}. The idea with bitcoded regular |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
273 |
expressions is to incrementally generate the value information (for |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
274 |
example @{text Left} and @{text Right}) as bitsequences |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
275 |
as part of the regular expression constructors. |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
276 |
Sulzmann and Lu then define a coding |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
277 |
function for how values can be coded into bitsequences. |
402
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
278 |
|
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
279 |
\begin{center} |
416
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
280 |
\begin{tabular}{cc} |
402
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
281 |
\begin{tabular}{lcl} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
282 |
@{thm (lhs) code.simps(1)} & $\dn$ & @{thm (rhs) code.simps(1)}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
283 |
@{thm (lhs) code.simps(2)} & $\dn$ & @{thm (rhs) code.simps(2)}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
284 |
@{thm (lhs) code.simps(3)} & $\dn$ & @{thm (rhs) code.simps(3)}\\ |
416
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
285 |
@{thm (lhs) code.simps(4)} & $\dn$ & @{thm (rhs) code.simps(4)} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
286 |
\end{tabular} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
287 |
& |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
288 |
\begin{tabular}{lcl} |
402
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
289 |
@{thm (lhs) code.simps(5)[of "v\<^sub>1" "v\<^sub>2"]} & $\dn$ & @{thm (rhs) code.simps(5)[of "v\<^sub>1" "v\<^sub>2"]}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
290 |
@{thm (lhs) code.simps(6)} & $\dn$ & @{thm (rhs) code.simps(6)}\\ |
416
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
291 |
@{thm (lhs) code.simps(7)} & $\dn$ & @{thm (rhs) code.simps(7)}\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
292 |
\mbox{\phantom{XX}}\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
293 |
\end{tabular} |
402
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
294 |
\end{tabular} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
295 |
\end{center} |
416
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
296 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
297 |
\noindent |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
298 |
As can be seen, this coding is ``lossy'' in the sense that we do not |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
299 |
record explicitly character values and also not sequence values (for |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
300 |
them we just append two bitsequences). We do, however, record the |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
301 |
different alternatives for @{text Left}, respectively @{text Right}, as @{text Z} and |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
302 |
@{text S} followed by some bitsequence. Similarly, we use @{text Z} to indicate |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
303 |
if there is still a value coming in the list of @{text Stars}, whereas @{text S} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
304 |
indicates the end of the list. The lossiness makes the process of |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
305 |
decoding a bit more involved, but the point is that if we have a |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
306 |
regular expression \emph{and} a bitsequence of a corresponding value, |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
307 |
then we can always decode the value accurately. The decoding can be |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
308 |
defined by using two functions called $\textit{decode}'$ and |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
309 |
\textit{decode}: |
402
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
310 |
|
416
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
311 |
\begin{center} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
312 |
\begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
313 |
$\textit{decode}'\,bs\,(\ONE)$ & $\dn$ & $(\Empty, bs)$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
314 |
$\textit{decode}'\,bs\,(c)$ & $\dn$ & $(\Char\,c, bs)$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
315 |
$\textit{decode}'\,(\Z\!::\!bs)\;(r_1 + r_2)$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
316 |
$\textit{let}\,(v, bs_1) = \textit{decode}'\,bs\,r_1\;\textit{in}\; |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
317 |
(\Left\,v, bs_1)$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
318 |
$\textit{decode}'\,(\S\!::\!bs)\;(r_1 + r_2)$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
319 |
$\textit{let}\,(v, bs_1) = \textit{decode}'\,bs\,r_2\;\textit{in}\; |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
320 |
(\Right\,v, bs_1)$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
321 |
$\textit{decode}'\,bs\;(r_1\cdot r_2)$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
322 |
$\textit{let}\,(v_1, bs_1) = \textit{decode}'\,bs\,r_1\;\textit{in}$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
323 |
& & $\textit{let}\,(v_2, bs_2) = \textit{decode}'\,bs_1\,r_2$ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
324 |
\hspace{2mm}$\textit{in}\;(\Seq\,v_1\,v_2, bs_2)$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
325 |
$\textit{decode}'\,(\Z\!::\!bs)\,(r^*)$ & $\dn$ & $(\Stars\,[], bs)$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
326 |
$\textit{decode}'\,(\S\!::\!bs)\,(r^*)$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
327 |
$\textit{let}\,(v, bs_1) = \textit{decode}'\,bs\,r\;\textit{in}$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
328 |
& & $\textit{let}\,(\Stars\,vs, bs_2) = \textit{decode}'\,bs_1\,r^*$ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
329 |
\hspace{2mm}$\textit{in}\;(\Stars\,v\!::\!vs, bs_2)$\bigskip\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
330 |
$\textit{decode}\,bs\,r$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
331 |
$\textit{let}\,(v, bs') = \textit{decode}'\,bs\,r\;\textit{in}$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
332 |
& & \hspace{7mm}$\textit{if}\;bs' = []\;\textit{then}\;\textit{Some}\,v\; |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
333 |
\textit{else}\;\textit{None}$ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
334 |
\end{tabular} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
335 |
\end{center} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
336 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
337 |
\noindent |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
338 |
The function \textit{decode} checks whether all of the bitsequence is |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
339 |
consumed and returns the corresponding value as @{term "Some v"}; otherwise |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
340 |
it fails with @{text "None"}. We can establish that for a value $v$ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
341 |
inhabited by a regular expression $r$, the decoding of its |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
342 |
bitsequence never fails. |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
343 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
344 |
\begin{lemma}\label{codedecode}\it |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
345 |
If $\;\vdash v : r$ then |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
346 |
$\;\textit{decode}\,(\textit{code}\, v)\,r = \textit{Some}\, v$. |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
347 |
\end{lemma} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
348 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
349 |
\begin{proof} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
350 |
This follows from the property that |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
351 |
$\textit{decode}'\,((\textit{code}\,v) \,@\, bs)\,r = (v, bs)$ holds |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
352 |
for any bit-sequence $bs$ and $\vdash v : r$. This property can be |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
353 |
easily proved by induction on $\vdash v : r$. |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
354 |
\end{proof} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
355 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
356 |
Sulzmann and Lu define the function \emph{internalise} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
357 |
in order to transform standard regular expressions into annotated |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
358 |
regular expressions. We write this operation as $r^\uparrow$. |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
359 |
This internalisation uses the following |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
360 |
\emph{fuse} function. |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
361 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
362 |
\begin{center} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
363 |
\begin{tabular}{lcl} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
364 |
$\textit{fuse}\,bs\,(\textit{ZERO})$ & $\dn$ & $\textit{ZERO}$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
365 |
$\textit{fuse}\,bs\,(\textit{ONE}\,bs')$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
366 |
$\textit{ONE}\,(bs\,@\,bs')$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
367 |
$\textit{fuse}\,bs\,(\textit{CHAR}\,bs'\,c)$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
368 |
$\textit{CHAR}\,(bs\,@\,bs')\,c$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
369 |
$\textit{fuse}\,bs\,(\textit{ALTs}\,bs'\,rs)$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
370 |
$\textit{ALTs}\,(bs\,@\,bs')\,rs$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
371 |
$\textit{fuse}\,bs\,(\textit{SEQ}\,bs'\,r_1\,r_2)$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
372 |
$\textit{SEQ}\,(bs\,@\,bs')\,r_1\,r_2$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
373 |
$\textit{fuse}\,bs\,(\textit{STAR}\,bs'\,r)$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
374 |
$\textit{STAR}\,(bs\,@\,bs')\,r$ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
375 |
\end{tabular} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
376 |
\end{center} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
377 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
378 |
\noindent |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
379 |
A regular expression can then be \emph{internalised} into a bitcoded |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
380 |
regular expression as follows. |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
381 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
382 |
\begin{center} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
383 |
\begin{tabular}{lcl} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
384 |
$(\ZERO)^\uparrow$ & $\dn$ & $\textit{ZERO}$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
385 |
$(\ONE)^\uparrow$ & $\dn$ & $\textit{ONE}\,[]$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
386 |
$(c)^\uparrow$ & $\dn$ & $\textit{CHAR}\,[]\,c$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
387 |
$(r_1 + r_2)^\uparrow$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
388 |
$\textit{ALT}\;[]\,(\textit{fuse}\,[\Z]\,r_1^\uparrow)\, |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
389 |
(\textit{fuse}\,[\S]\,r_2^\uparrow)$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
390 |
$(r_1\cdot r_2)^\uparrow$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
391 |
$\textit{SEQ}\;[]\,r_1^\uparrow\,r_2^\uparrow$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
392 |
$(r^*)^\uparrow$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
393 |
$\textit{STAR}\;[]\,r^\uparrow$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
394 |
\end{tabular} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
395 |
\end{center} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
396 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
397 |
\noindent |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
398 |
There is also an \emph{erase}-function, written $a^\downarrow$, which |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
399 |
transforms a bitcoded regular expression into a (standard) regular |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
400 |
expression by just erasing the annotated bitsequences. We omit the |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
401 |
straightforward definition. For defining the algorithm, we also need |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
402 |
the functions \textit{bnullable} and \textit{bmkeps}, which are the |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
403 |
``lifted'' versions of \textit{nullable} and \textit{mkeps} acting on |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
404 |
bitcoded regular expressions, instead of regular expressions. |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
405 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
406 |
\begin{center} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
407 |
\begin{tabular}{lcl} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
408 |
$\textit{bnullable}\,(\textit{ZERO})$ & $\dn$ & $\textit{false}$ \textbf{fix}\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
409 |
$\textit{bnullable}\,(\textit{ONE}\,bs)$ & $\dn$ & $\textit{true}$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
410 |
$\textit{bnullable}\,(\textit{CHAR}\,bs\,c)$ & $\dn$ & $\textit{false}$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
411 |
$\textit{bnullable}\,(\textit{ALT}\,bs\,a_1\,a_2)$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
412 |
$\textit{bnullable}\,a_1\vee \textit{bnullable}\,a_2$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
413 |
$\textit{bnullable}\,(\textit{SEQ}\,bs\,a_1\,a_2)$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
414 |
$\textit{bnullable}\,a_1\wedge \textit{bnullable}\,a_2$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
415 |
$\textit{bnullable}\,(\textit{STAR}\,bs\,a)$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
416 |
$\textit{true}$ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
417 |
\end{tabular} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
418 |
\end{center} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
419 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
420 |
\begin{center} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
421 |
\begin{tabular}{lcl} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
422 |
$\textit{bmkeps}\,(\textit{ONE}\,bs)$ & $\dn$ & $bs$ \textbf{fix}\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
423 |
$\textit{bmkeps}\,(\textit{ALT}\,bs\,a_1\,a_2)$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
424 |
$\textit{if}\;\textit{bnullable}\,a_1$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
425 |
& &$\textit{then}\;bs\,@\,\textit{bmkeps}\,a_1$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
426 |
& &$\textit{else}\;bs\,@\,\textit{bmkeps}\,a_2$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
427 |
$\textit{bmkeps}\,(\textit{SEQ}\,bs\,a_1\,a_2)$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
428 |
$bs \,@\,\textit{bmkeps}\,a_1\,@\, \textit{bmkeps}\,a_2$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
429 |
$\textit{bmkeps}\,(\textit{STAR}\,bs\,a)$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
430 |
$bs \,@\, [\S]$ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
431 |
\end{tabular} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
432 |
\end{center} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
433 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
434 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
435 |
\noindent |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
436 |
The key function in the bitcoded algorithm is the derivative of an |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
437 |
annotated regular expression. This derivative calculates the |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
438 |
derivative but at the same time also the incremental part that |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
439 |
contributes to constructing a value. |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
440 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
441 |
\begin{center} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
442 |
\begin{tabular}{@ {}lcl@ {}} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
443 |
$(\textit{ZERO})\backslash c$ & $\dn$ & $\textit{ZERO}$ \textbf{fix}\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
444 |
$(\textit{ONE}\;bs)\backslash c$ & $\dn$ & $\textit{ZERO}$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
445 |
$(\textit{CHAR}\;bs\,d)\backslash c$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
446 |
$\textit{if}\;c=d\; \;\textit{then}\; |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
447 |
\textit{ONE}\;bs\;\textit{else}\;\textit{ZERO}$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
448 |
$(\textit{ALT}\;bs\,a_1\,a_2)\backslash c$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
449 |
$\textit{ALT}\,bs\,(a_1\backslash c)\,(a_2\backslash c)$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
450 |
$(\textit{SEQ}\;bs\,a_1\,a_2)\backslash c$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
451 |
$\textit{if}\;\textit{bnullable}\,a_1$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
452 |
& &$\textit{then}\;\textit{ALT}\,bs\,(\textit{SEQ}\,[]\,(a_1\backslash c)\,a_2)$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
453 |
& &$\phantom{\textit{then}\;\textit{ALT}\,bs\,}(\textit{fuse}\,(\textit{bmkeps}\,a_1)\,(a_2\backslash c))$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
454 |
& &$\textit{else}\;\textit{SEQ}\,bs\,(a_1\backslash c)\,a_2$\\ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
455 |
$(\textit{STAR}\,bs\,a)\backslash c$ & $\dn$ & |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
456 |
$\textit{SEQ}\;bs\,(\textit{fuse}\, [\Z] (r\backslash c))\, |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
457 |
(\textit{STAR}\,[]\,r)$ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
458 |
\end{tabular} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
459 |
\end{center} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
460 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
461 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
462 |
\noindent |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
463 |
This function can also be extended to strings, written $a\backslash s$, |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
464 |
just like the standard derivative. We omit the details. Finally we |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
465 |
can define Sulzmann and Lu's bitcoded lexer, which we call \textit{blexer}: |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
466 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
467 |
\noindent |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
468 |
This bitcoded lexer first internalises the regular expression $r$ and then |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
469 |
builds the annotated derivative according to $s$. If the derivative is |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
470 |
nullable, then it extracts the bitcoded value using the |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
471 |
$\textit{bmkeps}$ function. Finally it decodes the bitcoded value. If |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
472 |
the derivative is \emph{not} nullable, then $\textit{None}$ is |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
473 |
returned. The task is to show that this way of calculating a value |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
474 |
generates the same result as with \textit{lexer}. |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
475 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
476 |
Before we can proceed we need to define a function, called |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
477 |
\textit{retrieve}, which Sulzmann and Lu introduced for the proof. |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
478 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
479 |
\textbf{fix} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
480 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
481 |
\noindent |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
482 |
The idea behind this function is to retrieve a possibly partial |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
483 |
bitcode from an annotated regular expression, where the retrieval is |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
484 |
guided by a value. For example if the value is $\Left$ then we |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
485 |
descend into the left-hand side of an alternative (annotated) regular |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
486 |
expression in order to assemble the bitcode. Similarly for |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
487 |
$\Right$. The property we can show is that for a given $v$ and $r$ |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
488 |
with $\vdash v : r$, the retrieved bitsequence from the internalised |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
489 |
regular expression is equal to the bitcoded version of $v$. |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
490 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
491 |
\begin{lemma}\label{retrievecode} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
492 |
If $\vdash v : r$ then $\textit{code}\, v = \textit{retrieve}\,(r^\uparrow)\,v$. |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
493 |
\end{lemma} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
494 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
495 |
*} |
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
496 |
|
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
497 |
text {* |
405
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
498 |
There is also a corresponding decoding function that takes a bitsequence |
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
499 |
and generates back a value. However, since the bitsequences are a ``lossy'' |
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
500 |
coding (@{term Seq}s are not coded) the decoding function depends also |
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
501 |
on a regular expression in order to decode values. |
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
502 |
|
416
57182b36ec01
more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
410
diff
changeset
|
503 |
|
405
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
504 |
|
402
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
505 |
|
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
506 |
The idea of the bitcodes is to annotate them to regular expressions and generate values |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
507 |
incrementally. The bitcodes can be read off from the @{text breg} and then decoded into a value. |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
508 |
|
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
509 |
\begin{center} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
510 |
\begin{tabular}{lcl} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
511 |
@{term breg} & $::=$ & @{term "AZERO"}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
512 |
& $\mid$ & @{term "AONE bs"}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
513 |
& $\mid$ & @{term "ACHAR bs c"}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
514 |
& $\mid$ & @{term "AALTs bs rs"}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
515 |
& $\mid$ & @{term "ASEQ bs r\<^sub>1 r\<^sub>2"}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
516 |
& $\mid$ & @{term "ASTAR bs r"} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
517 |
\end{tabular} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
518 |
\end{center} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
519 |
|
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
520 |
|
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
521 |
|
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
522 |
\begin{center} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
523 |
\begin{tabular}{lcl} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
524 |
@{thm (lhs) retrieve.simps(1)} & $\dn$ & @{thm (rhs) retrieve.simps(1)}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
525 |
@{thm (lhs) retrieve.simps(2)} & $\dn$ & @{thm (rhs) retrieve.simps(2)}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
526 |
@{thm (lhs) retrieve.simps(3)} & $\dn$ & @{thm (rhs) retrieve.simps(3)}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
527 |
@{thm (lhs) better_retrieve(1)} & $\dn$ & @{thm (rhs) better_retrieve(1)}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
528 |
@{thm (lhs) better_retrieve(2)} & $\dn$ & @{thm (rhs) better_retrieve(2)}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
529 |
@{thm (lhs) retrieve.simps(6)[of _ "r\<^sub>1" "r\<^sub>2" "v\<^sub>1" "v\<^sub>2"]} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
530 |
& $\dn$ & @{thm (rhs) retrieve.simps(6)[of _ "r\<^sub>1" "r\<^sub>2" "v\<^sub>1" "v\<^sub>2"]}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
531 |
@{thm (lhs) retrieve.simps(7)} & $\dn$ & @{thm (rhs) retrieve.simps(7)}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
532 |
@{thm (lhs) retrieve.simps(8)} & $\dn$ & @{thm (rhs) retrieve.simps(8)} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
533 |
\end{tabular} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
534 |
\end{center} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
535 |
|
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
536 |
|
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
537 |
\begin{theorem} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
538 |
@{thm blexer_correctness} |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
539 |
\end{theorem} |
405
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
540 |
|
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
541 |
|
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
542 |
bitcoded regexes / decoding / bmkeps |
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
543 |
gets rid of the second phase (only single phase) |
3cfea5bb5e23
updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents:
402
diff
changeset
|
544 |
correctness |
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
545 |
*} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
546 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
547 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
548 |
section {* Simplification *} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
549 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
550 |
text {* |
400 | 551 |
Sulzmann \& Lu apply simplification via a fixpoint operation; also does not use erase to filter out duplicates. |
402
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
552 |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
553 |
not direct correspondence with PDERs, because of example |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
554 |
problem with retrieve |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
555 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
556 |
correctness |
398 | 557 |
|
402
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
558 |
|
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
559 |
|
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
560 |
|
398 | 561 |
|
562 |
||
563 |
\begin{figure}[t] |
|
564 |
\begin{center} |
|
565 |
\begin{tabular}{c} |
|
402
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
566 |
@{thm[mode=Axiom] bs1[of _ "r\<^sub>2"]}\qquad |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
567 |
@{thm[mode=Axiom] bs2[of _ "r\<^sub>1"]}\qquad |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
568 |
@{thm[mode=Axiom] bs3[of "bs\<^sub>1" "bs\<^sub>2"]}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
569 |
@{thm[mode=Rule] bs4[of "r\<^sub>1" "r\<^sub>2" _ "r\<^sub>3"]}\qquad |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
570 |
@{thm[mode=Rule] bs5[of "r\<^sub>3" "r\<^sub>4" _ "r\<^sub>1"]}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
571 |
@{thm[mode=Axiom] bs6}\qquad |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
572 |
@{thm[mode=Axiom] bs7}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
573 |
@{thm[mode=Rule] bs8[of "rs\<^sub>1" "rs\<^sub>2"]}\\ |
410 | 574 |
%@ { t hm[mode=Axiom] ss1}\qquad |
402
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
575 |
@{thm[mode=Rule] ss2[of "rs\<^sub>1" "rs\<^sub>2"]}\qquad |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
576 |
@{thm[mode=Rule] ss3[of "r\<^sub>1" "r\<^sub>2"]}\\ |
398 | 577 |
@{thm[mode=Axiom] ss4}\qquad |
402
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
578 |
@{thm[mode=Axiom] ss5[of "bs" "rs\<^sub>1" "rs\<^sub>2"]}\\ |
1612f2a77bf6
more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
400
diff
changeset
|
579 |
@{thm[mode=Rule] ss6[of "r\<^sub>1" "r\<^sub>2" "rs\<^sub>1" "rs\<^sub>2" "rs\<^sub>3"]}\\ |
398 | 580 |
\end{tabular} |
581 |
\end{center} |
|
582 |
\caption{???}\label{SimpRewrites} |
|
583 |
\end{figure} |
|
397
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
584 |
*} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
585 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
586 |
section {* Bound - NO *} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
587 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
588 |
section {* Bounded Regex / Not *} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
589 |
|
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
590 |
section {* Conclusion *} |
e1b74d618f1b
updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents:
396
diff
changeset
|
591 |
|
396 | 592 |
text {* |
593 |
||
594 |
\cite{AusafDyckhoffUrban2016} |
|
595 |
||
596 |
%%\bibliographystyle{plain} |
|
597 |
\bibliography{root} |
|
598 |
*} |
|
599 |
||
600 |
(*<*) |
|
601 |
end |
|
602 |
(*>*) |