thys2/Paper/Paper.thy
author Christian Urban <christian.urban@kcl.ac.uk>
Wed, 09 Feb 2022 00:29:04 +0000
changeset 424 2416fdec6396
parent 423 b7199d6c672d
child 425 14c558ae0b09
permissions -rw-r--r--
updated
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"
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
    20
abbreviation 
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
    21
 "ders_syn r s \<equiv> ders s r"  
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
    22
abbreviation 
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
    23
  "bder_syn r c \<equiv> bder c r"  
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
    24
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
    25
notation (latex output)
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
    26
  der_syn ("_\\_" [79, 1000] 76) and
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
    27
  ders_syn ("_\\_" [79, 1000] 76) and
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
    28
  bder_syn ("_\\_" [79, 1000] 76) and
420
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
    29
  bders ("_\\_" [79, 1000] 76) and
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
    30
  bders_simp ("_\\\<^sub>s\<^sub>i\<^sub>m\<^sub>p _" [79, 1000] 76) and
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
    31
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
    32
  ZERO ("\<^bold>0" 81) and 
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
    33
  ONE ("\<^bold>1" 81) and 
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
    34
  CH ("_" [1000] 80) and
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
    35
  ALT ("_ + _" [77,77] 78) and
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
    36
  SEQ ("_ \<cdot> _" [77,77] 78) and
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
    37
  STAR ("_\<^sup>*" [79] 78) and
402
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    38
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    39
  val.Void ("Empty" 78) and
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    40
  val.Char ("Char _" [1000] 78) and
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    41
  val.Left ("Left _" [79] 78) and
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    42
  val.Right ("Right _" [1000] 78) and
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    43
  val.Seq ("Seq _ _" [79,79] 78) and
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    44
  val.Stars ("Stars _" [79] 78) and
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    45
420
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
    46
  Prf ("\<turnstile> _ : _" [75,75] 75) and  
402
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    47
  Posix ("'(_, _') \<rightarrow> _" [63,75,75] 75) and
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    48
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    49
  flat ("|_|" [75] 74) and
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    50
  flats ("|_|" [72] 74) and
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    51
  injval ("inj _ _ _" [79,77,79] 76) and 
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    52
  mkeps ("mkeps _" [79] 76) and 
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    53
  length ("len _" [73] 73) and
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    54
  set ("_" [73] 73) and
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    55
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    56
  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
    57
  AONE ("ONE _" [79] 78) and 
402
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    58
  ACHAR ("CHAR _ _" [79, 79] 80) and
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    59
  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
    60
  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
    61
  ASTAR ("STAR _ _" [79, 79] 78) and
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
  code ("code _" [79] 74) and
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    64
  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
    65
  erase ("_\<^latex>\<open>\\mbox{$^\\downarrow$}\<close>" [1000] 74) and
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
    66
  bnullable ("bnullable _" [1000] 80) and
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
    67
  bsimp_AALTs ("bsimpALT _ _" [10,1000] 80) and
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
    68
  bsimp_ASEQ ("bsimpSEQ _ _ _" [10,1000,1000] 80) and
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
    69
  bmkeps ("bmkeps _" [1000] 80) and
405
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
    70
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
    71
  srewrite ("_\<^latex>\<open>\\mbox{$\\,\\stackrel{s}{\\leadsto}$}\<close> _" [71, 71] 80) and
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
    72
  rrewrites ("_ \<^latex>\<open>\\mbox{$\\,\\leadsto^*$}\<close> _" [71, 71] 80) and
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
    73
  blexer_simp ("blexer\<^sup>+" 1000) 
402
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    74
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    75
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    76
lemma better_retrieve:
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    77
   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
    78
   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
    79
  apply (metis list.exhaust retrieve.simps(4))
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
    80
  by (metis list.exhaust retrieve.simps(5))
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
    81
396
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
    82
(*>*)
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
    83
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
    84
section {* Introduction *}
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
    85
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
    86
