thys2/Paper/Paper.thy
author Christian Urban <christian.urban@kcl.ac.uk>
Wed, 02 Feb 2022 14:52:41 +0000
changeset 405 3cfea5bb5e23
parent 402 1612f2a77bf6
child 410 9261d980225d
permissions -rw-r--r--
updated some of the text and cardinality proof
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
396
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
     1
(*<*)
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
     2
theory Paper
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
     3
imports 
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
     4
   "../Lexer"
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
     5
   "../Simplifying" 
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
     6
   "../Positions"
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
     7
   "../SizeBound4" 
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
     8
   "HOL-Library.LaTeXsugar"
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
     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
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
    13
notation (latex output)
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
    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
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
    15
  Cons ("_\<^latex>\<open>\\mbox{$\\,$}\<close>::\<^latex>\<open>\\mbox{$\\,$}\<close>_" [75,73] 73) 
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
    16
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
    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
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
    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
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    75
In the last fifteen or so years, Brzozowski's derivatives of regular
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    76
expressions have sparked quite a bit of interest in the functional
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    77
programming and theorem prover communities.  The beauty of
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    78
Brzozowski's derivatives \cite{Brzozowski1964} is that they are neatly
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    79
expressible in any functional language, and easily definable and
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    80
reasoned about in theorem provers---the definitions just consist of
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    81
inductive datatypes and simple recursive functions. A mechanised
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    82
correctness proof of Brzozowski's matcher in for example HOL4 has been
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    83
mentioned by Owens and Slind~\cite{Owens2008}. Another one in
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    84
Isabelle/HOL is part of the work by Krauss and Nipkow
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    85
\cite{Krauss2011}.  And another one in Coq is given by Coquand and
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    86
Siles \cite{Coquand2012}.
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    87
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    88
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    89
The notion of derivatives
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    90
\cite{Brzozowski1964}, written @{term "der c r"}, of a regular
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    91
expression give a simple solution to the problem of matching a string
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    92
@{term s} with a regular expression @{term r}: if the derivative of
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    93
@{term r} w.r.t.\ (in succession) all the characters of the string
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    94
matches the empty string, then @{term r} matches @{term s} (and {\em
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    95
vice versa}). The derivative has the property (which may almost be
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    96
regarded as its specification) that, for every string @{term s} and
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    97
regular expression @{term r} and character @{term c}, one has @{term
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    98
"cs \<in> L(r)"} if and only if \mbox{@{term "s \<in> L(der c r)"}}.
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    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
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   250
  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
   251
  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
   252
  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
   253
  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
   254
