thys2/Journal/Paper.thy~
author Chengsong
Sat, 08 Jan 2022 15:26:33 +0000
changeset 384 f5866f1d6a59
child 389 d4b3b0f942f4
permissions -rw-r--r--
from christian
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
384
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
     1
(*<*)
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
     2
theory Paper
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
     3
imports 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
     4
   "../Lexer"
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
     5
   "../Simplifying" 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
     6
   "../Positions"
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
     7
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
     8
   "../SizeBound" 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
     9
   "HOL-Library.LaTeXsugar"
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    10
begin
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    11
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    12
lemma Suc_0_fold:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    13
  "Suc 0 = 1"
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    14
by simp
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    15
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    16
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    17
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    18
declare [[show_question_marks = false]]
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    19
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    20
syntax (latex output)
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    21
  "_Collect" :: "pttrn => bool => 'a set"              ("(1{_ \<^latex>\<open>\\mbox{\\boldmath$\\mid$}\<close> _})")
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    22
  "_CollectIn" :: "pttrn => 'a set => bool => 'a set"   ("(1{_ \<in> _ |e _})")
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    23
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    24
syntax
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    25
  "_Not_Ex" :: "idts \<Rightarrow> bool \<Rightarrow> bool"  ("(3\<nexists>_.a./ _)" [0, 10] 10)
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    26
  "_Not_Ex1" :: "pttrn \<Rightarrow> bool \<Rightarrow> bool"  ("(3\<nexists>!_.a./ _)" [0, 10] 10)
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    27
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    28
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    29
abbreviation 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    30
  "der_syn r c \<equiv> der c r"
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    31
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    32
abbreviation 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    33
  "ders_syn r s \<equiv> ders s r"
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    34
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    35
  abbreviation 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    36
  "bder_syn r c \<equiv> bder c r"
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    37
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    38
abbreviation 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    39
  "bders_syn r s \<equiv> bders r s"
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    40
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    41
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    42
abbreviation
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    43
  "nprec v1 v2 \<equiv> \<not>(v1 :\<sqsubset>val v2)"
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    44
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    45
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    46
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    47
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    48
notation (latex output)
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    49
  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
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    50
  Cons ("_\<^latex>\<open>\\mbox{$\\,$}\<close>::\<^latex>\<open>\\mbox{$\\,$}\<close>_" [75,73] 73) and  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    51
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    52
  ZERO ("\<^bold>0" 81) and 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    53
  ONE ("\<^bold>1" 81) and 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    54
  CH ("_" [1000] 80) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    55
  ALT ("_ + _" [77,77] 78) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    56
  SEQ ("_ \<cdot> _" [77,77] 78) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    57
  STAR ("_\<^sup>\<star>" [79] 78) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    58
  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    59
  val.Void ("Empty" 78) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    60
  val.Char ("Char _" [1000] 78) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    61
  val.Left ("Left _" [79] 78) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    62
  val.Right ("Right _" [1000] 78) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    63
  val.Seq ("Seq _ _" [79,79] 78) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    64
  val.Stars ("Stars _" [79] 78) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    65
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    66
  L ("L'(_')" [10] 78) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    67
  LV ("LV _ _" [80,73] 78) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    68
  der_syn ("_\\_" [79, 1000] 76) and  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    69
  ders_syn ("_\\_" [79, 1000] 76) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    70
  flat ("|_|" [75] 74) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    71
  flats ("|_|" [72] 74) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    72
  Sequ ("_ @ _" [78,77] 63) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    73
  injval ("inj _ _ _" [79,77,79] 76) and 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    74
  mkeps ("mkeps _" [79] 76) and 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    75
  length ("len _" [73] 73) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    76
  intlen ("len _" [73] 73) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    77
  set ("_" [73] 73) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    78
 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    79
  Prf ("_ : _" [75,75] 75) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    80
  Posix ("'(_, _') \<rightarrow> _" [63,75,75] 75) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    81
 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    82
  lexer ("lexer _ _" [78,78] 77) and 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    83
  F_RIGHT ("F\<^bsub>Right\<^esub> _") and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    84
  F_LEFT ("F\<^bsub>Left\<^esub> _") and  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    85
  F_ALT ("F\<^bsub>Alt\<^esub> _ _") and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    86
  F_SEQ1 ("F\<^bsub>Seq1\<^esub> _ _") and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    87
  F_SEQ2 ("F\<^bsub>Seq2\<^esub> _ _") and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    88
  F_SEQ ("F\<^bsub>Seq\<^esub> _ _") and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    89
  simp_SEQ ("simp\<^bsub>Seq\<^esub> _ _" [1000, 1000] 1) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    90
  simp_ALT ("simp\<^bsub>Alt\<^esub> _ _" [1000, 1000] 1) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    91
  slexer ("lexer\<^sup>+" 1000) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    92
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    93
  at ("_\<^latex>\<open>\\mbox{$\\downharpoonleft$}\<close>\<^bsub>_\<^esub>") and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    94
  lex_list ("_ \<prec>\<^bsub>lex\<^esub> _") and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    95
  PosOrd ("_ \<prec>\<^bsub>_\<^esub> _" [77,77,77] 77) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    96
  PosOrd_ex ("_ \<prec> _" [77,77] 77) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    97
  PosOrd_ex_eq ("_ \<^latex>\<open>\\mbox{$\\preccurlyeq$}\<close> _" [77,77] 77) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    98
  pflat_len ("\<parallel>_\<parallel>\<^bsub>_\<^esub>") and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
    99
  nprec ("_ \<^latex>\<open>\\mbox{$\\not\\prec$}\<close> _" [77,77] 77) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   100
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   101
  bder_syn ("_\<^latex>\<open>\\mbox{$\\bbslash$}\<close>_" [79, 1000] 76) and  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   102
  bders_syn ("_\<^latex>\<open>\\mbox{$\\bbslash$}\<close>_" [79, 1000] 76) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   103
  intern ("_\<^latex>\<open>\\mbox{$^\\uparrow$}\<close>" [900] 80) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   104
  erase ("_\<^latex>\<open>\\mbox{$^\\downarrow$}\<close>" [1000] 74) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   105
  bnullable ("nullable\<^latex>\<open>\\mbox{$_b$}\<close> _" [1000] 80) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   106
  bmkeps ("mkeps\<^latex>\<open>\\mbox{$_b$}\<close> _" [1000] 80) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   107
  blexer ("lexer\<^latex>\<open>\\mbox{$_b$}\<close> _ _" [77, 77] 80) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   108
  code ("code _" [79] 74) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   109
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   110
  DUMMY ("\<^latex>\<open>\\underline{\\hspace{2mm}}\<close>")
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   111
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   112
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   113
definition 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   114
  "match r s \<equiv> nullable (ders s r)"
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   115
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   116
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   117
lemma LV_STAR_ONE_empty: 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   118
  shows "LV (STAR ONE) [] = {Stars []}"
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   119
by(auto simp add: LV_def elim: Prf.cases intro: Prf.intros)
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   120
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   121
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   122
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   123
(*
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   124
comments not implemented
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   125
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   126
p9. The condition "not exists s3 s4..." appears often enough (in particular in
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   127
the proof of Lemma 3) to warrant a definition.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   128
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   129
*)
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   130
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   131
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   132
(*>*)
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   133
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   134
section\<open>Core of the proof\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   135
text \<open>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   136
This paper builds on previous work by Ausaf and Urban using 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   137
regular expression'd bit-coded derivatives to do lexing that 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   138
is both fast and satisfies the POSIX specification.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   139
In their work, a bit-coded algorithm introduced by Sulzmann and Lu
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   140
was formally verified in Isabelle, by a very clever use of
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   141
flex function and retrieve to carefully mimic the way a value is 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   142
built up by the injection funciton.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   143
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   144
In the previous work, Ausaf and Urban established the below equality:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   145
\begin{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   146
@{thm [mode=IfThen] MAIN_decode}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   147
\end{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   148
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   149
This lemma establishes a link with the lexer without bit-codes.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   150
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   151
With it we get the correctness of bit-coded algorithm.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   152
\begin{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   153
@{thm [mode=IfThen] blexer_correctness}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   154
\end{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   155
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   156
However what is not certain is whether we can add simplification
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   157
to the bit-coded algorithm, without breaking the correct lexing output.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   158
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   159
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   160
The reason that we do need to add a simplification phase
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   161
after each derivative step of  $\textit{blexer}$ is
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   162
because it produces intermediate
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   163
regular expressions that can grow exponentially.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   164
For example, the regular expression $(a+aa)^*$ after taking
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   165
derivative against just 10 $a$s will have size 8192.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   166
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   167
%TODO: add figure for this?
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   168
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   169
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   170
Therefore, we insert a simplification phase
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   171
after each derivation step, as defined below:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   172
\begin{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   173
@{thm blexer_simp_def}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   174
\end{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   175
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   176
The simplification function is given as follows:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   177
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   178
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   179
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   180
  @{thm (lhs) bsimp.simps(1)[of  "bs" "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bsimp.simps(1)[of  "bs" "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   181
  @{thm (lhs) bsimp.simps(2)} & $\dn$ & @{thm (rhs) bsimp.simps(2)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   182
  @{thm (lhs) bsimp.simps(3)} & $\dn$ & @{thm (rhs) bsimp.simps(3)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   183
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   184
\end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   185
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   186
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   187
And the two helper functions are:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   188
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   189
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   190
  @{thm (lhs) bsimp_AALTs.simps(2)[of  "bs\<^sub>1" "r" ]} & $\dn$ & @{thm (rhs) bsimp.simps(1)[of  "bs\<^sub>1" "r" ]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   191
  @{thm (lhs) bsimp_AALTs.simps(2)} & $\dn$ & @{thm (rhs) bsimp.simps(2)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   192
  @{thm (lhs) bsimp_AALTs.simps(3)} & $\dn$ & @{thm (rhs) bsimp.simps(3)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   193
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   194
\end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   195
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   196
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   197
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   198
This might sound trivial in the case of producing a YES/NO answer,
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   199
but once we require a lexing output to be produced (which is required
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   200
in applications like compiler front-end, malicious attack domain extraction, 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   201
etc.), it is not straightforward if we still extract what is needed according
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   202
to the POSIX standard.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   203
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   204
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   205
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   206
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   207
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   208
By simplification, we mean specifically the following rules:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   209
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   210
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   211
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   212
  @{thm[mode=Axiom] rrewrite.intros(1)[of "bs" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   213
  @{thm[mode=Axiom] rrewrite.intros(2)[of "bs" "r\<^sub>1"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   214
  @{thm[mode=Axiom] rrewrite.intros(3)[of  "bs" "bs\<^sub>1" "r\<^sub>1"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   215
  @{thm[mode=Rule] rrewrite.intros(4)[of  "r\<^sub>1" "r\<^sub>2" "bs" "r\<^sub>3"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   216
  @{thm[mode=Rule] rrewrite.intros(5)[of "r\<^sub>3" "r\<^sub>4" "bs" "r\<^sub>1"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   217
  @{thm[mode=Rule] rrewrite.intros(6)[of "r" "r'" "bs" "rs\<^sub>1" "rs\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   218
  @{thm[mode=Axiom] rrewrite.intros(7)[of "bs" "rs\<^sub>a" "rs\<^sub>b"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   219
  @{thm[mode=Axiom] rrewrite.intros(8)[of "bs" "rs\<^sub>a" "bs\<^sub>1" "rs\<^sub>1" "rs\<^sub>b"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   220
  @{thm[mode=Axiom] rrewrite.intros(9)[of "bs" "bs\<^sub>1" "rs"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   221
  @{thm[mode=Axiom] rrewrite.intros(10)[of "bs" ]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   222
  @{thm[mode=Axiom] rrewrite.intros(11)[of "bs" "r\<^sub>1"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   223
  @{thm[mode=Rule] rrewrite.intros(12)[of "a\<^sub>1" "a\<^sub>2" "bs" "rs\<^sub>a" "rs\<^sub>b" "rs\<^sub>c"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   224
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   225
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   226
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   227
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   228
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   229
And these can be made compact by the following simplification function:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   230
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   231
where the function $\mathit{bsimp_AALTs}$
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   232
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   233
The core idea of the proof is that two regular expressions,
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   234
if "isomorphic" up to a finite number of rewrite steps, will
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   235
remain "isomorphic" when we take the same sequence of
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   236
derivatives on both of them.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   237
This can be expressed by the following rewrite relation lemma:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   238
\begin{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   239
@{thm [mode=IfThen] central}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   240
\end{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   241
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   242
This isomorphic relation implies a property that leads to the 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   243
correctness result: 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   244
if two (nullable) regular expressions are "rewritable" in many steps
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   245
from one another, 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   246
then a call to function $\textit{bmkeps}$ gives the same
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   247
bit-sequence :
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   248
\begin{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   249
@{thm [mode=IfThen] rewrites_bmkeps}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   250
\end{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   251
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   252
Given the same bit-sequence, the decode function
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   253
will give out the same value, which is the output
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   254
of both lexers:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   255
\begin{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   256
@{thm blexer_def}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   257
\end{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   258
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   259
\begin{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   260
@{thm blexer_simp_def}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   261
\end{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   262
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   263
And that yields the correctness result:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   264
\begin{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   265
@{thm blexersimp_correctness}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   266
\end{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   267
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   268
The nice thing about the above
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   269
\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   270
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   271
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   272
section \<open> Additional Simp Rules?\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   273
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   274
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   275
text  \<open>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   276
One question someone would ask is:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   277
can we add more "atomic" simplification/rewriting rules,
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   278
so the simplification is even more aggressive, making
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   279
the intermediate results smaller, and therefore more space-efficient? 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   280
For example, one might want to do open up alternatives who is a child
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   281
of a sequence:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   282
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   283
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   284
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   285
    @{thm[mode=Rule] aggressive.intros(1)[of "bs" "bs1" "rs" "r"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   286
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   287
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   288
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   289
This rule allows us to simplify \mbox{@{term "(ALT (SEQ (ALT a b) c)  (SEQ a c))"}}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   290
 into  \mbox{@{term "ALT (SEQ a c)  (SEQ b c)"}},
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   291
which is cannot be done under the \mbox{ \<leadsto>} rule because only alternatives which are
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   292
children of another alternative can be spilled out.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   293
\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   294
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   295
(*This rule allows us to simplify \mbox{@{term "ALT (SEQ (ALT a b) c) (SEQ a c)"}}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   296
 into  \mbox{@{term "ALT (SEQ a c)  (SEQ b c)"}},
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   297
which is cannot be done under the  \<leadsto> rule because only alternatives which are 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   298
children of another alternative can be spilled out.*)
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   299
section \<open>Introduction\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   300
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   301
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   302
text \<open>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   303
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   304
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   305
Brzozowski \cite{Brzozowski1964} introduced the notion of the {\em
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   306
derivative} @{term "der c r"} of a regular expression \<open>r\<close> w.r.t.\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   307
a character~\<open>c\<close>, and showed that it gave a simple solution to the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   308
problem of matching a string @{term s} with a regular expression @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   309
r}: if the derivative of @{term r} w.r.t.\ (in succession) all the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   310
characters of the string matches the empty string, then @{term r}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   311
matches @{term s} (and {\em vice versa}). The derivative has the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   312
property (which may almost be regarded as its specification) that, for
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   313
every string @{term s} and regular expression @{term r} and character
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   314
@{term c}, one has @{term "cs \<in> L(r)"} if and only if \mbox{@{term "s \<in> L(der c r)"}}. 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   315
The beauty of Brzozowski's derivatives is that
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   316
they are neatly expressible in any functional language, and easily
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   317
definable and reasoned about in theorem provers---the definitions just
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   318
consist of inductive datatypes and simple recursive functions. A
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   319
mechanised correctness proof of Brzozowski's matcher in for example HOL4
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   320
has been mentioned by Owens and Slind~\cite{Owens2008}. Another one in
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   321
Isabelle/HOL is part of the work by Krauss and Nipkow \cite{Krauss2011}.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   322
And another one in Coq is given by Coquand and Siles \cite{Coquand2012}.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   323
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   324
If a regular expression matches a string, then in general there is more
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   325
than one way of how the string is matched. There are two commonly used
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   326
disambiguation strategies to generate a unique answer: one is called
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   327
GREEDY matching \cite{Frisch2004} and the other is POSIX
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   328
matching~\cite{POSIX,Kuklewicz,OkuiSuzuki2010,Sulzmann2014,Vansummeren2006}.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   329
For example consider the string @{term xy} and the regular expression
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   330
\mbox{@{term "STAR (ALT (ALT x y) xy)"}}. Either the string can be
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   331
matched in two `iterations' by the single letter-regular expressions
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   332
@{term x} and @{term y}, or directly in one iteration by @{term xy}. The
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   333
first case corresponds to GREEDY matching, which first matches with the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   334
left-most symbol and only matches the next symbol in case of a mismatch
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   335
(this is greedy in the sense of preferring instant gratification to
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   336
delayed repletion). The second case is POSIX matching, which prefers the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   337
longest match.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   338
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   339
In the context of lexing, where an input string needs to be split up
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   340
into a sequence of tokens, POSIX is the more natural disambiguation
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   341
strategy for what programmers consider basic syntactic building blocks
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   342
in their programs.  These building blocks are often specified by some
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   343
regular expressions, say \<open>r\<^bsub>key\<^esub>\<close> and \<open>r\<^bsub>id\<^esub>\<close> for recognising keywords and identifiers,
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   344
respectively. There are a few underlying (informal) rules behind
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   345
tokenising a string in a POSIX \cite{POSIX} fashion:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   346
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   347
\begin{itemize} 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   348
\item[$\bullet$] \emph{The Longest Match Rule} (or \emph{``{M}aximal {M}unch {R}ule''}):
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   349
The longest initial substring matched by any regular expression is taken as
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   350
next token.\smallskip
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   351
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   352
\item[$\bullet$] \emph{Priority Rule:}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   353
For a particular longest initial substring, the first (leftmost) regular expression
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   354
that can match determines the token.\smallskip
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   355
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   356
\item[$\bullet$] \emph{Star Rule:} A subexpression repeated by ${}^\star$ shall 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   357
not match an empty string unless this is the only match for the repetition.\smallskip
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   358
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   359
\item[$\bullet$] \emph{Empty String Rule:} An empty string shall be considered to 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   360
be longer than no match at all.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   361
\end{itemize}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   362
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   363
\noindent Consider for example a regular expression \<open>r\<^bsub>key\<^esub>\<close> for recognising keywords such as \<open>if\<close>,
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   364
\<open>then\<close> and so on; and \<open>r\<^bsub>id\<^esub>\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   365
recognising identifiers (say, a single character followed by
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   366
characters or numbers).  Then we can form the regular expression
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   367
\<open>(r\<^bsub>key\<^esub> + r\<^bsub>id\<^esub>)\<^sup>\<star>\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   368
and use POSIX matching to tokenise strings, say \<open>iffoo\<close> and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   369
\<open>if\<close>.  For \<open>iffoo\<close> we obtain by the Longest Match Rule
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   370
a single identifier token, not a keyword followed by an
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   371
identifier. For \<open>if\<close> we obtain by the Priority Rule a keyword
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   372
token, not an identifier token---even if \<open>r\<^bsub>id\<^esub>\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   373
matches also. By the Star Rule we know \<open>(r\<^bsub>key\<^esub> +
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   374
r\<^bsub>id\<^esub>)\<^sup>\<star>\<close> matches \<open>iffoo\<close>,
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   375
respectively \<open>if\<close>, in exactly one `iteration' of the star. The
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   376
Empty String Rule is for cases where, for example, the regular expression 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   377
\<open>(a\<^sup>\<star>)\<^sup>\<star>\<close> matches against the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   378
string \<open>bc\<close>. Then the longest initial matched substring is the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   379
empty string, which is matched by both the whole regular expression
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   380
and the parenthesised subexpression.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   381
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   382
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   383
One limitation of Brzozowski's matcher is that it only generates a
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   384
YES/NO answer for whether a string is being matched by a regular
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   385
expression.  Sulzmann and Lu~\cite{Sulzmann2014} extended this matcher
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   386
to allow generation not just of a YES/NO answer but of an actual
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   387
matching, called a [lexical] {\em value}. Assuming a regular
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   388
expression matches a string, values encode the information of
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   389
\emph{how} the string is matched by the regular expression---that is,
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   390
which part of the string is matched by which part of the regular
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   391
expression. For this consider again the string \<open>xy\<close> and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   392
the regular expression \mbox{\<open>(x + (y + xy))\<^sup>\<star>\<close>}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   393
(this time fully parenthesised). We can view this regular expression
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   394
as tree and if the string \<open>xy\<close> is matched by two Star
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   395
`iterations', then the \<open>x\<close> is matched by the left-most
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   396
alternative in this tree and the \<open>y\<close> by the right-left alternative. This
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   397
suggests to record this matching as
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   398
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   399
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   400
@{term "Stars [Left(Char x), Right(Left(Char y))]"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   401
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   402
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   403
\noindent where @{const Stars}, \<open>Left\<close>, \<open>Right\<close> and \<open>Char\<close> are constructors for values. \<open>Stars\<close> records how many
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   404
iterations were used; \<open>Left\<close>, respectively \<open>Right\<close>, which
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   405
alternative is used. This `tree view' leads naturally to the idea that
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   406
regular expressions act as types and values as inhabiting those types
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   407
(see, for example, \cite{HosoyaVouillonPierce2005}).  The value for
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   408
matching \<open>xy\<close> in a single `iteration', i.e.~the POSIX value,
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   409
would look as follows
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   410
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   411
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   412
@{term "Stars [Seq (Char x) (Char y)]"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   413
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   414
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   415
\noindent where @{const Stars} has only a single-element list for the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   416
single iteration and @{const Seq} indicates that @{term xy} is matched 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   417
by a sequence regular expression.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   418
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   419
%, which we will in what follows 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   420
%write more formally as @{term "SEQ x y"}.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   421
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   422
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   423
Sulzmann and Lu give a simple algorithm to calculate a value that
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   424
appears to be the value associated with POSIX matching.  The challenge
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   425
then is to specify that value, in an algorithm-independent fashion,
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   426
and to show that Sulzmann and Lu's derivative-based algorithm does
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   427
indeed calculate a value that is correct according to the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   428
specification.  The answer given by Sulzmann and Lu
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   429
\cite{Sulzmann2014} is to define a relation (called an ``order
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   430
relation'') on the set of values of @{term r}, and to show that (once
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   431
a string to be matched is chosen) there is a maximum element and that
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   432
it is computed by their derivative-based algorithm. This proof idea is
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   433
inspired by work of Frisch and Cardelli \cite{Frisch2004} on a GREEDY
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   434
regular expression matching algorithm. However, we were not able to
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   435
establish transitivity and totality for the ``order relation'' by
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   436
Sulzmann and Lu.  There are some inherent problems with their approach
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   437
(of which some of the proofs are not published in
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   438
\cite{Sulzmann2014}); perhaps more importantly, we give in this paper
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   439
a simple inductive (and algorithm-independent) definition of what we
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   440
call being a {\em POSIX value} for a regular expression @{term r} and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   441
a string @{term s}; we show that the algorithm by Sulzmann and Lu
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   442
computes such a value and that such a value is unique. Our proofs are
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   443
both done by hand and checked in Isabelle/HOL.  The experience of
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   444
doing our proofs has been that this mechanical checking was absolutely
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   445
essential: this subject area has hidden snares. This was also noted by
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   446
Kuklewicz \cite{Kuklewicz} who found that nearly all POSIX matching
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   447
implementations are ``buggy'' \cite[Page 203]{Sulzmann2014} and by
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   448
Grathwohl et al \cite[Page 36]{CrashCourse2014} who wrote:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   449
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   450
\begin{quote}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   451
\it{}``The POSIX strategy is more complicated than the greedy because of 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   452
the dependence on information about the length of matched strings in the 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   453
various subexpressions.''
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   454
\end{quote}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   455
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   456
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   457
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   458
\noindent {\bf Contributions:} We have implemented in Isabelle/HOL the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   459
derivative-based regular expression matching algorithm of
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   460
Sulzmann and Lu \cite{Sulzmann2014}. We have proved the correctness of this
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   461
algorithm according to our specification of what a POSIX value is (inspired
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   462
by work of Vansummeren \cite{Vansummeren2006}). Sulzmann
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   463
and Lu sketch in \cite{Sulzmann2014} an informal correctness proof: but to
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   464
us it contains unfillable gaps.\footnote{An extended version of
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   465
\cite{Sulzmann2014} is available at the website of its first author; this
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   466
extended version already includes remarks in the appendix that their
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   467
informal proof contains gaps, and possible fixes are not fully worked out.}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   468
Our specification of a POSIX value consists of a simple inductive definition
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   469
that given a string and a regular expression uniquely determines this value.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   470
We also show that our definition is equivalent to an ordering 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   471
of values based on positions by Okui and Suzuki \cite{OkuiSuzuki2010}.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   472
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   473
%Derivatives as calculated by Brzozowski's method are usually more complex
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   474
%regular expressions than the initial one; various optimisations are
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   475
%possible. We prove the correctness when simplifications of @{term "ALT ZERO r"}, 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   476
%@{term "ALT r ZERO"}, @{term "SEQ ONE r"} and @{term "SEQ r ONE"} to
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   477
%@{term r} are applied. 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   478
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   479
We extend our results to ??? Bitcoded version??
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   480
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   481
\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   482
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   483
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   484
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   485
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   486
section \<open>Preliminaries\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   487
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   488
text \<open>\noindent Strings in Isabelle/HOL are lists of characters with
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   489
the empty string being represented by the empty list, written @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   490
"[]"}, and list-cons being written as @{term "DUMMY # DUMMY"}. Often
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   491
we use the usual bracket notation for lists also for strings; for
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   492
example a string consisting of just a single character @{term c} is
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   493
written @{term "[c]"}. We use the usual definitions for 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   494
\emph{prefixes} and \emph{strict prefixes} of strings.  By using the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   495
type @{type char} for characters we have a supply of finitely many
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   496
characters roughly corresponding to the ASCII character set. Regular
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   497
expressions are defined as usual as the elements of the following
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   498
inductive datatype:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   499
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   500
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   501
  \<open>r :=\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   502
  @{const "ZERO"} $\mid$
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   503
  @{const "ONE"} $\mid$
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   504
  @{term "CH c"} $\mid$
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   505
  @{term "ALT r\<^sub>1 r\<^sub>2"} $\mid$
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   506
  @{term "SEQ r\<^sub>1 r\<^sub>2"} $\mid$
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   507
  @{term "STAR r"} 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   508
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   509
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   510
  \noindent where @{const ZERO} stands for the regular expression that does
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   511
  not match any string, @{const ONE} for the regular expression that matches
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   512
  only the empty string and @{term c} for matching a character literal. The
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   513
  language of a regular expression is also defined as usual by the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   514
  recursive function @{term L} with the six clauses:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   515
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   516
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   517
  \begin{tabular}{l@ {\hspace{4mm}}rcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   518
  \textit{(1)} & @{thm (lhs) L.simps(1)} & $\dn$ & @{thm (rhs) L.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   519
  \textit{(2)} & @{thm (lhs) L.simps(2)} & $\dn$ & @{thm (rhs) L.simps(2)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   520
  \textit{(3)} & @{thm (lhs) L.simps(3)} & $\dn$ & @{thm (rhs) L.simps(3)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   521
  \textit{(4)} & @{thm (lhs) L.simps(4)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   522
        @{thm (rhs) L.simps(4)[of "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   523
  \textit{(5)} & @{thm (lhs) L.simps(5)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   524
        @{thm (rhs) L.simps(5)[of "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   525
  \textit{(6)} & @{thm (lhs) L.simps(6)} & $\dn$ & @{thm (rhs) L.simps(6)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   526
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   527
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   528
  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   529
  \noindent In clause \textit{(4)} we use the operation @{term "DUMMY ;;
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   530
  DUMMY"} for the concatenation of two languages (it is also list-append for
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   531
  strings). We use the star-notation for regular expressions and for
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   532
  languages (in the last clause above). The star for languages is defined
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   533
  inductively by two clauses: \<open>(i)\<close> the empty string being in
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   534
  the star of a language and \<open>(ii)\<close> if @{term "s\<^sub>1"} is in a
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   535
  language and @{term "s\<^sub>2"} in the star of this language, then also @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   536
  "s\<^sub>1 @ s\<^sub>2"} is in the star of this language. It will also be convenient
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   537
  to use the following notion of a \emph{semantic derivative} (or \emph{left
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   538
  quotient}) of a language defined as
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   539
  %
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   540
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   541
  @{thm Der_def}\;.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   542
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   543
 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   544
  \noindent
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   545
  For semantic derivatives we have the following equations (for example
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   546
  mechanically proved in \cite{Krauss2011}):
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   547
  %
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   548
  \begin{equation}\label{SemDer}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   549
  \begin{array}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   550
  @{thm (lhs) Der_null}  & \dn & @{thm (rhs) Der_null}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   551
  @{thm (lhs) Der_empty}  & \dn & @{thm (rhs) Der_empty}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   552
  @{thm (lhs) Der_char}  & \dn & @{thm (rhs) Der_char}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   553
  @{thm (lhs) Der_union}  & \dn & @{thm (rhs) Der_union}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   554
  @{thm (lhs) Der_Sequ}  & \dn & @{thm (rhs) Der_Sequ}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   555
  @{thm (lhs) Der_star}  & \dn & @{thm (rhs) Der_star}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   556
  \end{array}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   557
  \end{equation}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   558
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   559
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   560
  \noindent \emph{\Brz's derivatives} of regular expressions
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   561
  \cite{Brzozowski1964} can be easily defined by two recursive functions:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   562
  the first is from regular expressions to booleans (implementing a test
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   563
  when a regular expression can match the empty string), and the second
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   564
  takes a regular expression and a character to a (derivative) regular
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   565
  expression:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   566
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   567
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   568
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   569
  @{thm (lhs) nullable.simps(1)} & $\dn$ & @{thm (rhs) nullable.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   570
  @{thm (lhs) nullable.simps(2)} & $\dn$ & @{thm (rhs) nullable.simps(2)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   571
  @{thm (lhs) nullable.simps(3)} & $\dn$ & @{thm (rhs) nullable.simps(3)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   572
  @{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"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   573
  @{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"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   574
  @{thm (lhs) nullable.simps(6)} & $\dn$ & @{thm (rhs) nullable.simps(6)}\medskip\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   575
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   576
%  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   577
%  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   578
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   579
%  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   580
%  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   581
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   582
  @{thm (lhs) der.simps(1)} & $\dn$ & @{thm (rhs) der.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   583
  @{thm (lhs) der.simps(2)} & $\dn$ & @{thm (rhs) der.simps(2)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   584
  @{thm (lhs) der.simps(3)} & $\dn$ & @{thm (rhs) der.simps(3)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   585
  @{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"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   586
  @{thm (lhs) der.simps(5)[of c "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) der.simps(5)[of c "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   587
  @{thm (lhs) der.simps(6)} & $\dn$ & @{thm (rhs) der.simps(6)}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   588
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   589
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   590
 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   591
  \noindent
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   592
  We may extend this definition to give derivatives w.r.t.~strings:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   593
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   594
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   595
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   596
  @{thm (lhs) ders.simps(1)} & $\dn$ & @{thm (rhs) ders.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   597
  @{thm (lhs) ders.simps(2)} & $\dn$ & @{thm (rhs) ders.simps(2)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   598
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   599
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   600
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   601
  \noindent Given the equations in \eqref{SemDer}, it is a relatively easy
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   602
  exercise in mechanical reasoning to establish that
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   603
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   604
  \begin{proposition}\label{derprop}\mbox{}\\ 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   605
  \begin{tabular}{ll}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   606
  \textit{(1)} & @{thm (lhs) nullable_correctness} if and only if
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   607
  @{thm (rhs) nullable_correctness}, and \\ 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   608
  \textit{(2)} & @{thm[mode=IfThen] der_correctness}.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   609
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   610
  \end{proposition}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   611
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   612
  \noindent With this in place it is also very routine to prove that the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   613
  regular expression matcher defined as
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   614
  %
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   615
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   616
  @{thm match_def}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   617
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   618
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   619
  \noindent gives a positive answer if and only if @{term "s \<in> L r"}.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   620
  Consequently, this regular expression matching algorithm satisfies the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   621
  usual specification for regular expression matching. While the matcher
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   622
  above calculates a provably correct YES/NO answer for whether a regular
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   623
  expression matches a string or not, the novel idea of Sulzmann and Lu
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   624
  \cite{Sulzmann2014} is to append another phase to this algorithm in order
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   625
  to calculate a [lexical] value. We will explain the details next.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   626
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   627
\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   628
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   629
section \<open>POSIX Regular Expression Matching\label{posixsec}\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   630
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   631
text \<open>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   632
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   633
  There have been many previous works that use values for encoding 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   634
  \emph{how} a regular expression matches a string.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   635
  The clever idea by Sulzmann and Lu \cite{Sulzmann2014} is to 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   636
  define a function on values that mirrors (but inverts) the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   637
  construction of the derivative on regular expressions. \emph{Values}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   638
  are defined as the inductive datatype
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   639
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   640
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   641
  \<open>v :=\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   642
  @{const "Void"} $\mid$
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   643
  @{term "val.Char c"} $\mid$
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   644
  @{term "Left v"} $\mid$
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   645
  @{term "Right v"} $\mid$
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   646
  @{term "Seq v\<^sub>1 v\<^sub>2"} $\mid$ 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   647
  @{term "Stars vs"} 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   648
  \end{center}  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   649
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   650
  \noindent where we use @{term vs} to stand for a list of
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   651
  values. (This is similar to the approach taken by Frisch and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   652
  Cardelli for GREEDY matching \cite{Frisch2004}, and Sulzmann and Lu
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   653
  for POSIX matching \cite{Sulzmann2014}). The string underlying a
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   654
  value can be calculated by the @{const flat} function, written
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   655
  @{term "flat DUMMY"} and defined as:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   656
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   657
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   658
  \begin{tabular}[t]{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   659
  @{thm (lhs) flat.simps(1)} & $\dn$ & @{thm (rhs) flat.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   660
  @{thm (lhs) flat.simps(2)} & $\dn$ & @{thm (rhs) flat.simps(2)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   661
  @{thm (lhs) flat.simps(3)} & $\dn$ & @{thm (rhs) flat.simps(3)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   662
  @{thm (lhs) flat.simps(4)} & $\dn$ & @{thm (rhs) flat.simps(4)}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   663
  \end{tabular}\hspace{14mm}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   664
  \begin{tabular}[t]{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   665
  @{thm (lhs) flat.simps(5)[of "v\<^sub>1" "v\<^sub>2"]} & $\dn$ & @{thm (rhs) flat.simps(5)[of "v\<^sub>1" "v\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   666
  @{thm (lhs) flat.simps(6)} & $\dn$ & @{thm (rhs) flat.simps(6)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   667
  @{thm (lhs) flat.simps(7)} & $\dn$ & @{thm (rhs) flat.simps(7)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   668
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   669
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   670
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   671
  \noindent We will sometimes refer to the underlying string of a
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   672
  value as \emph{flattened value}.  We will also overload our notation and 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   673
  use @{term "flats vs"} for flattening a list of values and concatenating
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   674
  the resulting strings.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   675
  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   676
  Sulzmann and Lu define
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   677
  inductively an \emph{inhabitation relation} that associates values to
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   678
  regular expressions. We define this relation as
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   679
  follows:\footnote{Note that the rule for @{term Stars} differs from
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   680
  our earlier paper \cite{AusafDyckhoffUrban2016}. There we used the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   681
  original definition by Sulzmann and Lu which does not require that
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   682
  the values @{term "v \<in> set vs"} flatten to a non-empty
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   683
  string. The reason for introducing the more restricted version of
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   684
  lexical values is convenience later on when reasoning about an
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   685
  ordering relation for values.}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   686
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   687
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   688
  \begin{tabular}{c@ {\hspace{12mm}}c}\label{prfintros}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   689
  \\[-8mm]
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   690
  @{thm[mode=Axiom] Prf.intros(4)} & 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   691
  @{thm[mode=Axiom] Prf.intros(5)[of "c"]}\\[4mm]
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   692
  @{thm[mode=Rule] Prf.intros(2)[of "v\<^sub>1" "r\<^sub>1" "r\<^sub>2"]} &
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   693
  @{thm[mode=Rule] Prf.intros(3)[of "v\<^sub>2" "r\<^sub>1" "r\<^sub>2"]}\\[4mm]
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   694
  @{thm[mode=Rule] Prf.intros(1)[of "v\<^sub>1" "r\<^sub>1" "v\<^sub>2" "r\<^sub>2"]}  &
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   695
  @{thm[mode=Rule] Prf.intros(6)[of "vs"]}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   696
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   697
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   698
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   699
  \noindent where in the clause for @{const "Stars"} we use the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   700
  notation @{term "v \<in> set vs"} for indicating that \<open>v\<close> is a
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   701
  member in the list \<open>vs\<close>.  We require in this rule that every
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   702
  value in @{term vs} flattens to a non-empty string. The idea is that
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   703
  @{term "Stars"}-values satisfy the informal Star Rule (see Introduction)
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   704
  where the $^\star$ does not match the empty string unless this is
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   705
  the only match for the repetition.  Note also that no values are
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   706
  associated with the regular expression @{term ZERO}, and that the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   707
  only value associated with the regular expression @{term ONE} is
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   708
  @{term Void}.  It is routine to establish how values ``inhabiting''
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   709
  a regular expression correspond to the language of a regular
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   710
  expression, namely
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   711
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   712
  \begin{proposition}\label{inhabs}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   713
  @{thm L_flat_Prf}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   714
  \end{proposition}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   715
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   716
  \noindent
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   717
  Given a regular expression \<open>r\<close> and a string \<open>s\<close>, we define the 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   718
  set of all \emph{Lexical Values} inhabited by \<open>r\<close> with the underlying string 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   719
  being \<open>s\<close>:\footnote{Okui and Suzuki refer to our lexical values 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   720
  as \emph{canonical values} in \cite{OkuiSuzuki2010}. The notion of \emph{non-problematic
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   721
  values} by Cardelli and Frisch \cite{Frisch2004} is related, but not identical
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   722
  to our lexical values.}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   723
  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   724
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   725
  @{thm LV_def}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   726
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   727
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   728
  \noindent The main property of @{term "LV r s"} is that it is alway finite.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   729
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   730
  \begin{proposition}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   731
  @{thm LV_finite}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   732
  \end{proposition}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   733
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   734
  \noindent This finiteness property does not hold in general if we
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   735
  remove the side-condition about @{term "flat v \<noteq> []"} in the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   736
  @{term Stars}-rule above. For example using Sulzmann and Lu's
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   737
  less restrictive definition, @{term "LV (STAR ONE) []"} would contain
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   738
  infinitely many values, but according to our more restricted
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   739
  definition only a single value, namely @{thm LV_STAR_ONE_empty}.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   740
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   741
  If a regular expression \<open>r\<close> matches a string \<open>s\<close>, then
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   742
  generally the set @{term "LV r s"} is not just a singleton set.  In
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   743
  case of POSIX matching the problem is to calculate the unique lexical value
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   744
  that satisfies the (informal) POSIX rules from the Introduction.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   745
  Graphically the POSIX value calculation algorithm by Sulzmann and Lu
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   746
  can be illustrated by the picture in Figure~\ref{Sulz} where the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   747
  path from the left to the right involving @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   748
  derivatives}/@{const nullable} is the first phase of the algorithm
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   749
  (calculating successive \Brz's derivatives) and @{const
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   750
  mkeps}/\<open>inj\<close>, the path from right to left, the second
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   751
  phase. This picture shows the steps required when a regular
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   752
  expression, say \<open>r\<^sub>1\<close>, matches the string @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   753
  "[a,b,c]"}. We first build the three derivatives (according to
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   754
  @{term a}, @{term b} and @{term c}). We then use @{const nullable}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   755
  to find out whether the resulting derivative regular expression
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   756
  @{term "r\<^sub>4"} can match the empty string. If yes, we call the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   757
  function @{const mkeps} that produces a value @{term "v\<^sub>4"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   758
  for how @{term "r\<^sub>4"} can match the empty string (taking into
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   759
  account the POSIX constraints in case there are several ways). This
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   760
  function is defined by the clauses:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   761
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   762
\begin{figure}[t]
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   763
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   764
\begin{tikzpicture}[scale=2,node distance=1.3cm,
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   765
                    every node/.style={minimum size=6mm}]
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   766
\node (r1)  {@{term "r\<^sub>1"}};
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   767
\node (r2) [right=of r1]{@{term "r\<^sub>2"}};
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   768
\draw[->,line width=1mm](r1)--(r2) node[above,midway] {@{term "der a DUMMY"}};
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   769
\node (r3) [right=of r2]{@{term "r\<^sub>3"}};
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   770
\draw[->,line width=1mm](r2)--(r3) node[above,midway] {@{term "der b DUMMY"}};
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   771
\node (r4) [right=of r3]{@{term "r\<^sub>4"}};
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   772
\draw[->,line width=1mm](r3)--(r4) node[above,midway] {@{term "der c DUMMY"}};
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   773
\draw (r4) node[anchor=west] {\;\raisebox{3mm}{@{term nullable}}};
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   774
\node (v4) [below=of r4]{@{term "v\<^sub>4"}};
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   775
\draw[->,line width=1mm](r4) -- (v4);
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   776
\node (v3) [left=of v4] {@{term "v\<^sub>3"}};
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   777
\draw[->,line width=1mm](v4)--(v3) node[below,midway] {\<open>inj r\<^sub>3 c\<close>};
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   778
\node (v2) [left=of v3]{@{term "v\<^sub>2"}};
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   779
\draw[->,line width=1mm](v3)--(v2) node[below,midway] {\<open>inj r\<^sub>2 b\<close>};
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   780
\node (v1) [left=of v2] {@{term "v\<^sub>1"}};
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   781
\draw[->,line width=1mm](v2)--(v1) node[below,midway] {\<open>inj r\<^sub>1 a\<close>};
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   782
\draw (r4) node[anchor=north west] {\;\raisebox{-8mm}{@{term "mkeps"}}};
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   783
\end{tikzpicture}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   784
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   785
\mbox{}\\[-13mm]
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   786
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   787
\caption{The two phases of the algorithm by Sulzmann \& Lu \cite{Sulzmann2014},
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   788
matching the string @{term "[a,b,c]"}. The first phase (the arrows from 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   789
left to right) is \Brz's matcher building successive derivatives. If the 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   790
last regular expression is @{term nullable}, then the functions of the 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   791
second phase are called (the top-down and right-to-left arrows): first 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   792
@{term mkeps} calculates a value @{term "v\<^sub>4"} witnessing
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   793
how the empty string has been recognised by @{term "r\<^sub>4"}. After
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   794
that the function @{term inj} ``injects back'' the characters of the string into
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   795
the values.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   796
\label{Sulz}}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   797
\end{figure} 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   798
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   799
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   800
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   801
  @{thm (lhs) mkeps.simps(1)} & $\dn$ & @{thm (rhs) mkeps.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   802
  @{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"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   803
  @{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"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   804
  @{thm (lhs) mkeps.simps(4)} & $\dn$ & @{thm (rhs) mkeps.simps(4)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   805
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   806
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   807
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   808
  \noindent Note that this function needs only to be partially defined,
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   809
  namely only for regular expressions that are nullable. In case @{const
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   810
  nullable} fails, the string @{term "[a,b,c]"} cannot be matched by @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   811
  "r\<^sub>1"} and the null value @{term "None"} is returned. Note also how this function
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   812
  makes some subtle choices leading to a POSIX value: for example if an
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   813
  alternative regular expression, say @{term "ALT r\<^sub>1 r\<^sub>2"}, can
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   814
  match the empty string and furthermore @{term "r\<^sub>1"} can match the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   815
  empty string, then we return a \<open>Left\<close>-value. The \<open>Right\<close>-value will only be returned if @{term "r\<^sub>1"} cannot match the empty
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   816
  string.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   817
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   818
  The most interesting idea from Sulzmann and Lu \cite{Sulzmann2014} is
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   819
  the construction of a value for how @{term "r\<^sub>1"} can match the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   820
  string @{term "[a,b,c]"} from the value how the last derivative, @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   821
  "r\<^sub>4"} in Fig.~\ref{Sulz}, can match the empty string. Sulzmann and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   822
  Lu achieve this by stepwise ``injecting back'' the characters into the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   823
  values thus inverting the operation of building derivatives, but on the level
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   824
  of values. The corresponding function, called @{term inj}, takes three
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   825
  arguments, a regular expression, a character and a value. For example in
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   826
  the first (or right-most) @{term inj}-step in Fig.~\ref{Sulz} the regular
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   827
  expression @{term "r\<^sub>3"}, the character @{term c} from the last
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   828
  derivative step and @{term "v\<^sub>4"}, which is the value corresponding
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   829
  to the derivative regular expression @{term "r\<^sub>4"}. The result is
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   830
  the new value @{term "v\<^sub>3"}. The final result of the algorithm is
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   831
  the value @{term "v\<^sub>1"}. The @{term inj} function is defined by recursion on regular
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   832
  expressions and by analysing the shape of values (corresponding to 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   833
  the derivative regular expressions).
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   834
  %
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   835
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   836
  \begin{tabular}{l@ {\hspace{5mm}}lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   837
  \textit{(1)} & @{thm (lhs) injval.simps(1)} & $\dn$ & @{thm (rhs) injval.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   838
  \textit{(2)} & @{thm (lhs) injval.simps(2)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1"]} & $\dn$ & 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   839
      @{thm (rhs) injval.simps(2)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   840
  \textit{(3)} & @{thm (lhs) injval.simps(3)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]} & $\dn$ & 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   841
      @{thm (rhs) injval.simps(3)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   842
  \textit{(4)} & @{thm (lhs) injval.simps(4)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]} & $\dn$ 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   843
      & @{thm (rhs) injval.simps(4)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   844
  \textit{(5)} & @{thm (lhs) injval.simps(5)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]} & $\dn$ 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   845
      & @{thm (rhs) injval.simps(5)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   846
  \textit{(6)} & @{thm (lhs) injval.simps(6)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]} & $\dn$ 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   847
      & @{thm (rhs) injval.simps(6)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   848
  \textit{(7)} & @{thm (lhs) injval.simps(7)[of "r" "c" "v" "vs"]} & $\dn$ 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   849
      & @{thm (rhs) injval.simps(7)[of "r" "c" "v" "vs"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   850
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   851
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   852
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   853
  \noindent To better understand what is going on in this definition it
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   854
  might be instructive to look first at the three sequence cases (clauses
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   855
  \textit{(4)} -- \textit{(6)}). In each case we need to construct an ``injected value'' for
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   856
  @{term "SEQ r\<^sub>1 r\<^sub>2"}. This must be a value of the form @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   857
  "Seq DUMMY DUMMY"}\,. Recall the clause of the \<open>derivative\<close>-function
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   858
  for sequence regular expressions:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   859
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   860
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   861
  @{thm (lhs) der.simps(5)[of c "r\<^sub>1" "r\<^sub>2"]} $\dn$ @{thm (rhs) der.simps(5)[of c "r\<^sub>1" "r\<^sub>2"]}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   862
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   863
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   864
  \noindent Consider first the \<open>else\<close>-branch where the derivative is @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   865
  "SEQ (der c r\<^sub>1) r\<^sub>2"}. The corresponding value must therefore
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   866
  be of the form @{term "Seq v\<^sub>1 v\<^sub>2"}, which matches the left-hand
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   867
  side in clause~\textit{(4)} of @{term inj}. In the \<open>if\<close>-branch the derivative is an
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   868
  alternative, namely @{term "ALT (SEQ (der c r\<^sub>1) r\<^sub>2) (der c
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   869
  r\<^sub>2)"}. This means we either have to consider a \<open>Left\<close>- or
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   870
  \<open>Right\<close>-value. In case of the \<open>Left\<close>-value we know further it
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   871
  must be a value for a sequence regular expression. Therefore the pattern
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   872
  we match in the clause \textit{(5)} is @{term "Left (Seq v\<^sub>1 v\<^sub>2)"},
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   873
  while in \textit{(6)} it is just @{term "Right v\<^sub>2"}. One more interesting
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   874
  point is in the right-hand side of clause \textit{(6)}: since in this case the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   875
  regular expression \<open>r\<^sub>1\<close> does not ``contribute'' to
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   876
  matching the string, that means it only matches the empty string, we need to
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   877
  call @{const mkeps} in order to construct a value for how @{term "r\<^sub>1"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   878
  can match this empty string. A similar argument applies for why we can
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   879
  expect in the left-hand side of clause \textit{(7)} that the value is of the form
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   880
  @{term "Seq v (Stars vs)"}---the derivative of a star is @{term "SEQ (der c r)
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   881
  (STAR r)"}. Finally, the reason for why we can ignore the second argument
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   882
  in clause \textit{(1)} of @{term inj} is that it will only ever be called in cases
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   883
  where @{term "c=d"}, but the usual linearity restrictions in patterns do
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   884
  not allow us to build this constraint explicitly into our function
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   885
  definition.\footnote{Sulzmann and Lu state this clause as @{thm (lhs)
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   886
  injval.simps(1)[of "c" "c"]} $\dn$ @{thm (rhs) injval.simps(1)[of "c"]},
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   887
  but our deviation is harmless.}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   888
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   889
  The idea of the @{term inj}-function to ``inject'' a character, say
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   890
  @{term c}, into a value can be made precise by the first part of the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   891
  following lemma, which shows that the underlying string of an injected
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   892
  value has a prepended character @{term c}; the second part shows that
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   893
  the underlying string of an @{const mkeps}-value is always the empty
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   894
  string (given the regular expression is nullable since otherwise
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   895
  \<open>mkeps\<close> might not be defined).
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   896
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   897
  \begin{lemma}\mbox{}\smallskip\\\label{Prf_injval_flat}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   898
  \begin{tabular}{ll}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   899
  (1) & @{thm[mode=IfThen] Prf_injval_flat}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   900
  (2) & @{thm[mode=IfThen] mkeps_flat}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   901
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   902
  \end{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   903
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   904
  \begin{proof}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   905
  Both properties are by routine inductions: the first one can, for example,
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   906
  be proved by induction over the definition of @{term derivatives}; the second by
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   907
  an induction on @{term r}. There are no interesting cases.\qed
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   908
  \end{proof}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   909
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   910
  Having defined the @{const mkeps} and \<open>inj\<close> function we can extend
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   911
  \Brz's matcher so that a value is constructed (assuming the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   912
  regular expression matches the string). The clauses of the Sulzmann and Lu lexer are
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   913
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   914
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   915
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   916
  @{thm (lhs) lexer.simps(1)} & $\dn$ & @{thm (rhs) lexer.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   917
  @{thm (lhs) lexer.simps(2)} & $\dn$ & \<open>case\<close> @{term "lexer (der c r) s"} \<open>of\<close>\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   918
                     & & \phantom{$|$} @{term "None"}  \<open>\<Rightarrow>\<close> @{term None}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   919
                     & & $|$ @{term "Some v"} \<open>\<Rightarrow>\<close> @{term "Some (injval r c v)"}                          
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   920
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   921
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   922
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   923
  \noindent If the regular expression does not match the string, @{const None} is
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   924
  returned. If the regular expression \emph{does}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   925
  match the string, then @{const Some} value is returned. One important
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   926
  virtue of this algorithm is that it can be implemented with ease in any
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   927
  functional programming language and also in Isabelle/HOL. In the remaining
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   928
  part of this section we prove that this algorithm is correct.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   929
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   930
  The well-known idea of POSIX matching is informally defined by some
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   931
  rules such as the Longest Match and Priority Rules (see
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   932
  Introduction); as correctly argued in \cite{Sulzmann2014}, this
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   933
  needs formal specification. Sulzmann and Lu define an ``ordering
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   934
  relation'' between values and argue that there is a maximum value,
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   935
  as given by the derivative-based algorithm.  In contrast, we shall
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   936
  introduce a simple inductive definition that specifies directly what
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   937
  a \emph{POSIX value} is, incorporating the POSIX-specific choices
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   938
  into the side-conditions of our rules. Our definition is inspired by
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   939
  the matching relation given by Vansummeren~\cite{Vansummeren2006}. 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   940
  The relation we define is ternary and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   941
  written as \mbox{@{term "s \<in> r \<rightarrow> v"}}, relating
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   942
  strings, regular expressions and values; the inductive rules are given in 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   943
  Figure~\ref{POSIXrules}.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   944
  We can prove that given a string @{term s} and regular expression @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   945
   r}, the POSIX value @{term v} is uniquely determined by @{term "s \<in> r \<rightarrow> v"}.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   946
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   947
  %
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   948
  \begin{figure}[t]
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   949
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   950
  \begin{tabular}{c}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   951
  @{thm[mode=Axiom] Posix.intros(1)}\<open>P\<close>@{term "ONE"} \qquad
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   952
  @{thm[mode=Axiom] Posix.intros(2)}\<open>P\<close>@{term "c"}\medskip\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   953
  @{thm[mode=Rule] Posix.intros(3)[of "s" "r\<^sub>1" "v" "r\<^sub>2"]}\<open>P+L\<close>\qquad
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   954
  @{thm[mode=Rule] Posix.intros(4)[of "s" "r\<^sub>2" "v" "r\<^sub>1"]}\<open>P+R\<close>\medskip\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   955
  $\mprset{flushleft}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   956
   \inferrule
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   957
   {@{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
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   958
    @{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"]} \\\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   959
    @{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"]}}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   960
   {@{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>\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   961
  @{thm[mode=Axiom] Posix.intros(7)}\<open>P[]\<close>\medskip\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   962
  $\mprset{flushleft}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   963
   \inferrule
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   964
   {@{thm (prem 1) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]} \qquad
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   965
    @{thm (prem 2) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]} \qquad
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   966
    @{thm (prem 3) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]} \\\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   967
    @{thm (prem 4) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]}}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   968
   {@{thm (concl) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]}}$\<open>P\<star>\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   969
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   970
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   971
  \caption{Our inductive definition of POSIX values.}\label{POSIXrules}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   972
  \end{figure}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   973
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   974
   
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   975
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   976
  \begin{theorem}\mbox{}\smallskip\\\label{posixdeterm}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   977
  \begin{tabular}{ll}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   978
  (1) & If @{thm (prem 1) Posix1(1)} then @{thm (concl)
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   979
  Posix1(1)} and @{thm (concl) Posix1(2)}.\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   980
  (2) & @{thm[mode=IfThen] Posix_determ(1)[of _ _ "v" "v'"]}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   981
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   982
  \end{theorem}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   983
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   984
  \begin{proof} Both by induction on the definition of @{term "s \<in> r \<rightarrow> v"}. 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   985
  The second parts follows by a case analysis of @{term "s \<in> r \<rightarrow> v'"} and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   986
  the first part.\qed
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   987
  \end{proof}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   988
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   989
  \noindent
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   990
  We claim that our @{term "s \<in> r \<rightarrow> v"} relation captures the idea behind the four
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   991
  informal POSIX rules shown in the Introduction: Consider for example the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   992
  rules \<open>P+L\<close> and \<open>P+R\<close> where the POSIX value for a string
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   993
  and an alternative regular expression, that is @{term "(s, ALT r\<^sub>1 r\<^sub>2)"},
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   994
  is specified---it is always a \<open>Left\<close>-value, \emph{except} when the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   995
  string to be matched is not in the language of @{term "r\<^sub>1"}; only then it
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   996
  is a \<open>Right\<close>-value (see the side-condition in \<open>P+R\<close>).
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   997
  Interesting is also the rule for sequence regular expressions (\<open>PS\<close>). The first two premises state that @{term "v\<^sub>1"} and @{term "v\<^sub>2"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   998
  are the POSIX values for @{term "(s\<^sub>1, r\<^sub>1)"} and @{term "(s\<^sub>2, r\<^sub>2)"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
   999
  respectively. Consider now the third premise and note that the POSIX value
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1000
  of this rule should match the string \mbox{@{term "s\<^sub>1 @ s\<^sub>2"}}. According to the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1001
  Longest Match Rule, we want that the @{term "s\<^sub>1"} is the longest initial
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1002
  split of \mbox{@{term "s\<^sub>1 @ s\<^sub>2"}} such that @{term "s\<^sub>2"} is still recognised
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1003
  by @{term "r\<^sub>2"}. Let us assume, contrary to the third premise, that there
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1004
  \emph{exist} an @{term "s\<^sub>3"} and @{term "s\<^sub>4"} such that @{term "s\<^sub>2"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1005
  can be split up into a non-empty string @{term "s\<^sub>3"} and a possibly empty
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1006
  string @{term "s\<^sub>4"}. Moreover the longer string @{term "s\<^sub>1 @ s\<^sub>3"} can be
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1007
  matched by \<open>r\<^sub>1\<close> and the shorter @{term "s\<^sub>4"} can still be
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1008
  matched by @{term "r\<^sub>2"}. In this case @{term "s\<^sub>1"} would \emph{not} be the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1009
  longest initial split of \mbox{@{term "s\<^sub>1 @ s\<^sub>2"}} and therefore @{term "Seq v\<^sub>1
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1010
  v\<^sub>2"} cannot be a POSIX value for @{term "(s\<^sub>1 @ s\<^sub>2, SEQ r\<^sub>1 r\<^sub>2)"}. 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1011
  The main point is that our side-condition ensures the Longest 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1012
  Match Rule is satisfied.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1013
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1014
  A similar condition is imposed on the POSIX value in the \<open>P\<star>\<close>-rule. Also there we want that @{term "s\<^sub>1"} is the longest initial
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1015
  split of @{term "s\<^sub>1 @ s\<^sub>2"} and furthermore the corresponding value
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1016
  @{term v} cannot be flattened to the empty string. In effect, we require
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1017
  that in each ``iteration'' of the star, some non-empty substring needs to
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1018
  be ``chipped'' away; only in case of the empty string we accept @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1019
  "Stars []"} as the POSIX value. Indeed we can show that our POSIX values
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1020
  are lexical values which exclude those \<open>Stars\<close> that contain subvalues 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1021
  that flatten to the empty string.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1022
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1023
  \begin{lemma}\label{LVposix}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1024
  @{thm [mode=IfThen] Posix_LV}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1025
  \end{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1026
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1027
  \begin{proof}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1028
  By routine induction on @{thm (prem 1) Posix_LV}.\qed 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1029
  \end{proof}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1030
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1031
  \noindent
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1032
  Next is the lemma that shows the function @{term "mkeps"} calculates
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1033
  the POSIX value for the empty string and a nullable regular expression.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1034
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1035
  \begin{lemma}\label{lemmkeps}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1036
  @{thm[mode=IfThen] Posix_mkeps}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1037
  \end{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1038
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1039
  \begin{proof}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1040
  By routine induction on @{term r}.\qed 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1041
  \end{proof}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1042
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1043
  \noindent
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1044
  The central lemma for our POSIX relation is that the \<open>inj\<close>-function
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1045
  preserves POSIX values.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1046
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1047
  \begin{lemma}\label{Posix2}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1048
  @{thm[mode=IfThen] Posix_injval}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1049
  \end{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1050
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1051
  \begin{proof}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1052
  By induction on \<open>r\<close>. We explain two cases.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1053
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1054
  \begin{itemize}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1055
  \item[$\bullet$] Case @{term "r = ALT r\<^sub>1 r\<^sub>2"}. There are
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1056
  two subcases, namely \<open>(a)\<close> \mbox{@{term "v = Left v'"}} and @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1057
  "s \<in> der c r\<^sub>1 \<rightarrow> v'"}; and \<open>(b)\<close> @{term "v = Right v'"}, @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1058
  "s \<notin> L (der c r\<^sub>1)"} and @{term "s \<in> der c r\<^sub>2 \<rightarrow> v'"}. In \<open>(a)\<close> we
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1059
  know @{term "s \<in> der c r\<^sub>1 \<rightarrow> v'"}, from which we can infer @{term "(c # s)
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1060
  \<in> r\<^sub>1 \<rightarrow> injval r\<^sub>1 c v'"} by induction hypothesis and hence @{term "(c #
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1061
  s) \<in> ALT r\<^sub>1 r\<^sub>2 \<rightarrow> injval (ALT r\<^sub>1 r\<^sub>2) c (Left v')"} as needed. Similarly
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1062
  in subcase \<open>(b)\<close> where, however, in addition we have to use
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1063
  Proposition~\ref{derprop}(2) in order to infer @{term "c # s \<notin> L r\<^sub>1"} from @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1064
  "s \<notin> L (der c r\<^sub>1)"}.\smallskip
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1065
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1066
  \item[$\bullet$] Case @{term "r = SEQ r\<^sub>1 r\<^sub>2"}. There are three subcases:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1067
  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1068
  \begin{quote}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1069
  \begin{description}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1070
  \item[\<open>(a)\<close>] @{term "v = Left (Seq v\<^sub>1 v\<^sub>2)"} and @{term "nullable r\<^sub>1"} 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1071
  \item[\<open>(b)\<close>] @{term "v = Right v\<^sub>1"} and @{term "nullable r\<^sub>1"} 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1072
  \item[\<open>(c)\<close>] @{term "v = Seq v\<^sub>1 v\<^sub>2"} and @{term "\<not> nullable r\<^sub>1"} 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1073
  \end{description}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1074
  \end{quote}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1075
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1076
  \noindent For \<open>(a)\<close> we know @{term "s\<^sub>1 \<in> der c r\<^sub>1 \<rightarrow> v\<^sub>1"} and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1077
  @{term "s\<^sub>2 \<in> r\<^sub>2 \<rightarrow> v\<^sub>2"} as well as
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1078
  %
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1079
  \[@{term "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s\<^sub>2 \<and> s\<^sub>1 @ s\<^sub>3 \<in> L (der c r\<^sub>1) \<and> s\<^sub>4 \<in> L r\<^sub>2)"}\]
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1080
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1081
  \noindent From the latter we can infer by Proposition~\ref{derprop}(2):
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1082
  %
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1083
  \[@{term "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s\<^sub>2 \<and> (c # s\<^sub>1) @ s\<^sub>3 \<in> L r\<^sub>1 \<and> s\<^sub>4 \<in> L r\<^sub>2)"}\]
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1084
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1085
  \noindent We can use the induction hypothesis for \<open>r\<^sub>1\<close> to obtain
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1086
  @{term "(c # s\<^sub>1) \<in> r\<^sub>1 \<rightarrow> injval r\<^sub>1 c v\<^sub>1"}. Putting this all together allows us to infer
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1087
  @{term "((c # s\<^sub>1) @ s\<^sub>2) \<in> SEQ r\<^sub>1 r\<^sub>2 \<rightarrow> Seq (injval r\<^sub>1 c v\<^sub>1) v\<^sub>2"}. The case \<open>(c)\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1088
  is similar.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1089
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1090
  For \<open>(b)\<close> we know @{term "s \<in> der c r\<^sub>2 \<rightarrow> v\<^sub>1"} and 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1091
  @{term "s\<^sub>1 @ s\<^sub>2 \<notin> L (SEQ (der c r\<^sub>1) r\<^sub>2)"}. From the former
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1092
  we have @{term "(c # s) \<in> r\<^sub>2 \<rightarrow> (injval r\<^sub>2 c v\<^sub>1)"} by induction hypothesis
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1093
  for @{term "r\<^sub>2"}. From the latter we can infer
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1094
  %
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1095
  \[@{term "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = c # s \<and> s\<^sub>3 \<in> L r\<^sub>1 \<and> s\<^sub>4 \<in> L r\<^sub>2)"}\]
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1096
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1097
  \noindent By Lemma~\ref{lemmkeps} we know @{term "[] \<in> r\<^sub>1 \<rightarrow> (mkeps r\<^sub>1)"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1098
  holds. Putting this all together, we can conclude with @{term "(c #
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1099
  s) \<in> SEQ r\<^sub>1 r\<^sub>2 \<rightarrow> Seq (mkeps r\<^sub>1) (injval r\<^sub>2 c v\<^sub>1)"}, as required.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1100
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1101
  Finally suppose @{term "r = STAR r\<^sub>1"}. This case is very similar to the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1102
  sequence case, except that we need to also ensure that @{term "flat (injval r\<^sub>1
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1103
  c v\<^sub>1) \<noteq> []"}. This follows from @{term "(c # s\<^sub>1)
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1104
  \<in> r\<^sub>1 \<rightarrow> injval r\<^sub>1 c v\<^sub>1"}  (which in turn follows from @{term "s\<^sub>1 \<in> der c
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1105
  r\<^sub>1 \<rightarrow> v\<^sub>1"} and the induction hypothesis).\qed
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1106
  \end{itemize}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1107
  \end{proof}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1108
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1109
  \noindent
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1110
  With Lemma~\ref{Posix2} in place, it is completely routine to establish
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1111
  that the Sulzmann and Lu lexer satisfies our specification (returning
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1112
  the null value @{term "None"} iff the string is not in the language of the regular expression,
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1113
  and returning a unique POSIX value iff the string \emph{is} in the language):
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1114
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1115
  \begin{theorem}\mbox{}\smallskip\\\label{lexercorrect}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1116
  \begin{tabular}{ll}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1117
  (1) & @{thm (lhs) lexer_correct_None} if and only if @{thm (rhs) lexer_correct_None}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1118
  (2) & @{thm (lhs) lexer_correct_Some} if and only if @{thm (rhs) lexer_correct_Some}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1119
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1120
  \end{theorem}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1121
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1122
  \begin{proof}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1123
  By induction on @{term s} using Lemma~\ref{lemmkeps} and \ref{Posix2}.\qed  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1124
  \end{proof}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1125
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1126
  \noindent In \textit{(2)} we further know by Theorem~\ref{posixdeterm} that the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1127
  value returned by the lexer must be unique.   A simple corollary 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1128
  of our two theorems is:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1129
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1130
  \begin{corollary}\mbox{}\smallskip\\\label{lexercorrectcor}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1131
  \begin{tabular}{ll}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1132
  (1) & @{thm (lhs) lexer_correctness(2)} if and only if @{thm (rhs) lexer_correctness(2)}\\ 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1133
  (2) & @{thm (lhs) lexer_correctness(1)} if and only if @{thm (rhs) lexer_correctness(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1134
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1135
  \end{corollary}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1136
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1137
  \noindent This concludes our correctness proof. Note that we have
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1138
  not changed the algorithm of Sulzmann and Lu,\footnote{All
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1139
  deviations we introduced are harmless.} but introduced our own
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1140
  specification for what a correct result---a POSIX value---should be.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1141
  In the next section we show that our specification coincides with
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1142
  another one given by Okui and Suzuki using a different technique.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1143
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1144
\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1145
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1146
section \<open>Ordering of Values according to Okui and Suzuki\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1147
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1148
text \<open>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1149
  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1150
  While in the previous section we have defined POSIX values directly
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1151
  in terms of a ternary relation (see inference rules in Figure~\ref{POSIXrules}),
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1152
  Sulzmann and Lu took a different approach in \cite{Sulzmann2014}:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1153
  they introduced an ordering for values and identified POSIX values
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1154
  as the maximal elements.  An extended version of \cite{Sulzmann2014}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1155
  is available at the website of its first author; this includes more
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1156
  details of their proofs, but which are evidently not in final form
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1157
  yet. Unfortunately, we were not able to verify claims that their
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1158
  ordering has properties such as being transitive or having maximal
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1159
  elements. 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1160
 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1161
  Okui and Suzuki \cite{OkuiSuzuki2010,OkuiSuzukiTech} described
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1162
  another ordering of values, which they use to establish the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1163
  correctness of their automata-based algorithm for POSIX matching.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1164
  Their ordering resembles some aspects of the one given by Sulzmann
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1165
  and Lu, but overall is quite different. To begin with, Okui and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1166
  Suzuki identify POSIX values as minimal, rather than maximal,
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1167
  elements in their ordering. A more substantial difference is that
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1168
  the ordering by Okui and Suzuki uses \emph{positions} in order to
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1169
  identify and compare subvalues. Positions are lists of natural
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1170
  numbers. This allows them to quite naturally formalise the Longest
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1171
  Match and Priority rules of the informal POSIX standard.  Consider
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1172
  for example the value @{term v}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1173
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1174
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1175
  @{term "v == Stars [Seq (Char x) (Char y), Char z]"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1176
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1177
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1178
  \noindent
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1179
  At position \<open>[0,1]\<close> of this value is the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1180
  subvalue \<open>Char y\<close> and at position \<open>[1]\<close> the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1181
  subvalue @{term "Char z"}.  At the `root' position, or empty list
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1182
  @{term "[]"}, is the whole value @{term v}. Positions such as \<open>[0,1,0]\<close> or \<open>[2]\<close> are outside of \<open>v\<close>. If it exists, the subvalue of @{term v} at a position \<open>p\<close>, written @{term "at v p"}, can be recursively defined by
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1183
  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1184
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1185
  \begin{tabular}{r@ {\hspace{0mm}}lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1186
  @{term v} &  \<open>\<downharpoonleft>\<^bsub>[]\<^esub>\<close> & \<open>\<equiv>\<close>& @{thm (rhs) at.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1187
  @{term "Left v"} & \<open>\<downharpoonleft>\<^bsub>0::ps\<^esub>\<close> & \<open>\<equiv>\<close>& @{thm (rhs) at.simps(2)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1188
  @{term "Right v"} & \<open>\<downharpoonleft>\<^bsub>1::ps\<^esub>\<close> & \<open>\<equiv>\<close> & 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1189
  @{thm (rhs) at.simps(3)[simplified Suc_0_fold]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1190
  @{term "Seq v\<^sub>1 v\<^sub>2"} & \<open>\<downharpoonleft>\<^bsub>0::ps\<^esub>\<close> & \<open>\<equiv>\<close> & 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1191
  @{thm (rhs) at.simps(4)[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]} \\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1192
  @{term "Seq v\<^sub>1 v\<^sub>2"} & \<open>\<downharpoonleft>\<^bsub>1::ps\<^esub>\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1193
  & \<open>\<equiv>\<close> & 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1194
  @{thm (rhs) at.simps(5)[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2", simplified Suc_0_fold]} \\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1195
  @{term "Stars vs"} & \<open>\<downharpoonleft>\<^bsub>n::ps\<^esub>\<close> & \<open>\<equiv>\<close>& @{thm (rhs) at.simps(6)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1196
  \end{tabular} 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1197
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1198
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1199
  \noindent In the last clause we use Isabelle's notation @{term "vs ! n"} for the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1200
  \<open>n\<close>th element in a list.  The set of positions inside a value \<open>v\<close>,
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1201
  written @{term "Pos v"}, is given by 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1202
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1203
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1204
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1205
  @{thm (lhs) Pos.simps(1)} & \<open>\<equiv>\<close> & @{thm (rhs) Pos.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1206
  @{thm (lhs) Pos.simps(2)} & \<open>\<equiv>\<close> & @{thm (rhs) Pos.simps(2)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1207
  @{thm (lhs) Pos.simps(3)} & \<open>\<equiv>\<close> & @{thm (rhs) Pos.simps(3)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1208
  @{thm (lhs) Pos.simps(4)} & \<open>\<equiv>\<close> & @{thm (rhs) Pos.simps(4)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1209
  @{thm (lhs) Pos.simps(5)[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1210
  & \<open>\<equiv>\<close> 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1211
  & @{thm (rhs) Pos.simps(5)[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1212
  @{thm (lhs) Pos_stars} & \<open>\<equiv>\<close> & @{thm (rhs) Pos_stars}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1213
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1214
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1215
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1216
  \noindent 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1217
  whereby \<open>len\<close> in the last clause stands for the length of a list. Clearly
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1218
  for every position inside a value there exists a subvalue at that position.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1219
 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1220
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1221
  To help understanding the ordering of Okui and Suzuki, consider again 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1222
  the earlier value
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1223
  \<open>v\<close> and compare it with the following \<open>w\<close>:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1224
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1225
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1226
  \begin{tabular}{l}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1227
  @{term "v == Stars [Seq (Char x) (Char y), Char z]"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1228
  @{term "w == Stars [Char x, Char y, Char z]"}  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1229
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1230
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1231
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1232
  \noindent Both values match the string \<open>xyz\<close>, that means if
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1233
  we flatten these values at their respective root position, we obtain
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1234
  \<open>xyz\<close>. However, at position \<open>[0]\<close>, \<open>v\<close> matches
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1235
  \<open>xy\<close> whereas \<open>w\<close> matches only the shorter \<open>x\<close>. So
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1236
  according to the Longest Match Rule, we should prefer \<open>v\<close>,
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1237
  rather than \<open>w\<close> as POSIX value for string \<open>xyz\<close> (and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1238
  corresponding regular expression). In order to
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1239
  formalise this idea, Okui and Suzuki introduce a measure for
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1240
  subvalues at position \<open>p\<close>, called the \emph{norm} of \<open>v\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1241
  at position \<open>p\<close>. We can define this measure in Isabelle as an
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1242
  integer as follows
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1243
  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1244
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1245
  @{thm pflat_len_def}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1246
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1247
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1248
  \noindent where we take the length of the flattened value at
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1249
  position \<open>p\<close>, provided the position is inside \<open>v\<close>; if
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1250
  not, then the norm is \<open>-1\<close>. The default for outside
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1251
  positions is crucial for the POSIX requirement of preferring a
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1252
  \<open>Left\<close>-value over a \<open>Right\<close>-value (if they can match the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1253
  same string---see the Priority Rule from the Introduction). For this
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1254
  consider
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1255
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1256
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1257
  @{term "v == Left (Char x)"} \qquad and \qquad @{term "w == Right (Char x)"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1258
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1259
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1260
  \noindent Both values match \<open>x\<close>. At position \<open>[0]\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1261
  the norm of @{term v} is \<open>1\<close> (the subvalue matches \<open>x\<close>),
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1262
  but the norm of \<open>w\<close> is \<open>-1\<close> (the position is outside
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1263
  \<open>w\<close> according to how we defined the `inside' positions of
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1264
  \<open>Left\<close>- and \<open>Right\<close>-values).  Of course at position
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1265
  \<open>[1]\<close>, the norms @{term "pflat_len v [1]"} and @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1266
  "pflat_len w [1]"} are reversed, but the point is that subvalues
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1267
  will be analysed according to lexicographically ordered
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1268
  positions. According to this ordering, the position \<open>[0]\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1269
  takes precedence over \<open>[1]\<close> and thus also \<open>v\<close> will be 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1270
  preferred over \<open>w\<close>.  The lexicographic ordering of positions, written
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1271
  @{term "DUMMY \<sqsubset>lex DUMMY"}, can be conveniently formalised
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1272
  by three inference rules
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1273
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1274
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1275
  \begin{tabular}{ccc}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1276
  @{thm [mode=Axiom] lex_list.intros(1)}\hspace{1cm} &
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1277
  @{thm [mode=Rule] lex_list.intros(3)[where ?p1.0="p\<^sub>1" and ?p2.0="p\<^sub>2" and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1278
                                            ?ps1.0="ps\<^sub>1" and ?ps2.0="ps\<^sub>2"]}\hspace{1cm} &
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1279
  @{thm [mode=Rule] lex_list.intros(2)[where ?ps1.0="ps\<^sub>1" and ?ps2.0="ps\<^sub>2"]}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1280
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1281
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1282
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1283
  With the norm and lexicographic order in place,
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1284
  we can state the key definition of Okui and Suzuki
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1285
  \cite{OkuiSuzuki2010}: a value @{term "v\<^sub>1"} is \emph{smaller at position \<open>p\<close>} than
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1286
  @{term "v\<^sub>2"}, written @{term "v\<^sub>1 \<sqsubset>val p v\<^sub>2"}, 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1287
  if and only if  $(i)$ the norm at position \<open>p\<close> is
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1288
  greater in @{term "v\<^sub>1"} (that is the string @{term "flat (at v\<^sub>1 p)"} is longer 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1289
  than @{term "flat (at v\<^sub>2 p)"}) and $(ii)$ all subvalues at 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1290
  positions that are inside @{term "v\<^sub>1"} or @{term "v\<^sub>2"} and that are
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1291
  lexicographically smaller than \<open>p\<close>, we have the same norm, namely
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1292
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1293
 \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1294
 \begin{tabular}{c}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1295
 @{thm (lhs) PosOrd_def[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]} 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1296
 \<open>\<equiv>\<close> 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1297
 $\begin{cases}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1298
 (i) & @{term "pflat_len v\<^sub>1 p > pflat_len v\<^sub>2 p"}   \quad\text{and}\smallskip \\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1299
 (ii) & @{term "(\<forall>q \<in> Pos v\<^sub>1 \<union> Pos v\<^sub>2. q \<sqsubset>lex p --> pflat_len v\<^sub>1 q = pflat_len v\<^sub>2 q)"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1300
 \end{cases}$
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1301
 \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1302
 \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1303
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1304
 \noindent The position \<open>p\<close> in this definition acts as the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1305
  \emph{first distinct position} of \<open>v\<^sub>1\<close> and \<open>v\<^sub>2\<close>, where both values match strings of different length
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1306
  \cite{OkuiSuzuki2010}.  Since at \<open>p\<close> the values \<open>v\<^sub>1\<close> and \<open>v\<^sub>2\<close> match different strings, the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1307
  ordering is irreflexive. Derived from the definition above
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1308
  are the following two orderings:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1309
  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1310
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1311
  \begin{tabular}{l}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1312
  @{thm PosOrd_ex_def[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1313
  @{thm PosOrd_ex_eq_def[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1314
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1315
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1316
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1317
 While we encountered a number of obstacles for establishing properties like
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1318
 transitivity for the ordering of Sulzmann and Lu (and which we failed
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1319
 to overcome), it is relatively straightforward to establish this
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1320
 property for the orderings
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1321
 @{term "DUMMY :\<sqsubset>val DUMMY"} and @{term "DUMMY :\<sqsubseteq>val DUMMY"}  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1322
 by Okui and Suzuki.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1323
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1324
 \begin{lemma}[Transitivity]\label{transitivity}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1325
 @{thm [mode=IfThen] PosOrd_trans[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2" and ?v3.0="v\<^sub>3"]} 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1326
 \end{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1327
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1328
 \begin{proof} From the assumption we obtain two positions \<open>p\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1329
 and \<open>q\<close>, where the values \<open>v\<^sub>1\<close> and \<open>v\<^sub>2\<close> (respectively \<open>v\<^sub>2\<close> and \<open>v\<^sub>3\<close>) are `distinct'.  Since \<open>\<prec>\<^bsub>lex\<^esub>\<close> is trichotomous, we need to consider
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1330
 three cases, namely @{term "p = q"}, @{term "p \<sqsubset>lex q"} and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1331
 @{term "q \<sqsubset>lex p"}. Let us look at the first case.  Clearly
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1332
 @{term "pflat_len v\<^sub>2 p < pflat_len v\<^sub>1 p"} and @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1333
 "pflat_len v\<^sub>3 p < pflat_len v\<^sub>2 p"} imply @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1334
 "pflat_len v\<^sub>3 p < pflat_len v\<^sub>1 p"}.  It remains to show
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1335
 that for a @{term "p' \<in> Pos v\<^sub>1 \<union> Pos v\<^sub>3"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1336
 with @{term "p' \<sqsubset>lex p"} that @{term "pflat_len v\<^sub>1
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1337
 p' = pflat_len v\<^sub>3 p'"} holds.  Suppose @{term "p' \<in> Pos
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1338
 v\<^sub>1"}, then we can infer from the first assumption that @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1339
 "pflat_len v\<^sub>1 p' = pflat_len v\<^sub>2 p'"}.  But this means
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1340
 that @{term "p'"} must be in @{term "Pos v\<^sub>2"} too (the norm
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1341
 cannot be \<open>-1\<close> given @{term "p' \<in> Pos v\<^sub>1"}).  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1342
 Hence we can use the second assumption and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1343
 infer @{term "pflat_len v\<^sub>2 p' = pflat_len v\<^sub>3 p'"},
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1344
 which concludes this case with @{term "v\<^sub>1 :\<sqsubset>val
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1345
 v\<^sub>3"}.  The reasoning in the other cases is similar.\qed
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1346
 \end{proof}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1347
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1348
 \noindent 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1349
 The proof for $\preccurlyeq$ is similar and omitted.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1350
 It is also straightforward to show that \<open>\<prec>\<close> and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1351
 $\preccurlyeq$ are partial orders.  Okui and Suzuki furthermore show that they
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1352
 are linear orderings for lexical values \cite{OkuiSuzuki2010} of a given
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1353
 regular expression and given string, but we have not formalised this in Isabelle. It is
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1354
 not essential for our results. What we are going to show below is
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1355
 that for a given \<open>r\<close> and \<open>s\<close>, the orderings have a unique
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1356
 minimal element on the set @{term "LV r s"}, which is the POSIX value
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1357
 we defined in the previous section. We start with two properties that
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1358
 show how the length of a flattened value relates to the \<open>\<prec>\<close>-ordering.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1359
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1360
 \begin{proposition}\mbox{}\smallskip\\\label{ordlen}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1361
 \begin{tabular}{@ {}ll}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1362
 (1) &
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1363
 @{thm [mode=IfThen] PosOrd_shorterE[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1364
 (2) &
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1365
 @{thm [mode=IfThen] PosOrd_shorterI[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]} 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1366
 \end{tabular} 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1367
 \end{proposition}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1368
 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1369
 \noindent Both properties follow from the definition of the ordering. Note that
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1370
 \textit{(2)} entails that a value, say @{term "v\<^sub>2"}, whose underlying 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1371
 string is a strict prefix of another flattened value, say @{term "v\<^sub>1"}, then
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1372
 @{term "v\<^sub>1"} must be smaller than @{term "v\<^sub>2"}. For our proofs it
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1373
 will be useful to have the following properties---in each case the underlying strings 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1374
 of the compared values are the same: 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1375
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1376
  \begin{proposition}\mbox{}\smallskip\\\label{ordintros}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1377
  \begin{tabular}{ll}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1378
  \textit{(1)} & 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1379
  @{thm [mode=IfThen] PosOrd_Left_Right[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1380
  \textit{(2)} & If
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1381
  @{thm (prem 1) PosOrd_Left_eq[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]} \;then\;
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1382
  @{thm (lhs) PosOrd_Left_eq[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]} \;iff\;
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1383
  @{thm (rhs) PosOrd_Left_eq[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1384
  \textit{(3)} & If
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1385
  @{thm (prem 1) PosOrd_Right_eq[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]} \;then\;
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1386
  @{thm (lhs) PosOrd_Right_eq[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]} \;iff\;
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1387
  @{thm (rhs) PosOrd_Right_eq[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1388
  \textit{(4)} & If
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1389
  @{thm (prem 1) PosOrd_Seq_eq[where ?v2.0="v\<^sub>2" and ?w2.0="w\<^sub>2"]} \;then\;
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1390
  @{thm (lhs) PosOrd_Seq_eq[where ?v2.0="v\<^sub>2" and ?w2.0="w\<^sub>2"]} \;iff\;
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1391
  @{thm (rhs) PosOrd_Seq_eq[where ?v2.0="v\<^sub>2" and ?w2.0="w\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1392
  \textit{(5)} & If
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1393
  @{thm (prem 2) PosOrd_SeqI1[simplified, where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2" and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1394
                                    ?w1.0="w\<^sub>1" and ?w2.0="w\<^sub>2"]} \;and\;
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1395
  @{thm (prem 1) PosOrd_SeqI1[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2" and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1396
                                    ?w1.0="w\<^sub>1" and ?w2.0="w\<^sub>2"]} \;then\;
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1397
  @{thm (concl) PosOrd_SeqI1[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2" and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1398
                                   ?w1.0="w\<^sub>1" and ?w2.0="w\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1399
  \textit{(6)} & If
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1400
  @{thm (prem 1) PosOrd_Stars_append_eq[where ?vs1.0="vs\<^sub>1" and ?vs2.0="vs\<^sub>2"]} \;then\;
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1401
  @{thm (lhs) PosOrd_Stars_append_eq[where ?vs1.0="vs\<^sub>1" and ?vs2.0="vs\<^sub>2"]} \;iff\;
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1402
  @{thm (rhs) PosOrd_Stars_append_eq[where ?vs1.0="vs\<^sub>1" and ?vs2.0="vs\<^sub>2"]}\\  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1403
  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1404
  \textit{(7)} & If
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1405
  @{thm (prem 2) PosOrd_StarsI[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2" and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1406
                            ?vs1.0="vs\<^sub>1" and ?vs2.0="vs\<^sub>2"]} \;and\;
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1407
  @{thm (prem 1) PosOrd_StarsI[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2" and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1408
                            ?vs1.0="vs\<^sub>1" and ?vs2.0="vs\<^sub>2"]} \;then\;
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1409
   @{thm (concl) PosOrd_StarsI[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2" and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1410
                            ?vs1.0="vs\<^sub>1" and ?vs2.0="vs\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1411
  \end{tabular} 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1412
  \end{proposition}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1413
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1414
  \noindent One might prefer that statements \textit{(4)} and \textit{(5)} 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1415
  (respectively \textit{(6)} and \textit{(7)})
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1416
  are combined into a single \textit{iff}-statement (like the ones for \<open>Left\<close> and \<open>Right\<close>). Unfortunately this cannot be done easily: such
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1417
  a single statement would require an additional assumption about the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1418
  two values @{term "Seq v\<^sub>1 v\<^sub>2"} and @{term "Seq w\<^sub>1 w\<^sub>2"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1419
  being inhabited by the same regular expression. The
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1420
  complexity of the proofs involved seems to not justify such a
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1421
  `cleaner' single statement. The statements given are just the properties that
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1422
  allow us to establish our theorems without any difficulty. The proofs 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1423
  for Proposition~\ref{ordintros} are routine.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1424
 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1425
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1426
  Next we establish how Okui and Suzuki's orderings relate to our
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1427
  definition of POSIX values.  Given a \<open>POSIX\<close> value \<open>v\<^sub>1\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1428
  for \<open>r\<close> and \<open>s\<close>, then any other lexical value \<open>v\<^sub>2\<close> in @{term "LV r s"} is greater or equal than \<open>v\<^sub>1\<close>, namely:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1429
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1430
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1431
  \begin{theorem}\label{orderone}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1432
  @{thm [mode=IfThen] Posix_PosOrd[where ?v1.0="v\<^sub>1" and ?v2.0="v\<^sub>2"]}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1433
  \end{theorem}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1434
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1435
  \begin{proof} By induction on our POSIX rules. By
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1436
  Theorem~\ref{posixdeterm} and the definition of @{const LV}, it is clear
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1437
  that \<open>v\<^sub>1\<close> and \<open>v\<^sub>2\<close> have the same
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1438
  underlying string @{term s}.  The three base cases are
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1439
  straightforward: for example for @{term "v\<^sub>1 = Void"}, we have
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1440
  that @{term "v\<^sub>2 \<in> LV ONE []"} must also be of the form
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1441
  \mbox{@{term "v\<^sub>2 = Void"}}. Therefore we have @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1442
  "v\<^sub>1 :\<sqsubseteq>val v\<^sub>2"}.  The inductive cases for
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1443
  \<open>r\<close> being of the form @{term "ALT r\<^sub>1 r\<^sub>2"} and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1444
  @{term "SEQ r\<^sub>1 r\<^sub>2"} are as follows:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1445
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1446
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1447
  \begin{itemize} 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1448
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1449
  \item[$\bullet$] Case \<open>P+L\<close> with @{term "s \<in> (ALT r\<^sub>1 r\<^sub>2)
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1450
  \<rightarrow> (Left w\<^sub>1)"}: In this case the value 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1451
  @{term "v\<^sub>2"} is either of the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1452
  form @{term "Left w\<^sub>2"} or @{term "Right w\<^sub>2"}. In the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1453
  latter case we can immediately conclude with \mbox{@{term "v\<^sub>1
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1454
  :\<sqsubseteq>val v\<^sub>2"}} since a \<open>Left\<close>-value with the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1455
  same underlying string \<open>s\<close> is always smaller than a
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1456
  \<open>Right\<close>-value by Proposition~\ref{ordintros}\textit{(1)}.  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1457
  In the former case we have @{term "w\<^sub>2
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1458
  \<in> LV r\<^sub>1 s"} and can use the induction hypothesis to infer
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1459
  @{term "w\<^sub>1 :\<sqsubseteq>val w\<^sub>2"}. Because @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1460
  "w\<^sub>1"} and @{term "w\<^sub>2"} have the same underlying string
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1461
  \<open>s\<close>, we can conclude with @{term "Left w\<^sub>1
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1462
  :\<sqsubseteq>val Left w\<^sub>2"} using
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1463
  Proposition~\ref{ordintros}\textit{(2)}.\smallskip
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1464
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1465
  \item[$\bullet$] Case \<open>P+R\<close> with @{term "s \<in> (ALT r\<^sub>1 r\<^sub>2)
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1466
  \<rightarrow> (Right w\<^sub>1)"}: This case similar to the previous
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1467
  case, except that we additionally know @{term "s \<notin> L
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1468
  r\<^sub>1"}. This is needed when @{term "v\<^sub>2"} is of the form
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1469
  \mbox{@{term "Left w\<^sub>2"}}. Since \mbox{@{term "flat v\<^sub>2 = flat
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1470
  w\<^sub>2"} \<open>= s\<close>} and @{term "\<Turnstile> w\<^sub>2 :
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1471
  r\<^sub>1"}, we can derive a contradiction for \mbox{@{term "s \<notin> L
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1472
  r\<^sub>1"}} using
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1473
  Proposition~\ref{inhabs}. So also in this case \mbox{@{term "v\<^sub>1
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1474
  :\<sqsubseteq>val v\<^sub>2"}}.\smallskip
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1475
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1476
  \item[$\bullet$] Case \<open>PS\<close> with @{term "(s\<^sub>1 @
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1477
  s\<^sub>2) \<in> (SEQ r\<^sub>1 r\<^sub>2) \<rightarrow> (Seq
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1478
  w\<^sub>1 w\<^sub>2)"}: We can assume @{term "v\<^sub>2 = Seq
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1479
  (u\<^sub>1) (u\<^sub>2)"} with @{term "\<Turnstile> u\<^sub>1 :
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1480
  r\<^sub>1"} and \mbox{@{term "\<Turnstile> u\<^sub>2 :
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1481
  r\<^sub>2"}}. We have @{term "s\<^sub>1 @ s\<^sub>2 = (flat
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1482
  u\<^sub>1) @ (flat u\<^sub>2)"}.  By the side-condition of the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1483
  \<open>PS\<close>-rule we know that either @{term "s\<^sub>1 = flat
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1484
  u\<^sub>1"} or that @{term "flat u\<^sub>1"} is a strict prefix of
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1485
  @{term "s\<^sub>1"}. In the latter case we can infer @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1486
  "w\<^sub>1 :\<sqsubset>val u\<^sub>1"} by
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1487
  Proposition~\ref{ordlen}\textit{(2)} and from this @{term "v\<^sub>1
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1488
  :\<sqsubseteq>val v\<^sub>2"} by Proposition~\ref{ordintros}\textit{(5)}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1489
  (as noted above @{term "v\<^sub>1"} and @{term "v\<^sub>2"} must have the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1490
  same underlying string).
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1491
  In the former case we know
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1492
  @{term "u\<^sub>1 \<in> LV r\<^sub>1 s\<^sub>1"} and @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1493
  "u\<^sub>2 \<in> LV r\<^sub>2 s\<^sub>2"}. With this we can use the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1494
  induction hypotheses to infer @{term "w\<^sub>1 :\<sqsubseteq>val
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1495
  u\<^sub>1"} and @{term "w\<^sub>2 :\<sqsubseteq>val u\<^sub>2"}. By
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1496
  Proposition~\ref{ordintros}\textit{(4,5)} we can again infer 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1497
  @{term "v\<^sub>1 :\<sqsubseteq>val
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1498
  v\<^sub>2"}.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1499
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1500
  \end{itemize}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1501
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1502
  \noindent The case for \<open>P\<star>\<close> is similar to the \<open>PS\<close>-case and omitted.\qed
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1503
  \end{proof}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1504
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1505
  \noindent This theorem shows that our \<open>POSIX\<close> value for a
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1506
  regular expression \<open>r\<close> and string @{term s} is in fact a
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1507
  minimal element of the values in \<open>LV r s\<close>. By
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1508
  Proposition~\ref{ordlen}\textit{(2)} we also know that any value in
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1509
  \<open>LV r s'\<close>, with @{term "s'"} being a strict prefix, cannot be
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1510
  smaller than \<open>v\<^sub>1\<close>. The next theorem shows the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1511
  opposite---namely any minimal element in @{term "LV r s"} must be a
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1512
  \<open>POSIX\<close> value. This can be established by induction on \<open>r\<close>, but the proof can be drastically simplified by using the fact
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1513
  from the previous section about the existence of a \<open>POSIX\<close> value
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1514
  whenever a string @{term "s \<in> L r"}.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1515
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1516
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1517
  \begin{theorem}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1518
  @{thm [mode=IfThen] PosOrd_Posix[where ?v1.0="v\<^sub>1"]} 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1519
  \end{theorem}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1520
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1521
  \begin{proof} 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1522
  If @{thm (prem 1) PosOrd_Posix[where ?v1.0="v\<^sub>1"]} then 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1523
  @{term "s \<in> L r"} by Proposition~\ref{inhabs}. Hence by Theorem~\ref{lexercorrect}(2) 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1524
  there exists a
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1525
  \<open>POSIX\<close> value @{term "v\<^sub>P"} with @{term "s \<in> r \<rightarrow> v\<^sub>P"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1526
  and by Lemma~\ref{LVposix} we also have \mbox{@{term "v\<^sub>P \<in> LV r s"}}.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1527
  By Theorem~\ref{orderone} we therefore have 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1528
  @{term "v\<^sub>P :\<sqsubseteq>val v\<^sub>1"}. If @{term "v\<^sub>P = v\<^sub>1"} then
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1529
  we are done. Otherwise we have @{term "v\<^sub>P :\<sqsubset>val v\<^sub>1"}, which 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1530
  however contradicts the second assumption about @{term "v\<^sub>1"} being the smallest
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1531
  element in @{term "LV r s"}. So we are done in this case too.\qed
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1532
  \end{proof}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1533
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1534
  \noindent
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1535
  From this we can also show 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1536
  that if @{term "LV r s"} is non-empty (or equivalently @{term "s \<in> L r"}) then 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1537
  it has a unique minimal element:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1538
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1539
  \begin{corollary}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1540
  @{thm [mode=IfThen] Least_existence1}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1541
  \end{corollary}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1542
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1543
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1544
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1545
  \noindent To sum up, we have shown that the (unique) minimal elements 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1546
  of the ordering by Okui and Suzuki are exactly the \<open>POSIX\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1547
  values we defined inductively in Section~\ref{posixsec}. This provides
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1548
  an independent confirmation that our ternary relation formalises the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1549
  informal POSIX rules. 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1550
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1551
\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1552
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1553
section \<open>Bitcoded Lexing\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1554
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1555
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1556
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1557
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1558
text \<open>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1559
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1560
Incremental calculation of the value. To simplify the proof we first define the function
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1561
@{const flex} which calculates the ``iterated'' injection function. With this we can 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1562
rewrite the lexer as
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1563
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1564
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1565
@{thm lexer_flex}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1566
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1567
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1568
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1569
\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1570
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1571
section \<open>Optimisations\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1572
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1573
text \<open>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1574
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1575
  Derivatives as calculated by \Brz's method are usually more complex
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1576
  regular expressions than the initial one; the result is that the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1577
  derivative-based matching and lexing algorithms are often abysmally slow.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1578
  However, various optimisations are possible, such as the simplifications
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1579
  of @{term "ALT ZERO r"}, @{term "ALT r ZERO"}, @{term "SEQ ONE r"} and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1580
  @{term "SEQ r ONE"} to @{term r}. These simplifications can speed up the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1581
  algorithms considerably, as noted in \cite{Sulzmann2014}. One of the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1582
  advantages of having a simple specification and correctness proof is that
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1583
  the latter can be refined to prove the correctness of such simplification
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1584
  steps. While the simplification of regular expressions according to 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1585
  rules like
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1586
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1587
  \begin{equation}\label{Simpl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1588
  \begin{array}{lcllcllcllcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1589
  @{term "ALT ZERO r"} & \<open>\<Rightarrow>\<close> & @{term r} \hspace{8mm}%\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1590
  @{term "ALT r ZERO"} & \<open>\<Rightarrow>\<close> & @{term r} \hspace{8mm}%\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1591
  @{term "SEQ ONE r"}  & \<open>\<Rightarrow>\<close> & @{term r} \hspace{8mm}%\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1592
  @{term "SEQ r ONE"}  & \<open>\<Rightarrow>\<close> & @{term r}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1593
  \end{array}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1594
  \end{equation}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1595
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1596
  \noindent is well understood, there is an obstacle with the POSIX value
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1597
  calculation algorithm by Sulzmann and Lu: if we build a derivative regular
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1598
  expression and then simplify it, we will calculate a POSIX value for this
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1599
  simplified derivative regular expression, \emph{not} for the original (unsimplified)
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1600
  derivative regular expression. Sulzmann and Lu \cite{Sulzmann2014} overcome this obstacle by
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1601
  not just calculating a simplified regular expression, but also calculating
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1602
  a \emph{rectification function} that ``repairs'' the incorrect value.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1603
  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1604
  The rectification functions can be (slightly clumsily) implemented  in
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1605
  Isabelle/HOL as follows using some auxiliary functions:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1606
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1607
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1608
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1609
  @{thm (lhs) F_RIGHT.simps(1)} & $\dn$ & \<open>Right (f v)\<close>\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1610
  @{thm (lhs) F_LEFT.simps(1)} & $\dn$ & \<open>Left (f v)\<close>\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1611
  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1612
  @{thm (lhs) F_ALT.simps(1)} & $\dn$ & \<open>Right (f\<^sub>2 v)\<close>\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1613
  @{thm (lhs) F_ALT.simps(2)} & $\dn$ & \<open>Left (f\<^sub>1 v)\<close>\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1614
  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1615
  @{thm (lhs) F_SEQ1.simps(1)} & $\dn$ & \<open>Seq (f\<^sub>1 ()) (f\<^sub>2 v)\<close>\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1616
  @{thm (lhs) F_SEQ2.simps(1)} & $\dn$ & \<open>Seq (f\<^sub>1 v) (f\<^sub>2 ())\<close>\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1617
  @{thm (lhs) F_SEQ.simps(1)} & $\dn$ & \<open>Seq (f\<^sub>1 v\<^sub>1) (f\<^sub>2 v\<^sub>2)\<close>\medskip\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1618
  %\end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1619
  %
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1620
  %\begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1621
  @{term "simp_ALT (ZERO, DUMMY) (r\<^sub>2, f\<^sub>2)"} & $\dn$ & @{term "(r\<^sub>2, F_RIGHT f\<^sub>2)"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1622
  @{term "simp_ALT (r\<^sub>1, f\<^sub>1) (ZERO, DUMMY)"} & $\dn$ & @{term "(r\<^sub>1, F_LEFT f\<^sub>1)"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1623
  @{term "simp_ALT (r\<^sub>1, f\<^sub>1) (r\<^sub>2, f\<^sub>2)"} & $\dn$ & @{term "(ALT r\<^sub>1 r\<^sub>2, F_ALT f\<^sub>1 f\<^sub>2)"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1624
  @{term "simp_SEQ (ONE, f\<^sub>1) (r\<^sub>2, f\<^sub>2)"} & $\dn$ & @{term "(r\<^sub>2, F_SEQ1 f\<^sub>1 f\<^sub>2)"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1625
  @{term "simp_SEQ (r\<^sub>1, f\<^sub>1) (ONE, f\<^sub>2)"} & $\dn$ & @{term "(r\<^sub>1, F_SEQ2 f\<^sub>1 f\<^sub>2)"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1626
  @{term "simp_SEQ (r\<^sub>1, f\<^sub>1) (r\<^sub>2, f\<^sub>2)"} & $\dn$ & @{term "(SEQ r\<^sub>1 r\<^sub>2, F_SEQ f\<^sub>1 f\<^sub>2)"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1627
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1628
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1629
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1630
  \noindent
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1631
  The functions \<open>simp\<^bsub>Alt\<^esub>\<close> and \<open>simp\<^bsub>Seq\<^esub>\<close> encode the simplification rules
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1632
  in \eqref{Simpl} and compose the rectification functions (simplifications can occur
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1633
  deep inside the regular expression). The main simplification function is then 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1634
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1635
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1636
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1637
  @{term "simp (ALT r\<^sub>1 r\<^sub>2)"} & $\dn$ & @{term "simp_ALT (simp r\<^sub>1) (simp r\<^sub>2)"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1638
  @{term "simp (SEQ r\<^sub>1 r\<^sub>2)"} & $\dn$ & @{term "simp_SEQ (simp r\<^sub>1) (simp r\<^sub>2)"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1639
  @{term "simp r"} & $\dn$ & @{term "(r, id)"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1640
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1641
  \end{center} 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1642
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1643
  \noindent where @{term "id"} stands for the identity function. The
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1644
  function @{const simp} returns a simplified regular expression and a corresponding
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1645
  rectification function. Note that we do not simplify under stars: this
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1646
  seems to slow down the algorithm, rather than speed it up. The optimised
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1647
  lexer is then given by the clauses:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1648
  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1649
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1650
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1651
  @{thm (lhs) slexer.simps(1)} & $\dn$ & @{thm (rhs) slexer.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1652
  @{thm (lhs) slexer.simps(2)} & $\dn$ & 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1653
                         \<open>let (r\<^sub>s, f\<^sub>r) = simp (r \<close>$\backslash$\<open> c) in\<close>\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1654
                     & & \<open>case\<close> @{term "slexer r\<^sub>s s"} \<open>of\<close>\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1655
                     & & \phantom{$|$} @{term "None"}  \<open>\<Rightarrow>\<close> @{term None}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1656
                     & & $|$ @{term "Some v"} \<open>\<Rightarrow>\<close> \<open>Some (inj r c (f\<^sub>r v))\<close>                          
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1657
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1658
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1659
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1660
  \noindent
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1661
  In the second clause we first calculate the derivative @{term "der c r"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1662
  and then simpli
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1663
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1664
text \<open>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1665
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1666
Incremental calculation of the value. To simplify the proof we first define the function
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1667
@{const flex} which calculates the ``iterated'' injection function. With this we can 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1668
rewrite the lexer as
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1669
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1670
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1671
@{thm lexer_flex}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1672
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1673
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1674
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1675
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1676
  @{thm (lhs) code.simps(1)} & $\dn$ & @{thm (rhs) code.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1677
  @{thm (lhs) code.simps(2)} & $\dn$ & @{thm (rhs) code.simps(2)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1678
  @{thm (lhs) code.simps(3)} & $\dn$ & @{thm (rhs) code.simps(3)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1679
  @{thm (lhs) code.simps(4)} & $\dn$ & @{thm (rhs) code.simps(4)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1680
  @{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"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1681
  @{thm (lhs) code.simps(6)} & $\dn$ & @{thm (rhs) code.simps(6)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1682
  @{thm (lhs) code.simps(7)} & $\dn$ & @{thm (rhs) code.simps(7)}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1683
\end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1684
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1685
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1686
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1687
\begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1688
  @{term areg} & $::=$ & @{term "AZERO"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1689
               & $\mid$ & @{term "AONE bs"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1690
               & $\mid$ & @{term "ACHAR bs c"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1691
               & $\mid$ & @{term "AALT bs r1 r2"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1692
               & $\mid$ & @{term "ASEQ bs r\<^sub>1 r\<^sub>2"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1693
               & $\mid$ & @{term "ASTAR bs r"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1694
\end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1695
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1696
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1697
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1698
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1699
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1700
  @{thm (lhs) intern.simps(1)} & $\dn$ & @{thm (rhs) intern.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1701
  @{thm (lhs) intern.simps(2)} & $\dn$ & @{thm (rhs) intern.simps(2)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1702
  @{thm (lhs) intern.simps(3)} & $\dn$ & @{thm (rhs) intern.simps(3)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1703
  @{thm (lhs) intern.simps(4)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) intern.simps(4)[of "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1704
  @{thm (lhs) intern.simps(5)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) intern.simps(5)[of "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1705
  @{thm (lhs) intern.simps(6)} & $\dn$ & @{thm (rhs) intern.simps(6)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1706
\end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1707
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1708
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1709
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1710
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1711
  @{thm (lhs) erase.simps(1)} & $\dn$ & @{thm (rhs) erase.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1712
  @{thm (lhs) erase.simps(2)[of bs]} & $\dn$ & @{thm (rhs) erase.simps(2)[of bs]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1713
  @{thm (lhs) erase.simps(3)[of bs]} & $\dn$ & @{thm (rhs) erase.simps(3)[of bs]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1714
  @{thm (lhs) erase.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) erase.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1715
  @{thm (lhs) erase.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) erase.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1716
  @{thm (lhs) erase.simps(6)[of bs]} & $\dn$ & @{thm (rhs) erase.simps(6)[of bs]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1717
\end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1718
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1719
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1720
Some simple facts about erase
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1721
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1722
\begin{lemma}\mbox{}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1723
@{thm erase_bder}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1724
@{thm erase_intern}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1725
\end{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1726
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1727
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1728
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1729
  @{thm (lhs) bnullable.simps(1)} & $\dn$ & @{thm (rhs) bnullable.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1730
  @{thm (lhs) bnullable.simps(2)} & $\dn$ & @{thm (rhs) bnullable.simps(2)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1731
  @{thm (lhs) bnullable.simps(3)} & $\dn$ & @{thm (rhs) bnullable.simps(3)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1732
  @{thm (lhs) bnullable.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bnullable.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1733
  @{thm (lhs) bnullable.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bnullable.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1734
  @{thm (lhs) bnullable.simps(6)} & $\dn$ & @{thm (rhs) bnullable.simps(6)}\medskip\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1735
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1736
%  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1737
%  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1738
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1739
%  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1740
%  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1741
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1742
  @{thm (lhs) bder.simps(1)} & $\dn$ & @{thm (rhs) bder.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1743
  @{thm (lhs) bder.simps(2)} & $\dn$ & @{thm (rhs) bder.simps(2)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1744
  @{thm (lhs) bder.simps(3)} & $\dn$ & @{thm (rhs) bder.simps(3)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1745
  @{thm (lhs) bder.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bder.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1746
  @{thm (lhs) bder.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bder.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1747
  @{thm (lhs) bder.simps(6)} & $\dn$ & @{thm (rhs) bder.simps(6)}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1748
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1749
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1750
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1751
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1752
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1753
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1754
  @{thm (lhs) bmkeps.simps(1)} & $\dn$ & @{thm (rhs) bmkeps.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1755
  @{thm (lhs) bmkeps.simps(2)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bmkeps.simps(2)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1756
  @{thm (lhs) bmkeps.simps(3)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bmkeps.simps(3)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1757
  @{thm (lhs) bmkeps.simps(4)} & $\dn$ & @{thm (rhs) bmkeps.simps(4)}\medskip\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1758
\end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1759
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1760
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1761
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1762
@{thm [mode=IfThen] bder_retrieve}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1763
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1764
By induction on \<open>r\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1765
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1766
\begin{theorem}[Main Lemma]\mbox{}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1767
@{thm [mode=IfThen] MAIN_decode}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1768
\end{theorem}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1769
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1770
\noindent
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1771
Definition of the bitcoded lexer
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1772
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1773
@{thm blexer_def}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1774
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1775
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1776
\begin{theorem}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1777
@{thm blexer_correctness}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1778
\end{theorem}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1779
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1780
\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1781
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1782
section \<open>Optimisations\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1783
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1784
text \<open>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1785
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1786
  Derivatives as calculated by \Brz's method are usually more complex
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1787
  regular expressions than the initial one; the result is that the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1788
  derivative-based matching and lexing algorithms are often abysmally slow.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1789
  However, various optimisations are possible, such as the simplifications
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1790
  of @{term "ALT ZERO r"}, @{term "ALT r ZERO"}, @{term "SEQ ONE r"} and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1791
  @{term "SEQ r ONE"} to @{term r}. These simplifications can speed up the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1792
  algorithms considerably, as noted in \cite{Sulzmann2014}. One of the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1793
  advantages of having a simple specification and correctness proof is that
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1794
  the latter can be refined to prove the correctness of such simplification
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1795
  steps. While the simplification of regular expressions according to 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1796
  rules like
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1797
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1798
  \begin{equation}\label{Simpl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1799
  \begin{array}{lcllcllcllcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1800
  @{term "ALT ZERO r"} & \<open>\<Rightarrow>\<close> & @{term r} \hspace{8mm}%\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1801
  @{term "ALT r ZERO"} & \<open>\<Rightarrow>\<close> & @{term r} \hspace{8mm}%\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1802
  @{term "SEQ ONE r"}  & \<open>\<Rightarrow>\<close> & @{term r} \hspace{8mm}%\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1803
  @{term "SEQ r ONE"}  & \<open>\<Rightarrow>\<close> & @{term r}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1804
  \end{array}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1805
  \end{equation}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1806
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1807
  \noindent is well understood, there is an obstacle with the POSIX value
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1808
  calculation algorithm by Sulzmann and Lu: if we build a derivative regular
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1809
  expression and then simplify it, we will calculate a POSIX value for this
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1810
  simplified derivative regular expression, \emph{not} for the original (unsimplified)
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1811
  derivative regular expression. Sulzmann and Lu \cite{Sulzmann2014} overcome this obstacle by
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1812
  not just calculating a simplified regular expression, but also calculating
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1813
  a \emph{rectification function} that ``repairs'' the incorrect value.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1814
  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1815
  The rectification functions can be (slightly clumsily) implemented  in
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1816
  Isabelle/HOL as follows using some auxiliary functions:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1817
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1818
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1819
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1820
  @{thm (lhs) F_RIGHT.simps(1)} & $\dn$ & \<open>Right (f v)\<close>\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1821
  @{thm (lhs) F_LEFT.simps(1)} & $\dn$ & \<open>Left (f v)\<close>\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1822
  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1823
  @{thm (lhs) F_ALT.simps(1)} & $\dn$ & \<open>Right (f\<^sub>2 v)\<close>\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1824
  @{thm (lhs) F_ALT.simps(2)} & $\dn$ & \<open>Left (f\<^sub>1 v)\<close>\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1825
  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1826
  @{thm (lhs) F_SEQ1.simps(1)} & $\dn$ & \<open>Seq (f\<^sub>1 ()) (f\<^sub>2 v)\<close>\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1827
  @{thm (lhs) F_SEQ2.simps(1)} & $\dn$ & \<open>Seq (f\<^sub>1 v) (f\<^sub>2 ())\<close>\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1828
  @{thm (lhs) F_SEQ.simps(1)} & $\dn$ & \<open>Seq (f\<^sub>1 v\<^sub>1) (f\<^sub>2 v\<^sub>2)\<close>\medskip\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1829
  %\end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1830
  %
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1831
  %\begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1832
  @{term "simp_ALT (ZERO, DUMMY) (r\<^sub>2, f\<^sub>2)"} & $\dn$ & @{term "(r\<^sub>2, F_RIGHT f\<^sub>2)"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1833
  @{term "simp_ALT (r\<^sub>1, f\<^sub>1) (ZERO, DUMMY)"} & $\dn$ & @{term "(r\<^sub>1, F_LEFT f\<^sub>1)"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1834
  @{term "simp_ALT (r\<^sub>1, f\<^sub>1) (r\<^sub>2, f\<^sub>2)"} & $\dn$ & @{term "(ALT r\<^sub>1 r\<^sub>2, F_ALT f\<^sub>1 f\<^sub>2)"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1835
  @{term "simp_SEQ (ONE, f\<^sub>1) (r\<^sub>2, f\<^sub>2)"} & $\dn$ & @{term "(r\<^sub>2, F_SEQ1 f\<^sub>1 f\<^sub>2)"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1836
  @{term "simp_SEQ (r\<^sub>1, f\<^sub>1) (ONE, f\<^sub>2)"} & $\dn$ & @{term "(r\<^sub>1, F_SEQ2 f\<^sub>1 f\<^sub>2)"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1837
  @{term "simp_SEQ (r\<^sub>1, f\<^sub>1) (r\<^sub>2, f\<^sub>2)"} & $\dn$ & @{term "(SEQ r\<^sub>1 r\<^sub>2, F_SEQ f\<^sub>1 f\<^sub>2)"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1838
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1839
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1840
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1841
  \noindent
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1842
  The functions \<open>simp\<^bsub>Alt\<^esub>\<close> and \<open>simp\<^bsub>Seq\<^esub>\<close> encode the simplification rules
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1843
  in \eqref{Simpl} and compose the rectification functions (simplifications can occur
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1844
  deep inside the regular expression). The main simplification function is then 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1845
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1846
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1847
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1848
  @{term "simp (ALT r\<^sub>1 r\<^sub>2)"} & $\dn$ & @{term "simp_ALT (simp r\<^sub>1) (simp r\<^sub>2)"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1849
  @{term "simp (SEQ r\<^sub>1 r\<^sub>2)"} & $\dn$ & @{term "simp_SEQ (simp r\<^sub>1) (simp r\<^sub>2)"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1850
  @{term "simp r"} & $\dn$ & @{term "(r, id)"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1851
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1852
  \end{center} 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1853
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1854
  \noindent where @{term "id"} stands for the identity function. The
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1855
  function @{const simp} returns a simplified regular expression and a corresponding
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1856
  rectification function. Note that we do not simplify under stars: this
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1857
  seems to slow down the algorithm, rather than speed it up. The optimised
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1858
  lexer is then given by the clauses:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1859
  
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1860
  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1861
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1862
  @{thm (lhs) slexer.simps(1)} & $\dn$ & @{thm (rhs) slexer.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1863
  @{thm (lhs) slexer.simps(2)} & $\dn$ & 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1864
                         \<open>let (r\<^sub>s, f\<^sub>r) = simp (r \<close>$\backslash$\<open> c) in\<close>\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1865
                     & & \<open>case\<close> @{term "slexer r\<^sub>s s"} \<open>of\<close>\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1866
                     & & \phantom{$|$} @{term "None"}  \<open>\<Rightarrow>\<close> @{term None}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1867
                     & & $|$ @{term "Some v"} \<open>\<Rightarrow>\<close> \<open>Some (inj r c (f\<^sub>r v))\<close>                          
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1868
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1869
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1870
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1871
  \noindent
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1872
  In the second clause we first calculate the derivative @{term "der c r"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1873
  and then simplify the result. This gives us a simplified derivative
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1874
  \<open>r\<^sub>s\<close> and a rectification function \<open>f\<^sub>r\<close>. The lexer
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1875
  is then recursively called with the simplified derivative, but before
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1876
  we inject the character @{term c} into the value @{term v}, we need to rectify
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1877
  @{term v} (that is construct @{term "f\<^sub>r v"}). Before we can establish the correctness
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1878
  of @{term "slexer"}, we need to show that simplification preserves the language
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1879
  and simplification preserves our POSIX relation once the value is rectified
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1880
  (recall @{const "simp"} generates a (regular expression, rectification function) pair):
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1881
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1882
  \begin{lemma}\mbox{}\smallskip\\\label{slexeraux}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1883
  \begin{tabular}{ll}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1884
  (1) & @{thm L_fst_simp[symmetric]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1885
  (2) & @{thm[mode=IfThen] Posix_simp}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1886
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1887
  \end{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1888
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1889
  \begin{proof} Both are by induction on \<open>r\<close>. There is no
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1890
  interesting case for the first statement. For the second statement,
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1891
  of interest are the @{term "r = ALT r\<^sub>1 r\<^sub>2"} and @{term "r = SEQ r\<^sub>1
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1892
  r\<^sub>2"} cases. In each case we have to analyse four subcases whether
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1893
  @{term "fst (simp r\<^sub>1)"} and @{term "fst (simp r\<^sub>2)"} equals @{const
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1894
  ZERO} (respectively @{const ONE}). For example for @{term "r = ALT
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1895
  r\<^sub>1 r\<^sub>2"}, consider the subcase @{term "fst (simp r\<^sub>1) = ZERO"} and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1896
  @{term "fst (simp r\<^sub>2) \<noteq> ZERO"}. By assumption we know @{term "s \<in>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1897
  fst (simp (ALT r\<^sub>1 r\<^sub>2)) \<rightarrow> v"}. From this we can infer @{term "s \<in> fst (simp r\<^sub>2) \<rightarrow> v"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1898
  and by IH also (*) @{term "s \<in> r\<^sub>2 \<rightarrow> (snd (simp r\<^sub>2) v)"}. Given @{term "fst (simp r\<^sub>1) = ZERO"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1899
  we know @{term "L (fst (simp r\<^sub>1)) = {}"}. By the first statement
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1900
  @{term "L r\<^sub>1"} is the empty set, meaning (**) @{term "s \<notin> L r\<^sub>1"}.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1901
  Taking (*) and (**) together gives by the \mbox{\<open>P+R\<close>}-rule 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1902
  @{term "s \<in> ALT r\<^sub>1 r\<^sub>2 \<rightarrow> Right (snd (simp r\<^sub>2) v)"}. In turn this
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1903
  gives @{term "s \<in> ALT r\<^sub>1 r\<^sub>2 \<rightarrow> snd (simp (ALT r\<^sub>1 r\<^sub>2)) v"} as we need to show.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1904
  The other cases are similar.\qed
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1905
  \end{proof}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1906
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1907
  \noindent We can now prove relatively straightforwardly that the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1908
  optimised lexer produces the expected result:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1909
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1910
  \begin{theorem}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1911
  @{thm slexer_correctness}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1912
  \end{theorem}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1913
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1914
  \begin{proof} By induction on @{term s} generalising over @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1915
  r}. The case @{term "[]"} is trivial. For the cons-case suppose the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1916
  string is of the form @{term "c # s"}. By induction hypothesis we
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1917
  know @{term "slexer r s = lexer r s"} holds for all @{term r} (in
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1918
  particular for @{term "r"} being the derivative @{term "der c
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1919
  r"}). Let @{term "r\<^sub>s"} be the simplified derivative regular expression, that is @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1920
  "fst (simp (der c r))"}, and @{term "f\<^sub>r"} be the rectification
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1921
  function, that is @{term "snd (simp (der c r))"}.  We distinguish the cases
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1922
  whether (*) @{term "s \<in> L (der c r)"} or not. In the first case we
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1923
  have by Theorem~\ref{lexercorrect}(2) a value @{term "v"} so that @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1924
  "lexer (der c r) s = Some v"} and @{term "s \<in> der c r \<rightarrow> v"} hold.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1925
  By Lemma~\ref{slexeraux}(1) we can also infer from~(*) that @{term "s
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1926
  \<in> L r\<^sub>s"} holds.  Hence we know by Theorem~\ref{lexercorrect}(2) that
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1927
  there exists a @{term "v'"} with @{term "lexer r\<^sub>s s = Some v'"} and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1928
  @{term "s \<in> r\<^sub>s \<rightarrow> v'"}. From the latter we know by
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1929
  Lemma~\ref{slexeraux}(2) that @{term "s \<in> der c r \<rightarrow> (f\<^sub>r v')"} holds.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1930
  By the uniqueness of the POSIX relation (Theorem~\ref{posixdeterm}) we
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1931
  can infer that @{term v} is equal to @{term "f\<^sub>r v'"}---that is the 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1932
  rectification function applied to @{term "v'"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1933
  produces the original @{term "v"}.  Now the case follows by the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1934
  definitions of @{const lexer} and @{const slexer}.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1935
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1936
  In the second case where @{term "s \<notin> L (der c r)"} we have that
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1937
  @{term "lexer (der c r) s = None"} by Theorem~\ref{lexercorrect}(1).  We
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1938
  also know by Lemma~\ref{slexeraux}(1) that @{term "s \<notin> L r\<^sub>s"}. Hence
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1939
  @{term "lexer r\<^sub>s s = None"} by Theorem~\ref{lexercorrect}(1) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1940
  by IH then also @{term "slexer r\<^sub>s s = None"}. With this we can
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1941
  conclude in this case too.\qed   
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1942
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1943
  \end{proof} 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1944
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1945
\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1946
fy the result. This gives us a simplified derivative
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1947
  \<open>r\<^sub>s\<close> and a rectification function \<open>f\<^sub>r\<close>. The lexer
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1948
  is then recursively called with the simplified derivative, but before
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1949
  we inject the character @{term c} into the value @{term v}, we need to rectify
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1950
  @{term v} (that is construct @{term "f\<^sub>r v"}). Before we can establish the correctness
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1951
  of @{term "slexer"}, we need to show that simplification preserves the language
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1952
  and simplification preserves our POSIX relation once the value is rectified
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1953
  (recall @{const "simp"} generates a (regular expression, rectification function) pair):
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1954
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1955
  \begin{lemma}\mbox{}\smallskip\\\label{slexeraux}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1956
  \begin{tabular}{ll}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1957
  (1) & @{thm L_fst_simp[symmetric]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1958
  (2) & @{thm[mode=IfThen] Posix_simp}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1959
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1960
  \end{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1961
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1962
  \begin{proof} Both are by induction on \<open>r\<close>. There is no
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1963
  interesting case for the first statement. For the second statement,
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1964
  of interest are the @{term "r = ALT r\<^sub>1 r\<^sub>2"} and @{term "r = SEQ r\<^sub>1
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1965
  r\<^sub>2"} cases. In each case we have to analyse four subcases whether
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1966
  @{term "fst (simp r\<^sub>1)"} and @{term "fst (simp r\<^sub>2)"} equals @{const
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1967
  ZERO} (respectively @{const ONE}). For example for @{term "r = ALT
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1968
  r\<^sub>1 r\<^sub>2"}, consider the subcase @{term "fst (simp r\<^sub>1) = ZERO"} and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1969
  @{term "fst (simp r\<^sub>2) \<noteq> ZERO"}. By assumption we know @{term "s \<in>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1970
  fst (simp (ALT r\<^sub>1 r\<^sub>2)) \<rightarrow> v"}. From this we can infer @{term "s \<in> fst (simp r\<^sub>2) \<rightarrow> v"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1971
  and by IH also (*) @{term "s \<in> r\<^sub>2 \<rightarrow> (snd (simp r\<^sub>2) v)"}. Given @{term "fst (simp r\<^sub>1) = ZERO"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1972
  we know @{term "L (fst (simp r\<^sub>1)) = {}"}. By the first statement
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1973
  @{term "L r\<^sub>1"} is the empty set, meaning (**) @{term "s \<notin> L r\<^sub>1"}.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1974
  Taking (*) and (**) together gives by the \mbox{\<open>P+R\<close>}-rule 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1975
  @{term "s \<in> ALT r\<^sub>1 r\<^sub>2 \<rightarrow> Right (snd (simp r\<^sub>2) v)"}. In turn this
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1976
  gives @{term "s \<in> ALT r\<^sub>1 r\<^sub>2 \<rightarrow> snd (simp (ALT r\<^sub>1 r\<^sub>2)) v"} as we need to show.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1977
  The other cases are similar.\qed
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1978
  \end{proof}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1979
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1980
  \noindent We can now prove relatively straightforwardly that the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1981
  optimised lexer produces the expected result:
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1982
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1983
  \begin{theorem}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1984
  @{thm slexer_correctness}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1985
  \end{theorem}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1986
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1987
  \begin{proof} By induction on @{term s} generalising over @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1988
  r}. The case @{term "[]"} is trivial. For the cons-case suppose the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1989
  string is of the form @{term "c # s"}. By induction hypothesis we
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1990
  know @{term "slexer r s = lexer r s"} holds for all @{term r} (in
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1991
  particular for @{term "r"} being the derivative @{term "der c
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1992
  r"}). Let @{term "r\<^sub>s"} be the simplified derivative regular expression, that is @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1993
  "fst (simp (der c r))"}, and @{term "f\<^sub>r"} be the rectification
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1994
  function, that is @{term "snd (simp (der c r))"}.  We distinguish the cases
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1995
  whether (*) @{term "s \<in> L (der c r)"} or not. In the first case we
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1996
  have by Theorem~\ref{lexercorrect}(2) a value @{term "v"} so that @{term
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1997
  "lexer (der c r) s = Some v"} and @{term "s \<in> der c r \<rightarrow> v"} hold.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1998
  By Lemma~\ref{slexeraux}(1) we can also infer from~(*) that @{term "s
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  1999
  \<in> L r\<^sub>s"} holds.  Hence we know by Theorem~\ref{lexercorrect}(2) that
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2000
  there exists a @{term "v'"} with @{term "lexer r\<^sub>s s = Some v'"} and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2001
  @{term "s \<in> r\<^sub>s \<rightarrow> v'"}. From the latter we know by
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2002
  Lemma~\ref{slexeraux}(2) that @{term "s \<in> der c r \<rightarrow> (f\<^sub>r v')"} holds.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2003
  By the uniqueness of the POSIX relation (Theorem~\ref{posixdeterm}) we
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2004
  can infer that @{term v} is equal to @{term "f\<^sub>r v'"}---that is the 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2005
  rectification function applied to @{term "v'"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2006
  produces the original @{term "v"}.  Now the case follows by the
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2007
  definitions of @{const lexer} and @{const slexer}.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2008
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2009
  In the second case where @{term "s \<notin> L (der c r)"} we have that
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2010
  @{term "lexer (der c r) s = None"} by Theorem~\ref{lexercorrect}(1).  We
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2011
  also know by Lemma~\ref{slexeraux}(1) that @{term "s \<notin> L r\<^sub>s"}. Hence
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2012
  @{term "lexer r\<^sub>s s = None"} by Theorem~\ref{lexercorrect}(1) and
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2013
  by IH then also @{term "slexer r\<^sub>s s = None"}. With this we can
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2014
  conclude in this case too.\qed   
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2015
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2016
  \end{proof} 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2017
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2018
\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2019
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2020
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2021
section \<open>HERE\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2022
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2023
text \<open>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2024
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2025
  \begin{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2026
  @{thm [mode=IfThen] bder_retrieve}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2027
  \end{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2028
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2029
  \begin{proof}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2030
  By induction on the definition of @{term "erase r"}. The cases for rule 1) and 2) are
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2031
  straightforward as @{term "der c ZERO"} and @{term "der c ONE"} are both equal to 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2032
  @{term ZERO}. This means @{term "\<Turnstile> v : ZERO"} cannot hold. Similarly in case of rule 3)
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2033
  where @{term r} is of the form @{term "ACHAR d"} with @{term "c = d"}. Then by assumption
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2034
  we know @{term "\<Turnstile> v : ONE"}, which implies @{term "v = Void"}. The equation follows by 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2035
  simplification of left- and right-hand side. In  case @{term "c \<noteq> d"} we have again
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2036
  @{term "\<Turnstile> v : ZERO"}, which cannot  hold. 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2037
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2038
  For rule 4a) we have again @{term "\<Turnstile> v : ZERO"}. The property holds by IH for rule 4b).
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2039
  The  induction hypothesis is 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2040
  \[
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2041
  @{term "retrieve (bder c r) v = retrieve r (injval (erase r) c v)"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2042
  \]
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2043
  which is what left- and right-hand side simplify to.  The slightly more interesting case
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2044
  is for 4c). By assumption  we have 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2045
  @{term "\<Turnstile> v : ALT (der c (erase r\<^sub>1)) (der c (erase (AALTs bs (r\<^sub>2 # rs))))"}. This means we 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2046
  have either (*) @{term "\<Turnstile> v1 : der c (erase r\<^sub>1)"} with @{term "v = Left v1"} or
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2047
  (**) @{term "\<Turnstile> v2 : der c (erase (AALTs bs (r\<^sub>2 # rs)))"} with @{term "v = Right v2"}.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2048
  The former  case is straightforward by simplification. The second case is \ldots TBD.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2049
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2050
  Rule 5) TBD.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2051
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2052
  Finally for rule 6) the reasoning is as follows:   By assumption we  have
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2053
  @{term "\<Turnstile> v : SEQ (der c (erase r)) (STAR (erase r))"}. This means we also have
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2054
  @{term "v = Seq v1 v2"}, @{term "\<Turnstile> v1 : der c (erase r)"}  and @{term "v2 = Stars vs"}.
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2055
  We want to prove
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2056
  \begin{align}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2057
  & @{term "retrieve (ASEQ bs (fuse [Z] (bder c r)) (ASTAR [] r)) v"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2058
  &= @{term "retrieve (ASTAR bs r) (injval (STAR (erase r)) c v)"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2059
  \end{align}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2060
  The right-hand side @{term inj}-expression is equal to 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2061
  @{term "Stars (injval (erase r) c v1 # vs)"}, which means the @{term  retrieve}-expression
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2062
  simplifies to 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2063
  \[
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2064
  @{term "bs @ [Z] @ retrieve r (injval (erase r) c v1) @ retrieve (ASTAR [] r) (Stars vs)"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2065
  \]
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2066
  The left-hand side (3) above simplifies to 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2067
  \[
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2068
  @{term "bs @ retrieve (fuse [Z] (bder c r)) v1 @ retrieve (ASTAR [] r) (Stars vs)"} 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2069
  \]
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2070
  We can move out the @{term "fuse  [Z]"} and then use the IH to show that left-hand side
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2071
  and right-hand side are equal. This completes the proof. 
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2072
  \end{proof}   
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2073
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2074
   
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2075
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2076
  \bibliographystyle{plain}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2077
  \bibliography{root}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2078
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2079
\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2080
(*<*)
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2081
end
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2082
(*>*)
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2083
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2084
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2085
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2086
(*
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2087
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2088
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2089
  \begin{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2090
    @{thm[mode=Rule] aggressive.intros(1)[of "bs" "bs1" "rs" "r"]}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2091
  end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2092
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2093
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2094
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2095
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2096
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2097
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2098
  @{thm (lhs) code.simps(1)} & $\dn$ & @{thm (rhs) code.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2099
  @{thm (lhs) code.simps(2)} & $\dn$ & @{thm (rhs) code.simps(2)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2100
  @{thm (lhs) code.simps(3)} & $\dn$ & @{thm (rhs) code.simps(3)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2101
  @{thm (lhs) code.simps(4)} & $\dn$ & @{thm (rhs) code.simps(4)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2102
  @{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"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2103
  @{thm (lhs) code.simps(6)} & $\dn$ & @{thm (rhs) code.simps(6)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2104
  @{thm (lhs) code.simps(7)} & $\dn$ & @{thm (rhs) code.simps(7)}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2105
\end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2106
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2107
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2108
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2109
\begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2110
  @{term areg} & $::=$ & @{term "AZERO"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2111
               & $\mid$ & @{term "AONE bs"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2112
               & $\mid$ & @{term "ACHAR bs c"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2113
               & $\mid$ & @{term "AALT bs r1 r2"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2114
               & $\mid$ & @{term "ASEQ bs r\<^sub>1 r\<^sub>2"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2115
               & $\mid$ & @{term "ASTAR bs r"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2116
\end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2117
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2118
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2119
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2120
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2121
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2122
  @{thm (lhs) intern.simps(1)} & $\dn$ & @{thm (rhs) intern.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2123
  @{thm (lhs) intern.simps(2)} & $\dn$ & @{thm (rhs) intern.simps(2)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2124
  @{thm (lhs) intern.simps(3)} & $\dn$ & @{thm (rhs) intern.simps(3)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2125
  @{thm (lhs) intern.simps(4)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) intern.simps(4)[of "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2126
  @{thm (lhs) intern.simps(5)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) intern.simps(5)[of "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2127
  @{thm (lhs) intern.simps(6)} & $\dn$ & @{thm (rhs) intern.simps(6)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2128
\end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2129
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2130
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2131
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2132
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2133
  @{thm (lhs) erase.simps(1)} & $\dn$ & @{thm (rhs) erase.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2134
  @{thm (lhs) erase.simps(2)[of bs]} & $\dn$ & @{thm (rhs) erase.simps(2)[of bs]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2135
  @{thm (lhs) erase.simps(3)[of bs]} & $\dn$ & @{thm (rhs) erase.simps(3)[of bs]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2136
  @{thm (lhs) erase.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) erase.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2137
  @{thm (lhs) erase.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) erase.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2138
  @{thm (lhs) erase.simps(6)[of bs]} & $\dn$ & @{thm (rhs) erase.simps(6)[of bs]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2139
\end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2140
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2141
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2142
Some simple facts about erase
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2143
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2144
\begin{lemma}\mbox{}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2145
@{thm erase_bder}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2146
@{thm erase_intern}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2147
\end{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2148
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2149
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2150
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2151
  @{thm (lhs) bnullable.simps(1)} & $\dn$ & @{thm (rhs) bnullable.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2152
  @{thm (lhs) bnullable.simps(2)} & $\dn$ & @{thm (rhs) bnullable.simps(2)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2153
  @{thm (lhs) bnullable.simps(3)} & $\dn$ & @{thm (rhs) bnullable.simps(3)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2154
  @{thm (lhs) bnullable.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bnullable.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2155
  @{thm (lhs) bnullable.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bnullable.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2156
  @{thm (lhs) bnullable.simps(6)} & $\dn$ & @{thm (rhs) bnullable.simps(6)}\medskip\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2157
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2158
%  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2159
%  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2160
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2161
%  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2162
%  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2163
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2164
  @{thm (lhs) bder.simps(1)} & $\dn$ & @{thm (rhs) bder.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2165
  @{thm (lhs) bder.simps(2)} & $\dn$ & @{thm (rhs) bder.simps(2)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2166
  @{thm (lhs) bder.simps(3)} & $\dn$ & @{thm (rhs) bder.simps(3)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2167
  @{thm (lhs) bder.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bder.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2168
  @{thm (lhs) bder.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bder.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2169
  @{thm (lhs) bder.simps(6)} & $\dn$ & @{thm (rhs) bder.simps(6)}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2170
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2171
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2172
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2173
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2174
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2175
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2176
  @{thm (lhs) bmkeps.simps(1)} & $\dn$ & @{thm (rhs) bmkeps.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2177
  @{thm (lhs) bmkeps.simps(2)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bmkeps.simps(2)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2178
  @{thm (lhs) bmkeps.simps(3)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bmkeps.simps(3)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2179
  @{thm (lhs) bmkeps.simps(4)} & $\dn$ & @{thm (rhs) bmkeps.simps(4)}\medskip\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2180
\end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2181
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2182
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2183
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2184
@{thm [mode=IfThen] bder_retrieve}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2185
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2186
By induction on \<open>r\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2187
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2188
\begin{theorem}[Main Lemma]\mbox{}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2189
@{thm [mode=IfThen] MAIN_decode}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2190
\end{theorem}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2191
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2192
\noindent
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2193
Definition of the bitcoded lexer
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2194
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2195
@{thm blexer_def}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2196
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2197
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2198
\begin{theorem}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2199
@{thm blexer_correctness}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2200
\end{theorem}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2201
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2202
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2203
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2204
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2205
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2206
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2207
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2208
  @{thm (lhs) code.simps(1)} & $\dn$ & @{thm (rhs) code.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2209
  @{thm (lhs) code.simps(2)} & $\dn$ & @{thm (rhs) code.simps(2)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2210
  @{thm (lhs) code.simps(3)} & $\dn$ & @{thm (rhs) code.simps(3)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2211
  @{thm (lhs) code.simps(4)} & $\dn$ & @{thm (rhs) code.simps(4)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2212
  @{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"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2213
  @{thm (lhs) code.simps(6)} & $\dn$ & @{thm (rhs) code.simps(6)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2214
  @{thm (lhs) code.simps(7)} & $\dn$ & @{thm (rhs) code.simps(7)}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2215
\end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2216
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2217
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2218
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2219
\begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2220
  @{term areg} & $::=$ & @{term "AZERO"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2221
               & $\mid$ & @{term "AONE bs"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2222
               & $\mid$ & @{term "ACHAR bs c"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2223
               & $\mid$ & @{term "AALT bs r1 r2"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2224
               & $\mid$ & @{term "ASEQ bs r\<^sub>1 r\<^sub>2"}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2225
               & $\mid$ & @{term "ASTAR bs r"}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2226
\end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2227
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2228
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2229
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2230
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2231
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2232
  @{thm (lhs) intern.simps(1)} & $\dn$ & @{thm (rhs) intern.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2233
  @{thm (lhs) intern.simps(2)} & $\dn$ & @{thm (rhs) intern.simps(2)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2234
  @{thm (lhs) intern.simps(3)} & $\dn$ & @{thm (rhs) intern.simps(3)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2235
  @{thm (lhs) intern.simps(4)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) intern.simps(4)[of "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2236
  @{thm (lhs) intern.simps(5)[of "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) intern.simps(5)[of "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2237
  @{thm (lhs) intern.simps(6)} & $\dn$ & @{thm (rhs) intern.simps(6)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2238
\end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2239
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2240
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2241
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2242
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2243
  @{thm (lhs) erase.simps(1)} & $\dn$ & @{thm (rhs) erase.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2244
  @{thm (lhs) erase.simps(2)[of bs]} & $\dn$ & @{thm (rhs) erase.simps(2)[of bs]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2245
  @{thm (lhs) erase.simps(3)[of bs]} & $\dn$ & @{thm (rhs) erase.simps(3)[of bs]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2246
  @{thm (lhs) erase.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) erase.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2247
  @{thm (lhs) erase.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) erase.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2248
  @{thm (lhs) erase.simps(6)[of bs]} & $\dn$ & @{thm (rhs) erase.simps(6)[of bs]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2249
\end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2250
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2251
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2252
Some simple facts about erase
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2253
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2254
\begin{lemma}\mbox{}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2255
@{thm erase_bder}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2256
@{thm erase_intern}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2257
\end{lemma}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2258
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2259
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2260
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2261
  @{thm (lhs) bnullable.simps(1)} & $\dn$ & @{thm (rhs) bnullable.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2262
  @{thm (lhs) bnullable.simps(2)} & $\dn$ & @{thm (rhs) bnullable.simps(2)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2263
  @{thm (lhs) bnullable.simps(3)} & $\dn$ & @{thm (rhs) bnullable.simps(3)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2264
  @{thm (lhs) bnullable.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bnullable.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2265
  @{thm (lhs) bnullable.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bnullable.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2266
  @{thm (lhs) bnullable.simps(6)} & $\dn$ & @{thm (rhs) bnullable.simps(6)}\medskip\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2267
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2268
%  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2269
%  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2270
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2271
%  \begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2272
%  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2273
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2274
  @{thm (lhs) bder.simps(1)} & $\dn$ & @{thm (rhs) bder.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2275
  @{thm (lhs) bder.simps(2)} & $\dn$ & @{thm (rhs) bder.simps(2)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2276
  @{thm (lhs) bder.simps(3)} & $\dn$ & @{thm (rhs) bder.simps(3)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2277
  @{thm (lhs) bder.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bder.simps(4)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2278
  @{thm (lhs) bder.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bder.simps(5)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2279
  @{thm (lhs) bder.simps(6)} & $\dn$ & @{thm (rhs) bder.simps(6)}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2280
  \end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2281
  \end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2282
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2283
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2284
\begin{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2285
  \begin{tabular}{lcl}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2286
  @{thm (lhs) bmkeps.simps(1)} & $\dn$ & @{thm (rhs) bmkeps.simps(1)}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2287
  @{thm (lhs) bmkeps.simps(2)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bmkeps.simps(2)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2288
  @{thm (lhs) bmkeps.simps(3)[of bs "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{thm (rhs) bmkeps.simps(3)[of bs "r\<^sub>1" "r\<^sub>2"]}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2289
  @{thm (lhs) bmkeps.simps(4)} & $\dn$ & @{thm (rhs) bmkeps.simps(4)}\medskip\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2290
\end{tabular}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2291
\end{center}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2292
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2293
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2294
@{thm [mode=IfThen] bder_retrieve}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2295
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2296
By induction on \<open>r\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2297
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2298
\begin{theorem}[Main Lemma]\mbox{}\\
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2299
@{thm [mode=IfThen] MAIN_decode}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2300
\end{theorem}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2301
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2302
\noindent
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2303
Definition of the bitcoded lexer
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2304
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2305
@{thm blexer_def}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2306
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2307
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2308
\begin{theorem}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2309
@{thm blexer_correctness}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2310
\end{theorem}
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2311
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2312
\<close>
f5866f1d6a59 from christian
Chengsong
parents:
diff changeset
  2313
\<close>*)