text {*
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
    87
400
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    88
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
    89
expressions have sparked quite a bit of interest in the functional
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    90
programming and theorem prover communities.  The beauty of
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    91
Brzozowski's derivatives \cite{Brzozowski1964} is that they are neatly
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    92
expressible in any functional language, and easily definable and
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
    93
reasoned about in theorem provers---the definitions just consist of
420
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
    94
inductive datatypes and simple recursive functions.  Derivatives of a
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
    95
regular expression, written @{term "der c r"}, give a simple solution
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
    96
to the problem of matching a string @{term s} with a regular
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
    97
expression @{term r}: if the derivative of @{term r} w.r.t.\ (in
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
    98
succession) all the characters of the string matches the empty string,
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
    99
then @{term r} matches @{term s} (and {\em vice versa}).  We are aware
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   100
of a mechanised correctness proof of Brzozowski's derivative-based matcher in HOL4 by
420
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   101
Owens and Slind~\cite{Owens2008}. Another one in Isabelle/HOL is part
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   102
of the work by Krauss and Nipkow \cite{Krauss2011}.  And another one
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   103
in Coq is given by Coquand and Siles \cite{Coquand2012}.
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   104
Also Ribeiro and Du Bois give one in Agda \cite{RibeiroAgda2017}.
420
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   105
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   106
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   107
However, there are two difficulties with derivative-based matchers:
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   108
First, Brzozowski's original matcher only generates a yes/no answer
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   109
for whether a regular expression matches a string or not.  This is too
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   110
little information in the context of lexing where separate tokens must
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   111
be identified and also classified (for example as keywords
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   112
or identifiers).  Sulzmann and Lu~\cite{Sulzmann2014} overcome this
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   113
difficulty by cleverly extending Brzozowski's matching
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   114
algorithm. Their extended version generates additional information on
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   115
\emph{how} a regular expression matches a string following the POSIX
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   116
rules for regular expression matching. They achieve this by adding a
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   117
second ``phase'' to Brzozowski's algorithm involving an injection
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   118
function.  In our own earlier work we provided the formal
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   119
specification of what POSIX matching means and proved in Isabelle/HOL
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   120
the correctness
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   121
of Sulzmann and Lu's extended algorithm accordingly
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   122
\cite{AusafDyckhoffUrban2016}.
400
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
   123
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   124
The second difficulty is that Brzozowski's derivatives can 
420
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   125
grow to arbitrarily big sizes. For example if we start with the
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   126
regular expression \mbox{@{text "(a + aa)\<^sup>*"}} and take
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   127
successive derivatives according to the character $a$, we end up with
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   128
a sequence of ever-growing derivatives like 
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   129
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   130
\def\ll{\stackrel{\_\backslash{} a}{\longrightarrow}}
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   131
\begin{center}
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   132
\begin{tabular}{rll}
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   133
$(a + aa)^*$ & $\ll$ & $(\ONE + \ONE{}a) \cdot (a + aa)^*$\\
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   134
& $\ll$ & $(\ZERO + \ZERO{}a + \ONE) \cdot (a + aa)^* \;+\; (\ONE + \ONE{}a) \cdot (a + aa)^*$\\
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   135
& $\ll$ & $(\ZERO + \ZERO{}a + \ZERO) \cdot (a + aa)^* + (\ONE + \ONE{}a) \cdot (a + aa)^* \;+\; $\\
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   136
& & $\qquad(\ZERO + \ZERO{}a + \ONE) \cdot (a + aa)^* + (\ONE + \ONE{}a) \cdot (a + aa)^*$\\
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   137
& $\ll$ & \ldots \hspace{15mm}(regular expressions of sizes 98, 169, 283, 468, 767, \ldots)
420
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   138
\end{tabular}
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   139
\end{center}
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   140
 
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   141
\noindent where after around 35 steps we run out of memory on a
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   142
typical computer (we shall define shortly the precise details of our
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   143
regular expressions and the derivative operation).  Clearly, the
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   144
notation involving $\ZERO$s and $\ONE$s already suggests
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   145
simplification rules that can be applied to regular regular
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   146
expressions, for example $\ZERO{}\,r \Rightarrow \ZERO$, $\ONE{}\,r
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   147
\Rightarrow r$, $\ZERO{} + r \Rightarrow r$ and $r + r \Rightarrow
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   148
r$. While such simple-minded simplifications have been proved in our
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   149
earlier work to preserve the correctness of Sulzmann and Lu's
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   150
algorithm \cite{AusafDyckhoffUrban2016}, they unfortunately do
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   151
\emph{not} help with limiting the growth of the derivatives shown
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   152
above: the growth is slowed, but the derivatives can still grow rather
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   153
quickly beyond any finite bound.
420
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   154
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   155
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   156
Sulzmann and Lu overcome this ``growth problem'' in a second algorithm
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   157
\cite{Sulzmann2014} where they introduce bitcoded
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   158
regular expressions. In this version, POSIX values are
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   159
represented as bitsequences and such sequences are incrementally generated
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   160
when derivatives are calculated. The compact representation
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   161
of bitsequences and regular expressions allows them to define a more
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   162
``aggressive'' simplification method that keeps the size of the
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   163
derivatives finite no matter what the length of the string is.
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   164
They make some informal claims about the correctness and linear behaviour
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   165
of this version, but do not provide any supporting proof arguments, not
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   166
even ``pencil-and-paper'' arguments. They write about their bitcoded
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   167
\emph{incremental parsing method} (that is the algorithm to be formalised
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   168
in this paper):
420
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   169
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   170
\begin{quote}\it
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   171
  ``Correctness Claim: We further claim that the incremental parsing
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   172
  method [..] in combination with the simplification steps [..]
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   173
  yields POSIX parse trees. We have tested this claim
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   174
  extensively [..] but yet
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   175
  have to work out all proof details.''
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   176
\end{quote}  
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   177
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   178
\noindent{}\textbf{Contributions:} We have implemented in Isabelle/HOL
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   179
the derivative-based lexing algorithm of Sulzmann and Lu
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   180
\cite{Sulzmann2014} where regular expressions are annotated with
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   181
bitsequences. We define the crucial simplification function as a
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   182
recursive function, instead of a fix-point operation. One objective of
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   183
the simplification function is to remove duplicates of regular
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   184
expressions.  For this Sulzmann and Lu use in their paper the standard
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   185
@{text nub} function from Haskell's list library, but this function
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   186
does not achieve the intended objective with bitcoded regular expressions.  The
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   187
reason is that in the bitcoded setting, each copy generally has a
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   188
different bitcode annotation---so @{text nub} would never ``fire''.
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   189
Inspired by Scala's library for lists, we shall instead use a @{text
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   190
distinctBy} function that finds duplicates under an erasing function
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   191
that deletes bitcodes.
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   192
We shall also introduce our own argument and definitions for
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   193
establishing the correctness of the bitcoded algorithm when 
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   194
simplifications are included.\medskip
420
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   195
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   196
\noindent In this paper, we shall first briefly introduce the basic notions
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   197
of regular expressions and describe the basic definitions
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   198
of POSIX lexing from our earlier work \cite{AusafDyckhoffUrban2016}. This serves
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   199
as a reference point for what correctness means in our Isabelle/HOL proofs. We shall then prove
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   200
the correctness for the bitcoded algorithm without simplification, and
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   201
after that extend the proof to include simplification. 
400
46e5566ad4ba updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 398
diff changeset
   202
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   203
*}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   204
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   205
section {* Background *}
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   206
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   207
text {*
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   208
  In our Isabelle/HOL formalisation strings are lists of characters with
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   209
  the empty string being represented by the empty list, written $[]$,
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   210
  and list-cons being written as $\_\!::\!\_\,$; string
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   211
  concatenation is $\_ \,@\, \_\,$. We often use the usual
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   212
  bracket notation for lists also for strings; for example a string
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   213
  consisting of just a single character $c$ is written $[c]$.   
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   214
  Our regular expressions are defined as usual as the elements of the following inductive
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   215
  datatype:
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   216
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   217
  \begin{center}
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   218
  @{text "r ::="} \;
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   219
  @{const "ZERO"} $\mid$
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   220
  @{const "ONE"} $\mid$
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   221
  @{term "CH c"} $\mid$
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   222
  @{term "ALT r\<^sub>1 r\<^sub>2"} $\mid$
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   223
  @{term "SEQ r\<^sub>1 r\<^sub>2"} $\mid$
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   224
  @{term "STAR r"} 
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   225
  \end{center}
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   226
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   227
  \noindent where @{const ZERO} stands for the regular expression that does
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   228
  not match any string, @{const ONE} for the regular expression that matches
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   229
  only the empty string and @{term c} for matching a character literal.
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   230
  The constructors $+$ and $\cdot$ represent alternatives and sequences, respectively.
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   231
  The
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   232
  \emph{language} of a regular expression, written $L$, is defined as usual
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   233
  and we omit giving the definition here (see for example \cite{AusafDyckhoffUrban2016}).
420
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   234
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   235
  Central to Brzozowski's regular expression matcher are two functions
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   236
  called @{text nullable} and \emph{derivative}. The latter is written
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   237
  $r\backslash c$ for the derivative of the regular expression $r$
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   238
  w.r.t.~the character $c$. Both functions are defined by recursion over
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   239
  regular expressions.  
420
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   240
402
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   241
\begin{center}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   242
\begin{tabular}{cc}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   243
  \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
   244
  @{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
   245
  @{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
   246
  @{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
   247
  @{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
   248
  @{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
   249
  & & @{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
   250
  & & @{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
   251
  % & & @{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
   252
  @{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
   253
  \end{tabular}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   254
  &
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   255
  \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
   256
  @{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
   257
  @{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
   258
  @{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
   259
  @{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
   260
  @{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
   261
  @{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
   262
  \end{tabular}  
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   263
\end{tabular}  
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   264
\end{center}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   265
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   266
  \noindent
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   267
  We can extend this definition to give derivatives w.r.t.~strings:
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   268
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   269
  \begin{center}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   270
  \begin{tabular}{cc}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   271
  \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   272
  @{thm (lhs) ders.simps(1)} & $\dn$ & @{thm (rhs) ders.simps(1)}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   273
  \end{tabular}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   274
  &
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   275
  \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   276
  @{thm (lhs) ders.simps(2)} & $\dn$ & @{thm (rhs) ders.simps(2)}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   277
  \end{tabular}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   278
  \end{tabular}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   279
  \end{center}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   280
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   281
\noindent
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   282
Using @{text nullable} and the derivative operation, we can
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   283
define the following simple regular expression matcher:
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   284
%
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   285
\[
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   286
@{text "match s r"} \;\dn\; @{term nullable}(r\backslash s)
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   287
\]
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   288
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   289
\noindent This is essentially Brzozowski's algorithm from 1964. Its
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   290
main virtue is that the algorithm can be easily implemented as a
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   291
functional program (either in a functional programming language or in
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   292
a theorem prover). The correctness proof for @{text match} amounts to
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   293
establishing the property
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   294
%
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   295
\begin{proposition}\label{matchcorr} 
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   296
@{text "match s r"} \;\;\text{if and only if}\;\; $s \in L(r)$
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   297
\end{proposition}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   298
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   299
\noindent
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   300
It is a fun exercise to formaly prove this property in a theorem prover.
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   301
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   302
The novel idea of Sulzmann and Lu is to extend this algorithm for 
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   303
lexing, where it is important to find out which part of the string
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   304
is matched by which part of the regular expression.
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   305
For this Sulzmann and Lu presented two lexing algorithms in their paper
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   306
  \cite{Sulzmann2014}. The first algorithm consists of two phases: first a
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   307
  matching phase (which is Brzozowski's algorithm) and then a value
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   308
  construction phase. The values encode \emph{how} a regular expression
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   309
  matches a string. \emph{Values} are defined as the inductive datatype
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   310
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   311
  \begin{center}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   312
  @{text "v :="}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   313
  @{const "Void"} $\mid$
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   314
  @{term "val.Char c"} $\mid$
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   315
  @{term "Left v"} $\mid$
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   316
  @{term "Right v"} $\mid$
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   317
  @{term "Seq v\<^sub>1 v\<^sub>2"} $\mid$ 
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   318
  @{term "Stars vs"} 
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   319
  \end{center}  
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   320
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   321
  \noindent where we use @{term vs} to stand for a list of values. The
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   322
  string underlying a value can be calculated by a @{const flat}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   323
  function, written @{term "flat DUMMY"}. It traverses a value and
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   324
  collects the characters contained in it. Sulzmann and Lu also define inductively an
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   325
  inhabitation relation that associates values to regular expressions:
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   326
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   327
  \begin{center}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   328
  \begin{tabular}{c}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   329
  \\[-8mm]
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   330
  @{thm[mode=Axiom] Prf.intros(4)} \qquad
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   331
  @{thm[mode=Axiom] Prf.intros(5)[of "c"]}\\[4mm]
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   332
  @{thm[mode=Rule] Prf.intros(2)[of "v\<^sub>1" "r\<^sub>1" "r\<^sub>2"]} \qquad 
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   333
  @{thm[mode=Rule] Prf.intros(3)[of "v\<^sub>2" "r\<^sub>1" "r\<^sub>2"]}\\[4mm]
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   334
  @{thm[mode=Rule] Prf.intros(1)[of "v\<^sub>1" "r\<^sub>1" "v\<^sub>2" "r\<^sub>2"]} \qquad
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   335
  @{thm[mode=Rule] Prf.intros(6)[of "vs" "r"]}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   336
  \end{tabular}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   337
  \end{center}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   338
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   339
  \noindent Note that no values are associated with the regular expression
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   340
  @{term ZERO}, since it cannot match any string.
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   341
  It is routine to establish how values ``inhabiting'' a regular
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   342
  expression correspond to the language of a regular expression, namely
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   343
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   344
  \begin{proposition}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   345
  @{thm L_flat_Prf}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   346
  \end{proposition}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   347
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   348
  In general there is more than one value inhabited by a regular
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   349
  expression (meaning regular expressions can typically match more
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   350
  than one string). But even when fixing a string from the language of the
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   351
  regular expression, there are generally more than one way of how the
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   352
  regular expression can match this string. POSIX lexing is about
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   353
  identifying the unique value for a given regular expression and a
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   354
  string that satisfies the informal POSIX rules (see
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   355
  \cite{POSIX,Kuklewicz,OkuiSuzuki2010,Sulzmann2014,Vansummeren2006}).\footnote{POSIX
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   356
	lexing acquired its name from the fact that the corresponding
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   357
	rules were described as part of the POSIX specification for
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   358
	Unix-like operating systems \cite{POSIX}.} Sometimes these
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   359
  informal rules are called \emph{maximal much rule} and \emph{rule priority}.
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   360
  One contribution of our earlier paper is to give a convenient
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   361
 specification for what a POSIX value is (the inductive rules are shown in
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   362
  Figure~\ref{POSIXrules}).
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   363
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   364
\begin{figure}[t]
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   365
  \begin{center}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   366
  \begin{tabular}{c}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   367
  @{thm[mode=Axiom] Posix.intros(1)}\<open>P\<close>@{term "ONE"} \qquad
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   368
  @{thm[mode=Axiom] Posix.intros(2)}\<open>P\<close>@{term "c"}\medskip\\
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   369
  @{thm[mode=Rule] Posix.intros(3)[of "s" "r\<^sub>1" "v" "r\<^sub>2"]}\<open>P+L\<close>\qquad
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   370
  @{thm[mode=Rule] Posix.intros(4)[of "s" "r\<^sub>2" "v" "r\<^sub>1"]}\<open>P+R\<close>\medskip\\
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   371
  $\mprset{flushleft}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   372
   \inferrule
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   373
   {@{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
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   374
    @{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"]} \\\\
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   375
    @{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"]}}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   376
   {@{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>\medskip\smallskip\\
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   377
  @{thm[mode=Axiom] Posix.intros(7)}\<open>P[]\<close>\qquad
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   378
  $\mprset{flushleft}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   379
   \inferrule
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   380
   {@{thm (prem 1) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]} \qquad
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   381
    @{thm (prem 2) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]} \qquad
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   382
    @{thm (prem 3) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]} \\\\
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   383
    @{thm (prem 4) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]}}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   384
   {@{thm (concl) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]}}$\<open>P\<star>\<close>\\[-4mm]
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   385
  \end{tabular}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   386
  \end{center}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   387
  \caption{The inductive definition of POSIX values taken from our earlier paper \cite{AusafDyckhoffUrban2016}. The ternary relation, written $(s, r) \rightarrow v$, formalises the notion
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   388
  of given a string $s$ and a regular
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   389
  expression $r$ what is the unique value $v$ that satisfies the informal POSIX constraints for
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   390
  regular expression matching.}\label{POSIXrules}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   391
  \end{figure}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   392
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   393
  The clever idea by Sulzmann and Lu \cite{Sulzmann2014} in their first algorithm is to define
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   394
  an injection function on values that mirrors (but inverts) the
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   395
  construction of the derivative on regular expressions. Essentially it
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   396
  injects back a character into a value.
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   397
  For this they define two functions called @{text mkeps} and @{text inj}:
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   398
 
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   399
  \begin{center}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   400
  \begin{tabular}{l}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   401
  \begin{tabular}{lcl}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   402
  @{thm (lhs) mkeps.simps(1)} & $\dn$ & @{thm (rhs) mkeps.simps(1)}\\
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   403
  @{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"]}\\
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   404
  @{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"]}\\
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   405
  @{thm (lhs) mkeps.simps(4)} & $\dn$ & @{thm (rhs) mkeps.simps(4)}\\
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   406
  \end{tabular}\smallskip\\
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   407
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   408
  \begin{tabular}{lcl}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   409
  @{thm (lhs) injval.simps(1)} & $\dn$ & @{thm (rhs) injval.simps(1)}\\
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   410
  @{thm (lhs) injval.simps(2)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1"]} & $\dn$ & 
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   411
      @{thm (rhs) injval.simps(2)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1"]}\\
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   412
  @{thm (lhs) injval.simps(3)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]} & $\dn$ & 
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   413
      @{thm (rhs) injval.simps(3)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]}\\
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   414
  @{thm (lhs) injval.simps(4)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]} & $\dn$ 
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   415
      & @{thm (rhs) injval.simps(4)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]}\\
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   416
  @{thm (lhs) injval.simps(5)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]} & $\dn$ 
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   417
      & @{thm (rhs) injval.simps(5)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]}\\
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   418
  @{thm (lhs) injval.simps(6)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]} & $\dn$ 
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   419
      & @{thm (rhs) injval.simps(6)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]}\\
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   420
  @{thm (lhs) injval.simps(7)[of "r" "c" "v" "vs"]} & $\dn$ 
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   421
      & @{thm (rhs) injval.simps(7)[of "r" "c" "v" "vs"]}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   422
  \end{tabular}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   423
  \end{tabular}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   424
  \end{center}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   425
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   426
  \noindent
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   427
  The function @{text mkeps} is run when the last derivative is nullable, that is
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   428
  the string to be matched is in the language of the regular expression. It generates
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   429
  a value for how the last derivative can match the empty string. The injection function
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   430
  then calculates the corresponding value for each intermediate derivative until
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   431
  a value for the original regular expression is generated.
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   432
  Graphically the algorithm by
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   433
  Sulzmann and Lu can be illustrated by the picture in Figure~\ref{Sulz}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   434
  where the path from the left to the right involving @{term derivatives}/@{const
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   435
  nullable} is the first phase of the algorithm (calculating successive
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   436
  \Brz's derivatives) and @{const mkeps}/@{text inj}, the path from right to
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   437
  left, the second phase. The picture above shows the steps required when a
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   438
  regular expression, say @{text "r\<^sub>1"}, matches the string @{term
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   439
  "[a,b,c]"}. The lexing algorithm by Sulzmann and Lu can be defined as:
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   440
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   441
  \begin{figure}[t]
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   442
\begin{center}
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   443
\begin{tikzpicture}[scale=2,node distance=1.3cm,
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   444
                    every node/.style={minimum size=6mm}]
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   445
\node (r1)  {@{term "r\<^sub>1"}};
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   446
\node (r2) [right=of r1]{@{term "r\<^sub>2"}};
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   447
\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
   448
\node (r3) [right=of r2]{@{term "r\<^sub>3"}};
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   449
\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
   450
\node (r4) [right=of r3]{@{term "r\<^sub>4"}};
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   451
\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
   452
\draw (r4) node[anchor=west] {\;\raisebox{3mm}{@{term nullable}}};
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   453
\node (v4) [below=of r4]{@{term "v\<^sub>4"}};
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   454
\draw[->,line width=1mm](r4) -- (v4);
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   455
\node (v3) [left=of v4] {@{term "v\<^sub>3"}};
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   456
\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
   457
\node (v2) [left=of v3]{@{term "v\<^sub>2"}};
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   458
\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
   459
\node (v1) [left=of v2] {@{term "v\<^sub>1"}};
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   460
\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
   461
\draw (r4) node[anchor=north west] {\;\raisebox{-8mm}{@{term "mkeps"}}};
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   462
\end{tikzpicture}
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   463
\end{center}
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   464
\mbox{}\\[-13mm]
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   465
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   466
\caption{The two phases of the first algorithm by Sulzmann \& Lu \cite{Sulzmann2014},
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   467
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
   468
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
   469
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
   470
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
   471
@{term mkeps} calculates a value @{term "v\<^sub>4"} witnessing
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   472
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
   473
that the function @{term inj} ``injects back'' the characters of the string into
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   474
the values. The value @{term "v\<^sub>1"} is the result of the algorithm representing
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   475
the POSIX value for this string and
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   476
regular expression.
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   477
\label{Sulz}}
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   478
\end{figure} 
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   479
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   480
420
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   481
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   482
  \begin{center}
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   483
  \begin{tabular}{lcl}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   484
  @{thm (lhs) lexer.simps(1)} & $\dn$ & @{thm (rhs) lexer.simps(1)}\\
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   485
  @{thm (lhs) lexer.simps(2)} & $\dn$ & @{text "case"} @{term "lexer (der c r) s"} @{text of}\\
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   486
                     & & \phantom{$|$} @{term "None"}  @{text "\<Rightarrow>"} @{term None}\\
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   487
                     & & $|$ @{term "Some v"} @{text "\<Rightarrow>"} @{term "Some (injval r c v)"}                          
420
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   488
  \end{tabular}
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   489
  \end{center}
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   490
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   491
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   492
We have shown in our earlier paper \cite{AusafDyckhoffUrban2016} that
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   493
this algorithm is correct, that is it generates POSIX values. The
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   494
cenral property we established relates the derivative operation to the
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   495
injection function.
420
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   496
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   497
  \begin{proposition}\label{Posix2}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   498
	\textit{If} $(s,\; r\backslash c) \rightarrow v$ \textit{then} $(c :: s,\; r) \rightarrow$ \textit{inj} $r\; c\; v$. 
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   499
\end{proposition}
420
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   500
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   501
  \noindent
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   502
  With this in place we were able to prove:
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   503
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   504
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   505
  \begin{proposition}\mbox{}\smallskip\\\label{lexercorrect}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   506
  \begin{tabular}{ll}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   507
  (1) & @{thm (lhs) lexer_correct_None} if and only if @{thm (rhs) lexer_correct_None}\\
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   508
  (2) & @{thm (lhs) lexer_correct_Some} if and only if @{thm (rhs) lexer_correct_Some}\\
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   509
  \end{tabular}
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   510
  \end{proposition}
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   511
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   512
  \noindent
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   513
  In fact we have shown that in the success case the generated POSIX value $v$ is
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   514
  unique and in the failure case that there is no POSIX value $v$ that satisfies
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   515
  $(s, r) \rightarrow v$. While the algorithm is correct, it is excrutiatingly
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   516
  slow in cases where the derivatives grow arbitrarily (see example from the
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   517
  Introduction). However it can be used as a convenient reference point for the correctness
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   518
  proof of the second algorithm by Sulzmann and Lu, which we shall describe next.
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   519
  
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   520
*}
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   521
402
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   522
section {* Bitcoded Regular Expressions and Derivatives *}
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   523
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   524
text {*
405
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   525
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   526
  In the second part of their paper \cite{Sulzmann2014},
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   527
  Sulzmann and Lu describe another algorithm that also generates POSIX
405
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   528
  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
   529
  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
   530
  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
   531
405
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   532
  \begin{center}
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   533
  \begin{tabular}{lcl}
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   534
  @{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
   535
               & $\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
   536
               & $\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
   537
               & $\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
   538
               & $\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
   539
  \end{tabular}
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   540
  \end{center}
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   541
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   542
  \noindent where @{text bs} stands for bitsequences; @{text r},
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   543
  @{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
   544
  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
   545
  expressions. The binary alternative @{text "ALT bs r\<^sub>1 r\<^sub>2"}
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   546
  is just an abbreviation for \mbox{@{text "ALTs bs [r\<^sub>1, r\<^sub>2]"}}. 
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   547
  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
   548
  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
   549
  expressions is to incrementally generate the value information (for
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   550
  example @{text Left} and @{text Right}) as bitsequences. For this 
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   551
  Sulzmann and Lu define a coding
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   552
  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
   553
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   554
  \begin{center}
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   555
  \begin{tabular}{cc}
402
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   556
  \begin{tabular}{lcl}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   557
  @{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
   558
  @{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
   559
  @{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
   560
  @{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
   561
  \end{tabular}
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   562
  &
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   563
  \begin{tabular}{lcl}
402
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   564
  @{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
   565
  @{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
   566
  @{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
   567
  \mbox{\phantom{XX}}\\
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   568
  \end{tabular}
402
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   569
  \end{tabular}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   570
  \end{center}
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   571
   
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   572
  \noindent
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   573
  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
   574
  record explicitly character values and also not sequence values (for
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   575
  them we just append two bitsequences). However, the
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   576
  different alternatives for @{text Left}, respectively @{text Right}, are recorded as @{text Z} and
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   577
  @{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
   578
  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
   579
  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
   580
  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
   581
  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
   582
  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
   583
  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
   584
  \textit{decode}:
402
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   585
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   586
  \begin{center}
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   587
  \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
   588
  $\textit{decode}'\,bs\,(\ONE)$ & $\dn$ & $(\Empty, bs)$\\
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   589
  $\textit{decode}'\,bs\,(c)$ & $\dn$ & $(\Char\,c, bs)$\\
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   590
  $\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
   591
     $\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
   592
       (\Left\,v, bs_1)$\\
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   593
  $\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
   594
     $\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
   595
       (\Right\,v, bs_1)$\\                           
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   596
  $\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
   597
        $\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
   598
  & &   $\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
   599
        \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
   600
  $\textit{decode}'\,(\Z\!::\!bs)\,(r^*)$ & $\dn$ & $(\Stars\,[], bs)$\\
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   601
  $\textit{decode}'\,(\S\!::\!bs)\,(r^*)$ & $\dn$ & 
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   602
         $\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
   603
  & &   $\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
   604
        \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
   605
  $\textit{decode}\,bs\,r$ & $\dn$ &
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   606
     $\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
   607
  & & \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
   608
       \textit{else}\;\textit{None}$   
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   609
  \end{tabular}    
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   610
  \end{center}
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   611
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   612
  \noindent
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   613
  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
   614
  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
   615
  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
   616
  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
   617
  bitsequence never fails.
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   618
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   619
\begin{lemma}\label{codedecode}\it
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   620
  If $\;\vdash v : r$ then
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   621
  $\;\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
   622
\end{lemma}
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   623
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   624
\begin{proof}
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   625
  This follows from the property that
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   626
  $\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
   627
  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
   628
  easily proved by induction on $\vdash v : r$.
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   629
\end{proof}  
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   630
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   631
  Sulzmann and Lu define the function \emph{internalise}
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   632
  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
   633
  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
   634
  This internalisation uses the following
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   635
  \emph{fuse} function.	     
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   636
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   637
  \begin{center}
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   638
  \begin{tabular}{lcl}
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   639
  $\textit{fuse}\,bs\,(\textit{ZERO})$ & $\dn$ & $\textit{ZERO}$\\
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   640
  $\textit{fuse}\,bs\,(\textit{ONE}\,bs')$ & $\dn$ &
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   641
     $\textit{ONE}\,(bs\,@\,bs')$\\
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   642
  $\textit{fuse}\,bs\,(\textit{CHAR}\,bs'\,c)$ & $\dn$ &
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   643
     $\textit{CHAR}\,(bs\,@\,bs')\,c$\\
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   644
  $\textit{fuse}\,bs\,(\textit{ALTs}\,bs'\,rs)$ & $\dn$ &
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   645
     $\textit{ALTs}\,(bs\,@\,bs')\,rs$\\
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   646
  $\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
   647
     $\textit{SEQ}\,(bs\,@\,bs')\,r_1\,r_2$\\
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   648
  $\textit{fuse}\,bs\,(\textit{STAR}\,bs'\,r)$ & $\dn$ &
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   649
     $\textit{STAR}\,(bs\,@\,bs')\,r$
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   650
  \end{tabular}    
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   651
  \end{center}    
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   652
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   653
  \noindent
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   654
  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
   655
  regular expression as follows.
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   656
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   657
  \begin{center}
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   658
  \begin{tabular}{lcl}
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   659
  $(\ZERO)^\uparrow$ & $\dn$ & $\textit{ZERO}$\\
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   660
  $(\ONE)^\uparrow$ & $\dn$ & $\textit{ONE}\,[]$\\
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   661
  $(c)^\uparrow$ & $\dn$ & $\textit{CHAR}\,[]\,c$\\
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   662
  $(r_1 + r_2)^\uparrow$ & $\dn$ &
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   663
         $\textit{ALT}\;[]\,(\textit{fuse}\,[\Z]\,r_1^\uparrow)\,
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   664
                            (\textit{fuse}\,[\S]\,r_2^\uparrow)$\\
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   665
  $(r_1\cdot r_2)^\uparrow$ & $\dn$ &
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   666
         $\textit{SEQ}\;[]\,r_1^\uparrow\,r_2^\uparrow$\\
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   667
  $(r^*)^\uparrow$ & $\dn$ &
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   668
         $\textit{STAR}\;[]\,r^\uparrow$\\
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   669
  \end{tabular}    
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   670
  \end{center}    
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   671
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   672
  \noindent
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   673
  There is also an \emph{erase}-function, written $r^\downarrow$, which
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   674
  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
   675
  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
   676
  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
   677
  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
   678
  ``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
   679
  bitcoded regular expressions, instead of regular expressions.
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   680
  %
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   681
  \begin{center}
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   682
  \begin{tabular}{@ {}c@ {}c@ {}}
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   683
  \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   684
  $\textit{bnullable}\,(\textit{ZERO})$ & $\dn$ & $\textit{false}$\\
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   685
  $\textit{bnullable}\,(\textit{ONE}\,bs)$ & $\dn$ & $\textit{true}$\\
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   686
  $\textit{bnullable}\,(\textit{CHAR}\,bs\,c)$ & $\dn$ & $\textit{false}$\\
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   687
  $\textit{bnullable}\,(\textit{ALTs}\,bs\,\rs)$ & $\dn$ &
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   688
     $\exists\, r \in \rs. \,\textit{bnullable}\,r$\\
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   689
  $\textit{bnullable}\,(\textit{SEQ}\,bs\,r_1\,r_2)$ & $\dn$ &
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   690
     $\textit{bnullable}\,r_1\wedge \textit{bnullable}\,r_2$\\
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   691
  $\textit{bnullable}\,(\textit{STAR}\,bs\,r)$ & $\dn$ &
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   692
     $\textit{true}$
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   693
  \end{tabular}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   694
  &
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   695
  \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}}
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   696
  $\textit{bmkeps}\,(\textit{ONE}\,bs)$ & $\dn$ & $bs$\\
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   697
  $\textit{bmkeps}\,(\textit{ALTs}\,bs\,r\!::\!\rs)$ & $\dn$ &
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   698
     $\textit{if}\;\textit{bnullable}\,r$\\
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   699
  & &$\textit{then}\;bs\,@\,\textit{bmkeps}\,r$\\
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   700
  & &$\textit{else}\;bs\,@\,\textit{bmkeps}\,\rs$\\
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   701
  $\textit{bmkeps}\,(\textit{SEQ}\,bs\,r_1\,r_2)$ & $\dn$ &\\
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   702
  \multicolumn{3}{r}{$bs \,@\,\textit{bmkeps}\,r_1\,@\, \textit{bmkeps}\,r_2$}\\
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   703
  $\textit{bmkeps}\,(\textit{STAR}\,bs\,r)$ & $\dn$ &
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   704
     $bs \,@\, [\S]$
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   705
  \end{tabular}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   706
  \end{tabular}
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   707
  \end{center}    
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   708
 
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   709
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   710
  \noindent
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   711
  The key function in the bitcoded algorithm is the derivative of a
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   712
  bitcoded regular expression. This derivative calculates the
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   713
  derivative but at the same time also the incremental part of bitsequences
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   714
  that contribute to constructing a POSIX value.	
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   715
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   716
  \begin{center}
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   717
  \begin{tabular}{@ {}lcl@ {}}
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   718
  $(\textit{ZERO})\backslash c$ & $\dn$ & $\textit{ZERO}$ \\  
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   719
  $(\textit{ONE}\;bs)\backslash c$ & $\dn$ & $\textit{ZERO}$\\  
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   720
  $(\textit{CHAR}\;bs\,d)\backslash c$ & $\dn$ &
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   721
        $\textit{if}\;c=d\; \;\textit{then}\;
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   722
         \textit{ONE}\;bs\;\textit{else}\;\textit{ZERO}$\\  
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   723
  $(\textit{ALTs}\;bs\,\rs)\backslash c$ & $\dn$ &
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   724
        $\textit{ALTs}\,bs\,(\mathit{map}\,(\_\backslash c)\,\rs)$\\
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   725
  $(\textit{SEQ}\;bs\,r_1\,r_2)\backslash c$ & $\dn$ &
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   726
     $\textit{if}\;\textit{bnullable}\,r_1$\\
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   727
  & &$\textit{then}\;\textit{ALT}\,bs\,(\textit{SEQ}\,[]\,(r_1\backslash c)\,r_2)$\\
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   728
  & &$\phantom{\textit{then}\;\textit{ALT}\,bs\,}(\textit{fuse}\,(\textit{bmkeps}\,r_1)\,(r_2\backslash c))$\\
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   729
  & &$\textit{else}\;\textit{SEQ}\,bs\,(r_1\backslash c)\,r_2$\\
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   730
  $(\textit{STAR}\,bs\,r)\backslash c$ & $\dn$ &
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   731
      $\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
   732
       (\textit{STAR}\,[]\,r)$
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   733
  \end{tabular}    
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   734
  \end{center}
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   735
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   736
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   737
  \noindent
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   738
  This function can also be extended to strings, written $r\backslash s$,
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   739
  just like the standard derivative.  We omit the details. Finally we
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   740
  can define Sulzmann and Lu's bitcoded lexer, which we call \textit{blexer}:
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   741
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   742
  \begin{center}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   743
\begin{tabular}{lcl}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   744
  $\textit{blexer}\;r\,s$ & $\dn$ &
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   745
      $\textit{let}\;r_{der} = (r^\uparrow)\backslash s\;\textit{in}$\\                
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   746
  & & $\;\;\;\;\textit{if}\; \textit{bnullable}(r_{der}) \;\;\textit{then}\;\textit{decode}\,(\textit{bmkeps}\,r_{der})\,r
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   747
       \;\;\textit{else}\;\textit{None}$
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   748
\end{tabular}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   749
\end{center}
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   750
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   751
  \noindent
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   752
This bitcoded lexer first internalises the regular expression $r$ and then
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   753
builds the bitcoded derivative according to $s$. If the derivative is
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   754
(b)nullable the string is in the language of $r$ and it extracts the bitsequence using the
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   755
$\textit{bmkeps}$ function. Finally it decodes the bitsequence into a value.  If
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   756
the derivative is \emph{not} nullable, then $\textit{None}$ is
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   757
returned. We can show that this way of calculating a value
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   758
generates the same result as \textit{lexer}.
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   759
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   760
Before we can proceed we need to define a helper function, called
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   761
\textit{retrieve}, which Sulzmann and Lu introduced for the correctness proof.
416
57182b36ec01 more with the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 410
diff changeset
   762
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   763
\begin{center}
402
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   764
  \begin{tabular}{lcl}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   765
  @{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
   766
  @{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
   767
  @{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
   768
  @{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
   769
  @{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
   770
  @{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
   771
      & $\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
   772
  @{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
   773
  @{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
   774
  \end{tabular}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   775
  \end{center}
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   776
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   777
\noindent
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   778
The idea behind this function is to retrieve a possibly partial
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   779
bitsequence from a bitcoded regular expression, where the retrieval is
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   780
guided by a value.  For example if the value is $\Left$ then we
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   781
descend into the left-hand side of an alternative in order to
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   782
assemble the bitcode. Similarly for
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   783
$\Right$. The property we can show is that for a given $v$ and $r$
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   784
with $\vdash v : r$, the retrieved bitsequence from the internalised
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   785
regular expression is equal to the bitcoded version of $v$.
402
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
   786
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   787
\begin{lemma}\label{retrievecode}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   788
If $\vdash v : r$ then $\textit{code}\, v = \textit{retrieve}\,(r^\uparrow)\,v$.
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   789
\end{lemma}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   790
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   791
\noindent
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   792
We also need some auxiliary facts about how the bitcoded operations
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   793
relate to the ``standard'' operations on regular expressions. For
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   794
example if we build a bitcoded derivative and erase the result, this
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   795
is the same as if we first erase the bitcoded regular expression and
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   796
then perform the ``standard'' derivative operation.
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   797
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   798
\begin{lemma}\label{bnullable}\mbox{}\smallskip\\
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   799
  \begin{tabular}{ll}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   800
\textit{(1)} & $(a\backslash s)^\downarrow = (a^\downarrow)\backslash s$\\    
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   801
\textit{(2)} & $\textit{bnullable}(a)$ iff $\textit{nullable}(a^\downarrow)$\\
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   802
\textit{(3)} & $\textit{bmkeps}(a) = \textit{retrieve}\,a\,(\textit{mkeps}\,(a^\downarrow))$ provided $\textit{nullable}(a^\downarrow)$.
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   803
\end{tabular}  
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   804
\end{lemma}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   805
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   806
\begin{proof}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   807
  All properties are by induction on annotated regular expressions. There are no
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   808
  interesting cases.
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   809
\end{proof}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   810
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   811
\noindent
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   812
The only difficulty left for the correctness proof is that the bitcoded algorithm
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   813
has only a ``forward phase'' where POSIX values are generated incrementally.
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   814
We can achieve the same effect with @{text lexer} by stacking up injection
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   815
functions during the forward phase. An auxiliary function, called $\textit{flex}$,
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   816
allows us to recast the rules of $\lexer$ (with its two phases) in terms of a single
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   817
phase and stacked up injection functions.
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   818
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   819
\begin{center}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   820
\begin{tabular}{lcl}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   821
  $\textit{flex}\;r\,f\,[]$ & $\dn$ & $f$\\
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   822
  $\textit{flex}\;r\,f\,(c\!::\!s)$ & $\dn$ &
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   823
  $\textit{flex}\,(r\backslash c)\,(\lambda v.\,f\,(\inj\,r\,c\,v))\,s$\\
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   824
\end{tabular}    
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   825
\end{center}    
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   826
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   827
\noindent
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   828
The point of this function is that when
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   829
reaching the end of the string, we just need to apply the stacked
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   830
injection functions to the value generated by @{text mkeps}.
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   831
Using this function we can recast the success case in @{text lexer} 
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   832
as follows:
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   833
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   834
\begin{proposition}\label{flex}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   835
If @{text "lexer r s = Some v"} \;then\; @{text "v = "}$\,\textit{flex}\,r\,id\,s\,
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   836
      (\mkeps (r\backslash s))$.
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   837
\end{proposition}
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   838
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   839
\noindent
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   840
Note we did not redefine \textit{lexer}, we just established that the
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   841
value generated by \textit{lexer} can also be obtained by a different
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   842
method. While this different method is not efficient (we essentially
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   843
need to traverse the string $s$ twice, once for building the
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   844
derivative $r\backslash s$ and another time for stacking up injection
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   845
functions using \textit{flex}), it helps us with proving
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   846
that incrementally building up values generates the same result.
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   847
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   848
This brings us to our main lemma in this section: if we calculate a
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   849
derivative, say $r\backslash s$ and have a value, say $v$, inhabited
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   850
by this derivative, then we can produce the result $\lexer$ generates
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   851
by applying this value to the stacked-up injection functions
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   852
that $\textit{flex}$ assembles. The lemma establishes that this is the same
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   853
value as if we build the annotated derivative $r^\uparrow\backslash s$
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   854
and then retrieve the corresponding bitcoded version, followed by a
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   855
decoding step.
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   856
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   857
\begin{lemma}[Main Lemma]\label{mainlemma}\it
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   858
If $\vdash v : r\backslash s$ then 
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   859
\[\textit{Some}\,(\textit{flex}\,r\,\textit{id}\,s\,v) =
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   860
  \textit{decode}(\textit{retrieve}\,(r^\uparrow \backslash s)\,v)\,r\]
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   861
\end{lemma}  
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   862
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   863
\begin{proof}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   864
  This can be proved by induction on $s$ and generalising over
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   865
  $v$. The interesting point is that we need to prove this in the
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   866
  reverse direction for $s$. This means instead of cases $[]$ and
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   867
  $c\!::\!s$, we have cases $[]$ and $s\,@\,[c]$ where we unravel the
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   868
  string from the back.\footnote{Isabelle/HOL provides an induction principle
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   869
    for this way of performing the induction.}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   870
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   871
  The case for $[]$ is routine using Lemmas~\ref{codedecode}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   872
  and~\ref{retrievecode}. In the case $s\,@\,[c]$, we can infer from
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   873
  the assumption that $\vdash v : (r\backslash s)\backslash c$
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   874
  holds. Hence by Prop.~\ref{Posix2} we know that 
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   875
  (*) $\vdash \inj\,(r\backslash s)\,c\,v : r\backslash s$ holds too.
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   876
  By definition of $\textit{flex}$ we can unfold the left-hand side
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   877
  to be
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   878
  \[
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   879
    \textit{Some}\,(\textit{flex}\;r\,\textit{id}\,(s\,@\,[c])\,v) =
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   880
    \textit{Some}\,(\textit{flex}\;r\,\textit{id}\,s\,(\inj\,(r\backslash s)\,c\,v))  
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   881
  \]  
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   882
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   883
  \noindent
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   884
  By induction hypothesis and (*) we can rewrite the right-hand side to
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   885
  %
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   886
  \[
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   887
    \textit{decode}\,(\textit{retrieve}\,(r^\uparrow\backslash s)\;
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   888
    (\inj\,(r\backslash s)\,c\,\,v))\,r
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   889
  \]
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   890
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   891
  \noindent
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   892
  which is equal to
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   893
  $\textit{decode}\,(\textit{retrieve}\, (r^\uparrow\backslash
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   894
  (s\,@\,[c]))\,v)\,r$ as required. The last rewrite step is possible
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   895
  because we generalised over $v$ in our induction.
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   896
\end{proof}  
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   897
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   898
\noindent
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   899
With this lemma in place, we can prove the correctness of \textit{blexer}---it indeed
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   900
produces the same result as \textit{lexer}.
405
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   901
3cfea5bb5e23 updated some of the text and cardinality proof
Christian Urban <christian.urban@kcl.ac.uk>
parents: 402
diff changeset
   902
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   903
\begin{theorem}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   904
$\textit{lexer}\,r\,s = \textit{blexer}\,r\,s$
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   905
\end{theorem}  
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   906
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   907
\begin{proof}
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   908
  We can first expand both sides using Prop.~\ref{flex} and the
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   909
  definition of \textit{blexer}. This gives us two
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   910
  \textit{if}-statements, which we need to show to be equal. By 
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   911
  Lemma~\ref{bnullable}\textit{(2)} we know the \textit{if}-tests coincide:
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   912
  \[
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   913
    \textit{bnullable}(r^\uparrow\backslash s) \;\textit{iff}\;
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   914
    \nullable(r\backslash s)
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   915
  \]
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   916
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   917
  \noindent
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   918
  For the \textit{if}-branch suppose $r_d \dn r^\uparrow\backslash s$ and
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   919
  $d \dn r\backslash s$. We have (*) @{text "nullable d"}. We can then show
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   920
  by Lemma~\ref{bnullable}\textit{(3)} that
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   921
  %
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   922
  \[
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   923
    \textit{decode}(\textit{bmkeps}\,r_d)\,r =
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   924
    \textit{decode}(\textit{retrieve}\,a\,(\textit{mkeps}\,d))\,r
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   925
  \]
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   926
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   927
  \noindent
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   928
  where the right-hand side is equal to
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   929
  $\textit{Some}\,(\textit{flex}\,r\,\textit{id}\,s\,(\textit{mkeps}\,
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   930
  d))$ by Lemma~\ref{mainlemma} (we know
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   931
  $\vdash \textit{mkeps}\,d : d$ by (*)).  This shows the
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   932
  \textit{if}-branches return the same value. In the
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   933
  \textit{else}-branches both \textit{lexer} and \textit{blexer} return
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   934
  \textit{None}. Therefore we can conclude the proof.
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   935
\end{proof}  
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   936
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   937
\noindent This establishes that the bitcoded algorithm by Sulzmann and
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   938
Lu \emph{without} simplification produces correct results. This was
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   939
only conjectured by Sulzmann and Lu in their paper
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   940
\cite{Sulzmann2014}. The next step is to add simplifications.
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   941
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   942
*}
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   943
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   944
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   945
section {* Simplification *}
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   946
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
   947
text {*
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
   948
420
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   949
     Derivatives as calculated by Brzozowski’s method are usually more
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
   950
     complex regular expressions than the initial one; the result is
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   951
     that derivative-based matching and lexing algorithms are
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   952
     often abysmally slow if the ``growth problem'' is not addressed. As Sulzmann and Lu wrote, various
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   953
     optimisations are possible, such as the simplifications
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   954
     $\ZERO{}\,r \Rightarrow \ZERO$, $\ONE{}\,r \Rightarrow r$,
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
   955
     $\ZERO{} + r \Rightarrow r$ and $r + r \Rightarrow r$. While these
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   956
     simplifications can considerably speed up the two algorithms  in many
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   957
     cases, they do not solve fundamentally the growth problem with
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   958
     derivatives. To see this let us return to the example from the
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   959
     Introduction that shows the derivatives for \mbox{@{text "(a + aa)\<^sup>*"}}.
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   960
     If we delete in the 3rd step all $\ZERO{}s$ and $\ONE$s according to
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   961
     the simplification rules shown above we obtain
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   962
     %
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   963
     \def\xll{\xrightarrow{\_\backslash{} [a, a, a]}}%%
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   964
     %
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   965
     \begin{equation}\label{derivex}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   966
     (a + aa)^* \quad\xll\quad
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   967
      \underbrace{\mbox{$(\ONE + \ONE{}a) \cdot (a + aa)^*$}}_{r} \;+\;
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   968
     ((a + aa)^* + \underbrace{\mbox{$(\ONE + \ONE{}a) \cdot (a + aa)^*$}}_{r})
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   969
     \end{equation}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   970
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   971
     \noindent This is a simpler derivative, but unfortunately we
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   972
     cannot make further simplifications. This is a problem because
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   973
     the outermost alternatives contains two copies of the same
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   974
     regular expression (underlined with $r$). The copies will
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   975
     spawn new copies in later steps and they in turn further copies. This
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   976
     destroys an hope of taming the size of the derivatives.  But the
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   977
     second copy of $r$ in \eqref{derivex} will never contribute to a
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   978
     value, because POSIX lexing will always prefer matching a string
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   979
     with the first copy. So in principle it could be removed.
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   980
     The dilemma with the simple-minded
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   981
     simplification rules above is that the rule $r + r \Rightarrow r$
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   982
     will never be applicable because as can be seen in this example the
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   983
     regular expressions are separated by another sub-regular expression.
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   984
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   985
     But here is where Sulzmann and Lu's representation of generalised
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   986
     alternatives in the bitcoded algorithm shines: in @{term
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   987
     "ALTs bs rs"} we can define a more aggressive simplification by
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   988
     recursively simplifying all regular expressions in @{text rs} and
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   989
     then analyse the resulting list and remove any duplicates.
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   990
     Another advantage is that the bitsequences in  bitcoded regular
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   991
     expressions can be easily modified such that simplification does not
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   992
     interfere with the value constructions. For example we can ``flatten'', or
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   993
     de-nest, @{text ALTs} as follows
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   994
     %
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   995
     \[
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   996
     @{term "ALTs bs\<^sub>1 ((ALTs bs\<^sub>2 rs\<^sub>2) # rs\<^sub>1)"}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   997
     \quad\xrightarrow{bsimp}\quad
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   998
     @{term "ALTs bs\<^sub>1 ((map (fuse bs\<^sub>2) rs\<^sub>2) # rs\<^sub>1)"}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
   999
     \]
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1000
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1001
     \noindent
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1002
     where we just need to fuse the bitsequence that has accumulated in @{text "bs\<^sub>2"}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1003
     to the alternatives in @{text "rs\<^sub>2"}. As we shall show below this will
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1004
     ensure that the correct value corresponding to the original (unsimplified)
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1005
     regular expression can still be extracted. %In this way the value construction
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1006
     %is not affected by simplification. 
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1007
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1008
     However there is one problem with the definition for the more
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1009
     aggressive simlification rules by Sulzmann and Lu. Recasting
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1010
     their definition with our syntax they define the step of removing
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1011
     duplicates as
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1012
     %
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1013
     \[ @{text "bsimp (ALTs bs rs)"} \dn @{text "ALTs
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1014
     bs (nup (map bsimp rs))"}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1015
     \]
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1016
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1017
     \noindent where they first recursively simplify the regular
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1018
     expressions in @{text rs} (using @{text map}) and then use
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1019
     Haskell's @{text nub}-function to remove potential
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1020
     duplicates. While this makes sense when considering the example
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1021
     shown in \eqref{derivex}, @{text nub} is the inappropriate
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1022
     function in the case of bitcoded regular expressions. The reason
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1023
     is that in general the n elements in @{text rs} will have a
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1024
     different bitsequence annotated to it and in this way @{text nub}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1025
     will never find a duplicate to be removed. The correct way to
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1026
     handle this situation is to first \emph{erase} the regular
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1027
     expressions when comparing potential duplicates. This is inspired
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1028
     by Scala's list functions of the form \mbox{@{text "distinctBy rs f
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1029
     acc"}} where a function is applied first before two elements
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1030
     are compared. We define this function in Isabelle/HOL as
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1031
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1032
     \begin{center}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1033
     \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1034
     @{thm (lhs) distinctBy.simps(1)} & $\dn$ & @{thm (rhs) distinctBy.simps(1)}\\
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1035
     @{thm (lhs) distinctBy.simps(2)} & $\dn$ & @{thm (rhs) distinctBy.simps(2)}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1036
     \end{tabular}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1037
     \end{center}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1038
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1039
     \noindent where we scan the list from left to right (because we
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1040
     have to remove later copies). In this function, @{text f} is a
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1041
     function and @{text acc} is an accumulator for regular
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1042
     expressions---essentially a set of elements we have already seen
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1043
     while scanning the list. Therefore we delete an element, say @{text x},
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1044
     from the list provided @{text "f x"} is already in the accumulator;
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1045
     otherwise we keep @{text x} and scan the rest of the list but now
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1046
     add @{text "f x"} as another element to @{text acc}. We will use
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1047
     @{term distinctBy} where @{text f} is our erase functions, @{term "erase (DUMMY)"},
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1048
     that deletes bitsequences from bitcoded regular expressions.
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1049
     This is clearly a computationally more expensive operation, than @{text nub},
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1050
     but is needed in order to make the removal of unnecessary copies
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1051
     to work.
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1052
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1053
     Our simplification function depends on three helper functions, one is called
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1054
     @{text flts} and defined as follows:
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1055
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1056
     \begin{center}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1057
     \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1058
     @{thm (lhs) flts.simps(1)} & $\dn$ & @{thm (rhs) flts.simps(1)}\\
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1059
     @{thm (lhs) flts.simps(2)} & $\dn$ & @{thm (rhs) flts.simps(2)}\\
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1060
     @{thm (lhs) flts.simps(3)[of "bs'" "rs'"]} & $\dn$ & @{thm (rhs) flts.simps(3)[of "bs'" "rs'"]}\\
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1061
     \end{tabular}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1062
     \end{center}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1063
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1064
     \noindent
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1065
     The second clause removes all instances of @{text ZERO} in alternatives and
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1066
     the second ``spills'' out nested alternatives (but retaining the
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1067
     bitsequence @{text "bs'"} accumulated in the inner alternative). There are
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1068
     some corner cases to be considered when the resulting list inside an alternative is
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1069
     empty or a singleton list. We take care of those cases in the
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1070
     @{text "bsimpALTs"} function; similarly we define a helper function that simplifies
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1071
     sequences according to the usual rules about @{text ZERO}s and @{text ONE}s:
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1072
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1073
     \begin{center}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1074
     \begin{tabular}{c@ {\hspace{5mm}}c}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1075
     \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1076
     @{text "bsimpALTs bs []"}  & $\dn$ & @{text "ZERO"}\\
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1077
     @{text "bsimpALTs bs [r]"} & $\dn$ & @{text "fuse bs r"}\\
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1078
     @{text "bsimpALTs bs rs"}  & $\dn$ & @{text "ALTs bs rs"}\\
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1079
     \mbox{}\\
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1080
     \end{tabular}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1081
     &
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1082
     \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1083
     @{text "bsimpSEQ bs _ ZERO"}  & $\dn$ & @{text "ZERO"}\\
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1084
     @{text "bsimpSEQ bs ZERO _"} & $\dn$ & @{text "ZERO"}\\
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1085
     @{text "bsimpSEQ bs\<^sub>1 (ONE bs\<^sub>2) r\<^sub>2"}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1086
         & $\dn$ & @{text "fuse (bs\<^sub>1 @ bs\<^sub>2) r\<^sub>2"}\\
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1087
     @{text "bsimpSEQ bs r\<^sub>1 r\<^sub>2"} & $\dn$ &  @{text "SEQ bs r\<^sub>1 r\<^sub>2"}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1088
     \end{tabular}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1089
     \end{tabular}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1090
     \end{center}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1091
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1092
     \noindent
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1093
     With this in place we can define our simlification function as
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1094
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1095
     \begin{center}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1096
     \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1097
     @{thm (lhs) bsimp.simps(1)[of "bs" "r\<^sub>1" "r\<^sub>2"]} & $\dn$ &
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1098
         @{thm (rhs) bsimp.simps(1)[of "bs" "r\<^sub>1" "r\<^sub>2"]}\\
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1099
     @{thm (lhs) bsimp.simps(2)[of "bs" _]} & $\dn$ & @{thm (rhs) bsimp.simps(2)[of "bs" _]}\\
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1100
     @{text "bsimp r"} & $\dn$ & @{text r}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1101
     \end{tabular}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1102
     \end{center}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1103
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1104
     \noindent
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1105
     As far as we can see, our recursive function @{term bsimp} simplifies regular
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1106
     expressions as intended by Sulzmann and Lu. There is no point to apply the
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1107
     @{text bsimp}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1108
     function repeatedly (like the simplification in their paper which is applied
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1109
     until a fixpoint is reached), because we can show that it is idempotent, that is
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1110
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1111
     \begin{proposition}
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1112
     ???
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1113
     \end{proposition}
420
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
  1114
b66a4305749c updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 418
diff changeset
  1115
418
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1116
     \begin{lemma}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1117
     @{thm[mode=IfThen] bnullable0(1)[of "r\<^sub>1" "r\<^sub>2"]}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1118
     \end{lemma}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1119
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1120
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1121
     \begin{lemma}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1122
     @{thm[mode=IfThen] rewrite_bmkeps_aux(1)[of "r\<^sub>1" "r\<^sub>2"]}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1123
     \end{lemma}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1124
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1125
     \begin{lemma}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1126
     @{thm[mode=IfThen] rewrites_to_bsimp}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1127
     \end{lemma}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1128
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1129
     \begin{lemma}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1130
     @{thm[mode=IfThen] rewrite_preserves_bder(1)[of "r\<^sub>1" "r\<^sub>2"]}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1131
     \end{lemma}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1132
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1133
     \begin{lemma}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1134
     @{thm[mode=IfThen] central}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1135
     \end{lemma}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1136
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1137
     \begin{theorem}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1138
     @{thm[mode=IfThen] main_blexer_simp}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1139
     \end{theorem}
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1140
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1141
     Sulzmann \& Lu apply simplification via a fixpoint operation
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1142
41a2a3b63853 more of the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 416
diff changeset
  1143
     ; 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
  1144
  
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
  1145
   not direct correspondence with PDERs, because of example
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
  1146
   problem with retrieve 
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
  1147
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
  1148
   correctness
398
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
  1149
402
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
  1150
   
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
  1151
    
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
  1152
398
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
  1153
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
  1154
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
  1155
   \begin{figure}[t]
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
  1156
  \begin{center}
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
  1157
  \begin{tabular}{c}
402
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
  1158
  @{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
  1159
  @{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
  1160
  @{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
  1161
  @{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
  1162
  @{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
  1163
  @{thm[mode=Axiom] bs6}\qquad
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
  1164
  @{thm[mode=Axiom] bs7}\\
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
  1165
  @{thm[mode=Rule] bs8[of "rs\<^sub>1" "rs\<^sub>2"]}\\
410
9261d980225d updated papers
Christian Urban <christian.urban@kcl.ac.uk>
parents: 405
diff changeset
  1166
  %@ { t hm[mode=Axiom] ss1}\qquad
402
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
  1167
  @{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
  1168
  @{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
  1169
  @{thm[mode=Axiom] ss4}\qquad
402
1612f2a77bf6 more definitions in the paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 400
diff changeset
  1170
  @{thm[mode=Axiom] ss5[of "bs" "rs\<^sub>1" "rs\<^sub>2"]}\\
423
b7199d6c672d updated paper
Christian Urban <christian.urban@kcl.ac.uk>
parents: 420
diff changeset
  1171
  @{thm[mode=Rule] ss6[of "r\<^sub>2" "r\<^sub>1" "rs\<^sub>1" "rs\<^sub>2" "rs\<^sub>3"]}\\
398
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
  1172
  \end{tabular}
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
  1173
  \end{center}
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
  1174
  \caption{???}\label{SimpRewrites}
dac6d27c99c6 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 397
diff changeset
  1175
  \end{figure}
397
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
  1176
*}
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
  1177
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
  1178
section {* Bound - NO *}
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
  1179
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
  1180
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
  1181
section {* Conclusion *}
e1b74d618f1b updated Sizebound4
Christian Urban <christian.urban@kcl.ac.uk>
parents: 396
diff changeset
  1182
396
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
  1183
text {*
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
  1184
424
2416fdec6396 updated
Christian Urban <christian.urban@kcl.ac.uk>
parents: 423
diff changeset
  1185
396
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
  1186
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
  1187
%%\bibliographystyle{plain}
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
  1188
\bibliography{root}
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
  1189
*}
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
  1190
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
  1191
(*<*)
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
  1192
end
cc8e231529fb added ITP paper
Christian Urban <christian.urban@kcl.ac.uk>
parents:
diff changeset
  1193
(*>*)