405
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   255
  \begin{center}
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   256
  \begin{tabular}{lcl}
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   257
  @{term breg} & $::=$ & @{term "AZERO"}\\
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   258
               & $\mid$ & @{term "AONE bs"}\\
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},
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   267
  @{text "r\<^sub>1"} and @{text "r\<^sub>2"} for annotated regular
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   268
  expressions; and @{text rs} for a list of annotated regular
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   269
  expressions.  In contrast to Sulzmann and Lu we generalise the
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   270
  alternative regular expressions to lists, instead of just having
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   271
  binary regular expressions.  The idea with annotated regular
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   272
  expressions is to incrementally generate the value information by
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   273
  recording bitsequences. Sulzmann and Lu then  
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   274
  define a coding 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
   275
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   276
  \begin{center}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   277
  \begin{tabular}{lcl}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   278
  @{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
   279
  @{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
   280
  @{thm (lhs) code.simps(3)} & $\dn$ & @{thm (rhs) code.simps(3)}\\
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   281
  @{thm (lhs) code.simps(4)} & $\dn$ & @{thm (rhs) code.simps(4)}\\
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   282
  @{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
   283
  @{thm (lhs) code.simps(6)} & $\dn$ & @{thm (rhs) code.simps(6)}\\
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   284
  @{thm (lhs) code.simps(7)} & $\dn$ & @{thm (rhs) code.simps(7)}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   285
  \end{tabular}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   286
  \end{center}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   287
405
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   288
  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
   289
  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
   290
  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
   291
  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
   292
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   293
  \begin{center}
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   294
  \begin{tabular}{lcll}
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   295
  %@{thm (lhs) decode'.simps(1)} & $\dn$ & @{thm (rhs) decode'.simps(1)}\\
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   296
  @{thm (lhs) decode'.simps(2)} & $\dn$ & @{thm (rhs) decode'.simps(2)}\\
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   297
  @{thm (lhs) decode'.simps(3)} & $\dn$ & @{thm (rhs) decode'.simps(3)}\\
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   298
  @{thm (lhs) decode'.simps(4)} & $\dn$ & @{thm (rhs) decode'.simps(4)}\\
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   299
  @{thm (lhs) decode'.simps(5)} & $\dn$ & @{thm (rhs) decode'.simps(5)}\\
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   300
  @{thm (lhs) decode'.simps(6)} & $\dn$ & @{thm (rhs) decode'.simps(6)}\\
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   301
  @{thm (lhs) decode'.simps(7)} & $\dn$ & @{thm (rhs) decode'.simps(7)}\\
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   302
  @{thm (lhs) decode'.simps(8)} & $\dn$ & @{thm (rhs) decode'.simps(8)}\\
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   303
  @{thm (lhs) decode'.simps(9)} & $\dn$ & @{thm (rhs) decode'.simps(9)}\\
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   304
  @{thm (lhs) decode'.simps(10)} & $\dn$ & @{thm (rhs) decode'.simps(10)} & fix
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   305
  \end{tabular}
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   306
  \end{center}
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   307
402
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   308
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   309
  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
   310
  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
   311
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   312
  \begin{center}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   313
  \begin{tabular}{lcl}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   314
  @{term breg} & $::=$ & @{term "AZERO"}\\
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   315
               & $\mid$ & @{term "AONE bs"}\\
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   316
               & $\mid$ & @{term "ACHAR bs c"}\\
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   317
               & $\mid$ & @{term "AALTs bs rs"}\\
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   318
               & $\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
   319
               & $\mid$ & @{term "ASTAR bs r"}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   320
  \end{tabular}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   321
  \end{center}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   322
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   323
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   324
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   325
  \begin{center}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   326
  \begin{tabular}{lcl}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   327
  @{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
   328
  @{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
   329
  @{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
   330
  @{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
   331
  @{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
   332
  @{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
   333
      & $\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
   334
  @{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
   335
  @{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
   336
  \end{tabular}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   337
  \end{center}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   338
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   339
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   340
  \begin{theorem}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   341
  @{thm blexer_correctness} 
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   342
  \end{theorem}
405
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   343
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   344
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   345
  bitcoded regexes / decoding / bmkeps
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   346
  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
   347
  correctness
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   348
*}
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   349
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   350
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   351
section {* Simplification *}
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   352
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   353
text {*
400
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
   354
     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
   355
  
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   356
   not direct correspondence with PDERs, because of example
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   357
   problem with retrieve 
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   358
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   359
   correctness
398
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
   360
402
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   361
   
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   362
    
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   363
398
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
   364
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
   365
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
   366
   \begin{figure}[t]
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
   367
  \begin{center}
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
   368
  \begin{tabular}{c}
402
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   369
  @{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
   370
  @{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
   371
  @{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
   372
  @{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
   373
  @{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
   374
  @{thm[mode=Axiom] bs6}\qquad
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   375
  @{thm[mode=Axiom] bs7}\\
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   376
  @{thm[mode=Rule] bs8[of "rs\<^sub>1" "rs\<^sub>2"]}\\
398
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
   377
  @{thm[mode=Axiom] ss1}\qquad
402
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   378
  @{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
   379
  @{thm[mode=Rule] ss3[of "r\<^sub>1" "r\<^sub>2"]}\\
398
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
   380
  @{thm[mode=Axiom] ss4}\qquad
402
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   381
  @{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
   382
  @{thm[mode=Rule] ss6[of "r\<^sub>1" "r\<^sub>2" "rs\<^sub>1" "rs\<^sub>2" "rs\<^sub>3"]}\\
398
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
   383
  \end{tabular}
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
   384
  \end{center}
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
   385
  \caption{???}\label{SimpRewrites}
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
   386
  \end{figure}
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   387
*}
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   388
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   389
section {* Bound - NO *}
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   390
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   391
section {* Bounded Regex / Not *}
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   392
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   393
section {* Conclusion *}
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   394
396
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
   395
text {*
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
   396
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
   397
\cite{AusafDyckhoffUrban2016}
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
   398
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
   399
%%\bibliographystyle{plain}
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
   400
\bibliography{root}
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
   401
*}
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
   402
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
   403
(*<*)
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
   404
end
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
   405
(*>*)