# HG changeset patch # User Christian Urban # Date 1651276208 -3600 # Node ID f493a20feeb365200b6ea2674db28b678494c998 # Parent f9cdc295ccf77bb38e4fa20a42c0c706d0927bc2 updated to include the paper diff -r f9cdc295ccf7 -r f493a20feeb3 thys3/Paper.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys3/Paper.thy Sat Apr 30 00:50:08 2022 +0100 @@ -0,0 +1,1545 @@ +(*<*) +theory Paper + imports + "Posix.FBound" + "HOL-Library.LaTeXsugar" +begin + +declare [[show_question_marks = false]] + +notation (latex output) + If ("(\<^latex>\\\textrm{\if\<^latex>\}\ (_)/ \<^latex>\\\textrm{\then\<^latex>\}\ (_)/ \<^latex>\\\textrm{\else\<^latex>\}\ (_))" 10) and + Cons ("_\<^latex>\\\mbox{$\\,$}\::\<^latex>\\\mbox{$\\,$}\_" [75,73] 73) + + +abbreviation + "der_syn r c \ der c r" +abbreviation + "ders_syn r s \ ders s r" +abbreviation + "bder_syn r c \ bder c r" + +notation (latex output) + der_syn ("_\\_" [79, 1000] 76) and + ders_syn ("_\\_" [79, 1000] 76) and + bder_syn ("_\\_" [79, 1000] 76) and + bders ("_\\_" [79, 1000] 76) and + bders_simp ("_\\\<^sub>b\<^sub>s\<^sub>i\<^sub>m\<^sub>p _" [79, 1000] 76) and + + ZERO ("\<^bold>0" 81) and + ONE ("\<^bold>1" 81) and + CH ("_" [1000] 80) and + ALT ("_ + _" [77,77] 78) and + SEQ ("_ \ _" [77,77] 78) and + STAR ("_\<^sup>*" [79] 78) and + + val.Void ("Empty" 78) and + val.Char ("Char _" [1000] 78) and + val.Left ("Left _" [79] 78) and + val.Right ("Right _" [1000] 78) and + val.Seq ("Seq _ _" [79,79] 78) and + val.Stars ("Stars _" [79] 78) and + + Prf ("\ _ : _" [75,75] 75) and + Posix ("'(_, _') \ _" [63,75,75] 75) and + + flat ("|_|" [75] 74) and + flats ("|_|" [72] 74) and + injval ("inj _ _ _" [79,77,79] 76) and + mkeps ("mkeps _" [79] 76) and + length ("len _" [73] 73) and + set ("_" [73] 73) and + + AZERO ("ZERO" 81) and + AONE ("ONE _" [79] 78) and + ACHAR ("CHAR _ _" [79, 79] 80) and + AALTs ("ALTs _ _" [77,77] 78) and + ASEQ ("SEQ _ _ _" [79, 79,79] 78) and + ASTAR ("STAR _ _" [79, 79] 78) and + + code ("code _" [79] 74) and + intern ("_\<^latex>\\\mbox{$^\\uparrow$}\" [900] 80) and + erase ("_\<^latex>\\\mbox{$^\\downarrow$}\" [1000] 74) and + bnullable ("bnullable _" [1000] 80) and + bsimp_AALTs ("bsimpALT _ _" [10,1000] 80) and + bsimp_ASEQ ("bsimpSEQ _ _ _" [10,1000,1000] 80) and + bmkeps ("bmkeps _" [1000] 80) and + + eq1 ("_ \ _" [77,77] 76) and + + srewrite ("_\<^latex>\\\mbox{$\\,\\stackrel{s}{\\leadsto}$}\ _" [71, 71] 80) and + rrewrites ("_ \<^latex>\\\mbox{$\\,\\leadsto^*$}\ _" [71, 71] 80) and + srewrites ("_ \<^latex>\\\mbox{$\\,\\stackrel{s}{\\leadsto}^*$}\ _" [71, 71] 80) and + blexer_simp ("blexer\<^sup>+" 1000) + + +lemma better_retrieve: + shows "rs \ Nil ==> retrieve (AALTs bs (r#rs)) (Left v) = bs @ retrieve r v" + and "rs \ Nil ==> retrieve (AALTs bs (r#rs)) (Right v) = bs @ retrieve (AALTs [] rs) v" + apply (metis list.exhaust retrieve.simps(4)) + by (metis list.exhaust retrieve.simps(5)) + + + + +(*>*) + +section {* Introduction *} + +text {* + +In the last fifteen or so years, Brzozowski's derivatives of regular +expressions have sparked quite a bit of interest in the functional +programming and theorem prover communities. The beauty of +Brzozowski's derivatives \cite{Brzozowski1964} is that they are neatly +expressible in any functional language, and easily definable and +reasoned about in theorem provers---the definitions just consist of +inductive datatypes and simple recursive functions. Derivatives of a +regular expression, written @{term "der c r"}, give a simple solution +to the problem of matching a string @{term s} with a regular +expression @{term r}: if the derivative of @{term r} w.r.t.\ (in +succession) all the characters of the string matches the empty string, +then @{term r} matches @{term s} (and {\em vice versa}). We are aware +of a mechanised correctness proof of Brzozowski's derivative-based matcher in HOL4 by +Owens and Slind~\cite{Owens2008}. Another one in Isabelle/HOL is part +of the work by Krauss and Nipkow~\cite{Krauss2011}. And another one +in Coq is given by Coquand and Siles \cite{Coquand2012}. +Also Ribeiro and Du Bois give one in Agda~\cite{RibeiroAgda2017}. + + +However, there are two difficulties with derivative-based matchers: +First, Brzozowski's original matcher only generates a yes/no answer +for whether a regular expression matches a string or not. This is too +little information in the context of lexing where separate tokens must +be identified and also classified (for example as keywords +or identifiers). Sulzmann and Lu~\cite{Sulzmann2014} overcome this +difficulty by cleverly extending Brzozowski's matching +algorithm. Their extended version generates additional information on +\emph{how} a regular expression matches a string following the POSIX +rules for regular expression matching. They achieve this by adding a +second ``phase'' to Brzozowski's algorithm involving an injection +function. In our own earlier work we provided the formal +specification of what POSIX matching means and proved in Isabelle/HOL +the correctness +of Sulzmann and Lu's extended algorithm accordingly +\cite{AusafDyckhoffUrban2016}. + +The second difficulty is that Brzozowski's derivatives can +grow to arbitrarily big sizes. For example if we start with the +regular expression \mbox{@{text "(a + aa)\<^sup>*"}} and take +successive derivatives according to the character $a$, we end up with +a sequence of ever-growing derivatives like + +\def\ll{\stackrel{\_\backslash{} a}{\longrightarrow}} +\begin{center} +\begin{tabular}{rll} +$(a + aa)^*$ & $\ll$ & $(\ONE + \ONE{}a) \cdot (a + aa)^*$\\ +& $\ll$ & $(\ZERO + \ZERO{}a + \ONE) \cdot (a + aa)^* \;+\; (\ONE + \ONE{}a) \cdot (a + aa)^*$\\ +& $\ll$ & $(\ZERO + \ZERO{}a + \ZERO) \cdot (a + aa)^* + (\ONE + \ONE{}a) \cdot (a + aa)^* \;+\; $\\ +& & $\qquad(\ZERO + \ZERO{}a + \ONE) \cdot (a + aa)^* + (\ONE + \ONE{}a) \cdot (a + aa)^*$\\ +& $\ll$ & \ldots \hspace{15mm}(regular expressions of sizes 98, 169, 283, 468, 767, \ldots) +\end{tabular} +\end{center} + +\noindent where after around 35 steps we run out of memory on a +typical computer (we shall define shortly the precise details of our +regular expressions and the derivative operation). Clearly, the +notation involving $\ZERO$s and $\ONE$s already suggests +simplification rules that can be applied to regular regular +expressions, for example $\ZERO{}\,r \Rightarrow \ZERO$, $\ONE{}\,r +\Rightarrow r$, $\ZERO{} + r \Rightarrow r$ and $r + r \Rightarrow +r$. While such simple-minded simplifications have been proved in our +earlier work to preserve the correctness of Sulzmann and Lu's +algorithm \cite{AusafDyckhoffUrban2016}, they unfortunately do +\emph{not} help with limiting the growth of the derivatives shown +above: the growth is slowed, but the derivatives can still grow rather +quickly beyond any finite bound. + + +Sulzmann and Lu overcome this ``growth problem'' in a second algorithm +\cite{Sulzmann2014} where they introduce bitcoded +regular expressions. In this version, POSIX values are +represented as bitsequences and such sequences are incrementally generated +when derivatives are calculated. The compact representation +of bitsequences and regular expressions allows them to define a more +``aggressive'' simplification method that keeps the size of the +derivatives finitely bounded no matter what the length of the string is. +They make some informal claims about the correctness and linear behaviour +of this version, but do not provide any supporting proof arguments, not +even ``pencil-and-paper'' arguments. They write about their bitcoded +\emph{incremental parsing method} (that is the algorithm to be formalised +in this paper): +% +\begin{quote}\it + ``Correctness Claim: We further claim that the incremental parsing + method [..] in combination with the simplification steps [..] + yields POSIX parse trees. We have tested this claim + extensively [..] but yet + have to work out all proof details.'' \cite[Page 14]{Sulzmann2014} +\end{quote} + +\noindent{}\textbf{Contributions:} We have implemented in Isabelle/HOL +the derivative-based lexing algorithm of Sulzmann and Lu +\cite{Sulzmann2014} where regular expressions are annotated with +bitsequences. We define the crucial simplification function as a +recursive function, without the need of a fix-point operation. One objective of +the simplification function is to remove duplicates of regular +expressions. For this Sulzmann and Lu use in their paper the standard +@{text nub} function from Haskell's list library, but this function +does not achieve the intended objective with bitcoded regular expressions. The +reason is that in the bitcoded setting, each copy generally has a +different bitcode annotation---so @{text nub} would never ``fire''. +Inspired by Scala's library for lists, we shall instead use a @{text +distinctWith} function that finds duplicates under an ``erasing'' function +which deletes bitcodes before comparing regular expressions. +We shall also introduce our own argument and definitions for +establishing the correctness of the bitcoded algorithm when +simplifications are included. Finally we +establish that the size of derivatives can be finitely bounded.\medskip%\footnote{ In this paper, we shall first briefly introduce the basic notions +%of regular expressions and describe the definition +%of POSIX lexing from our earlier work \cite{AusafDyckhoffUrban2016}. This serves +%as a reference point for what correctness means in our Isabelle/HOL proofs. We shall then prove +%the correctness for the bitcoded algorithm without simplification, and +%after that extend the proof to include simplification.}\smallskip +%\mbox{}\\[-10mm] + +\noindent In this paper, we shall first briefly introduce the basic notions +of regular expressions and describe the definition +of POSIX lexing from our earlier work \cite{AusafDyckhoffUrban2016}. This serves +as a reference point for what correctness means in our Isabelle/HOL proofs. We shall then prove +the correctness for the bitcoded algorithm without simplification, and +after that extend the proof to include simplification.\mbox{}\\[-6mm] + +*} + +section {* Background *} + +text {* + In our Isabelle/HOL formalisation strings are lists of characters with + the empty string being represented by the empty list, written $[]$, + and list-cons being written as $\_\!::\!\_\,$; string + concatenation is $\_ \,@\, \_\,$. We often use the usual + bracket notation for lists also for strings; for example a string + consisting of just a single character $c$ is written $[c]$. + Our regular expressions are defined as usual as the following inductive + datatype:\\[-4mm] + % + \begin{center} + @{text "r ::="} \; + @{const "ZERO"} $\mid$ + @{const "ONE"} $\mid$ + @{term "CH c"} $\mid$ + @{term "ALT r\<^sub>1 r\<^sub>2"} $\mid$ + @{term "SEQ r\<^sub>1 r\<^sub>2"} $\mid$ + @{term "STAR r"} + \end{center} + + \noindent where @{const ZERO} stands for the regular expression that does + not match any string, @{const ONE} for the regular expression that matches + only the empty string and @{term c} for matching a character literal. + The constructors $+$ and $\cdot$ represent alternatives and sequences, respectively. + We sometimes omit the $\cdot$ in a sequence regular expression for brevity. + The + \emph{language} of a regular expression, written $L(r)$, is defined as usual + and we omit giving the definition here (see for example \cite{AusafDyckhoffUrban2016}). + + Central to Brzozowski's regular expression matcher are two functions + called @{text nullable} and \emph{derivative}. The latter is written + $r\backslash c$ for the derivative of the regular expression $r$ + w.r.t.~the character $c$. Both functions are defined by recursion over + regular expressions. +% +\begin{center} +\begin{tabular}{cc} + \begin{tabular}{r@ {\hspace{2mm}}c@ {\hspace{2mm}}l} + @{thm (lhs) der.simps(1)} & $\dn$ & @{thm (rhs) der.simps(1)}\\ + @{thm (lhs) der.simps(2)} & $\dn$ & @{thm (rhs) der.simps(2)}\\ + @{thm (lhs) der.simps(3)} & $\dn$ & @{thm (rhs) der.simps(3)}\\ + @{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"]}\\ + @{thm (lhs) der.simps(5)[of c "r\<^sub>1" "r\<^sub>2"]} & $\dn$ & @{text "if"} @{term "nullable(r\<^sub>1)"}\\ + & & @{text "then"} @{term "ALT (SEQ (der c r\<^sub>1) r\<^sub>2) (der c r\<^sub>2)"}\\ + & & @{text "else"} @{term "SEQ (der c r\<^sub>1) r\<^sub>2"}\\ + % & & @{thm (rhs) der.simps(5)[of c "r\<^sub>1" "r\<^sub>2"]}\\ + @{thm (lhs) der.simps(6)} & $\dn$ & @{thm (rhs) der.simps(6)} + \end{tabular} + & + \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l} + @{thm (lhs) nullable.simps(1)} & $\dn$ & @{thm (rhs) nullable.simps(1)}\\ + @{thm (lhs) nullable.simps(2)} & $\dn$ & @{thm (rhs) nullable.simps(2)}\\ + @{thm (lhs) nullable.simps(3)} & $\dn$ & @{thm (rhs) nullable.simps(3)}\\ + @{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"]}\\ + @{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"]}\\ + @{thm (lhs) nullable.simps(6)} & $\dn$ & @{thm (rhs) nullable.simps(6)}%\medskip\\ + \end{tabular} +\end{tabular} +\end{center} + + \noindent + We can extend this definition to give derivatives w.r.t.~strings, + namely @{thm (lhs) ders.simps(1)}$\dn$@{thm (rhs) ders.simps(1)} + and @{thm (lhs) ders.simps(2)}$\dn$@{thm (rhs) ders.simps(2)}. + Using @{text nullable} and the derivative operation, we can +define a simple regular expression matcher, namely +$@{text "match s r"} \;\dn\; @{term nullable}(r\backslash s)$. +This is essentially Brzozowski's algorithm from 1964. Its +main virtue is that the algorithm can be easily implemented as a +functional program (either in a functional programming language or in +a theorem prover). The correctness of @{text match} amounts to +establishing the property:%\footnote{It is a fun exercise to formally prove this property in a theorem prover.} +% +\begin{proposition}\label{matchcorr} +@{text "match s r"} \;\;\text{if and only if}\;\; $s \in L(r)$ +\end{proposition} + +\noindent +It is a fun exercise to formally prove this property in a theorem prover. + +The novel idea of Sulzmann and Lu is to extend this algorithm for +lexing, where it is important to find out which part of the string +is matched by which part of the regular expression. +For this Sulzmann and Lu presented two lexing algorithms in their paper + \cite{Sulzmann2014}. The first algorithm consists of two phases: first a + matching phase (which is Brzozowski's algorithm) and then a value + construction phase. The values encode \emph{how} a regular expression + matches a string. \emph{Values} are defined as the inductive datatype + % + \begin{center} + @{text "v ::="} + @{const "Void"} $\mid$ + @{term "val.Char c"} $\mid$ + @{term "Left v"} $\mid$ + @{term "Right v"} $\mid$ + @{term "Seq v\<^sub>1 v\<^sub>2"} $\mid$ + @{term "Stars vs"} + \end{center} + + \noindent where we use @{term vs} to stand for a list of values. The + string underlying a value can be calculated by a @{const flat} + function, written @{term "flat DUMMY"}. It traverses a value and + collects the characters contained in it \cite{AusafDyckhoffUrban2016}. + + + Sulzmann and Lu also define inductively an + inhabitation relation that associates values to regular expressions. This + is defined by the following six rules for the values: + % + \begin{center} + \begin{tabular}{@ {\hspace{-12mm}}c@ {}} + \begin{tabular}{@ {}c@ {}} + \\[-8mm] + @{thm[mode=Axiom] Prf.intros(4)}\\ + @{thm[mode=Axiom] Prf.intros(5)[of "c"]} + \end{tabular} + \quad + @{thm[mode=Rule] Prf.intros(2)[of "v\<^sub>1" "r\<^sub>1" "r\<^sub>2"]}\quad + @{thm[mode=Rule] Prf.intros(3)[of "v\<^sub>2" "r\<^sub>2" "r\<^sub>1"]}\quad + @{thm[mode=Rule] Prf.intros(1)[of "v\<^sub>1" "r\<^sub>1" "v\<^sub>2" "r\<^sub>2"]}\quad + @{thm[mode=Rule] Prf.intros(6)[of "vs" "r"]}\mbox{}\hspace{-10mm}\mbox{} + \end{tabular} + \end{center} + + \noindent Note that no values are associated with the regular expression + @{term ZERO}, since it cannot match any string. + It is routine to establish how values ``inhabiting'' a regular + expression correspond to the language of a regular expression, namely + @{thm L_flat_Prf}. + + In general there is more than one value inhabited by a regular + expression (meaning regular expressions can typically match more + than one string). But even when fixing a string from the language of the + regular expression, there are generally more than one way of how the + regular expression can match this string. POSIX lexing is about + identifying the unique value for a given regular expression and a + string that satisfies the informal POSIX rules (see + \cite{POSIX,Kuklewicz,OkuiSuzuki2010,Sulzmann2014,Vansummeren2006}).\footnote{POSIX + lexing acquired its name from the fact that the corresponding + rules were described as part of the POSIX specification for + Unix-like operating systems \cite{POSIX}.} Sometimes these + informal rules are called \emph{maximal munch rule} and \emph{rule priority}. + One contribution of our earlier paper is to give a convenient + specification for what POSIX values are (the inductive rules are shown in + Figure~\ref{POSIXrules}). + +\begin{figure}[t] + \begin{center} + \begin{tabular}{@ {}c@ {}} + @{thm[mode=Axiom] Posix.intros(1)}\P\@{term "ONE"} \quad + @{thm[mode=Axiom] Posix.intros(2)}\P\@{term "c"}\quad + @{thm[mode=Rule] Posix.intros(3)[of "s" "r\<^sub>1" "v" "r\<^sub>2"]}\P+L\\quad + @{thm[mode=Rule] Posix.intros(4)[of "s" "r\<^sub>2" "v" "r\<^sub>1"]}\P+R\\medskip\\ + $\mprset{flushleft} + \inferrule + {@{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 + @{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"]} \\\\ + @{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"]}} + {@{thm (concl) Posix.intros(5)[of "s\<^sub>1" "r\<^sub>1" "v\<^sub>1" "s\<^sub>2" "r\<^sub>2" "v\<^sub>2"]}}$\PS\\medskip\smallskip\\ + @{thm[mode=Axiom] Posix.intros(7)}\P[]\\qquad + $\mprset{flushleft} + \inferrule + {@{thm (prem 1) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]} \qquad + @{thm (prem 2) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]} \qquad + @{thm (prem 3) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]} \\\\ + @{thm (prem 4) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]}} + {@{thm (concl) Posix.intros(6)[of "s\<^sub>1" "r" "v" "s\<^sub>2" "vs"]}}$\P\\\\[-4mm] + \end{tabular} + \end{center} + \caption{The inductive definition of POSIX values taken from our earlier paper \cite{AusafDyckhoffUrban2016}. The ternary relation, written $(s, r) \rightarrow v$, formalises the notion + of given a string $s$ and a regular + expression $r$ what is the unique value $v$ that satisfies the informal POSIX constraints for + regular expression matching.}\label{POSIXrules} + \end{figure} + + The clever idea by Sulzmann and Lu \cite{Sulzmann2014} in their first algorithm is to define + an injection function on values that mirrors (but inverts) the + construction of the derivative on regular expressions. Essentially it + injects back a character into a value. + For this they define two functions called @{text mkeps} and @{text inj}: + % + \begin{center} + \begin{tabular}{@ {}l@ {}} + \begin{tabular}{@ {}lcl@ {}} + @{thm (lhs) mkeps.simps(1)} & $\dn$ & @{thm (rhs) mkeps.simps(1)}\\ + @{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"]}\\ + @{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"]}\\ + @{thm (lhs) mkeps.simps(4)} & $\dn$ & @{thm (rhs) mkeps.simps(4)}\\ + \end{tabular}\smallskip\\ + + \begin{tabular}{@ {}cc@ {}} + \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}@ {}} + @{thm (lhs) injval.simps(1)} & $\dn$ & @{thm (rhs) injval.simps(1)}\\ + @{thm (lhs) injval.simps(2)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1"]} & $\dn$ & + @{thm (rhs) injval.simps(2)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1"]}\\ + @{thm (lhs) injval.simps(3)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]} & $\dn$ & + @{thm (rhs) injval.simps(3)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]}\\ + @{thm (lhs) injval.simps(7)[of "r" "c" "v" "vs"]} & $\dn$ + & @{thm (rhs) injval.simps(7)[of "r" "c" "v" "vs"]} + \end{tabular} + & + \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}@ {}} + @{thm (lhs) injval.simps(4)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]} & $\dn$ + & @{thm (rhs) injval.simps(4)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]}\\ + @{thm (lhs) injval.simps(5)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]} & $\dn$ + & @{thm (rhs) injval.simps(5)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>1" "v\<^sub>2"]}\\ + @{thm (lhs) injval.simps(6)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]} & $\dn$\\ + \multicolumn{3}{@ {}r@ {}}{@{thm (rhs) injval.simps(6)[of "r\<^sub>1" "r\<^sub>2" "c" "v\<^sub>2"]}} + \end{tabular} + \end{tabular} + \end{tabular}\smallskip + \end{center} + + \noindent + The function @{text mkeps} is run when the last derivative is nullable, that is + the string to be matched is in the language of the regular expression. It generates + a value for how the last derivative can match the empty string. The injection function + then calculates the corresponding value for each intermediate derivative until + a value for the original regular expression is generated. + Graphically the algorithm by + Sulzmann and Lu can be illustrated by the following picture %in Figure~\ref{Sulz} + where the path from the left to the right involving @{term derivatives}/@{const + nullable} is the first phase of the algorithm (calculating successive + \Brz's derivatives) and @{const mkeps}/@{text inj}, the path from right to + left, the second phase. +% +\begin{center} +\begin{tikzpicture}[scale=0.99,node distance=9mm, + every node/.style={minimum size=6mm}] +\node (r1) {@{term "r\<^sub>1"}}; +\node (r2) [right=of r1]{@{term "r\<^sub>2"}}; +\draw[->,line width=1mm](r1)--(r2) node[above,midway] {@{term "der a DUMMY"}}; +\node (r3) [right=of r2]{@{term "r\<^sub>3"}}; +\draw[->,line width=1mm](r2)--(r3) node[above,midway] {@{term "der b DUMMY"}}; +\node (r4) [right=of r3]{@{term "r\<^sub>4"}}; +\draw[->,line width=1mm](r3)--(r4) node[above,midway] {@{term "der c DUMMY"}}; +\draw (r4) node[anchor=west] {\;\raisebox{3mm}{@{term nullable}}}; +\node (v4) [below=of r4]{@{term "v\<^sub>4"}}; +\draw[->,line width=1mm](r4) -- (v4); +\node (v3) [left=of v4] {@{term "v\<^sub>3"}}; +\draw[->,line width=1mm](v4)--(v3) node[below,midway] {\inj r\<^sub>3 c\}; +\node (v2) [left=of v3]{@{term "v\<^sub>2"}}; +\draw[->,line width=1mm](v3)--(v2) node[below,midway] {\inj r\<^sub>2 b\}; +\node (v1) [left=of v2] {@{term "v\<^sub>1"}}; +\draw[->,line width=1mm](v2)--(v1) node[below,midway] {\inj r\<^sub>1 a\}; +\draw (r4) node[anchor=north west] {\;\raisebox{-8mm}{@{term "mkeps"}}}; +\end{tikzpicture} +\end{center} +% +\noindent + The picture shows the steps required when a + regular expression, say @{text "r\<^sub>1"}, matches the string @{term + "[a,b,c]"}. The first lexing algorithm by Sulzmann and Lu can be defined as:\\[-8mm] + +% \begin{figure}[t] +%\begin{center} +%\begin{tikzpicture}[scale=1,node distance=1cm, +% every node/.style={minimum size=6mm}] +%\node (r1) {@{term "r\<^sub>1"}}; +%\node (r2) [right=of r1]{@{term "r\<^sub>2"}}; +%\draw[->,line width=1mm](r1)--(r2) node[above,midway] {@{term "der a DUMMY"}}; +%\node (r3) [right=of r2]{@{term "r\<^sub>3"}}; +%\draw[->,line width=1mm](r2)--(r3) node[above,midway] {@{term "der b DUMMY"}}; +%\node (r4) [right=of r3]{@{term "r\<^sub>4"}}; +%\draw[->,line width=1mm](r3)--(r4) node[above,midway] {@{term "der c DUMMY"}}; +%\draw (r4) node[anchor=west] {\;\raisebox{3mm}{@{term nullable}}}; +%\node (v4) [below=of r4]{@{term "v\<^sub>4"}}; +%\draw[->,line width=1mm](r4) -- (v4); +%\node (v3) [left=of v4] {@{term "v\<^sub>3"}}; +%\draw[->,line width=1mm](v4)--(v3) node[below,midway] {\inj r\<^sub>3 c\}; +%\node (v2) [left=of v3]{@{term "v\<^sub>2"}}; +%\draw[->,line width=1mm](v3)--(v2) node[below,midway] {\inj r\<^sub>2 b\}; +%\node (v1) [left=of v2] {@{term "v\<^sub>1"}}; +%\draw[->,line width=1mm](v2)--(v1) node[below,midway] {\inj r\<^sub>1 a\}; +%\draw (r4) node[anchor=north west] {\;\raisebox{-8mm}{@{term "mkeps"}}}; +%\end{tikzpicture} +%\end{center} +%\mbox{}\\[-13mm] +% +%\caption{The two phases of the first algorithm by Sulzmann \& Lu \cite{Sulzmann2014}, +%matching the string @{term "[a,b,c]"}. The first phase (the arrows from +%left to right) is \Brz's matcher building successive derivatives. If the +%last regular expression is @{term nullable}, then the functions of the +%second phase are called (the top-down and right-to-left arrows): first +%@{term mkeps} calculates a value @{term "v\<^sub>4"} witnessing +%how the empty string has been recognised by @{term "r\<^sub>4"}. After +%that the function @{term inj} ``injects back'' the characters of the string into +%the values. The value @{term "v\<^sub>1"} is the result of the algorithm representing +%the POSIX value for this string and +%regular expression. +%\label{Sulz}} +%\end{figure} + + + + \begin{center} + \begin{tabular}{lcl} + @{thm (lhs) lexer.simps(1)} & $\dn$ & @{thm (rhs) lexer.simps(1)}\\ + @{thm (lhs) lexer.simps(2)} & $\dn$ & @{text "case"} @{term "lexer (der c r) s"} @{text of} + @{term "None"} @{text "\"} @{term None}\\ + & & \hspace{24mm}$|\;$ @{term "Some v"} @{text "\"} @{term "Some (injval r c v)"} + \end{tabular} + \end{center} + + +We have shown in our earlier paper \cite{AusafDyckhoffUrban2016} that +this algorithm is correct, that is it generates POSIX values. The +central property we established relates the derivative operation to the +injection function. + + \begin{proposition}\label{Posix2} + \textit{If} $(s,\; r\backslash c) \rightarrow v$ \textit{then} $(c :: s,\; r) \rightarrow$ \textit{inj} $r\; c\; v$. +\end{proposition} + + \noindent + With this in place we were able to prove: + + + \begin{proposition}\mbox{}\smallskip\\\label{lexercorrect} + \begin{tabular}{ll} + (1) & @{thm (lhs) lexer_correct_None} if and only if @{thm (rhs) lexer_correct_None}\\ + (2) & @{thm (lhs) lexer_correct_Some} if and only if @{thm (rhs) lexer_correct_Some}\\ + \end{tabular} + \end{proposition} + + \noindent + In fact we have shown that, in the success case, the generated POSIX value $v$ is + unique and in the failure case that there is no POSIX value $v$ that satisfies + $(s, r) \rightarrow v$. While the algorithm is correct, it is excruciatingly + slow in cases where the derivatives grow arbitrarily (recall the example from the + Introduction). However it can be used as a convenient reference point for the correctness + proof of the second algorithm by Sulzmann and Lu, which we shall describe next. + +*} + +section {* Bitcoded Regular Expressions and Derivatives *} + +text {* + + In the second part of their paper \cite{Sulzmann2014}, + Sulzmann and Lu describe another algorithm that also generates POSIX + values but dispenses with the second phase where characters are + injected ``back'' into values. For this they annotate bitcodes to + regular expressions, which we define in Isabelle/HOL as the datatype + + \begin{center} + @{term breg} $\;::=\;$ @{term "AZERO"} + $\;\mid\;$ @{term "AONE bs"} + $\;\mid\;$ @{term "ACHAR bs c"} + $\;\mid\;$ @{term "AALTs bs rs"} + $\;\mid\;$ @{term "ASEQ bs r\<^sub>1 r\<^sub>2"} + $\;\mid\;$ @{term "ASTAR bs r"} + \end{center} + + \noindent where @{text bs} stands for bitsequences; @{text r}, + @{text "r\<^sub>1"} and @{text "r\<^sub>2"} for bitcoded regular + expressions; and @{text rs} for lists of bitcoded regular + expressions. The binary alternative @{text "ALT bs r\<^sub>1 r\<^sub>2"} + is just an abbreviation for \mbox{@{text "ALTs bs [r\<^sub>1, r\<^sub>2]"}}. + For bitsequences we use lists made up of the + constants @{text Z} and @{text S}. The idea with bitcoded regular + expressions is to incrementally generate the value information (for + example @{text Left} and @{text Right}) as bitsequences. For this + Sulzmann and Lu define a coding + function for how values can be coded into bitsequences. + + \begin{center} + \begin{tabular}{cc} + \begin{tabular}{lcl} + @{thm (lhs) code.simps(1)} & $\dn$ & @{thm (rhs) code.simps(1)}\\ + @{thm (lhs) code.simps(2)} & $\dn$ & @{thm (rhs) code.simps(2)}\\ + @{thm (lhs) code.simps(3)} & $\dn$ & @{thm (rhs) code.simps(3)}\\ + @{thm (lhs) code.simps(4)} & $\dn$ & @{thm (rhs) code.simps(4)} + \end{tabular} + & + \begin{tabular}{lcl} + @{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"]}\\ + @{thm (lhs) code.simps(6)} & $\dn$ & @{thm (rhs) code.simps(6)}\\ + @{thm (lhs) code.simps(7)} & $\dn$ & @{thm (rhs) code.simps(7)}\\ + \mbox{\phantom{XX}}\\ + \end{tabular} + \end{tabular} + \end{center} + + \noindent + As can be seen, this coding is ``lossy'' in the sense that we do not + record explicitly character values and also not sequence values (for + them we just append two bitsequences). However, the + different alternatives for @{text Left}, respectively @{text Right}, are recorded as @{text Z} and + @{text S} followed by some bitsequence. Similarly, we use @{text Z} to indicate + if there is still a value coming in the list of @{text Stars}, whereas @{text S} + indicates the end of the list. The lossiness makes the process of + decoding a bit more involved, but the point is that if we have a + regular expression \emph{and} a bitsequence of a corresponding value, + then we can always decode the value accurately (see Fig.~\ref{decode}). + The function \textit{decode} checks whether all of the bitsequence is + consumed and returns the corresponding value as @{term "Some v"}; otherwise + it fails with @{text "None"}. We can establish that for a value $v$ + inhabited by a regular expression $r$, the decoding of its + bitsequence never fails. + + %The decoding can be + %defined by using two functions called $\textit{decode}'$ and + %\textit{decode}: + + \begin{figure} + \begin{center} + \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}} + \multicolumn{3}{@ {}l}{$\textit{decode}'\,bs\,(\ONE)$ $\;\dn\;$ $(\Empty, bs)$ \quad\qquad + $\textit{decode}'\,bs\,(c)$ $\;\dn\;$ $(\Char\,c, bs)$}\\ + $\textit{decode}'\,(\Z\!::\!bs)\;(r_1 + r_2)$ & $\dn$ & + $\textit{let}\,(v, bs_1) = \textit{decode}'\,bs\,r_1\;\textit{in}\; + (\Left\,v, bs_1)$\\ + $\textit{decode}'\,(\S\!::\!bs)\;(r_1 + r_2)$ & $\dn$ & + $\textit{let}\,(v, bs_1) = \textit{decode}'\,bs\,r_2\;\textit{in}\; + (\Right\,v, bs_1)$\\ + $\textit{decode}'\,bs\;(r_1\cdot r_2)$ & $\dn$ & + $\textit{let}\,(v_1, bs_1) = \textit{decode}'\,bs\,r_1\;\textit{in}$\\ + & & $\textit{let}\,(v_2, bs_2) = \textit{decode}'\,bs_1\,r_2$ + \hspace{2mm}$\textit{in}\;(\Seq\,v_1\,v_2, bs_2)$\\ + $\textit{decode}'\,(\Z\!::\!bs)\,(r^*)$ & $\dn$ & $(\Stars\,[], bs)$\\ + $\textit{decode}'\,(\S\!::\!bs)\,(r^*)$ & $\dn$ & + $\textit{let}\,(v, bs_1) = \textit{decode}'\,bs\,r\;\textit{in}$\\ + & & $\textit{let}\,(\Stars\,vs, bs_2) = \textit{decode}'\,bs_1\,r^*$ + \hspace{2mm}$\textit{in}\;(\Stars\,v\!::\!vs, bs_2)$\medskip\\ + \multicolumn{3}{l}{$\textit{decode}\,bs\,r$ $\dn$ + $\textit{let}\,(v, bs') = \textit{decode}'\,bs\,r\;\textit{in}$ + $\;\,\textit{if}\;bs' = []\;\textit{then}\;\textit{Some}\,v\; + \textit{else}\;\textit{None}$}\\[-4mm] + \end{tabular} + \end{center} + \caption{Two functions, called $\textit{decode}'$ and \textit{decode}, for decoding a value from a bitsequence with the help of a regular expression.\\[-5mm]}\label{decode} + \end{figure} + + %\noindent + %The function \textit{decode} checks whether all of the bitsequence is + %consumed and returns the corresponding value as @{term "Some v"}; otherwise + %it fails with @{text "None"}. We can establish that for a value $v$ + %inhabited by a regular expression $r$, the decoding of its + %bitsequence never fails. + +\begin{lemma}\label{codedecode}\it + If $\;\vdash v : r$ then + $\;\textit{decode}\,(\textit{code}\, v)\,r = \textit{Some}\, v$. +\end{lemma} + + + + Sulzmann and Lu define the function \emph{internalise} + in order to transform (standard) regular expressions into annotated + regular expressions. We write this operation as $r^\uparrow$. + This internalisation uses the following + \emph{fuse} function. + % + \begin{center} + \begin{tabular}{@ {}cc@ {}} + \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l} + $\textit{fuse}\,bs\,(\textit{ZERO})$ & $\dn$ & $\textit{ZERO}$\\ + $\textit{fuse}\,bs\,(\textit{ONE}\,bs')$ & $\dn$ & + $\textit{ONE}\,(bs\,@\,bs')$\\ + $\textit{fuse}\,bs\,(\textit{CHAR}\,bs'\,c)$ & $\dn$ & + $\textit{CHAR}\,(bs\,@\,bs')\,c$ + \end{tabular} + & + \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l} + $\textit{fuse}\,bs\,(\textit{ALTs}\,bs'\,rs)$ & $\dn$ & + $\textit{ALTs}\,(bs\,@\,bs')\,rs$\\ + $\textit{fuse}\,bs\,(\textit{SEQ}\,bs'\,r_1\,r_2)$ & $\dn$ & + $\textit{SEQ}\,(bs\,@\,bs')\,r_1\,r_2$\\ + $\textit{fuse}\,bs\,(\textit{STAR}\,bs'\,r)$ & $\dn$ & + $\textit{STAR}\,(bs\,@\,bs')\,r$ + \end{tabular} + \end{tabular} + \end{center} + + \noindent + A regular expression can then be \emph{internalised} into a bitcoded + regular expression as follows: + % + \begin{center} + \begin{tabular}{@ {}ccc@ {}} + \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l} + $(\ZERO)^\uparrow$ & $\dn$ & $\textit{ZERO}$\\ + $(\ONE)^\uparrow$ & $\dn$ & $\textit{ONE}\,[]$ + \end{tabular} + & + \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l} + $(c)^\uparrow$ & $\dn$ & $\textit{CHAR}\,[]\,c$\\ + $(r^*)^\uparrow$ & $\dn$ & $\textit{STAR}\;[]\,r^\uparrow$ + \end{tabular} + & + \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l} + $(r_1 + r_2)^\uparrow$ & $\dn$ & + $\textit{ALT}\;[]\,(\textit{fuse}\,[\Z]\,r_1^\uparrow)\, + (\textit{fuse}\,[\S]\,r_2^\uparrow)$\\ + $(r_1\cdot r_2)^\uparrow$ & $\dn$ & + $\textit{SEQ}\;[]\,r_1^\uparrow\,r_2^\uparrow$ + \end{tabular} + \end{tabular} + \end{center} + + \noindent + There is also an \emph{erase}-function, written $r^\downarrow$, which + transforms a bitcoded regular expression into a (standard) regular + expression by just erasing the annotated bitsequences. We omit the + straightforward definition. For defining the algorithm, we also need + the functions \textit{bnullable} and \textit{bmkeps}(\textit{s}), which are the + ``lifted'' versions of \textit{nullable} and \textit{mkeps} acting on + bitcoded regular expressions. + % + \begin{center} + \begin{tabular}{@ {}c@ {\hspace{2mm}}c@ {}} + \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l} + $\textit{bnullable}\,(\textit{ZERO})$ & $\dn$ & $\textit{False}$\\ + $\textit{bnullable}\,(\textit{ONE}\,bs)$ & $\dn$ & $\textit{True}$\\ + $\textit{bnullable}\,(\textit{CHAR}\,bs\,c)$ & $\dn$ & $\textit{False}$\\ + $\textit{bnullable}\,(\textit{ALTs}\,bs\,\rs)$ & $\dn$ & + $\exists\, r \in \rs. \,\textit{bnullable}\,r$\\ + $\textit{bnullable}\,(\textit{SEQ}\,bs\,r_1\,r_2)$ & $\dn$ & + $\textit{bnullable}\,r_1\wedge \textit{bnullable}\,r_2$\\ + $\textit{bnullable}\,(\textit{STAR}\,bs\,r)$ & $\dn$ & + $\textit{True}$ + \end{tabular} + & + \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}} + $\textit{bmkeps}\,(\textit{ONE}\,bs)$ & $\dn$ & $bs$\\ + $\textit{bmkeps}\,(\textit{ALTs}\,bs\,\rs)$ & $\dn$ & + $bs\,@\,\textit{bmkepss}\,\rs$\\ + $\textit{bmkeps}\,(\textit{SEQ}\,bs\,r_1\,r_2)$ & $\dn$ &\\ + \multicolumn{3}{r}{$bs \,@\,\textit{bmkeps}\,r_1\,@\, \textit{bmkeps}\,r_2$}\\ + $\textit{bmkeps}\,(\textit{STAR}\,bs\,r)$ & $\dn$ & + $bs \,@\, [\S]$\\ + $\textit{bmkepss}\,(r\!::\!\rs)$ & $\dn$ & + $\textit{if}\;\textit{bnullable}\,r$\\ + & &$\textit{then}\;\textit{bmkeps}\,r$\\ + & &$\textit{else}\;\textit{bmkepss}\,\rs$ + \end{tabular} + \end{tabular} + \end{center} + + + \noindent + The key function in the bitcoded algorithm is the derivative of a + bitcoded regular expression. This derivative function calculates the + derivative but at the same time also the incremental part of the bitsequences + that contribute to constructing a POSIX value. + % + \begin{center} + \begin{tabular}{@ {}lcl@ {}} + \multicolumn{3}{@ {}l}{$(\textit{ZERO})\backslash c$ $\;\dn\;$ $\textit{ZERO}$ \quad\qquad + $(\textit{ONE}\;bs)\backslash c$ $\;\dn\;$ $\textit{ZERO}$}\\ + $(\textit{CHAR}\;bs\,d)\backslash c$ & $\dn$ & + $\textit{if}\;c=d\; \;\textit{then}\; + \textit{ONE}\;bs\;\textit{else}\;\textit{ZERO}$\\ + $(\textit{ALTs}\;bs\,\rs)\backslash c$ & $\dn$ & + $\textit{ALTs}\,bs\,(\mathit{map}\,(\_\backslash c)\,\rs)$\\ + $(\textit{SEQ}\;bs\,r_1\,r_2)\backslash c$ & $\dn$ & + $\textit{if}\;\textit{bnullable}\,r_1$\\ + & &$\textit{then}\;\textit{ALT}\,bs\;(\textit{SEQ}\,[]\,(r_1\backslash c)\,r_2)$ + $\;(\textit{fuse}\,(\textit{bmkeps}\,r_1)\,(r_2\backslash c))$\\ + & &$\textit{else}\;\textit{SEQ}\,bs\,(r_1\backslash c)\,r_2$\\ + $(\textit{STAR}\,bs\,r)\backslash c$ & $\dn$ & + $\textit{SEQ}\;bs\,(\textit{fuse}\, [\Z] (r\backslash c))\, + (\textit{STAR}\,[]\,r)$ + \end{tabular} + \end{center} + + + \noindent + This function can also be extended to strings, written $r\backslash s$, + just like the standard derivative. We omit the details. Finally we + can define Sulzmann and Lu's bitcoded lexer, which we call \textit{blexer}: + % + \begin{center} +\begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}} + $\textit{blexer}\;r\,s$ & $\dn$ & + $\textit{let}\;r_{der} = (r^\uparrow)\backslash s\;\textit{in}$ + $\;\;\textit{if}\; \textit{bnullable}(r_{der}) \;\;\textit{then}\;\textit{decode}\,(\textit{bmkeps}\,r_{der})\,r + \;\;\textit{else}\;\textit{None}$ +\end{tabular} +\end{center} + + \noindent +This bitcoded lexer first internalises the regular expression $r$ and then +builds the bitcoded derivative according to $s$. If the derivative is +(b)nullable the string is in the language of $r$ and it extracts the bitsequence using the +$\textit{bmkeps}$ function. Finally it decodes the bitsequence into a value. If +the derivative is \emph{not} nullable, then $\textit{None}$ is +returned. We can show that this way of calculating a value +generates the same result as \textit{lexer}. + +Before we can proceed we need to define a helper function, called +\textit{retrieve}, which Sulzmann and Lu introduced for the correctness proof. +% +\begin{center} + \begin{tabular}{lcl} + @{thm (lhs) retrieve.simps(1)} & $\dn$ & @{thm (rhs) retrieve.simps(1)}\\ + @{thm (lhs) retrieve.simps(2)} & $\dn$ & @{thm (rhs) retrieve.simps(2)}\\ + @{thm (lhs) retrieve.simps(3)} & $\dn$ & @{thm (rhs) retrieve.simps(3)}\\ + @{thm (lhs) better_retrieve(1)} & $\dn$ & @{thm (rhs) better_retrieve(1)}\\ + @{thm (lhs) better_retrieve(2)} & $\dn$ & @{thm (rhs) better_retrieve(2)}\\ + @{thm (lhs) retrieve.simps(6)[of _ "r\<^sub>1" "r\<^sub>2" "v\<^sub>1" "v\<^sub>2"]} + & $\dn$ & @{thm (rhs) retrieve.simps(6)[of _ "r\<^sub>1" "r\<^sub>2" "v\<^sub>1" "v\<^sub>2"]}\\ + @{thm (lhs) retrieve.simps(7)} & $\dn$ & @{thm (rhs) retrieve.simps(7)}\\ + @{thm (lhs) retrieve.simps(8)} & $\dn$ & @{thm (rhs) retrieve.simps(8)} + \end{tabular} + \end{center} + +\noindent +The idea behind this function is to retrieve a possibly partial +bitsequence from a bitcoded regular expression, where the retrieval is +guided by a value. For example if the value is $\Left$ then we +descend into the left-hand side of an alternative in order to +assemble the bitcode. Similarly for +$\Right$. The property we can show is that for a given $v$ and $r$ +with $\vdash v : r$, the retrieved bitsequence from the internalised +regular expression is equal to the bitcoded version of $v$. + +\begin{lemma}\label{retrievecode} +If $\vdash v : r$ then $\textit{code}\, v = \textit{retrieve}\,(r^\uparrow)\,v$. +\end{lemma} + +\noindent +We also need some auxiliary facts about how the bitcoded operations +relate to the ``standard'' operations on regular expressions. For +example if we build a bitcoded derivative and erase the result, this +is the same as if we first erase the bitcoded regular expression and +then perform the ``standard'' derivative operation. + +\begin{lemma}\label{bnullable}\mbox{}\smallskip\\ + \begin{tabular}{ll} +\textit{(1)} & $(r\backslash s)^\downarrow = (r^\downarrow)\backslash s$\\ +\textit{(2)} & $\textit{bnullable}(r)$ iff $\textit{nullable}(r^\downarrow)$\\ +\textit{(3)} & $\textit{bmkeps}(r) = \textit{retrieve}\,r\,(\textit{mkeps}\,(r^\downarrow))$ provided $\textit{nullable}(r^\downarrow)$. +\end{tabular} +\end{lemma} + +%\begin{proof} +% All properties are by induction on annotated regular expressions. +% %There are no interesting cases. +%\end{proof} + +\noindent +The only difficulty left for the correctness proof is that the bitcoded algorithm +has only a ``forward phase'' where POSIX values are generated incrementally. +We can achieve the same effect with @{text lexer} (which has two phases) by stacking up injection +functions during the forward phase. An auxiliary function, called $\textit{flex}$, +allows us to recast the rules of $\lexer$ in terms of a single +phase and stacked up injection functions. +% +\begin{center} +\begin{tabular}{@ {}l@ {}c@ {}l@ {\hspace{10mm}}l@ {}c@ {}l@ {}} + $\textit{flex}\;r\,f\,[]$ & $\dn$ & $f$ & + $\textit{flex}\;r\,f\,(c\!::\!s)$ & $\dn$ & + $\textit{flex}\,(r\backslash c)\,(\lambda v.\,f\,(\inj\,r\,c\,v))\,s$ +\end{tabular} +\end{center} + +\noindent +The point of this function is that when +reaching the end of the string, we just need to apply the stacked up +injection functions to the value generated by @{text mkeps}. +Using this function we can recast the success case in @{text lexer} +as follows: + +\begin{lemma}\label{flex} +If @{text "lexer r s = Some v"} \;then\; @{text "v = "}$\,\textit{flex}\,r\,id\,s\, + (\mkeps (r\backslash s))$. +\end{lemma} + +\noindent +Note we did not redefine \textit{lexer}, we just established that the +value generated by \textit{lexer} can also be obtained by a different +method. While this different method is not efficient (we essentially +need to traverse the string $s$ twice, once for building the +derivative $r\backslash s$ and another time for stacking up injection +functions), it helps us with proving +that incrementally building up values in @{text blexer} generates the same result. + +This brings us to our main lemma in this section: if we calculate a +derivative, say $r\backslash s$, and have a value, say $v$, inhabited +by this derivative, then we can produce the result @{text lexer} generates +by applying this value to the stacked-up injection functions +that $\textit{flex}$ assembles. The lemma establishes that this is the same +value as if we build the annotated derivative $r^\uparrow\backslash s$ +and then retrieve the corresponding bitcoded version, followed by a +decoding step. + +\begin{lemma}[Main Lemma]\label{mainlemma}\it +If $\vdash v : r\backslash s$ then +$\textit{Some}\,(\textit{flex}\,r\,\textit{id}\,s\,v) = + \textit{decode}(\textit{retrieve}\,(r^\uparrow \backslash s)\,v)\,r$ +\end{lemma} + + + +\noindent +%With this lemma in place, +We can then prove the correctness of \textit{blexer}---it indeed +produces the same result as \textit{lexer}. + + +\begin{theorem}\label{thmone} +$\textit{lexer}\,r\,s = \textit{blexer}\,r\,s$ +\end{theorem} + + + +\noindent This establishes that the bitcoded algorithm \emph{without} +simplification produces correct results. This was +only conjectured by Sulzmann and Lu in their paper +\cite{Sulzmann2014}. The next step is to add simplifications. + +*} + + +section {* Simplification *} + +text {* + + Derivatives as calculated by Brzozowski’s method are usually more + complex regular expressions than the initial one; the result is + that derivative-based matching and lexing algorithms are + often abysmally slow if the ``growth problem'' is not addressed. As Sulzmann and Lu wrote, various + optimisations are possible, such as the simplifications + $\ZERO{}\,r \Rightarrow \ZERO$, $\ONE{}\,r \Rightarrow r$, + $\ZERO{} + r \Rightarrow r$ and $r + r \Rightarrow r$. While these + simplifications can considerably speed up the two algorithms in many + cases, they do not solve fundamentally the growth problem with + derivatives. To see this let us return to the example from the + Introduction that shows the derivatives for \mbox{@{text "(a + aa)\<^sup>*"}}. + If we delete in the 3rd step all $\ZERO{}s$ and $\ONE$s according to + the simplification rules shown above we obtain + % + %% + % + \begin{equation}\def\xll{\xrightarrow{\_\backslash{} [a, a, a]}}\label{derivex} + (a + aa)^* \quad\xll\quad + \underbrace{\mbox{$(\ONE + a) \cdot (a + aa)^*$}}_{r} \;+\; + ((a + aa)^* + \underbrace{\mbox{$(\ONE + a) \cdot (a + aa)^*$}}_{r}) + \end{equation} + + \noindent This is a simpler derivative, but unfortunately we + cannot make any further simplifications. This is a problem because + the outermost alternatives contains two copies of the same + regular expression (underlined with $r$). These copies will + spawn new copies in later derivative steps and they in turn even more copies. This + destroys any hope of taming the size of the derivatives. But the + second copy of $r$ in \eqref{derivex} will never contribute to a + value, because POSIX lexing will always prefer matching a string + with the first copy. So it could be safely removed without affecting the correctness of the algorithm. + The dilemma with the simple-minded + simplification rules above is that the rule $r + r \Rightarrow r$ + will never be applicable because as can be seen in this example the + regular expressions are not next to each other but separated by another regular expression. + + But here is where Sulzmann and Lu's representation of generalised + alternatives in the bitcoded algorithm shines: in @{term + "ALTs bs rs"} we can define a more aggressive simplification by + recursively simplifying all regular expressions in @{text rs} and + then analyse the resulting list and remove any duplicates. + Another advantage with the bitsequences in bitcoded regular + expressions is that they can be easily modified such that simplification does not + interfere with the value constructions. For example we can ``flatten'', or + de-nest, or spill out, @{text ALTs} as follows + % + \[ + @{term "ALTs bs\<^sub>1 ((ALTs bs\<^sub>2 rs\<^sub>2) # rs\<^sub>1)"} + \quad\xrightarrow{bsimp}\quad + @{term "ALTs bs\<^sub>1 ((map (fuse bs\<^sub>2) rs\<^sub>2) # rs\<^sub>1)"} + \] + + \noindent + where we just need to fuse the bitsequence that has accumulated in @{text "bs\<^sub>2"} + to the alternatives in @{text "rs\<^sub>2"}. As we shall show below this will + ensure that the correct value corresponding to the original (unsimplified) + regular expression can still be extracted. %In this way the value construction + %is not affected by simplification. + + However there is one problem with the definition for the more + aggressive simplification rules described by Sulzmann and Lu. Recasting + their definition with our syntax they define the step of removing + duplicates as + % + \[ @{text "bsimp (ALTs bs rs)"} \dn @{text "ALTs + bs (nub (map bsimp rs))"} + \] + + \noindent where they first recursively simplify the regular + expressions in @{text rs} (using @{text map}) and then use + Haskell's @{text nub}-function to remove potential + duplicates. While this makes sense when considering the example + shown in \eqref{derivex}, @{text nub} is the inappropriate + function in the case of bitcoded regular expressions. The reason + is that in general the elements in @{text rs} will have a + different annotated bitsequence and in this way @{text nub} + will never find a duplicate to be removed. One correct way to + handle this situation is to first \emph{erase} the regular + expressions when comparing potential duplicates. This is inspired + by Scala's list functions of the form \mbox{@{text "distinctWith rs eq + acc"}} where @{text eq} is an user-defined equivalence relation that + compares two elements in @{text rs}. We define this function in + Isabelle/HOL as + + \begin{center} + \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {\hspace{1mm}}l@ {}} + @{thm (lhs) distinctWith.simps(1)} & $\dn$ & @{thm (rhs) distinctWith.simps(1)}\\ + @{thm (lhs) distinctWith.simps(2)} & $\dn$ & + @{text "if (\ y \ acc. eq x y)"} & @{text "then distinctWith xs eq acc"}\\ + & & & @{text "else x :: distinctWith xs eq ({x} \ acc)"} + \end{tabular} + \end{center} + + \noindent where we scan the list from left to right (because we + have to remove later copies). In @{text distinctWith}, @{text eq} is intended to be an + equivalence relation for annotated regular expressions and @{text acc} is an accumulator for annotated regular + expressions---essentially a set of regular expressions that we have already seen + while scanning the list. Therefore we delete an element, say @{text x}, + from the list provided a @{text "y"} with @{text "y"} being equivalent to @{text x} is already in the accumulator; + otherwise we keep @{text x} and scan the rest of the list but + add @{text "x"} as another ``seen'' element to @{text acc}. We will use + @{term distinctWith} where @{text eq} is an equivalence that deletes bitsequences from bitcoded regular expressions + before comparing the components. One way to define this in Isabelle/HOL is by the following recursive function from + annotated regular expressions to @{text bool}: + % + \begin{center} + \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {\hspace{1mm}}l@ {}} + @{thm (lhs) eq1.simps(1)} & $\dn$ & @{thm (rhs) eq1.simps(1)}\\ + @{thm (lhs) eq1.simps(2)[of DUMMY DUMMY]} & $\dn$ & @{thm (rhs) eq1.simps(2)[of DUMMY DUMMY]}\\ + @{thm (lhs) eq1.simps(3)[of DUMMY "c" DUMMY "d"]} & $\dn$ & @{thm (rhs) eq1.simps(3)[of DUMMY "c" DUMMY "d"]}\\ + @{thm (lhs) eq1.simps(4)[of DUMMY "r\<^sub>1\<^sub>1" "r\<^sub>1\<^sub>2" DUMMY "r\<^sub>2\<^sub>1" "r\<^sub>2\<^sub>2"]} & $\dn$ & + @{thm (rhs) eq1.simps(4)[of DUMMY "r\<^sub>1\<^sub>1" "r\<^sub>1\<^sub>2" DUMMY "r\<^sub>2\<^sub>1" "r\<^sub>2\<^sub>2"]}\\ + @{thm (lhs) eq1.simps(5)[of DUMMY DUMMY]} & $\dn$ & @{thm (rhs) eq1.simps(5)[of DUMMY DUMMY]}\\ + @{thm (lhs) eq1.simps(6)[of DUMMY "r\<^sub>1" "rs\<^sub>1" DUMMY "r\<^sub>2" "rs\<^sub>2"]} & $\dn$ & + @{thm (rhs) eq1.simps(6)[of DUMMY "r\<^sub>1" "rs\<^sub>1" DUMMY "r\<^sub>2" "rs\<^sub>2"]}\\ + @{thm (lhs) eq1.simps(7)[of DUMMY "r\<^sub>1" DUMMY "r\<^sub>2"]} & $\dn$ & @{thm (rhs) eq1.simps(7)[of DUMMY "r\<^sub>1" DUMMY "r\<^sub>2"]}\\ + \end{tabular} + \end{center} + + \noindent + where all other cases are set to @{text False}. + This equivalence is clearly a computationally more expensive operation than @{text nub}, + but is needed in order to make the removal of unnecessary copies + to work properly. + + Our simplification function depends on three more helper functions, one is called + @{text flts} and analyses lists of regular expressions coming from alternatives. + It is defined as follows: + + \begin{center} + \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l} + @{thm (lhs) flts.simps(1)} & $\dn$ & @{thm (rhs) flts.simps(1)}\\ + @{thm (lhs) flts.simps(2)} & $\dn$ & @{thm (rhs) flts.simps(2)}\\ + @{thm (lhs) flts.simps(3)[of "bs'" "rs'"]} & $\dn$ & @{thm (rhs) flts.simps(3)[of "bs'" "rs'"]}\\ + \end{tabular} + \end{center} + + \noindent + The second clause of @{text flts} removes all instances of @{text ZERO} in alternatives and + the third ``de-nests'' alternatives (but retaining the + bitsequence @{text "bs'"} accumulated in the inner alternative). There are + some corner cases to be considered when the resulting list inside an alternative is + empty or a singleton list. We take care of those cases in the + @{text "bsimpALTs"} function; similarly we define a helper function that simplifies + sequences according to the usual rules about @{text ZERO}s and @{text ONE}s: + + \begin{center} + \begin{tabular}{c@ {\hspace{5mm}}c} + \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l} + @{text "bsimpALTs bs []"} & $\dn$ & @{text "ZERO"}\\ + @{text "bsimpALTs bs [r]"} & $\dn$ & @{text "fuse bs r"}\\ + @{text "bsimpALTs bs rs"} & $\dn$ & @{text "ALTs bs rs"}\\ + \mbox{}\\ + \end{tabular} + & + \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l} + @{text "bsimpSEQ bs _ ZERO"} & $\dn$ & @{text "ZERO"}\\ + @{text "bsimpSEQ bs ZERO _"} & $\dn$ & @{text "ZERO"}\\ + @{text "bsimpSEQ bs\<^sub>1 (ONE bs\<^sub>2) r\<^sub>2"} + & $\dn$ & @{text "fuse (bs\<^sub>1 @ bs\<^sub>2) r\<^sub>2"}\\ + @{text "bsimpSEQ bs r\<^sub>1 r\<^sub>2"} & $\dn$ & @{text "SEQ bs r\<^sub>1 r\<^sub>2"} + \end{tabular} + \end{tabular} + \end{center} + + \noindent + With this in place we can define our simplification function as + + \begin{center} + \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l} + @{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"]}\\ + @{thm (lhs) bsimp.simps(2)[of "bs" _]} & $\dn$ & @{text "bsimpALTs bs (distinctWith (flts (map bsimp rs)) \ \)"}\\ + @{text "bsimp r"} & $\dn$ & @{text r} + \end{tabular} + \end{center} + + \noindent + We believe our recursive function @{term bsimp} simplifies regular + expressions as intended by Sulzmann and Lu. There is no point in applying the + @{text bsimp} function repeatedly (like the simplification in their paper which needs to be + applied until a fixpoint is reached) because we can show that @{term bsimp} is idempotent, + that is + + \begin{proposition} + @{term "bsimp (bsimp r) = bsimp r"} + \end{proposition} + + \noindent + This can be proved by induction on @{text r} but requires a detailed analysis + that the de-nesting of alternatives always results in a flat list of regular + expressions. We omit the details since it does not concern the correctness proof. + + Next we can include simplification after each derivative step leading to the + following notion of bitcoded derivatives: + + \begin{center} + \begin{tabular}{cc} + \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l} + @{thm (lhs) bders_simp.simps(1)} & $\dn$ & @{thm (rhs) bders_simp.simps(1)} + \end{tabular} + & + \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l} + @{thm (lhs) bders_simp.simps(2)} & $\dn$ & @{thm (rhs) bders_simp.simps(2)} + \end{tabular} + \end{tabular} + \end{center} + + \noindent + and use it in the improved lexing algorithm defined as + + \begin{center} +\begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}} + $\textit{blexer}^+\;r\,s$ & $\dn$ & + $\textit{let}\;r_{der} = (r^\uparrow)\backslash_{bsimp}\, s\;\textit{in}$ + $\;\;\textit{if}\; \textit{bnullable}(r_{der}) \;\;\textit{then}\;\textit{decode}\,(\textit{bmkeps}\,r_{der})\,r + \;\;\textit{else}\;\textit{None}$ +\end{tabular} +\end{center} + + \noindent The remaining task is to show that @{term blexer} and + @{term "blexer_simp"} generate the same answers. + + When we first + attempted this proof we encountered a problem with the idea + in Sulzmann and Lu's paper where the argument seems to be to appeal + again to the @{text retrieve}-function defined for the unsimplified version + of the algorithm. But + this does not work, because desirable properties such as + % + \[ + @{text "retrieve r v = retrieve (bsimp r) v"} + \] + + \noindent do not hold under simplification---this property + essentially purports that we can retrieve the same value from a + simplified version of the regular expression. To start with @{text retrieve} + depends on the fact that the value @{text v} corresponds to the + structure of the regular expression @{text r}---but the whole point of simplification + is to ``destroy'' this structure by making the regular expression simpler. + To see this consider the regular expression @{term "r = ALT r' ZERO"} and a corresponding + value @{text "v = Left v'"}. If we annotate bitcodes to @{text "r"}, then + we can use @{text retrieve} with @{text r} and @{text v} in order to extract a corresponding + bitsequence. The reason that this works is that @{text r} is an alternative + regular expression and @{text v} a corresponding @{text "Left"}-value. However, if we simplify + @{text r}, then @{text v} does not correspond to the shape of the regular + expression anymore. So unless one can somehow + synchronise the change in the simplified regular expressions with + the original POSIX value, there is no hope of appealing to @{text retrieve} in the + correctness argument for @{term blexer_simp}. + + We found it more helpful to introduce the rewriting systems shown in + Figure~\ref{SimpRewrites}. The idea is to generate + simplified regular expressions in small steps (unlike the @{text bsimp}-function which + does the same in a big step), and show that each of + the small steps preserves the bitcodes that lead to the final POSIX value. + The rewrite system is organised such that $\leadsto$ is for bitcoded regular + expressions and $\stackrel{s}{\leadsto}$ for lists of bitcoded regular + expressions. The former essentially implements the simplifications of + @{text "bsimpSEQ"} and @{text flts}; while the latter implements the + simplifications in @{text "bsimpALTs"}. We can show that any bitcoded + regular expression reduces in zero or more steps to the simplified + regular expression generated by @{text bsimp}: + + \begin{lemma}\label{lemone} + @{thm[mode=IfThen] rewrites_to_bsimp} + \end{lemma} + + + + \noindent + We can show that this rewrite system preserves @{term bnullable}, that + is simplification does not affect nullability: + + \begin{lemma}\label{lembnull} + @{thm[mode=IfThen] bnullable0(1)[of "r\<^sub>1" "r\<^sub>2"]} + \end{lemma} + + + + \noindent + From this, we can show that @{text bmkeps} will produce the same bitsequence + as long as one of the bitcoded regular expressions in $\leadsto$ is nullable (this lemma + establishes the missing fact we were not able to establish using @{text retrieve}, as suggested + in the paper by Sulzmannn and Lu). + + + \begin{lemma}\label{lemthree} + @{thm[mode=IfThen] rewrite_bmkeps_aux(1)[of "r\<^sub>1" "r\<^sub>2"]} + \end{lemma} + + + + \noindent + Crucial is also the fact that derivative steps and simplification steps can be interleaved, + which is shown by the fact that $\leadsto$ is preserved under derivatives. + + \begin{lemma}\label{bderlem} + @{thm[mode=IfThen] rewrite_preserves_bder(1)[of "r\<^sub>1" "r\<^sub>2"]} + \end{lemma} + + + + + \noindent + Using this fact together with Lemma~\ref{lemone} allows us to prove the central lemma + that the unsimplified + derivative (with a string @{term s}) reduces to the simplified derivative (with the same string). + + \begin{lemma}\label{lemtwo} + @{thm[mode=IfThen] central} + \end{lemma} + + %\begin{proof} + %By reverse induction on @{term s} generalising over @{text r}. + %\end{proof} + + \noindent + With these lemmas in place we can finally establish that @{term "blexer_simp"} and @{term "blexer"} + generate the same value, and using Theorem~\ref{thmone} from the previous section that this value + is indeed the POSIX value. + + \begin{theorem} + @{thm[mode=IfThen] main_blexer_simp} + \end{theorem} + + %\begin{proof} + %By unfolding the definitions and using Lemmas~\ref{lemtwo} and \ref{lemthree}. + %\end{proof} + + \noindent + This means that if the algorithm is called with a regular expression @{term r} and a string + @{term s} with $@{term s} \in L(@{term r})$, it will return @{term "Some v"} for the + @{term v} we defined by the POSIX relation $(@{term s}, @{term r}) \rightarrow @{term v}$; otherwise + the algorithm returns @{term "None"} when $s \not\in L(r)$ and no such @{text v} exists. + This completes the correctness proof for the second POSIX lexing algorithm by Sulzmann and Lu. + The interesting point of this algorithm is that the sizes of derivatives do not grow arbitrarily big but + can be finitely bounded, which + we shall show next. + + \begin{figure}[t] + \begin{center} + \begin{tabular}{@ {\hspace{-8mm}}c@ {}} + @{thm[mode=Axiom] bs1[of _ "r\<^sub>2"]}$S\ZERO{}_l$\quad + @{thm[mode=Axiom] bs2[of _ "r\<^sub>1"]}$S\ZERO{}_r$\quad + @{thm[mode=Axiom] bs3[of "bs\<^sub>1" "bs\<^sub>2"]}$S\ONE$\\ + @{thm[mode=Rule] bs4[of "r\<^sub>1" "r\<^sub>2" _ "r\<^sub>3"]}SL\qquad + @{thm[mode=Rule] bs5[of "r\<^sub>3" "r\<^sub>4" _ "r\<^sub>1"]}SR\\ + @{thm[mode=Axiom] bs6}$A0$\quad + @{thm[mode=Axiom] bs7}$A1$\quad + @{thm[mode=Rule] bs10[of "rs\<^sub>1" "rs\<^sub>2"]}$AL$\\ + @{thm[mode=Rule] ss2[of "rs\<^sub>1" "rs\<^sub>2"]}$LT$\quad + @{thm[mode=Rule] ss3[of "r\<^sub>1" "r\<^sub>2"]}$LH$\quad + @{thm[mode=Axiom] ss4}$L\ZERO$\quad + @{thm[mode=Axiom] ss5[of "bs" "rs\<^sub>1" "rs\<^sub>2"]}$LS$\medskip\\ + @{thm[mode=Rule] ss6[of "r\<^sub>2" "r\<^sub>1" "rs\<^sub>1" "rs\<^sub>2" "rs\<^sub>3"]}$LD$\\[-5mm] + \end{tabular} + \end{center} + \caption{The rewrite rules that generate simplified regular expressions + in small steps: @{term "rrewrite r\<^sub>1 r\<^sub>2"} is for bitcoded regular + expressions and @{term "srewrite rs\<^sub>1 rs\<^sub>2"} for \emph{lists} of bitcoded + expressions. Interesting is the $LD$ rule that allows copies of regular + expressions to be removed provided a regular expression earlier in the list can + match the same strings.}\label{SimpRewrites} + \end{figure} +*} + +section {* Finiteness of Derivatives *} + +text {* + +In this section let us sketch our argument for why the size of the simplified +derivatives with the aggressive simplification function can be finitely bounded. Suppose +we have a size function for bitcoded regular expressions, written +$\llbracket r\rrbracket$, which counts the number of nodes if we regard $r$ as a tree +(we omit the precise definition; ditto for lists $\llbracket r\!s\rrbracket$). For this we show that for every $r$ +there exists a bound $N$ +such that + +\begin{center} +$\forall s. \; \llbracket@{term "bders_simp r s"}\rrbracket \leq N$ +\end{center} + +\noindent +We prove this by induction on $r$. The base cases for @{term AZERO}, +@{term "AONE bs"} and @{term "ACHAR bs c"} are straightforward. The interesting case is +for sequences of the form @{term "ASEQ bs r\<^sub>1 r\<^sub>2"}. In this case our induction +hypotheses state $\exists N_1. \forall s. \; \llbracket{}@{term "bders_simp r\<^sub>1 s"}\rrbracket \leq N_1$ and +$\exists N_2. \forall s. \; \llbracket@{term "bders_simp r\<^sub>2 s"}\rrbracket \leq N_2$. We can reason as follows +% +\begin{center} +\begin{tabular}{lcll} +& & $ \llbracket@{term "bders_simp (ASEQ bs r\<^sub>1 r\<^sub>2) s"}\rrbracket$\\ +& $ = $ & $\llbracket bsimp\,(\textit{ALTs}\;bs\;((@{term "ASEQ [] (bders_simp r\<^sub>1 s) r\<^sub>2"}) :: + [@{term "bders_simp r\<^sub>2 s'"} \;|\; s' \in \textit{Suffix}(@{text r}_1, s)]))\rrbracket $ & (1) \\ +& $\leq$ & + $\llbracket\textit{distinctWith}\,((@{term "ASEQ [] (bders_simp r\<^sub>1 s) r\<^sub>2"}) :: + [@{term "bders_simp r\<^sub>2 s'"} \;|\; s' \in \textit{Suffix}(@{text r}_1, s)])\,\approx\;@{term "{}"}\rrbracket + 1 $ & (2) \\ +& $\leq$ & $\llbracket@{term "ASEQ [] (bders_simp r\<^sub>1 s) r\<^sub>2"}\rrbracket + + \llbracket\textit{distinctWith}\,[@{term "bders_simp r\<^sub>2 s'"} \;|\; s' \in \textit{Suffix}(@{text r}_1, s)]\,\approx\;@{term "{}"}\rrbracket + 1 $ & (3) \\ +& $\leq$ & $N_1 + \llbracket r_2\rrbracket + 2 + + \llbracket\textit{distinctWith}\,[@{term "bders_simp r\<^sub>2 s'"} \;|\; s' \in \textit{Suffix}(@{text r}_1, s)]\,\approx\;@{term "{}"}\rrbracket$ & (4)\\ +& $\leq$ & $N_1 + \llbracket r_2\rrbracket + 2 + l_{N_{2}} * N_{2}$ & (5) +\end{tabular} +\end{center} + +% tell Chengsong about Indian paper of closed forms of derivatives + +\noindent +where in (1) the $\textit{Suffix}(@{text "r"}_1, s)$ are the all the suffixes of $s$ where @{term "bders_simp r\<^sub>1 s'"} is nullable ($s'$ being a suffix of $s$). +In (3) we know that $\llbracket@{term "ASEQ [] (bders_simp r\<^sub>1 s) r\<^sub>2"}\rrbracket$ is +bounded by $N_1 + \llbracket{}r_2\rrbracket + 1$. In (5) we know the list comprehension contains only regular expressions of size smaller +than $N_2$. The list length after @{text distinctWith} is bounded by a number, which we call $l_{N_2}$. It stands +for the number of distinct regular expressions smaller than $N_2$ (there can only be finitely many of them). +We reason similarly for @{text STAR}.\medskip + +\noindent +Clearly we give in this finiteness argument (Step (5)) a very loose bound that is +far from the actual bound we can expect. We can do better than this, but this does not improve +the finiteness property we are proving. If we are interested in a polynomial bound, +one would hope to obtain a similar tight bound as for partial +derivatives introduced by Antimirov \cite{Antimirov95}. After all the idea with +@{text distinctWith} is to maintain a ``set'' of alternatives (like the sets in +partial derivatives). Unfortunately to obtain the exact same bound would mean +we need to introduce simplifications, such as + $(r_1 + r_2) \cdot r_3 \longrightarrow (r_1 \cdot r_3) + (r_2 \cdot r_3)$, +which exist for partial derivatives. However, if we introduce them in our +setting we would lose the POSIX property of our calculated values. We leave better +bounds for future work.\\[-6.5mm] + +*} + + +section {* Conclusion *} + +text {* + + We set out in this work to prove in Isabelle/HOL the correctness of + the second POSIX lexing algorithm by Sulzmann and Lu + \cite{Sulzmann2014}. This follows earlier work where we established + the correctness of the first algorithm + \cite{AusafDyckhoffUrban2016}. In the earlier work we needed to + introduce our own specification about what POSIX values are, + because the informal definition given by Sulzmann and Lu did not + stand up to a formal proof. Also for the second algorithm we needed + to introduce our own definitions and proof ideas in order to establish the + correctness. Our interest in the second algorithm + lies in the fact that by using bitcoded regular expressions and an aggressive + simplification method there is a chance that the derivatives + can be kept universally small (we established in this paper that + they can be kept finitely bounded for any string). This is important if one is after + an efficient POSIX lexing algorithm based on derivatives. + + Having proved the correctness of the POSIX lexing algorithm, which + lessons have we learned? Well, we feel this is a very good example + where formal proofs give further insight into the matter at + hand. For example it is very hard to see a problem with @{text nub} + vs @{text distinctWith} with only experimental data---one would still + see the correct result but find that simplification does not simplify in well-chosen, but not + obscure, examples. We found that from an implementation + point-of-view it is really important to have the formal proofs of + the corresponding properties at hand. + + We have also developed a + healthy suspicion when experimental data is used to back up + efficiency claims. For example Sulzmann and Lu write about their + equivalent of @{term blexer_simp} \textit{``...we can incrementally compute + bitcoded parse trees in linear time in the size of the input''} + \cite[Page 14]{Sulzmann2014}. + Given the growth of the + derivatives in some cases even after aggressive simplification, this + is a hard to believe claim. A similar claim about a theoretical runtime + of @{text "O(n\<^sup>2)"} is made for the Verbatim lexer, which calculates + tokens according to POSIX rules~\cite{verbatim}. For this Verbatim uses Brzozowski's + derivatives like in our work. + The authors write: \textit{``The results of our empirical tests [..] confirm that Verbatim has + @{text "O(n\<^sup>2)"} time complexity.''} \cite[Section~VII]{verbatim}. + While their correctness proof for Verbatim is formalised in Coq, the claim about + the runtime complexity is only supported by some emperical evidence obtained + by using the code extraction facilities of Coq. + Given our observation with the ``growth problem'' of derivatives, + we + tried out their extracted OCaml code with the example + \mbox{@{text "(a + aa)\<^sup>*"}} as a single lexing rule, and it took for us around 5 minutes to tokenise a + string of 40 $a$'s and that increased to approximately 19 minutes when the + string is 50 $a$'s long. Taking into account that derivatives are not simplified in the Verbatim + lexer, such numbers are not surprising. + Clearly our result of having finite + derivatives might sound rather weak in this context but we think such effeciency claims + really require further scrutiny. + + The contribution of this paper is to make sure + derivatives do not grow arbitrarily big (universially) In the example \mbox{@{text "(a + aa)\<^sup>*"}}, + \emph{all} derivatives have a size of 17 or less. The result is that + lexing a string of, say, 50\,000 a's with the regular expression \mbox{@{text "(a + aa)\<^sup>*"}} takes approximately + 10 seconds with our Scala implementation + of the presented algorithm. + \smallskip + + \noindent + Our Isabelle code including the results from Sec.~5 is available from \url{https://github.com/urbanchr/posix}.\\[-10mm] + + +%%\bibliographystyle{plain} +\bibliography{root} + +\appendix + +\section{Some Proofs} + +While we have formalised \emph{all} results in Isabelle/HOL, below we shall give some +rough details of our reasoning in ``English''. + +\begin{proof}[Proof of Lemma~\ref{codedecode}] + This follows from the property that + $\textit{decode}'\,((\textit{code}\,v) \,@\, bs)\,r = (v, bs)$ holds + for any bit-sequence $bs$ and $\vdash v : r$. This property can be + easily proved by induction on $\vdash v : r$. +\end{proof} + +\begin{proof}[Proof of Lemma~\ref{mainlemma}] + This can be proved by induction on $s$ and generalising over + $v$. The interesting point is that we need to prove this in the + reverse direction for $s$. This means instead of cases $[]$ and + $c\!::\!s$, we have cases $[]$ and $s\,@\,[c]$ where we unravel the + string from the back.\footnote{Isabelle/HOL provides an induction principle + for this way of performing the induction.} + + The case for $[]$ is routine using Lemmas~\ref{codedecode} + and~\ref{retrievecode}. In the case $s\,@\,[c]$, we can infer from + the assumption that $\vdash v : (r\backslash s)\backslash c$ + holds. Hence by Prop.~\ref{Posix2} we know that + (*) $\vdash \inj\,(r\backslash s)\,c\,v : r\backslash s$ holds too. + By definition of $\textit{flex}$ we can unfold the left-hand side + to be + \[ + \textit{Some}\,(\textit{flex}\;r\,\textit{id}\,(s\,@\,[c])\,v) = + \textit{Some}\,(\textit{flex}\;r\,\textit{id}\,s\,(\inj\,(r\backslash s)\,c\,v)) + \] + + \noindent + By IH and (*) we can rewrite the right-hand side to + $\textit{decode}\,(\textit{retrieve}\,(r^\uparrow\backslash s)\; + (\inj\,(r\backslash s)\,c\,\,v))\,r$ which is equal to + $\textit{decode}\,(\textit{retrieve}\, (r^\uparrow\backslash + (s\,@\,[c]))\,v)\,r$ as required. The last rewrite step is possible + because we generalised over $v$ in our induction. +\end{proof} + +\begin{proof}[Proof of Theorem~\ref{thmone}] + We can first expand both sides using Lem.~\ref{flex} and the + definition of \textit{blexer}. This gives us two + \textit{if}-statements, which we need to show to be equal. By + Lemma~\ref{bnullable}\textit{(2)} we know the \textit{if}-tests coincide: + $\textit{bnullable}(r^\uparrow\backslash s) \;\textit{iff}\; + \nullable(r\backslash s)$. + For the \textit{if}-branch suppose $r_d \dn r^\uparrow\backslash s$ and + $d \dn r\backslash s$. We have (*) @{text "nullable d"}. We can then show + by Lemma~\ref{bnullable}\textit{(3)} that + % + \[ + \textit{decode}(\textit{bmkeps}\:r_d)\,r = + \textit{decode}(\textit{retrieve}\,r_d\,(\textit{mkeps}\,d))\,r + \] + + \noindent + where the right-hand side is equal to + $\textit{Some}\,(\textit{flex}\,r\,\textit{id}\,s\,(\textit{mkeps}\, + d))$ by Lemma~\ref{mainlemma} (we know + $\vdash \textit{mkeps}\,d : d$ by (*)). This shows the + \textit{if}-branches return the same value. In the + \textit{else}-branches both \textit{lexer} and \textit{blexer} return + \textit{None}. Therefore we can conclude the proof. +\end{proof} + +\begin{proof}[Proof of Lemma~\ref{lemone}] + By induction on @{text r}. For this we can use the properties + @{thm fltsfrewrites} and @{text "rs"}$\;\stackrel{s}{\leadsto}^*$@{text "distinctWith rs \"} @{term "{}"}. The latter uses + repeated applications of the $LD$ rule which allows the removal + of duplicates that can recognise the same strings. + \end{proof} + + \begin{proof}[Proof of Lemma~\ref{lembnull}] + Straightforward mutual induction on the definition of $\leadsto$ and $\stackrel{s}{\leadsto}$. + The only interesting case is the rule $LD$ where the property holds since by the side-conditions of that rule the empty string will + be in both @{text "L (rs\<^sub>a @ [r\<^sub>1] @ rs\<^sub>b @ [r\<^sub>2] @ rs\<^sub>c)"} and + @{text "L (rs\<^sub>a @ [r\<^sub>1] @ rs\<^sub>b @ rs\<^sub>c)"}. + \end{proof} + + \begin{proof}[Proof of Lemma \ref{lemthree}] + By straightforward mutual induction on the definition of $\leadsto$ and $\stackrel{s}{\leadsto}$. + Again the only interesting case is the rule $LD$ where we need to ensure that + @{text "bmkeps (rs\<^sub>a @ [r\<^sub>1] @ rs\<^sub>b @ [r\<^sub>2] @ rs\<^sub>c) = + bmkeps (rs\<^sub>a @ [r\<^sub>1] @ rs\<^sub>b @ rs\<^sub>c)"} holds. + This is indeed the case because according to the POSIX rules the + generated bitsequence is determined by the first alternative that can match the + string (in this case being nullable). + \end{proof} + + \begin{proof}[Proof of Lemma~\ref{bderlem}] + By straightforward mutual induction on the definition of $\leadsto$ and $\stackrel{s}{\leadsto}$. + The case for $LD$ holds because @{term "L (erase (bder c r\<^sub>2)) \ L (erase (bder c r\<^sub>1))"} + if and only if @{term "L (erase (r\<^sub>2)) \ L (erase (r\<^sub>1))"}. + \end{proof} +*} + +(*<*) +end +(*>*) \ No newline at end of file diff -r f9cdc295ccf7 -r f493a20feeb3 thys3/README.md --- a/thys3/README.md Thu Apr 28 15:56:22 2022 +0100 +++ b/thys3/README.md Sat Apr 30 00:50:08 2022 +0100 @@ -4,6 +4,10 @@ ```isabelle build -c -v -d . Posix``` +Generate the paper with + +```isabelle build -c -v -d . Paper``` + Tested with Isabelle2021-1. diff -r f9cdc295ccf7 -r f493a20feeb3 thys3/ROOT --- a/thys3/ROOT Thu Apr 28 15:56:22 2022 +0100 +++ b/thys3/ROOT Sat Apr 30 00:50:08 2022 +0100 @@ -1,4 +1,4 @@ -session Posix = "HOL-Library" + +session Posix in src = "HOL-Library" + theories[document = false] "HOL-Library.Sublist" "RegLangs" @@ -13,4 +13,21 @@ "ClosedForms" "GeneralRegexBound" "ClosedFormsBounds" - "FBound" \ No newline at end of file + "FBound" + + +session Paper = Posix + + options [document_variants="paper", + document = pdf, + document_output = ".", + document_comment_latex, + document_build = "pdflatex", + document_heading_prefix = ""] +theories + Paper +document_files + "root.tex" + "lipics-v2021.cls" + "cc-by.pdf" + "lipics-logo-bw.pdf" + "root.bib" \ No newline at end of file diff -r f9cdc295ccf7 -r f493a20feeb3 thys3/document/cc-by.pdf Binary file thys3/document/cc-by.pdf has changed diff -r f9cdc295ccf7 -r f493a20feeb3 thys3/document/lipics-logo-bw.pdf Binary file thys3/document/lipics-logo-bw.pdf has changed diff -r f9cdc295ccf7 -r f493a20feeb3 thys3/document/lipics-v2021.cls --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys3/document/lipics-v2021.cls Sat Apr 30 00:50:08 2022 +0100 @@ -0,0 +1,1249 @@ +%% +%% This is file `lipics-v2021.cls'. +%% +%% ----------------------------------------------------------------- +%% Author: Dagstuhl Publishing & le-tex publishing services +%% +%% This file is part of the lipics package for preparing +%% LIPICS articles. +%% +%% Copyright (C) 2021 Schloss Dagstuhl +%% +%% This work may be distributed and/or modified under the +%% conditions of the LaTeX Project Public License, either version 1.3 +%% of this license or (at your option) any later version. +%% The latest version of this license is in +%% http://www.latex-project.org/lppl.txt +%% and version 1.3 or later is part of all distributions of LaTeX +%% version 2005/12/01 or later. +%% +%% This work has the LPPL maintenance status `maintained'. +%% +%% The Current Maintainer of this work is +%% Schloss Dagstuhl (publishing@dagstuhl.de). +%% ----------------------------------------------------------------- +%% +\ProvidesClass{lipics-v2021} + [2021/05/04 v3.1.2 LIPIcs articles] +\NeedsTeXFormat{LaTeX2e}[2015/01/01] +\emergencystretch1em +\advance\hoffset-1in +\advance\voffset-1in +\advance\hoffset2.95mm +\newif\if@nobotseplist \@nobotseplistfalse +\def\@endparenv{% + \addpenalty\@endparpenalty\if@nobotseplist\else\addvspace\@topsepadd\fi\@endpetrue} +\def\@doendpe{% + \@endpetrue + \def\par{\@restorepar + \everypar{}% + \par + \if@nobotseplist + \addvspace\topsep + \addvspace\partopsep + \global\@nobotseplistfalse + \fi + \@endpefalse}% + \everypar{{\setbox\z@\lastbox}% + \everypar{}% + \if@nobotseplist\global\@nobotseplistfalse\fi + \@endpefalse}} +\def\enumerate{% + \ifnum \@enumdepth >\thr@@\@toodeep\else + \advance\@enumdepth\@ne + \edef\@enumctr{enum\romannumeral\the\@enumdepth}% + \expandafter + \list + \csname label\@enumctr\endcsname + {\advance\partopsep\topsep + \topsep\z@\@plus\p@ + \ifnum\@listdepth=\@ne + \labelsep0.72em + \else + \ifnum\@listdepth=\tw@ + \labelsep0.3em + \else + \labelsep0.5em + \fi + \fi + \usecounter\@enumctr\def\makelabel##1{\hss\llap{##1}}}% + \fi} +\def\endenumerate{\ifnum\@listdepth=\@ne\global\@nobotseplisttrue\fi\endlist} +\def\itemize{% + \ifnum \@itemdepth >\thr@@\@toodeep\else + \advance\@itemdepth\@ne + \edef\@itemitem{labelitem\romannumeral\the\@itemdepth}% + \expandafter + \list + \csname\@itemitem\endcsname + {\advance\partopsep\topsep + \topsep\z@\@plus\p@ + \ifnum\@listdepth=\@ne + \labelsep0.83em + \else + \ifnum\@listdepth=\tw@ + \labelsep0.75em + \else + \labelsep0.5em + \fi + \fi + \def\makelabel##1{\hss\llap{##1}}}% + \fi} +\def\enditemize{\ifnum\@listdepth=\@ne\global\@nobotseplisttrue\fi\endlist} +\def\@title{\textcolor{red}{Author: Please provide a title}} +\let\@subtitle\@empty +\def\subtitle#1{\gdef\@subtitle{#1}} +\def\subtitleseperator{: } +\def\@sect#1#2#3#4#5#6[#7]#8{% + \ifnum #2>\c@secnumdepth + \let\@svsec\@empty + \else + \refstepcounter{#1}% + \protected@edef\@svsec{\@seccntformat{#1}\relax}% + \fi + \@tempskipa #5\relax + \ifdim \@tempskipa>\z@ + \begingroup + #6{% + \@hangfrom{\hskip #3\relax + \ifnum #2=1 + \colorbox{lipicsYellow}{\kern0.15em\@svsec\kern0.15em}\quad + \else + \@svsec\quad + \fi}% + \interlinepenalty \@M #8\@@par}% + \endgroup + \csname #1mark\endcsname{#7}% + \addcontentsline{toc}{#1}{% + \ifnum #2>\c@secnumdepth \else + \protect\numberline{\csname the#1\endcsname}% + \fi + #7}% + \else + \def\@svsechd{% + #6{\hskip #3\relax + \@svsec #8}% + \csname #1mark\endcsname{#7}% + \addcontentsline{toc}{#1}{% + \ifnum #2>\c@secnumdepth \else + \protect\numberline{\csname the#1\endcsname}% + \fi + #7}}% + \fi + \@xsect{#5}} +\def\@seccntformat#1{\csname the#1\endcsname} +\def\@biblabel#1{\textcolor{lipicsGray}{\sffamily\bfseries#1}} +\def\EventLogoHeight{25} +\def\copyrightline{% + \ifx\@hideLIPIcs\@undefined + \ifx\@EventLogo\@empty + \else + \setbox\@tempboxa\hbox{\includegraphics[height=\EventLogoHeight\p@]{\@EventLogo}}% + \rlap{\hspace\textwidth\hspace{-\wd\@tempboxa}\hspace{\z@}% + \vtop to\z@{\vskip-0mm\unhbox\@tempboxa\vss}}% + \fi + \scriptsize + \vtop{\hsize\textwidth + \nobreakspace\par + \@Copyright + \ifx\@EventLongTitle\@empty\else\@EventLongTitle.\\\fi + \ifx\@EventEditors\@empty\else + \@Eds: \@EventEditors + ; Article~No.\,\@ArticleNo; pp.\,\@ArticleNo:\thepage--\@ArticleNo:\number\numexpr\getpagerefnumber{TotPages}% + \\ + \fi + \setbox\@tempboxa\hbox{\IfFileExists{lipics-logo-bw.pdf}{\includegraphics[height=14\p@,trim=0 15 0 0]{lipics-logo-bw}}{\includegraphics[height=14\p@, width=62pt]{example-image-plain}}}% + \hspace*{\wd\@tempboxa}\enskip + \href{https://www.dagstuhl.de/lipics/}% + {Leibniz International Proceedings in Informatics}\\ + \smash{\unhbox\@tempboxa}\enskip + \href{https://www.dagstuhl.de}% + {Schloss Dagstuhl -- Leibniz-Zentrum f{\"u}r Informatik, Dagstuhl Publishing, Germany}}% + \fi} +\def\ps@plain{\let\@mkboth\@gobbletwo + \let\@oddhead\@empty + \let\@evenhead\@empty + \let\@evenfoot\copyrightline + \let\@oddfoot\copyrightline} +\def\lipics@opterrshort{Option "\CurrentOption" not supported} +\def\lipics@opterrlong{The option "\CurrentOption" from article.cls is not supported by lipics.cls.} +\DeclareOption{a5paper}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} +\DeclareOption{b5paper}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} +\DeclareOption{legalpaper}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} +\DeclareOption{executivepaper}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} +\DeclareOption{landscape}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} +\DeclareOption{10pt}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} +\DeclareOption{11pt}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} +\DeclareOption{12pt}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} +\DeclareOption{oneside}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} +\DeclareOption{twoside}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} +\DeclareOption{titlepage}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} +\DeclareOption{notitlepage}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} +\DeclareOption{onecolumn}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} +\DeclareOption{twocolumn}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} +\DeclareOption{fleqn}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} +\DeclareOption{openbib}{\@latexerr{\lipics@opterrshort}{\lipics@opterrlong}} +\DeclareOption{a4paper}{\PassOptionsToClass{\CurrentOption}{article} + \advance\hoffset-2.95mm + \advance\voffset8.8mm} +\DeclareOption{numberwithinsect}{\let\numberwithinsect\relax} +\DeclareOption{cleveref}{\let\usecleveref\relax} +\DeclareOption{autoref}{\let\useautoref\relax} +\DeclareOption{anonymous}{\let\authoranonymous\relax} +\DeclareOption{thm-restate}{\let\usethmrestate\relax} +\DeclareOption{authorcolumns}{\let\authorcolumns\relax} +\let\compactauthor\relax +\DeclareOption{oldauthorstyle}{\let\compactauthor\@empty} +\DeclareOption{compactauthor}{\let\compactauthor\relax} +\DeclareOption{pdfa}{\let\pdfa\relax} +\DeclareOption*{\PassOptionsToClass{\CurrentOption}{article}} +\ProcessOptions +\LoadClass[twoside,notitlepage,fleqn]{article} +\renewcommand\normalsize{% + \@setfontsize\normalsize\@xpt{13}% + \abovedisplayskip 10\p@ \@plus2\p@ \@minus5\p@ + \abovedisplayshortskip \z@ \@plus3\p@ + \belowdisplayshortskip 6\p@ \@plus3\p@ \@minus3\p@ + \belowdisplayskip \abovedisplayskip + \let\@listi\@listI} +\normalsize +\renewcommand\small{% + \@setfontsize\small\@ixpt{11.5}% + \abovedisplayskip 8.5\p@ \@plus3\p@ \@minus4\p@ + \abovedisplayshortskip \z@ \@plus2\p@ + \belowdisplayshortskip 4\p@ \@plus2\p@ \@minus2\p@ + \def\@listi{\leftmargin\leftmargini + \topsep 4\p@ \@plus2\p@ \@minus2\p@ + \parsep 2\p@ \@plus\p@ \@minus\p@ + \itemsep \parsep}% + \belowdisplayskip \abovedisplayskip +} +\renewcommand\footnotesize{% + \@setfontsize\footnotesize{8.5}{9.5}% + \abovedisplayskip 6\p@ \@plus2\p@ \@minus4\p@ + \abovedisplayshortskip \z@ \@plus\p@ + \belowdisplayshortskip 3\p@ \@plus\p@ \@minus2\p@ + \def\@listi{\leftmargin\leftmargini + \topsep 3\p@ \@plus\p@ \@minus\p@ + \parsep 2\p@ \@plus\p@ \@minus\p@ + \itemsep \parsep}% + \belowdisplayskip \abovedisplayskip +} +\renewcommand\large{\@setfontsize\large{10.5}{13}} +\renewcommand\Large{\@setfontsize\Large{12}{14}} +\setlength\parindent{1.5em} +\setlength\headheight{3mm} +\setlength\headsep {10mm} +\setlength\footskip{3mm} +\setlength\textwidth{140mm} +\setlength\textheight{222mm} +\setlength\oddsidemargin{32mm} +\setlength\evensidemargin{38mm} +\setlength\marginparwidth{25mm} +\setlength\topmargin{13mm} +\setlength{\skip\footins}{2\baselineskip \@plus 4\p@ \@minus 2\p@} +\def\@listi{\leftmargin\leftmargini + \parsep\z@ \@plus\p@ + \topsep 8\p@ \@plus2\p@ \@minus4\p@ + \itemsep \parsep} +\let\@listI\@listi +\@listi +\def\@listii {\leftmargin\leftmarginii + \labelwidth\leftmarginii + \advance\labelwidth-\labelsep + \topsep 4\p@ \@plus2\p@ \@minus\p@ + \parsep\z@ \@plus\p@ + \itemsep \parsep} +\def\@listiii{\leftmargin\leftmarginiii + \labelwidth\leftmarginiii + \advance\labelwidth-\labelsep + \topsep 2\p@ \@plus\p@\@minus\p@ + \parsep \z@ + \partopsep \p@ \@plus\z@ \@minus\p@ + \itemsep \z@ \@plus\p@} +\def\ps@headings{% + \def\@evenhead{\large\sffamily\bfseries + \llap{\hbox to0.5\oddsidemargin{ \ifx\@hideLIPIcs\@undefined\ifx\@ArticleNo\@empty\textcolor{red}{XX}\else\@ArticleNo\fi:\fi\thepage\hss}}\leftmark\hfil}% + \def\@oddhead{\large\sffamily\bfseries\rightmark\hfil + \rlap{\hbox to0.5\oddsidemargin{\hss \ifx\@hideLIPIcs\@undefined\ifx\@ArticleNo\@empty\textcolor{red}{XX}\else\@ArticleNo\fi:\fi\thepage}}}% + \def\@oddfoot{\hfil + \rlap{% + \vtop{% + \vskip10mm + \colorbox{lipicsYellow} + {\@tempdima\evensidemargin + \advance\@tempdima1in + \advance\@tempdima\hoffset + \hb@xt@\@tempdima{% + \ifx\@hideLIPIcs\@undefined + \textcolor{lipicsGray}{\normalsize\sffamily + \bfseries\quad + \expandafter\textsolittle + \expandafter{\@EventShortTitle}}% + \fi + \strut\hss}}}}} + \let\@evenfoot\@empty + \let\@mkboth\markboth + \let\sectionmark\@gobble + \let\subsectionmark\@gobble} +\pagestyle{headings} +\renewcommand\maketitle{\par + \begingroup + \thispagestyle{plain} + \renewcommand\thefootnote{\@fnsymbol\c@footnote}% + \if@twocolumn + \ifnum \col@number=\@ne + \@maketitle + \else + \twocolumn[\@maketitle]% + \fi + \else + \newpage + \global\@topnum\z@ % Prevents figures from going at top of page. + \@maketitle + \fi + \thispagestyle{plain}\@thanks + \endgroup + \global\let\thanks\relax + \global\let\maketitle\relax + \global\let\@maketitle\relax + \global\let\@thanks\@empty + \global\let\@author\@empty + \global\let\@date\@empty + \global\let\@title\@empty + \global\let\@subtitle\@empty + \global\let\title\relax + \global\let\author\relax + \global\let\date\relax + \global\let\and\relax +} +\newwrite\tocfile +\def\@maketitle{% + \newpage + \null\vskip-\baselineskip + \vskip-\headsep + \@titlerunning + \@authorrunning + %%\let \footnote \thanks + \parindent\z@ \raggedright + \if!\@title!\def\@title{\textcolor{red}{Author: Please fill in a title}}\fi + {\LARGE\sffamily\bfseries\mathversion{bold}\@title \if!\@subtitle!\else{\\\Large\sffamily\bfseries\mathversion{bold}\@subtitle}\fi \par}% + \vskip 1em + \ifx\@author\orig@author + \textcolor{red}{Author: Please provide author information}% + \else + {\def\thefootnote{\@arabic\c@footnote}% + \setcounter{footnote}{0}% + \fontsize{9.5}{12}\selectfont\@author}% + \fi + \bgroup + \immediate\openout\tocfile=\jobname.vtc + \protected@write\tocfile{ + \let\footnote\@gobble + \let\thanks\@gobble + \def\footnotemark{} + \def\and{and }% + \def\,{ } + \def\\{ } + }{% + \string\contitem + \string\title{\@title \if!\@subtitle!\else\subtitleseperator \@subtitle\fi}% + \string\author{\@authorsfortoc}% + \string\page{\@ArticleNo:\thepage--\@ArticleNo:\number\numexpr\getpagerefnumber{TotPages}}}% + \closeout\tocfile + \egroup + \par} +\renewcommand\tableofcontents{% + \section*{\contentsname}% + \@starttoc{toc}} +\setcounter{secnumdepth}{4} +\renewcommand\section{\@startsection {section}{1}{\z@}% + {-3.5ex \@plus -1ex \@minus -.2ex}% + {2.3ex \@plus.2ex}% + {\sffamily\Large\bfseries\raggedright}} +\renewcommand\subsection{\@startsection{subsection}{2}{\z@}% + {-3.25ex\@plus -1ex \@minus -.2ex}% + {1.5ex \@plus .2ex}% + {\sffamily\Large\bfseries\raggedright}} +\renewcommand\subsubsection{\@startsection{subsubsection}{3}{\z@}% + {-3.25ex\@plus -1ex \@minus -.2ex}% + {1.5ex \@plus .2ex}% + {\sffamily\Large\bfseries\raggedright}} +\renewcommand\paragraph{\@startsection{paragraph}{4}{\z@}% + {-3.25ex \@plus-1ex \@minus-.2ex}% + {1.5ex \@plus .2ex}% + {\sffamily\large\bfseries\raggedright}} +\renewcommand\subparagraph{\@startsection{subparagraph}{5}{\z@}% + {3.25ex \@plus1ex \@minus .2ex}% + {-1em}% + {\sffamily\normalsize\bfseries}} +\newcommand{\proofsubparagraph}{\@startsection{subparagraph}{5}{\z@}% + {3.25ex \@plus1ex \@minus .2ex}% + {-1em}% + {\color{lipicsGray}\sffamily\normalsize\bfseries}} +\setlength\leftmargini \parindent +\setlength\leftmarginii {1.2em} +\setlength\leftmarginiii{1.2em} +\setlength\leftmarginiv {1.2em} +\setlength\leftmarginv {1.2em} +\setlength\leftmarginvi {1.2em} +\renewcommand\labelenumi{% + \textcolor{lipicsGray}{\sffamily\bfseries\upshape\mathversion{bold}\theenumi.}} +\renewcommand\labelenumii{% + \textcolor{lipicsGray}{\sffamily\bfseries\upshape\mathversion{bold}\theenumii.}} +\renewcommand\labelenumiii{% + \textcolor{lipicsGray}{\sffamily\bfseries\upshape\mathversion{bold}\theenumiii.}} +\renewcommand\labelenumiv{% + \textcolor{lipicsGray}{\sffamily\bfseries\upshape\mathversion{bold}\theenumiv.}} +\renewcommand\labelitemi{% + \textcolor{lipicsBulletGray}{\ifnum\@listdepth=\@ne + \rule{0.67em}{0.33em}% + \else + \rule{0.45em}{0.225em}% + \fi}} +\renewcommand\labelitemii{% + \textcolor{lipicsBulletGray}{\rule{0.45em}{0.225em}}} +\renewcommand\labelitemiii{% + \textcolor{lipicsBulletGray}{\sffamily\bfseries\textasteriskcentered}} +\renewcommand\labelitemiv{% + \textcolor{lipicsBulletGray}{\sffamily\bfseries\textperiodcentered}} +\renewenvironment{description} + {\list{}{\advance\partopsep\topsep\topsep\z@\@plus\p@ + \labelwidth\z@ \itemindent-\leftmargin + \let\makelabel\descriptionlabel}} + {\ifnum\@listdepth=\@ne\global\@nobotseplisttrue\fi\endlist} +\renewcommand*\descriptionlabel[1]{% + \hspace\labelsep\textcolor{lipicsGray}{\sffamily\bfseries\mathversion{bold}#1}} +\def\topmattervskip{0.7} +\renewenvironment{abstract}{% + \vskip\topmattervskip\bigskipamount + \noindent + \rlap{\color{lipicsLineGray}\vrule\@width\textwidth\@height1\p@}% + \hspace*{7mm}\fboxsep1.5mm\colorbox[rgb]{1,1,1}{\raisebox{-0.4ex}{% + \large\selectfont\sffamily\bfseries\abstractname}}% + \vskip3\p@ + \fontsize{9}{12}\selectfont + \noindent\ignorespaces} + {\vskip\topmattervskip\baselineskip\noindent + \subjclassHeading + \ifx\@ccsdescString\@empty + \textcolor{red}{Author: Please fill in 1 or more \string\ccsdesc\space macro}% + \else + \@ccsdescString + \fi + \vskip\topmattervskip\baselineskip + \noindent\keywordsHeading + \ifx\@keywords\@empty + \textcolor{red}{Author: Please fill in \string\keywords\space macro}% + \else + \@keywords + \fi + \ifx\@hideLIPIcs\@undefined + \ifx\@DOIPrefix\@empty\else + \vskip\topmattervskip\baselineskip\noindent + \doiHeading\href{https://doi.org/\@lipicsdoi}{\@lipicsdoi}% + \fi + \fi + \ifx\@category\@empty\else + \vskip\topmattervskip\baselineskip\noindent + \categoryHeading\@category + \fi + \ifx\@relatedversion\@empty\else + \vskip\topmattervskip\baselineskip\noindent + \relatedversionHeading\ifx\authoranonymous\relax\textcolor{red}{Anonymous related version(s)}\else\@relatedversion\fi + \fi + \ifx\@supplement\@empty\else + \vskip\topmattervskip\baselineskip\noindent + \supplementHeading\ifx\authoranonymous\relax\textcolor{red}{Anonymous supplementary material}\else\@supplement\fi + \fi + \ifx\@funding\@empty\else + \vskip\topmattervskip\baselineskip\noindent + \fundingHeading\ifx\authoranonymous\relax\textcolor{red}{Anonymous funding}\else\@funding\fi + \fi + \ifx\@acknowledgements\@empty\else + \vskip\topmattervskip\baselineskip\noindent + \acknowledgementsHeading\ifx\authoranonymous\relax\textcolor{red}{Anonymous acknowledgements} \else\@acknowledgements\fi + \fi + \protected@write\@auxout{}{\string\gdef\string\@pageNumberEndAbstract{\thepage}}% + }% end abstract +\renewenvironment{thebibliography}[1] + {\if@noskipsec \leavevmode \fi + \par + \@tempskipa-3.5ex \@plus -1ex \@minus -.2ex\relax + \@afterindenttrue + \@tempskipa -\@tempskipa \@afterindentfalse + \if@nobreak + \everypar{}% + \else + \addpenalty\@secpenalty\addvspace\@tempskipa + \fi + \noindent + \rlap{\color{lipicsLineGray}\vrule\@width\textwidth\@height1\p@}% + \hspace*{7mm}\fboxsep1.5mm\colorbox[rgb]{1,1,1}{\raisebox{-0.4ex}{% + \normalsize\sffamily\bfseries\refname}}% + \@xsect{1ex \@plus.2ex}% + \list{\@biblabel{\@arabic\c@enumiv}}% + {\leftmargin8.5mm + \labelsep\leftmargin + \settowidth\labelwidth{\@biblabel{#1}}% + \advance\labelsep-\labelwidth + \usecounter{enumiv}% + \let\p@enumiv\@empty + \renewcommand\theenumiv{\@arabic\c@enumiv}}% + \fontsize{9}{12}\selectfont + \sloppy + \clubpenalty4000 + \@clubpenalty \clubpenalty + \widowpenalty4000% + \sfcode`\.\@m\protected@write\@auxout{}{\string\gdef\string\@pageNumberStartBibliography{\thepage}}} + {\def\@noitemerr + {\@latex@warning{Empty `thebibliography' environment}}% + \protected@write\@auxout{}{\string\gdef\string\@pageNumberEndBibliography{\thepage}}% + \endlist} +\g@addto@macro\appendix{\immediate\write\@auxout{\string\gdef\string\@pageNumberStartAppendix{\thepage}}}% +\renewcommand\footnoterule{% + \kern-8\p@ + {\color{lipicsBulletGray}\hrule\@width40mm\@height1\p@}% + \kern6.6\p@} +\renewcommand\@makefntext[1]{% + \parindent\z@\hangindent1em + \leavevmode + \hb@xt@1em{\@makefnmark\hss}#1} +\usepackage{microtype} +\usepackage[utf8]{inputenc} +\ifx\pdfa\relax% + \IfFileExists{glyphtounicode.tex}{ + \input glyphtounicode + \pdfgentounicode=1 + }{}% +\fi +\IfFileExists{lmodern.sty}{\RequirePackage{lmodern}}{} +\IfFileExists{fontawesome5.sty}{% +\RequirePackage{fontawesome5}% +\IfFileExists{orcid.pdf}{% +\def\orcidsymbol{\includegraphics[height=9\p@]{orcid}} +}{ +\def\orcidsymbol{\textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries \faOrcid}}% +} +\def\mailsymbol{\textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries \faIcon[regular]{envelope}}}% +\def\homesymbol{\textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries \faHome}}% +}{% +\ClassWarning{Package fontawesome5 not installed}{Please install package fontawesome5} +\def\orcidsymbol{\textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries ORCID}} +\def\mailsymbol{\textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries @}}% +\def\homesymbol{\textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries H}}% +}% +\RequirePackage[T1]{fontenc} +\RequirePackage{textcomp} +\RequirePackage[mathscr]{eucal} +\RequirePackage{amssymb} +\PassOptionsToPackage{retainmissing}{MnSymbol} +\AtBeginDocument{\@ifpackageloaded{MnSymbol}% + {\expandafter\let\csname ver@amssymb.sty\endcsname\relax + \let\complement\@undefined + \RequirePackage{amssymb}}{}} +\RequirePackage{soul} +\sodef\textsolittle{}{.12em}{.5em\@plus.08em\@minus.06em}% + {.4em\@plus.275em\@minus.183em} +\RequirePackage{color} %kept for backward compatibility +\AtBeginDocument{ + \@ifpackageloaded{xcolor}{ + }{ + \RequirePackage{xcolor} + } + \definecolor{darkgray}{rgb}{0.31,0.31,0.33} + \definecolor[named]{lipicsGray}{rgb}{0.31,0.31,0.33} + \definecolor[named]{lipicsBulletGray}{rgb}{0.60,0.60,0.61} + \definecolor[named]{lipicsLineGray}{rgb}{0.51,0.50,0.52} + \definecolor[named]{lipicsLightGray}{rgb}{0.85,0.85,0.86} + \definecolor[named]{lipicsYellow}{rgb}{0.99,0.78,0.07} +} +\RequirePackage{babel} +\RequirePackage[tbtags,fleqn]{amsmath} +\AtBeginDocument{ + \@ifpackageloaded{enumitem}{\ClassWarning{Package 'enumitem' incompatible}{Don't use package 'enumitem'; Package enumerate preloaded!}}{} + \@ifpackageloaded{paralist}{\ClassWarning{Package 'paralist' incompatible}{Don't use package 'paralist'; Package enumerate preloaded!}}{} +} +\RequirePackage{enumerate} +\def\@enum@{\list{\textcolor{lipicsGray}{\sffamily\bfseries\upshape\mathversion{bold}\csname label\@enumctr\endcsname}}% + {\advance\partopsep\topsep + \topsep\z@\@plus\p@ + \usecounter{\@enumctr}\def\makelabel##1{\hss\llap{##1}}}} +\def\romanenumerate{\enumerate[(i)]} +\let\endromanenumerate\endenumerate +\def\alphaenumerate{\enumerate[(a)]} +\let\endalphaenumerate\endenumerate +\def\bracketenumerate{\enumerate[(1)]} +\let\endbracketenumerate\endenumerate +\RequirePackage{graphicx} +\RequirePackage{array} +\let\@classzold\@classz +\def\@classz{% + \expandafter\ifx\d@llarbegin\begingroup + \toks \count@ = + \expandafter{\expandafter\small\the\toks\count@}% + \fi + \@classzold} +\RequirePackage{multirow} +\RequirePackage{tabularx} +\RequirePackage[online]{threeparttable} +\def\TPTtagStyle#1{#1)} +\def\tablenotes{\small\TPT@defaults + \@ifnextchar[\TPT@setuptnotes\TPTdoTablenotes} % ] +\RequirePackage{listings} +\lstset{basicstyle=\small\ttfamily,% + backgroundcolor=\color{lipicsLightGray},% + frame=single,framerule=0pt,xleftmargin=\fboxsep,xrightmargin=\fboxsep} +\RequirePackage[left,mathlines]{lineno} +\linenumbers +\renewcommand\linenumberfont{\normalfont\tiny\sffamily} +%%%% patch to cope with amsmath +%%%% http://phaseportrait.blogspot.de/2007/08/lineno-and-amsmath-compatibility.html +\newcommand*\patchAmsMathEnvironmentForLineno[1]{% + \expandafter\let\csname old#1\expandafter\endcsname\csname #1\endcsname + \expandafter\let\csname oldend#1\expandafter\endcsname\csname end#1\endcsname + \renewenvironment{#1}% + {\linenomath\csname old#1\endcsname}% + {\csname oldend#1\endcsname\endlinenomath}}% +\newcommand*\patchBothAmsMathEnvironmentsForLineno[1]{% + \patchAmsMathEnvironmentForLineno{#1}% + \patchAmsMathEnvironmentForLineno{#1*}}% +\AtBeginDocument{% + \patchBothAmsMathEnvironmentsForLineno{equation}% + \patchBothAmsMathEnvironmentsForLineno{align}% + \patchBothAmsMathEnvironmentsForLineno{flalign}% + \patchBothAmsMathEnvironmentsForLineno{alignat}% + \patchBothAmsMathEnvironmentsForLineno{gather}% + \patchBothAmsMathEnvironmentsForLineno{multline}} +\let\usehyperxmp\@empty% +\ifx\pdfa\relax% + \IfFileExists{hyperxmp.sty}{% + \RequirePackage{hyperxmp}% + \@ifpackagelater{hyperxmp}{2019/04/05}{% + \let\usehyperxmp\relax% + }{% + \ClassWarning{Package hyperxmp outdated}{You are using an outdated version of the package hyperxmp. Please update!}% + }}{}% +\fi% +\IfFileExists{totpages.sty}{ + \RequirePackage{totpages} +}{ + \ClassWarning{Package totpages not installed}{Please install package totpages} + \newcounter{TotPages} + \setcounter{TotPages}{99} +} +\ifx\usehyperxmp\relax% + \RequirePackage[pdfa,unicode]{hyperref}% +\else% + \RequirePackage[unicode]{hyperref}% +\fi% +\let\C\relax% +\let\G\relax% +\let\F\relax% +\let\U\relax% +\pdfstringdefDisableCommands{% + \let\thanks\@gobble% + \let\footnote\@gobble% + \def\footnotemark{}% + \def\cs#1{\textbackslash #1}% + \let\normalfont\@empty% + \let\scshape\@empty% + \def\and{and }% + \def\,{ }% + \def\textrightarrow{ -> }% + \let\mathsf\@empty% +}% +\hypersetup{ + breaklinks=true, + pdfencoding=unicode, + bookmarksnumbered, + pdfborder={0 0 0}, + pdfauthor={ } +}% +\AtBeginDocument{ +\ifx\usehyperxmp\relax +\hypersetup{ +pdftitle={\@title \if!\@subtitle!\else\subtitleseperator \@subtitle\fi}, +pdfauthor={\ifx\authoranonymous\relax Anonymous author(s) \else \@authorsforpdf \fi}, +pdfkeywords={\@keywords}, +pdfproducer={LaTeX with lipics-v2021.cls}, +pdfsubject={LIPIcs, Vol.\@SeriesVolume, \@EventShortTitle}, +pdfcopyright = { Copyright (C) \ifx\authoranonymous\relax Anonymous author(s) \else \@copyrightholder; \fi licensed under Creative Commons License CC-BY 4.0}, +pdflang={en}, +pdfmetalang={en}, +pdfpublisher={Schloss Dagstuhl -- Leibniz-Zentrum für Informatik, Dagstuhl Publishing, Germany}, +pdflicenseurl={https://creativecommons.org/licenses/by/4.0/}, +pdfpubtype={LIPIcs}, +pdfvolumenum={\@SeriesVolume}, +pdfpagerange={\@ArticleNo:\thepage-\@ArticleNo:\theTotPages}, +pdfdoi={}, +pdfapart=3, +pdfaconformance=B +} +\else% +\hypersetup{ +pdftitle={\@title \if!\@subtitle!\else\subtitleseperator \@subtitle\fi}, +pdfauthor={\ifx\authoranonymous\relax Anonymous author(s) \else \@authorsforpdf \fi}, +pdfkeywords={\@keywords}, +pdfcreator={LaTeX with lipics-v2021.cls}, +pdfsubject={LIPIcs, Vol.\@SeriesVolume, \@EventShortTitle; Copyright (C) \ifx\authoranonymous\relax Anonymous author(s) \else \@copyrightholder; \fi licensed under Creative Commons License CC-BY 4.0} +}% +\fi % +} +\ifx\usehyperxmp\relax +\pdfobjcompresslevel=0 +\pdfinclusioncopyfonts=1 +\IfFileExists{colorprofiles.tex}{ +\RequirePackage{colorprofiles}% +\IfFileExists{sRGB.icc}{ +\immediate\pdfobj stream attr{/N 3} file{sRGB.icc} +\pdfcatalog{% +/OutputIntents [ +<< +/Type /OutputIntent +/S /GTS_PDFA1 +/DestOutputProfile \the\pdflastobj\space 0 R +/OutputConditionIdentifier (sRGB) +/Info (sRGB) +>> +] +}}{} +}{\ClassWarning{Package colorprofiles not installed}{Please install package colorprofiles}} +\fi +\RequirePackage[labelsep=space,singlelinecheck=false,% + font={up,small},labelfont={sf,bf},% + listof=false]{caption}%"listof" instead of "list" for backward compatibility +\@ifpackagelater{hyperref}{2009/12/09} + {\captionsetup{compatibility=false}}%cf. http://groups.google.de/group/comp.text.tex/browse_thread/thread/db9310eb540fbbd8/42e30f3b7b3aa17a?lnk=raot + {} +\DeclareCaptionLabelFormat{boxed}{% + \kern0.05em{\color[rgb]{0.99,0.78,0.07}\rule{0.73em}{0.73em}}% + \hspace*{0.67em}\bothIfFirst{#1}{~}#2} +\captionsetup{labelformat=boxed} +\captionsetup[table]{position=top} +\RequirePackage[figuresright]{rotating} +\caption@AtBeginDocument{\@ifpackageloaded{subfig}{\ClassError{lipics}{% + Do not load the subfig package}{The more recent subcaption package is already loaded}}{}} +\RequirePackage{subcaption} +\def\titlerunning#1{\gdef\@titlerunning{{\let\footnote\@gobble\markboth{#1}{#1}}}} +\def\authorrunning#1{% + \gdef\@authorrunning{\markright{\ifx\authoranonymous\relax\textcolor{red}{Anonymous author(s)} \else\if!#1!\textcolor{red}{Author: Please fill in the \string\authorrunning\space macro}\else#1\fi\fi}}} +\titlerunning{\@title \if!\@subtitle!\else\subtitleseperator \@subtitle\fi} +\authorrunning{\textcolor{red}{Author: Please use the \string\authorrunning\space macro}} +\def\EventLongTitle#1{\gdef\@EventLongTitle{#1}} +\EventLongTitle{} +\def\EventShortTitle#1{\gdef\@EventShortTitle{#1}} +\EventShortTitle{} +\def\EventEditors#1{\gdef\@EventEditors{#1}} +\EventEditors{} +\def\EventNoEds#1{\gdef\@EventNoEds{#1}\xdef\@Eds{Editor\ifnum#1>1s\fi}} +\EventNoEds{1} +\def\EventLogo#1{\gdef\@EventLogo{#1}} +\EventLogo{} +\def\EventAcronym#1{\gdef\@EventAcronym{#1}} +\EventAcronym{} +\def\EventYear#1{\gdef\@EventYear{#1}} +\EventYear{} +\def\EventDate#1{\gdef\@EventDate{#1}} +\EventDate{} +\def\EventLocation#1{\gdef\@EventLocation{#1}} +\EventLocation{} +\def\SeriesVolume#1{\gdef\@SeriesVolume{#1}} +\SeriesVolume{} +\def\ArticleNo#1{\gdef\@ArticleNo{#1}} +\ArticleNo{} +\def\DOIPrefix#1{\gdef\@DOIPrefix{#1}} +\DOIPrefix{} +\def\@lipicsdoi{\@DOIPrefix.\@EventAcronym.\@EventYear.\@ArticleNo} +\def\and{\newline} +\let\orig@author\@author +\let\@authorsfortoc\@empty +\let\@authorsforpdf\@empty +\newcount\c@author +\newcounter{currentauthor} +\def\authorcolumnsMin{6} +\def\@authornum{0} +\def\author#1#2#3#4#5{% + \ifx\@author\orig@author\let\@author\@empty\fi + \g@addto@macro\@author{% + \noexpandarg\StrBehind{#2}{\and \url}[\homepageTemp]\IfSubStr{#2}{\and \url}{\StrBefore{#2}{\and \url}[\affiliation]}{\def\affiliation{#2}}% + \expandarg\exploregroups\StrRemoveBraces{\homepageTemp}[\homepage]% + \ifx\authorcolumns\relax + \ifnum\c@author>\authorcolumnsMin + \stepcounter{currentauthor} + \ifodd\value{currentauthor} + \begin{minipage}[t]{\textwidth} + \begin{minipage}[t]{0.49\textwidth} + \else + \hfill \begin{minipage}[t]{0.49\textwidth} + \fi + \else + \ClassWarning{Option 'authorcolumns' only applicable for > 6 authors}{Option 'authorcolumns' only applicable for >6 authors!} + \addvspace{0.5\baselineskip} + \fi + \else + \addvspace{0.5\baselineskip} + \fi + {\Large\bfseries + \if!#1! + \textcolor{red}{Author: Please enter author name}% + \else + \ifx\authoranonymous\relax + \textcolor{red}{Anonymous author} + \else + #1\,% + \ifx\compactauthor\relax\if!#3!\else{\,\href{mailto:#3}{\mailsymbol}}\fi% + \ifx\homepage\@empty\else{\,\href{\homepage}{\homesymbol}}\fi\fi% + \if!#4!\else{\,\href{#4}{\orcidsymbol}}\fi% + \if!#5!\else + \ifx\@funding\@empty + \expandafter\g@addto@macro\expandafter\@funding{\textit{\expandafter{\let\footnote\@gobble #1}}:\space{#5}} + \else + \expandafter\g@addto@macro\expandafter\@funding{\\\textit{\expandafter{\let\footnote\@gobble #1}}:\space{#5}} + \fi + \fi + \fi + \fi + } + {\small + \if!#2!\textcolor{red}{Author: Please enter affiliation as second parameter of the author macro}\else{\\* \ifx\authoranonymous\relax\textcolor{red}{Anonymous affiliation}\else\ifx\compactauthor\relax \affiliation \else#2\fi\fi}\fi + \ifx\compactauthor\relax\else\if!#3!\else{\ifx\authoranonymous\relax\else\\*\href{mailto:#3}{#3}\fi}\fi\fi + }\par + \ifx\authorcolumns\relax + \ifnum\c@author>\authorcolumnsMin + \end{minipage} + \ifnum\c@author=\value{currentauthor} + \end{minipage} + \else + \ifodd\value{currentauthor} + \else + \end{minipage}% + \medskip + \fi + \fi + \fi + \fi}% + \global\advance\c@author\@ne + \protected@write\@auxout{}{\string\gdef\string\@authornum{\the\c@author}} + \ifnum\c@author=\@ne + \gdef\@authorsfortoc{#1}% + \gdef\@authorsforpdf{#1} + \else + \expandafter\g@addto@macro\expandafter\@authorsforpdf\expandafter{, #1} + \expandafter\g@addto@macro\expandafter\@authorsfortoc\expandafter{\expandafter\csname\the\c@author authand\endcsname#1}% + \@namedef{\the\c@author authand}{,\space}% + \AtBeginDocument{% + \expandafter\ifnum\@authornum=2 + \@namedef{2authand}{\space and\space}% + \else + \@namedef{\@authornum authand}{,\space and\space}% + \fi} + \fi} +\newcommand*\affil[2][]{% + \ClassError{lipics} + {\string\affil\space deprecated: Please enter affiliation as second parameter of the author macro} + {Since 2017, \string\affil\space is obsolete in lipics.}} +\newcommand*\Copyright[1]{% + \def\@copyrightholder{#1} + \def\@Copyright{% + \setbox\@tempboxa\hbox{\IfFileExists{cc-by.pdf}{\includegraphics[height=14\p@,clip]{cc-by}}{\includegraphics[height=14\p@, width=40pt]{example-image-plain}}}% + \@rightskip\@flushglue \rightskip\@rightskip + \hangindent\dimexpr\wd\@tempboxa+0.5em\relax + \href{https://creativecommons.org/licenses/by/4.0/}% + {\smash{\lower\baselineskip\hbox{\unhcopy\@tempboxa}}}\enskip + \textcopyright\ % + \ifx!#1!\textcolor{red}{Author: Please fill in the \string\Copyright\space macro}\else\ifx\authoranonymous\relax\textcolor{red}{Anonymous author(s)}\else#1\fi\fi + ;\\% + licensed under Creative Commons License CC-BY 4.0\ifx!#1!\\\null\fi\par}} +\Copyright{\textcolor{red}{Author: Please provide a copyright holder}} +\let\@copyrightholder\@empty +\def\hideLIPIcs{\let\@hideLIPIcs\relax} +\usepackage{xstring} +\def\keywords#1{\def\@keywords{#1}} +\let\@keywords\@empty +\def\keywordsHeading{% + \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries + Keywords and phrases\enskip}} +\RequirePackage{comment} +\excludecomment{CCSXML} +% inspired by https://tex.stackexchange.com/questions/12810/how-do-i-split-a-string +\global\newcommand\ccsdesc[2][100]{\@ccsdesc#1~#2~~\relax} +\let\orig@ccsdesc\@ccsdesc +\let\@ccsdesc\@empty +\let\@ccsdescString\@empty +\gdef\@ccsdesc#1~#2~#3~{ + \ifx\@ccsdesc\orig@ccsdesc\let\@ccsdesc\@empty\fi + \ifx!#3! + \ifx\@ccsdescString\@empty + \g@addto@macro\@ccsdescString{{#2}} + \else + \g@addto@macro\@ccsdescString{; {#2}} + \fi + \else + \ifx\@ccsdescString\@empty + \g@addto@macro\@ccsdescString{{#2} $\rightarrow$ {#3}} + \else + \g@addto@macro\@ccsdescString{; {#2} $\rightarrow$ {#3}} + \fi + \fi +\ccsdescEnd +} +\def\ccsdescEnd#1\relax{} +\def\subjclass#1{ + \ClassError{lipics} + {\string\subjclass\space deprecated: Please enter subject classification in 1 or more ccsdesc macros} + {Since 2019, \string\subjclass\space is obsolete in lipics.}} +\let\@subjclass\@empty +\def\subjclassHeading{% + \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries + 2012 ACM Subject Classification\enskip}} +\def\doiHeading{% + \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries + Digital Object Identifier\enskip}} +\def\category#1{\def\@category{#1}} +\let\@category\@empty +\def\categoryHeading{% + \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries + Category\enskip}} +\def\relatedversion#1{\def\@relatedversion{#1}} +\let\@relatedversion\@empty +\define@key{relatedversiondetails}{linktext}{\def\relatedversiondetails@linktext{#1}} +\define@key{relatedversiondetails}{cite}{\def\relatedversiondetails@cite{#1}} +\newcommand*\addtorelatedversionmacro[2]{% + \ifx\@relatedversion\@empty% + \g@addto@macro\@relatedversion{#1}% + \else% + \g@addto@macro\@relatedversion{\\#1}% + \fi% +}% +\newcommand{\relatedversiondetails}[3][]{% + \begingroup% + \let\relatedversiondetails@linktext\@empty + \let\relatedversiondetails@cite\@empty + \setkeys{relatedversiondetails}{#1}% + \ifx\relatedversiondetails@linktext\@empty% + \protected@edef\tmp{\textit{#2}:\space{\url{#3}}}% + \else% + \protected@edef\tmp{\textit{#2}:\space{\href{#3}{\texttt{\relatedversiondetails@linktext}}}}% + \fi% + \ifx\relatedversiondetails@cite\@empty% + \else% + \protected@edef\tmp{\tmp\nobreakspace\cite{\relatedversiondetails@cite}}% + \fi% + \expandafter\addtorelatedversionmacro\expandafter{\tmp}{#1}% + \endgroup% +}% +\def\relatedversionHeading{% + \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries + Related Version\enskip}} +\def\supplement#1{\def\@supplement{#1}} +\let\@supplement\@empty +\define@key{supplementdetails}{linktext}{\def\supplementdetails@linktext{#1}} +\define@key{supplementdetails}{cite}{\def\supplementdetails@cite{#1}} +\define@key{supplementdetails}{subcategory}{\def\supplementdetails@subcategory{#1}} +\define@key{supplementdetails}{swhlinktext}{\def\supplementdetails@swhlinktext{#1}} +\let\supplementdetails@swhlinktext\@empty +\define@key{supplementdetails}{swhid}{ + \ifx\supplementdetails@swhlinktext\@empty% + \StrBefore{#1}{;}[\supplementdetails@swhlinktext]% + \fi% + \def\supplementdetails@swhid{#1}% +} + +\define@key{supplementdetails}{swhdelimiter}{\def\supplementdetails@swhdelimiter{#1}} +\def\supplementdetails@swhdelimiter{\\ \hspace*{1.2em}} +\newcommand*\addtosupplementmacro[2]{% + \ifx\@supplement\@empty% + \g@addto@macro\@supplement{#1}% + \else% + \g@addto@macro\@supplement{\\#1}% + \fi% +}% +\newcommand{\supplementdetails}[3][]{% + \begingroup% + \let\supplementdetails@linktext\@empty + \let\supplementdetails@cite\@empty + \let\supplementdetails@subcategory\@empty + \let\supplementdetails@swhid\@empty + \setkeys{supplementdetails}{#1}% + \ifx\supplementdetails@subcategory\@empty% + \protected@edef\tmp{\textit{#2}} + \else + \protected@edef\tmp{\textit{#2\,\,(\supplementdetails@subcategory)}}% + \fi + \ifx\supplementdetails@linktext\@empty% + \protected@edef\tmp{\tmp:\space{\url{#3}}}% + \else% + \protected@edef\tmp{\tmp:\space{\href{#3}{\texttt{\supplementdetails@linktext}}}}% + \fi% + \ifx\supplementdetails@cite\@empty% + \else% + \protected@edef\tmp{\tmp\nobreakspace\cite{\supplementdetails@cite}}% + \fi + \ifx\supplementdetails@swhid\@empty% + \else% + \ifx\supplementdetails@swhlinktext\@empty% + \protected@edef\tmp{\tmp \supplementdetails@swhdelimiter{} archived at % + \href{https://archive.softwareheritage.org/\supplementdetails@swhid}{\nolinkurl{\supplementdetails@swhid}}}% + \else% + \protected@edef\tmp{\tmp \supplementdetails@swhdelimiter{} archived at % + \href{https://archive.softwareheritage.org/\supplementdetails@swhid}{\nolinkurl{\supplementdetails@swhlinktext}}}% + \fi% + \fi% + \expandafter\addtosupplementmacro\expandafter{\tmp}{#1}% + \endgroup% +}% +\def\supplementHeading{% + \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries + Supplementary Material\enskip}} +\newcommand\flag[2][0.9cm]{% + \leavevmode\marginpar{% + \raisebox{\dimexpr-\totalheight+\ht\strutbox\relax}% + [\dimexpr\ht\strutbox+3mm][\dp\strutbox]{\expandafter\includegraphics[width=#1]{#2}}% +}} +\def\funding#1{\def\@funding{#1}} +\let\@funding\@empty +\def\fundingHeading{% + \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries + Funding\enskip}} +\def\acknowledgements#1{\def\@acknowledgements{#1}} +\let\@acknowledgements\@empty +\def\acknowledgementsHeading{% + \textcolor{lipicsGray}{\fontsize{9}{12}\sffamily\bfseries + Acknowledgements\enskip}} +\RequirePackage{amsthm} +\ifx\usethmrestate\relax + \RequirePackage{thm-restate} +\fi +\thm@headfont{% + \textcolor{lipicsGray}{$\blacktriangleright$}\nobreakspace\sffamily\bfseries} +\def\th@remark{% + \thm@headfont{% + \textcolor{lipicsGray}{$\blacktriangleright$}\nobreakspace\sffamily}% + \normalfont % body font + \thm@preskip\topsep \divide\thm@preskip\tw@ + \thm@postskip\thm@preskip +} +\def\@endtheorem{\endtrivlist}%\@endpefalse +\renewcommand\qedsymbol{\textcolor{lipicsGray}{\ensuremath{\blacktriangleleft}}} +\renewenvironment{proof}[1][\proofname]{\par + \pushQED{\qed}% + \normalfont \topsep6\p@\@plus6\p@\relax + \trivlist + \item[\hskip\labelsep + \color{lipicsGray}\sffamily\bfseries + #1\@addpunct{.}]\ignorespaces +}{% + \popQED\endtrivlist%\@endpefalse +} +\newcommand{\claimqedhere}{\renewcommand\qedsymbol{\textcolor{lipicsGray}{\ensuremath{\vartriangleleft}}}% +\qedhere% +\renewcommand\qedsymbol{\textcolor{lipicsGray}{\ensuremath{\blacktriangleleft}}}} +\newenvironment{claimproof}[1][\proofname]{ + \pushQED{\qed}% + \normalfont \topsep6\p@\@plus6\p@\relax + \trivlist + \item[\hskip\labelsep + \color{lipicsGray}\sffamily + #1\@addpunct{.}]\ignorespaces +}{% + \renewcommand\qedsymbol{\textcolor{lipicsGray}{\ensuremath{\vartriangleleft}}} + \popQED\endtrivlist%\@endpefalse + \renewcommand\qedsymbol{\textcolor{lipicsGray}{\ensuremath{\blacktriangleleft}}} +} +% inspired by qed of amsthm class +\DeclareRobustCommand{\lipicsEnd}{% + \leavevmode\unskip\penalty9999 \hbox{}\nobreak\hfill + \quad\hbox{$\lrcorner$}% +} +\AtBeginDocument{ + \@ifpackageloaded{algorithm2e}{ + \@ifpackagelater{algorithm2e}{2009/11/17}{ + \renewcommand{\algorithmcfname}{\sffamily\bfseries{}Algorithm}% + \renewcommand{\@algocf@procname}{\sffamily\bfseries{}Procedure}% + \SetAlgoCaptionSeparator{~} + \SetAlCapHSkip{0pt} + \renewcommand{\algocf@captiontext}[2]{% + \kern0.05em{\color{lipicsYellow}\rule{0.73em}{0.73em}}% + \hspace*{0.67em}\small #1\algocf@capseparator\nobreakspace#2} + \renewcommand{\algocf@makecaption}[2]{% + \parbox[t]{\textwidth}{\algocf@captiontext{#1}{#2}}% + }% + \renewcommand{\algocf@captionproctext}[2]{% + {% + \kern0.05em{\color{lipicsYellow}\rule{0.73em}{0.73em}}% + \hspace*{0.67em}\small% +\ProcSty{\ProcFnt\algocf@procname\ifthenelse{\boolean{algocf@procnumbered}}{\nobreakspace\thealgocf\algocf@typo\algocf@capseparator}{\relax}}% + \nobreakspace\ProcNameSty{\ProcNameFnt\algocf@captname #2@}% Name of the procedure in ProcName Style. + \ifthenelse{\equal{\algocf@captparam #2@}{\arg@e}}{}{% if no argument, write nothing + \ProcNameSty{\ProcNameFnt(}\ProcArgSty{\ProcArgFnt\algocf@captparam #2@}\ProcNameSty{\ProcNameFnt)}%else put arguments in ProcArgSty: + }% endif + \algocf@captother #2@% + }% +}% + \renewcommand{\@algocf@capt@boxed}{above} + \renewcommand{\@algocf@capt@ruled}{above} + \setlength\algotitleheightrule{0pt} + }{\ClassWarning{% + Package algorithm2e outdated}{You are using an outdated version of the package algorithm2e. Please update!}} + }{} + \@ifpackageloaded{algorithm}{ + \captionsetup[algorithm]{name=Algorithm, labelformat=boxed, position=top} + \newcommand\fs@ruled@notop{\def\@fs@cfont{\bfseries}\let\@fs@capt\floatc@ruled + \def\@fs@pre{}% + \def\@fs@post{\kern2pt\hrule\relax}% + \def\@fs@mid{\kern2pt\hrule\kern2pt}% + \let\@fs@iftopcapt\iftrue} + \@ifundefined{fst@algorithm}{}{ + \renewcommand\fst@algorithm{\fs@ruled@notop} + } + }{} + \ifx\usecleveref\relax\else + \@ifpackageloaded{cleveref}{\ClassWarning{Use document option 'cleveref' instead}{Use document option 'cleveref' instead directly loading package 'cleveref'}}{} + \fi + \ifx\usethmrestate\relax\else + \@ifpackageloaded{thm-restate}{\ClassWarning{Use document option 'thm-restate' instead}{Use document option 'thm-restate' instead directly loading package 'thm-restate'}}{} + \fi + \ifx\useautoref\relax + \@ifundefined{algorithmautorefname}{\newcommand{\algorithmautorefname}{Algorithm}}{\renewcommand{\algorithmautorefname}{Algorithm}}% + \fi +} + +\ifx\usecleveref\relax + \RequirePackage[capitalise, noabbrev]{cleveref} + \crefname{algocf}{Algorithm}{Algorithms} + \Crefname{algocf}{Algorithm}{Algorithms} + \newcommand{\crefrangeconjunction}{--} + \newcommand{\creflastconjunction}{, and\nobreakspace} +\fi +\ifx\useautoref\relax + \RequirePackage{aliascnt} +\fi +\newtheoremstyle{claimstyle}{\topsep}{\topsep}{}{0pt}{\sffamily}{. }{5pt plus 1pt minus 1pt}% + {$\vartriangleright$ \thmname{#1}\thmnumber{ #2}\thmnote{ (#3)}} +\theoremstyle{plain} +\newtheorem{theorem}{Theorem} +\ifx\numberwithinsect\relax + \@addtoreset{theorem}{section} + \expandafter\def\expandafter\thetheorem\expandafter{% + \expandafter\thesection\expandafter\@thmcountersep\thetheorem} +\fi + +\ifx\useautoref\relax + \addto\extrasenglish{% + \def\chapterautorefname{Chapter}% + \def\sectionautorefname{Section}% + \def\subsectionautorefname{Subsection}% + \def\subsubsectionautorefname{Subsubsection}% + \def\paragraphautorefname{Paragraph}% + \def\subparagraphautorefname{Subparagraph}% + } + \addto\extrasUKenglish{% + \def\chapterautorefname{Chapter}% + \def\sectionautorefname{Section}% + \def\subsectionautorefname{Subsection}% + \def\subsubsectionautorefname{Subsubsection}% + \def\paragraphautorefname{Paragraph}% + \def\subparagraphautorefname{Subparagraph}% + } + \addto\extrasUSenglish{% + \def\chapterautorefname{Chapter}% + \def\sectionautorefname{Section}% + \def\subsectionautorefname{Subsection}% + \def\subsubsectionautorefname{Subsubsection}% + \def\paragraphautorefname{Paragraph}% + \def\subparagraphautorefname{Subparagraph}% + } + \ifx\usethmrestate\relax + \newtheorem{lemma}[theorem]{Lemma} + \newtheorem{corollary}[theorem]{Corollary} + \newtheorem{proposition}[theorem]{Proposition} + \newtheorem{exercise}[theorem]{Exercise} + \newtheorem{definition}[theorem]{Definition} + \newtheorem{conjecture}[theorem]{Conjecture} + \newtheorem{observation}[theorem]{Observation} + \theoremstyle{definition} + \newtheorem{example}[theorem]{Example} + \theoremstyle{remark} + \newtheorem{note}[theorem]{Note} + \newtheorem*{note*}{Note} + \newtheorem{remark}[theorem]{Remark} + \newtheorem*{remark*}{Remark} + \theoremstyle{claimstyle} + \newtheorem{claim}[theorem]{Claim} + \newtheorem*{claim*}{Claim} + \else + \newaliascnt{lemma}{theorem} + \newtheorem{lemma}[lemma]{Lemma} + \aliascntresetthe{lemma} + \newcommand{\lemmaautorefname}{Lemma} + \newaliascnt{corollary}{theorem} + \newtheorem{corollary}[corollary]{Corollary} + \aliascntresetthe{corollary} + \newcommand{\corollaryautorefname}{Corollary} + \newaliascnt{proposition}{theorem} + \newtheorem{proposition}[proposition]{Proposition} + \aliascntresetthe{proposition} + \newcommand{\propositionautorefname}{Proposition} + \newaliascnt{exercise}{theorem} + \newtheorem{exercise}[exercise]{Exercise} + \aliascntresetthe{exercise} + \newcommand{\exerciseautorefname}{Exercise} + \newaliascnt{definition}{theorem} + \newtheorem{definition}[definition]{Definition} + \aliascntresetthe{definition} + \newcommand{\definitionautorefname}{Definition} + \newaliascnt{conjecture}{theorem} + \newtheorem{conjecture}[conjecture]{Conjecture} + \aliascntresetthe{conjecture} + \newcommand{\conjectureautorefname}{Conjecture} + \newaliascnt{observation}{theorem} + \newtheorem{observation}[observation]{Observation} + \aliascntresetthe{observation} + \newcommand{\observationautorefname}{Observation} + \theoremstyle{definition} + \newaliascnt{example}{theorem} + \newtheorem{example}[example]{Example} + \aliascntresetthe{example} + \newcommand{\exampleautorefname}{Example} + \theoremstyle{remark} + \newaliascnt{note}{theorem} + \newtheorem{note}[note]{Note} + \aliascntresetthe{note} + \newcommand{\noteautorefname}{Note} + \newtheorem*{note*}{Note} + \newaliascnt{remark}{theorem} + \newtheorem{remark}[remark]{Remark} + \aliascntresetthe{remark} + \newcommand{\remarkautorefname}{Remark} + \newtheorem*{remark*}{Remark} + \theoremstyle{claimstyle} + \newaliascnt{claim}{theorem} + \newtheorem{claim}[claim]{Claim} + \aliascntresetthe{claim} + \newcommand{\claimautorefname}{Claim} + \newtheorem*{claim*}{Claim} + \fi +\else + \newtheorem{lemma}[theorem]{Lemma} + \newtheorem{corollary}[theorem]{Corollary} + \newtheorem{proposition}[theorem]{Proposition} + \newtheorem{exercise}[theorem]{Exercise} + \newtheorem{definition}[theorem]{Definition} + \newtheorem{conjecture}[theorem]{Conjecture} + \newtheorem{observation}[theorem]{Observation} + \theoremstyle{definition} + \newtheorem{example}[theorem]{Example} + \theoremstyle{remark} + \newtheorem{note}[theorem]{Note} + \newtheorem*{note*}{Note} + \newtheorem{remark}[theorem]{Remark} + \newtheorem*{remark*}{Remark} + \theoremstyle{claimstyle} + \newtheorem{claim}[theorem]{Claim} + \newtheorem*{claim*}{Claim} +\fi +\theoremstyle{plain} +\endinput +%% +%% End of file `lipics-v2021.cls'. diff -r f9cdc295ccf7 -r f493a20feeb3 thys3/document/root.bib --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys3/document/root.bib Sat Apr 30 00:50:08 2022 +0100 @@ -0,0 +1,361 @@ +@article{HosoyaVouillonPierce2005, + author = {H.~Hosoya and J.~Vouillon and B.~C.~Pierce}, + title = {{R}egular {E}xpression {T}ypes for {XML}}, + journal = {ACM Transactions on Programming Languages and Systems (TOPLAS)}, + year = {2005}, + volume = 27, + number = 1, + pages = {46--90} +} + + +@Misc{POSIX, + title = {{T}he {O}pen {G}roup {B}ase {S}pecification {I}ssue 6 {IEEE} {S}td 1003.1 2004 {E}dition}, + year = {2004}, + note = {\url{http://pubs.opengroup.org/onlinepubs/009695399/basedefs/xbd_chap09.html}} +} + + + +@InProceedings{AusafDyckhoffUrban2016, + author = {F.~Ausaf and R.~Dyckhoff and C.~Urban}, + title = {{POSIX} {L}exing with {D}erivatives of {R}egular {E}xpressions ({P}roof {P}earl)}, + year = {2016}, + booktitle = {Proc.~of the 7th International Conference on Interactive Theorem Proving (ITP)}, + volume = {9807}, + series = {LNCS}, + pages = {69--86} +} + +@article{aduAFP16, + author = {F.~Ausaf and R.~Dyckhoff and C.~Urban}, + title = {{POSIX} {L}exing with {D}erivatives of {R}egular {E}xpressions}, + journal = {Archive of Formal Proofs}, + year = 2016, + note = {\url{http://www.isa-afp.org/entries/Posix-Lexing.shtml}, Formal proof development}, + ISSN = {2150-914x} +} + + +@TechReport{CrashCourse2014, + author = {N.~B.~B.~Grathwohl and F.~Henglein and U.~T.~Rasmussen}, + title = {{A} {C}rash-{C}ourse in {R}egular {E}xpression {P}arsing and {R}egular + {E}xpressions as {T}ypes}, + institution = {University of Copenhagen}, + year = {2014}, + annote = {draft report} +} + +@inproceedings{Sulzmann2014, + author = {M.~Sulzmann and K.~Lu}, + title = {{POSIX} {R}egular {E}xpression {P}arsing with {D}erivatives}, + booktitle = {Proc.~of the 12th International Conference on Functional and Logic Programming (FLOPS)}, + pages = {203--220}, + year = {2014}, + volume = {8475}, + series = {LNCS} +} + +@inproceedings{Sulzmann2014b, + author = {M.~Sulzmann and P.~van Steenhoven}, + title = {{A} {F}lexible and {E}fficient {ML} {L}exer {T}ool {B}ased on {E}xtended {R}egular + {E}xpression {S}ubmatching}, + booktitle = {Proc.~of the 23rd International Conference on Compiler Construction (CC)}, + pages = {174--191}, + year = {2014}, + volume = {8409}, + series = {LNCS} +} + +@book{Pierce2015, + author = {B.~C.~Pierce and C.~Casinghino and M.~Gaboardi and + M.~Greenberg and C.~Hri\c{t}cu and + V.~Sj\"{o}berg and B.~Yorgey}, + title = {{S}oftware {F}oundations}, + year = {2015}, + publisher = {Electronic textbook}, + note = {\url{http://www.cis.upenn.edu/~bcpierce/sf}} +} + +@Misc{Kuklewicz, + author = {C.~Kuklewicz}, + title = {{R}egex {P}osix}, + howpublished = "\url{https://wiki.haskell.org/Regex_Posix}" +} + +@article{Vansummeren2006, + author = {S.~Vansummeren}, + title = {{T}ype {I}nference for {U}nique {P}attern {M}atching}, + year = {2006}, + journal = {ACM Transactions on Programming Languages and Systems}, + volume = {28}, + number = {3}, + pages = {389--428} +} + +@InProceedings{Asperti12, + author = {A.~Asperti}, + title = {{A} {C}ompact {P}roof of {D}ecidability for {R}egular {E}xpression {E}quivalence}, + booktitle = {Proc.~of the 3rd International Conference on Interactive Theorem Proving (ITP)}, + pages = {283--298}, + year = {2012}, + volume = {7406}, + series = {LNCS} +} + +@inproceedings{Frisch2004, + author = {A.~Frisch and L.~Cardelli}, + title = {{G}reedy {R}egular {E}xpression {M}atching}, + booktitle = {Proc.~of the 31st International Conference on Automata, Languages and Programming (ICALP)}, + pages = {618--629}, + year = {2004}, + volume = {3142}, + series = {LNCS} +} + +@ARTICLE{Antimirov95, + author = {V.~Antimirov}, + title = {{P}artial {D}erivatives of {R}egular {E}xpressions and + {F}inite {A}utomata {C}onstructions}, + journal = {Theoretical Computer Science}, + year = {1995}, + volume = {155}, + pages = {291--319} +} + +@inproceedings{Nipkow98, + author={T.~Nipkow}, + title={{V}erified {L}exical {A}nalysis}, + booktitle={Proc.~of the 11th International Conference on Theorem Proving in Higher Order Logics (TPHOLs)}, + series={LNCS}, + volume=1479, + pages={1--15}, + year=1998 +} + +@article{Brzozowski1964, + author = {J.~A.~Brzozowski}, + title = {{D}erivatives of {R}egular {E}xpressions}, + journal = {Journal of the {ACM}}, + volume = {11}, + number = {4}, + pages = {481--494}, + year = {1964} +} + +@article{Leroy2009, + author = {X.~Leroy}, + title = {{F}ormal {V}erification of a {R}ealistic {C}ompiler}, + journal = {Communications of the ACM}, + year = 2009, + volume = 52, + number = 7, + pages = {107--115} +} + +@InProceedings{Paulson2015, + author = {L.~C.~Paulson}, + title = {{A} {F}ormalisation of {F}inite {A}utomata {U}sing {H}ereditarily {F}inite {S}ets}, + booktitle = {Proc.~of the 25th International Conference on Automated Deduction (CADE)}, + pages = {231--245}, + year = {2015}, + volume = {9195}, + series = {LNAI} +} + +@Article{Wu2014, + author = {C.~Wu and X.~Zhang and C.~Urban}, + title = {{A} {F}ormalisation of the {M}yhill-{N}erode {T}heorem based on {R}egular {E}xpressions}, + journal = {Journal of Automatic Reasoning}, + year = {2014}, + volume = {52}, + number = {4}, + pages = {451--480} +} + +@InProceedings{Regehr2011, + author = {X.~Yang and Y.~Chen and E.~Eide and J.~Regehr}, + title = {{F}inding and {U}nderstanding {B}ugs in {C} {C}ompilers}, + pages = {283--294}, + year = {2011}, + booktitle = {Proc.~of the 32nd ACM SIGPLAN Conference on Programming Language Design and + Implementation (PLDI)} +} + +@Article{Norrish2014, + author = {A.~Barthwal and M.~Norrish}, + title = {{A} {M}echanisation of {S}ome {C}ontext-{F}ree {L}anguage {T}heory in {HOL4}}, + journal = {Journal of Computer and System Sciences}, + year = {2014}, + volume = {80}, + number = {2}, + pages = {346--362} +} + +@Article{Thompson1968, + author = {K.~Thompson}, + title = {{P}rogramming {T}echniques: {R}egular {E}xpression {S}earch {A}lgorithm}, + journal = {Communications of the ACM}, + issue_date = {June 1968}, + volume = {11}, + number = {6}, + year = {1968}, + pages = {419--422} +} + +@article{Owens2009, + author = {S.~Owens and J.~H.~Reppy and A.~Turon}, + title = {{R}egular-{E}xpression {D}erivatives {R}e-{E}xamined}, + journal = {Journal of Functinal Programming}, + volume = {19}, + number = {2}, + pages = {173--190}, + year = {2009} +} + +@inproceedings{Sulzmann2015, + author = {M.~Sulzmann and P.~Thiemann}, + title = {{D}erivatives for {R}egular {S}huffle {E}xpressions}, + booktitle = {Proc.~of the 9th International Conference on Language and Automata Theory + and Applications (LATA)}, + pages = {275--286}, + year = {2015}, + volume = {8977}, + series = {LNCS} +} + +@inproceedings{Chen2012, + author = {H.~Chen and S.~Yu}, + title = {{D}erivatives of {R}egular {E}xpressions and an {A}pplication}, + booktitle = {Proc.~in the International Workshop on Theoretical + Computer Science (WTCS)}, + pages = {343--356}, + year = {2012}, + volume = {7160}, + series = {LNCS} +} + +@article{Krauss2011, + author={A.~Krauss and T.~Nipkow}, + title={{P}roof {P}earl: {R}egular {E}xpression {E}quivalence and {R}elation {A}lgebra}, + journal={Journal of Automated Reasoning}, + volume=49, + pages={95--106}, + year=2012 +} + +@InProceedings{Traytel2015, + author = {D.~Traytel}, + title = {{A} {C}oalgebraic {D}ecision {P}rocedure for {WS1S}}, + booktitle = {Proc.~of the 24th Annual Conference on Computer Science Logic (CSL)}, + pages = {487--503}, + series = {LIPIcs}, + year = {2015}, + volume = {41} +} + +@inproceedings{Traytel2013, + author={D.~Traytel and T.~Nipkow}, + title={{A} {V}erified {D}ecision {P}rocedure for {MSO} on + {W}ords {B}ased on {D}erivatives of {R}egular {E}xpressions}, + booktitle={Proc.~of the 18th ACM SIGPLAN International Conference on Functional Programming (ICFP)}, + pages={3-12}, + year=2013 +} + +@InProceedings{Coquand2012, + author = {T.~Coquand and V.~Siles}, + title = {{A} {D}ecision {P}rocedure for {R}egular {E}xpression {E}quivalence in {T}ype {T}heory}, + booktitle = {Proc.~of the 1st International Conference on Certified Programs and Proofs (CPP)}, + pages = {119--134}, + year = {2011}, + volume = {7086}, + series = {LNCS} +} + +@InProceedings{Almeidaetal10, + author = {J.~B.~Almeida and N.~Moriera and D.~Pereira and S.~M.~de Sousa}, + title = {{P}artial {D}erivative {A}utomata {F}ormalized in {C}oq}, + booktitle = {Proc.~of the 15th International Conference on Implementation + and Application of Automata (CIAA)}, + pages = {59-68}, + year = {2010}, + volume = {6482}, + series = {LNCS} +} + +@article{Owens2008, + author = {S.~Owens and K.~Slind}, + title = {{A}dapting {F}unctional {P}rograms to {H}igher {O}rder {L}ogic}, + journal = {Higher-Order and Symbolic Computation}, + volume = {21}, + number = {4}, + year = {2008}, + pages = {377--409} +} + + +@article{Owens2, + author = {S.~Owens and K.~Slind}, + title = {Adapting functional programs to higher order logic}, + journal = {Higher-Order and Symbolic Computation}, + volume = {21}, + number = {4}, + pages = {377--409}, + year = {2008}, + url = {http://dx.doi.org/10.1007/s10990-008-9038-0}, + doi = {10.1007/s10990-008-9038-0}, + timestamp = {Wed, 16 Dec 2009 13:51:02 +0100}, + biburl = {http://dblp.uni-trier.de/rec/bib/journals/lisp/OwensS08}, + bibsource = {dblp computer science bibliography, http://dblp.org} +} + +@misc{PCRE, + title = "{PCRE - Perl Compatible Regular Expressions}", + url = {http://www.pcre.org}, +} + + + +@InProceedings{OkuiSuzuki2010, + author = {S.~Okui and T.~Suzuki}, + title = {{D}isambiguation in {R}egular {E}xpression {M}atching via + {P}osition {A}utomata with {A}ugmented {T}ransitions}, + booktitle = {Proc.~of the 15th International Conference on Implementation + and Application of Automata (CIAA)}, + year = {2010}, + volume = {6482}, + series = {LNCS}, + pages = {231--240} +} + + + +@TechReport{OkuiSuzukiTech, + author = {S.~Okui and T.~Suzuki}, + title = {{D}isambiguation in {R}egular {E}xpression {M}atching via + {P}osition {A}utomata with {A}ugmented {T}ransitions}, + institution = {University of Aizu}, + year = {2013} +} + +@inproceedings{RibeiroAgda2017, +author = {R.~Ribeiro and A.~Du Bois}, +title = {{C}ertified {B}it-{C}oded {R}egular {E}xpression {P}arsing}, +year = {2017}, +publisher = {Association for Computing Machinery}, +address = {New York, NY, USA}, +booktitle = {Proc.~of the 21st Brazilian Symposium on Programming Languages}, +articleno = {4}, +numpages = {8} +} + + +@INPROCEEDINGS{verbatim, + author={D.~Egolf and S.~Lasser and K.~Fisher}, + booktitle={2021 IEEE Security and Privacy Workshops (SPW)}, + title={{V}erbatim: {A} {V}erified {L}exer {G}enerator}, + year={2021}, + volume={}, + number={}, + pages={92--100}} diff -r f9cdc295ccf7 -r f493a20feeb3 thys3/document/root.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys3/document/root.tex Sat Apr 30 00:50:08 2022 +0100 @@ -0,0 +1,140 @@ +\documentclass[runningheads]{lipics-v2021} +\usepackage{times} +\usepackage{isabelle} +\usepackage{isabellesym} +\usepackage{amsmath} +\usepackage{amssymb} +\usepackage{mathpartir} +\usepackage{tikz} +\usepackage{pgf} +\usetikzlibrary{positioning} +%\usepackage{pdfsetup} +\usepackage{stmaryrd} +%\usepackage{url} +%\usepackage{color} +%\usepackage[safe]{tipa} +%\usepackage[sc]{mathpazo} +%\usepackage{fontspec} +%\setmainfont[Ligatures=TeX]{Palatino Linotype} + + + + +\urlstyle{rm} +\isabellestyle{it} +\renewcommand{\isastyleminor}{\it}% +\renewcommand{\isastyle}{\normalsize\it}% + + +\def\dn{\,\stackrel{\mbox{\scriptsize def}}{=}\,} +\renewcommand{\isasymequiv}{$\dn$} +\renewcommand{\isasymemptyset}{$\varnothing$} +\renewcommand{\isacharunderscore}{\mbox{$\_\!\_$}} +\renewcommand{\isasymiota}{\makebox[0mm]{${}^{\prime}$}} +\renewcommand{\isasymin}{\ensuremath{\,\in\,}} + +\addtolength{\oddsidemargin}{-1.5mm} +\addtolength{\evensidemargin}{-1.5mm} +\addtolength{\textwidth}{4mm} +\addtolength{\textheight}{1.5mm} + +\def\lexer{\mathit{lexer}} +\def\mkeps{\mathit{mkeps}} +\def\inj{\mathit{inj}} +\def\Empty{\mathit{Empty}} +\def\Left{\mathit{Left}} +\def\Right{\mathit{Right}} +\def\Stars{\mathit{Stars}} +\def\Char{\mathit{Char}} +\def\Seq{\mathit{Seq}} +\def\Der{\mathit{Der}} +\def\nullable{\mathit{nullable}} +\def\Z{\mathit{Z}} +\def\S{\mathit{S}} +\newcommand{\ZERO}{\mbox{\bf 0}} +\newcommand{\ONE}{\mbox{\bf 1}} +\def\rs{\mathit{rs}} + +\def\Brz{Brzozowski} +\def\der{\backslash} +\newtheorem{falsehood}{Falsehood} +\newtheorem{conject}{Conjecture} + +\bibliographystyle{plainurl} + +\title{{POSIX} {L}exing with {B}itcoded {D}erivatives} +\titlerunning{POSIX Lexing with Bitcoded Derivatives} +\author{Chengsong Tan}{King's College London}{chengsong.tan@kcl.ac.uk}{}{} +\author{Christian Urban}{King's College London}{christian.urban@kcl.ac.uk}{}{} +\authorrunning{C.~Tan and C.~Urban} +\keywords{POSIX matching and lexing, derivatives of regular expressions, Isabelle/HOL} +\category{} +\ccsdesc[100]{Design and analysis of algorithms} +\ccsdesc[100]{Formal languages and automata theory} +\Copyright{\mbox{}} +\renewcommand{\DOIPrefix}{} +\nolinenumbers + + +\begin{document} +\maketitle + +\begin{abstract} + Sulzmann and Lu describe a lexing algorithm that calculates + Brzozowski derivatives using bitcodes annotated to regular + expressions. Their algorithm generates POSIX values which encode + the information of \emph{how} a regular expression matches a + string---that is, which part of the string is matched by which part + of the regular expression. This information is needed in the + context of lexing in order to extract and to classify tokens. + The purpose of the bitcodes is to generate POSIX values incrementally while + derivatives are calculated. They also help with designing + an ``aggressive'' simplification function that keeps the size of + derivatives finitely bounded. Without simplification the size of some derivatives can grow + arbitrarily big, resulting in an extremely slow lexing algorithm. In this + paper we describe a variant of Sulzmann and Lu's algorithm: Our + variant is a recursive functional program, whereas Sulzmann + and Lu's version involves a fixpoint construction. We \textit{(i)} + prove in Isabelle/HOL that our algorithm is correct and generates + unique POSIX values; we also \textit{(ii)} establish finite + bounds for the size of the derivatives. + + %The size can be seen as a + %proxy measure for the efficiency of the lexing algorithm: because of + %the polynomial bound our algorithm does not suffer from + %the exponential blowup in earlier works. + + % Brzozowski introduced the notion of derivatives for regular + % expressions. They can be used for a very simple regular expression + % matching algorithm. Sulzmann and Lu cleverly extended this + % algorithm in order to deal with POSIX matching, which is the + % underlying disambiguation strategy for regular expressions needed + % in lexers. Their algorithm generates POSIX values which encode + % the information of \emph{how} a regular expression matches a + % string---that is, which part of the string is matched by which + % part of the regular expression. In this paper we give our + % inductive definition of what a POSIX value is and show $(i)$ that + % such a value is unique (for given regular expression and string + % being matched) and $(ii)$ that Sulzmann and Lu's algorithm always + % generates such a value (provided that the regular expression + % matches the string). We show that $(iii)$ our inductive definition + % of a POSIX value is equivalent to an alternative definition by + % Okui and Suzuki which identifies POSIX values as least elements + % according to an ordering of values. We also prove the correctness + % of Sulzmann's bitcoded version of the POSIX matching algorithm and + % extend the results to additional constructors for regular + % expressions. \smallskip +\end{abstract} + + + +\input{session} + + + +\end{document} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff -r f9cdc295ccf7 -r f493a20feeb3 thys3/src/BasicIdentities.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys3/src/BasicIdentities.thy Sat Apr 30 00:50:08 2022 +0100 @@ -0,0 +1,1175 @@ +theory BasicIdentities + imports "Lexer" +begin + +datatype rrexp = + RZERO +| RONE +| RCHAR char +| RSEQ rrexp rrexp +| RALTS "rrexp list" +| RSTAR rrexp + +abbreviation + "RALT r1 r2 \ RALTS [r1, r2]" + + +fun + rnullable :: "rrexp \ bool" +where + "rnullable (RZERO) = False" +| "rnullable (RONE) = True" +| "rnullable (RCHAR c) = False" +| "rnullable (RALTS rs) = (\r \ set rs. rnullable r)" +| "rnullable (RSEQ r1 r2) = (rnullable r1 \ rnullable r2)" +| "rnullable (RSTAR r) = True" + + +fun + rder :: "char \ rrexp \ rrexp" +where + "rder c (RZERO) = RZERO" +| "rder c (RONE) = RZERO" +| "rder c (RCHAR d) = (if c = d then RONE else RZERO)" +| "rder c (RALTS rs) = RALTS (map (rder c) rs)" +| "rder c (RSEQ r1 r2) = + (if rnullable r1 + then RALT (RSEQ (rder c r1) r2) (rder c r2) + else RSEQ (rder c r1) r2)" +| "rder c (RSTAR r) = RSEQ (rder c r) (RSTAR r)" + + +fun + rders :: "rrexp \ string \ rrexp" +where + "rders r [] = r" +| "rders r (c#s) = rders (rder c r) s" + +fun rdistinct :: "'a list \'a set \ 'a list" + where + "rdistinct [] acc = []" +| "rdistinct (x#xs) acc = + (if x \ acc then rdistinct xs acc + else x # (rdistinct xs ({x} \ acc)))" + +lemma rdistinct1: + assumes "a \ acc" + shows "a \ set (rdistinct rs acc)" + using assms + apply(induct rs arbitrary: acc a) + apply(auto) + done + + +lemma rdistinct_does_the_job: + shows "distinct (rdistinct rs s)" + apply(induct rs s rule: rdistinct.induct) + apply(auto simp add: rdistinct1) + done + + + +lemma rdistinct_concat: + assumes "set rs \ rset" + shows "rdistinct (rs @ rsa) rset = rdistinct rsa rset" + using assms + apply(induct rs) + apply simp+ + done + +lemma distinct_not_exist: + assumes "a \ set rs" + shows "rdistinct rs rset = rdistinct rs (insert a rset)" + using assms + apply(induct rs arbitrary: rset) + apply(auto) + done + +lemma rdistinct_on_distinct: + shows "distinct rs \ rdistinct rs {} = rs" + apply(induct rs) + apply simp + using distinct_not_exist by fastforce + +lemma distinct_rdistinct_append: + assumes "distinct rs1" "\r \ set rs1. r \ acc" + shows "rdistinct (rs1 @ rsa) acc = rs1 @ (rdistinct rsa (acc \ set rs1))" + using assms + apply(induct rs1 arbitrary: rsa acc) + apply(auto)[1] + apply(auto)[1] + apply(drule_tac x="rsa" in meta_spec) + apply(drule_tac x="{a} \ acc" in meta_spec) + apply(simp) + apply(drule meta_mp) + apply(auto)[1] + apply(simp) + done + + +lemma rdistinct_set_equality1: + shows "set (rdistinct rs acc) = set rs - acc" + apply(induct rs acc rule: rdistinct.induct) + apply(auto) + done + + +lemma rdistinct_set_equality: + shows "set (rdistinct rs {}) = set rs" + by (simp add: rdistinct_set_equality1) + + +fun rflts :: "rrexp list \ rrexp list" + where + "rflts [] = []" +| "rflts (RZERO # rs) = rflts rs" +| "rflts ((RALTS rs1) # rs) = rs1 @ rflts rs" +| "rflts (r1 # rs) = r1 # rflts rs" + + +lemma rflts_def_idiot: + shows "\ a \ RZERO; \rs1. a = RALTS rs1\ \ rflts (a # rs) = a # rflts rs" + apply(case_tac a) + apply simp_all + done + +lemma rflts_def_idiot2: + shows "\a \ RZERO; \rs1. a = RALTS rs1; a \ set rs\ \ a \ set (rflts rs)" + apply(induct rs rule: rflts.induct) + apply(auto) + done + +lemma flts_append: + shows "rflts (rs1 @ rs2) = rflts rs1 @ rflts rs2" + apply(induct rs1) + apply simp + apply(case_tac a) + apply simp+ + done + + +fun rsimp_ALTs :: " rrexp list \ rrexp" + where + "rsimp_ALTs [] = RZERO" +| "rsimp_ALTs [r] = r" +| "rsimp_ALTs rs = RALTS rs" + +lemma rsimpalts_conscons: + shows "rsimp_ALTs (r1 # rsa @ r2 # rsb) = RALTS (r1 # rsa @ r2 # rsb)" + by (metis Nil_is_append_conv list.exhaust rsimp_ALTs.simps(3)) + +lemma rsimp_alts_equal: + shows "rsimp_ALTs (rsa @ a # rsb @ a # rsc) = RALTS (rsa @ a # rsb @ a # rsc) " + by (metis append_Cons append_Nil neq_Nil_conv rsimpalts_conscons) + + +fun rsimp_SEQ :: " rrexp \ rrexp \ rrexp" + where + "rsimp_SEQ RZERO _ = RZERO" +| "rsimp_SEQ _ RZERO = RZERO" +| "rsimp_SEQ RONE r2 = r2" +| "rsimp_SEQ r1 r2 = RSEQ r1 r2" + + +fun rsimp :: "rrexp \ rrexp" + where + "rsimp (RSEQ r1 r2) = rsimp_SEQ (rsimp r1) (rsimp r2)" +| "rsimp (RALTS rs) = rsimp_ALTs (rdistinct (rflts (map rsimp rs)) {}) " +| "rsimp r = r" + + +fun + rders_simp :: "rrexp \ string \ rrexp" +where + "rders_simp r [] = r" +| "rders_simp r (c#s) = rders_simp (rsimp (rder c r)) s" + +fun rsize :: "rrexp \ nat" where + "rsize RZERO = 1" +| "rsize (RONE) = 1" +| "rsize (RCHAR c) = 1" +| "rsize (RALTS rs) = Suc (sum_list (map rsize rs))" +| "rsize (RSEQ r1 r2) = Suc (rsize r1 + rsize r2)" +| "rsize (RSTAR r) = Suc (rsize r)" + +abbreviation rsizes where + "rsizes rs \ sum_list (map rsize rs)" + + +lemma rder_rsimp_ALTs_commute: + shows " (rder x (rsimp_ALTs rs)) = rsimp_ALTs (map (rder x) rs)" + apply(induct rs) + apply simp + apply(case_tac rs) + apply simp + apply auto + done + + +lemma rsimp_aalts_smaller: + shows "rsize (rsimp_ALTs rs) \ rsize (RALTS rs)" + apply(induct rs) + apply simp + apply simp + apply(case_tac "rs = []") + apply simp + apply(subgoal_tac "\rsp ap. rs = ap # rsp") + apply(erule exE)+ + apply simp + apply simp + by(meson neq_Nil_conv) + + + + + +lemma rSEQ_mono: + shows "rsize (rsimp_SEQ r1 r2) \rsize (RSEQ r1 r2)" + apply auto + apply(induct r1) + apply auto + apply(case_tac "r2") + apply simp_all + apply(case_tac r2) + apply simp_all + apply(case_tac r2) + apply simp_all + apply(case_tac r2) + apply simp_all + apply(case_tac r2) + apply simp_all + done + +lemma ralts_cap_mono: + shows "rsize (RALTS rs) \ Suc (rsizes rs)" + by simp + + + + +lemma rflts_mono: + shows "rsizes (rflts rs) \ rsizes rs" + apply(induct rs) + apply simp + apply(case_tac "a = RZERO") + apply simp + apply(case_tac "\rs1. a = RALTS rs1") + apply(erule exE) + apply simp + apply(subgoal_tac "rflts (a # rs) = a # (rflts rs)") + prefer 2 + + using rflts_def_idiot apply blast + apply simp + done + +lemma rdistinct_smaller: + shows "rsizes (rdistinct rs ss) \ rsizes rs" + apply (induct rs arbitrary: ss) + apply simp + by (simp add: trans_le_add2) + + +lemma rsimp_alts_mono : + shows "\x. (\xa. xa \ set x \ rsize (rsimp xa) \ rsize xa) \ + rsize (rsimp_ALTs (rdistinct (rflts (map rsimp x)) {})) \ Suc (rsizes x)" + apply(subgoal_tac "rsize (rsimp_ALTs (rdistinct (rflts (map rsimp x)) {} )) + \ rsize (RALTS (rdistinct (rflts (map rsimp x)) {} ))") + prefer 2 + using rsimp_aalts_smaller apply auto[1] + apply(subgoal_tac "rsize (RALTS (rdistinct (rflts (map rsimp x)) {})) \Suc (rsizes (rdistinct (rflts (map rsimp x)) {}))") + prefer 2 + using ralts_cap_mono apply blast + apply(subgoal_tac "rsizes (rdistinct (rflts (map rsimp x)) {}) \ rsizes (rflts (map rsimp x))") + prefer 2 + using rdistinct_smaller apply presburger + apply(subgoal_tac "rsizes (rflts (map rsimp x)) \ rsizes (map rsimp x)") + prefer 2 + using rflts_mono apply blast + apply(subgoal_tac "rsizes (map rsimp x) \ rsizes x") + prefer 2 + + apply (simp add: sum_list_mono) + by linarith + + + + + +lemma rsimp_mono: + shows "rsize (rsimp r) \ rsize r" + apply(induct r) + apply simp_all + apply(subgoal_tac "rsize (rsimp_SEQ (rsimp r1) (rsimp r2)) \ rsize (RSEQ (rsimp r1) (rsimp r2))") + apply force + using rSEQ_mono + apply presburger + using rsimp_alts_mono by auto + +lemma idiot: + shows "rsimp_SEQ RONE r = r" + apply(case_tac r) + apply simp_all + done + + + + + +lemma idiot2: + shows " \r1 \ RZERO; r1 \ RONE;r2 \ RZERO\ + \ rsimp_SEQ r1 r2 = RSEQ r1 r2" + apply(case_tac r1) + apply(case_tac r2) + apply simp_all + apply(case_tac r2) + apply simp_all + apply(case_tac r2) + apply simp_all + apply(case_tac r2) + apply simp_all + apply(case_tac r2) + apply simp_all + done + +lemma rders__onechar: + shows " (rders_simp r [c]) = (rsimp (rders r [c]))" + by simp + +lemma rders_append: + "rders c (s1 @ s2) = rders (rders c s1) s2" + apply(induct s1 arbitrary: c s2) + apply(simp_all) + done + +lemma rders_simp_append: + "rders_simp c (s1 @ s2) = rders_simp (rders_simp c s1) s2" + apply(induct s1 arbitrary: c s2) + apply(simp_all) + done + + +lemma rders_simp_one_char: + shows "rders_simp r [c] = rsimp (rder c r)" + apply auto + done + + + +fun nonalt :: "rrexp \ bool" + where + "nonalt (RALTS rs) = False" +| "nonalt r = True" + + +fun good :: "rrexp \ bool" where + "good RZERO = False" +| "good (RONE) = True" +| "good (RCHAR c) = True" +| "good (RALTS []) = False" +| "good (RALTS [r]) = False" +| "good (RALTS (r1 # r2 # rs)) = ((distinct ( (r1 # r2 # rs))) \(\r' \ set (r1 # r2 # rs). good r' \ nonalt r'))" +| "good (RSEQ RZERO _) = False" +| "good (RSEQ RONE _) = False" +| "good (RSEQ _ RZERO) = False" +| "good (RSEQ r1 r2) = (good r1 \ good r2)" +| "good (RSTAR r) = True" + + +lemma k0a: + shows "rflts [RALTS rs] = rs" + apply(simp) + done + +lemma bbbbs: + assumes "good r" "r = RALTS rs" + shows "rsimp_ALTs (rflts [r]) = RALTS rs" + using assms + by (metis good.simps(4) good.simps(5) k0a rsimp_ALTs.elims) + +lemma bbbbs1: + shows "nonalt r \ (\ rs. r = RALTS rs)" + by (meson nonalt.elims(3)) + + + +lemma good0: + assumes "rs \ Nil" "\r \ set rs. nonalt r" "distinct rs" + shows "good (rsimp_ALTs rs) \ (\r \ set rs. good r)" + using assms + apply(induct rs rule: rsimp_ALTs.induct) + apply(auto) + done + +lemma flts1: + assumes "good r" + shows "rflts [r] \ []" + using assms + apply(induct r) + apply(simp_all) + using good.simps(4) by blast + +lemma flts2: + assumes "good r" + shows "\r' \ set (rflts [r]). good r' \ nonalt r'" + using assms + apply(induct r) + apply(simp) + apply(simp) + apply(simp) + prefer 2 + apply(simp) + apply(auto)[1] + + apply (metis flts1 good.simps(5) good.simps(6) k0a neq_Nil_conv) + apply (metis flts1 good.simps(5) good.simps(6) k0a neq_Nil_conv) + apply fastforce + apply(simp) + done + + + +lemma flts3: + assumes "\r \ set rs. good r \ r = RZERO" + shows "\r \ set (rflts rs). good r" + using assms + apply(induct rs arbitrary: rule: rflts.induct) + apply(simp_all) + by (metis UnE flts2 k0a) + + +lemma k0: + shows "rflts (r # rs1) = rflts [r] @ rflts rs1" + apply(induct r arbitrary: rs1) + apply(auto) + done + + +lemma good_SEQ: + assumes "r1 \ RZERO" "r2 \ RZERO" " r1 \ RONE" + shows "good (RSEQ r1 r2) = (good r1 \ good r2)" + using assms + apply(case_tac r1) + apply(simp_all) + apply(case_tac r2) + apply(simp_all) + apply(case_tac r2) + apply(simp_all) + apply(case_tac r2) + apply(simp_all) + apply(case_tac r2) + apply(simp_all) + done + +lemma rsize0: + shows "0 < rsize r" + apply(induct r) + apply(auto) + done + + +fun nonnested :: "rrexp \ bool" + where + "nonnested (RALTS []) = True" +| "nonnested (RALTS ((RALTS rs1) # rs2)) = False" +| "nonnested (RALTS (r # rs2)) = nonnested (RALTS rs2)" +| "nonnested r = True" + + + +lemma k00: + shows "rflts (rs1 @ rs2) = rflts rs1 @ rflts rs2" + apply(induct rs1 arbitrary: rs2) + apply(auto) + by (metis append.assoc k0) + + + + +lemma k0b: + assumes "nonalt r" "r \ RZERO" + shows "rflts [r] = [r]" + using assms + apply(case_tac r) + apply(simp_all) + done + +lemma nn1qq: + assumes "nonnested (RALTS rs)" + shows "\ rs1. RALTS rs1 \ set rs" + using assms + apply(induct rs rule: rflts.induct) + apply(auto) + done + + + +lemma n0: + shows "nonnested (RALTS rs) \ (\r \ set rs. nonalt r)" + apply(induct rs ) + apply(auto) + apply (metis list.set_intros(1) nn1qq nonalt.elims(3)) + apply (metis nonalt.elims(2) nonnested.simps(3) nonnested.simps(4) nonnested.simps(5) nonnested.simps(6) nonnested.simps(7)) + using bbbbs1 apply fastforce + by (metis bbbbs1 list.set_intros(2) nn1qq) + + + + +lemma nn1c: + assumes "\r \ set rs. nonnested r" + shows "\r \ set (rflts rs). nonalt r" + using assms + apply(induct rs rule: rflts.induct) + apply(auto) + using n0 by blast + +lemma nn1bb: + assumes "\r \ set rs. nonalt r" + shows "nonnested (rsimp_ALTs rs)" + using assms + apply(induct rs rule: rsimp_ALTs.induct) + apply(auto) + using nonalt.simps(1) nonnested.elims(3) apply blast + using n0 by auto + +lemma bsimp_ASEQ0: + shows "rsimp_SEQ r1 RZERO = RZERO" + apply(induct r1) + apply(auto) + done + +lemma nn1b: + shows "nonnested (rsimp r)" + apply(induct r) + apply(simp_all) + apply(case_tac "rsimp r1 = RZERO") + apply(simp) + apply(case_tac "rsimp r2 = RZERO") + apply(simp) + apply(subst bsimp_ASEQ0) + apply(simp) + apply(case_tac "\bs. rsimp r1 = RONE") + apply(auto)[1] + using idiot apply fastforce + using idiot2 nonnested.simps(11) apply presburger + by (metis (mono_tags, lifting) Diff_empty image_iff list.set_map nn1bb nn1c rdistinct_set_equality1) + +lemma nonalt_flts_rd: + shows "\xa \ set (rdistinct (rflts (map rsimp rs)) {})\ + \ nonalt xa" + by (metis Diff_empty ex_map_conv nn1b nn1c rdistinct_set_equality1) + + +lemma rsimpalts_implies1: + shows " rsimp_ALTs (a # rdistinct rs {a}) = RZERO \ a = RZERO" + using rsimp_ALTs.elims by auto + + +lemma rsimpalts_implies2: + shows "rsimp_ALTs (a # rdistinct rs rset) = RZERO \ rdistinct rs rset = []" + by (metis append_butlast_last_id rrexp.distinct(7) rsimpalts_conscons) + +lemma rsimpalts_implies21: + shows "rsimp_ALTs (a # rdistinct rs {a}) = RZERO \ rdistinct rs {a} = []" + using rsimpalts_implies2 by blast + + +lemma bsimp_ASEQ2: + shows "rsimp_SEQ RONE r2 = r2" + apply(induct r2) + apply(auto) + done + +lemma elem_smaller_than_set: + shows "xa \ set list \ rsize xa < Suc (rsizes list)" + apply(induct list) + apply simp + by (metis image_eqI le_imp_less_Suc list.set_map member_le_sum_list) + +lemma rsimp_list_mono: + shows "rsizes (map rsimp rs) \ rsizes rs" + apply(induct rs) + apply simp+ + by (simp add: add_mono_thms_linordered_semiring(1) rsimp_mono) + + +(*says anything coming out of simp+flts+db will be good*) +lemma good2_obv_simplified: + shows " \\y. rsize y < Suc (rsizes rs) \ good (rsimp y) \ rsimp y = RZERO; + xa \ set (rdistinct (rflts (map rsimp rs)) {}); good (rsimp xa) \ rsimp xa = RZERO\ \ good xa" + apply(subgoal_tac " \xa' \ set (map rsimp rs). good xa' \ xa' = RZERO") + prefer 2 + apply (simp add: elem_smaller_than_set) + by (metis Diff_empty flts3 rdistinct_set_equality1) + + +lemma good1: + shows "good (rsimp a) \ rsimp a = RZERO" + apply(induct a taking: rsize rule: measure_induct) + apply(case_tac x) + apply(simp) + apply(simp) + apply(simp) + prefer 3 + apply(simp) + prefer 2 + apply(simp only:) + apply simp + apply (smt (verit, ccfv_threshold) add_mono_thms_linordered_semiring(1) elem_smaller_than_set good0 good2_obv_simplified le_eq_less_or_eq nonalt_flts_rd order_less_trans plus_1_eq_Suc rdistinct_does_the_job rdistinct_smaller rflts_mono rsimp_ALTs.simps(1) rsimp_list_mono) + apply simp + apply(subgoal_tac "good (rsimp x41) \ rsimp x41 = RZERO") + apply(subgoal_tac "good (rsimp x42) \ rsimp x42 = RZERO") + apply(case_tac "rsimp x41 = RZERO") + apply simp + apply(case_tac "rsimp x42 = RZERO") + apply simp + using bsimp_ASEQ0 apply blast + apply(subgoal_tac "good (rsimp x41)") + apply(subgoal_tac "good (rsimp x42)") + apply simp + apply (metis bsimp_ASEQ2 good_SEQ idiot2) + apply blast + apply fastforce + using less_add_Suc2 apply blast + using less_iff_Suc_add by blast + + + +fun + RL :: "rrexp \ string set" +where + "RL (RZERO) = {}" +| "RL (RONE) = {[]}" +| "RL (RCHAR c) = {[c]}" +| "RL (RSEQ r1 r2) = (RL r1) ;; (RL r2)" +| "RL (RALTS rs) = (\ (set (map RL rs)))" +| "RL (RSTAR r) = (RL r)\" + + +lemma RL_rnullable: + shows "rnullable r = ([] \ RL r)" + apply(induct r) + apply(auto simp add: Sequ_def) + done + +lemma RL_rder: + shows "RL (rder c r) = Der c (RL r)" + apply(induct r) + apply(auto simp add: Sequ_def Der_def) + apply (metis append_Cons) + using RL_rnullable apply blast + apply (metis append_eq_Cons_conv) + apply (metis append_Cons) + apply (metis RL_rnullable append_eq_Cons_conv) + apply (metis Star.step append_Cons) + using Star_decomp by auto + + + + +lemma RL_rsimp_RSEQ: + shows "RL (rsimp_SEQ r1 r2) = (RL r1 ;; RL r2)" + apply(induct r1 r2 rule: rsimp_SEQ.induct) + apply(simp_all) + done + +lemma RL_rsimp_RALTS: + shows "RL (rsimp_ALTs rs) = (\ (set (map RL rs)))" + apply(induct rs rule: rsimp_ALTs.induct) + apply(simp_all) + done + +lemma RL_rsimp_rdistinct: + shows "(\ (set (map RL (rdistinct rs {})))) = (\ (set (map RL rs)))" + apply(auto) + apply (metis Diff_iff rdistinct_set_equality1) + by (metis Diff_empty rdistinct_set_equality1) + +lemma RL_rsimp_rflts: + shows "(\ (set (map RL (rflts rs)))) = (\ (set (map RL rs)))" + apply(induct rs rule: rflts.induct) + apply(simp_all) + done + +lemma RL_rsimp: + shows "RL r = RL (rsimp r)" + apply(induct r rule: rsimp.induct) + apply(auto simp add: Sequ_def RL_rsimp_RSEQ) + using RL_rsimp_RALTS RL_rsimp_rdistinct RL_rsimp_rflts apply auto[1] + by (smt (verit, del_insts) RL_rsimp_RALTS RL_rsimp_rdistinct RL_rsimp_rflts UN_E image_iff list.set_map) + + +lemma qqq1: + shows "RZERO \ set (rflts (map rsimp rs))" + by (metis ex_map_conv flts3 good.simps(1) good1) + + +fun nonazero :: "rrexp \ bool" + where + "nonazero RZERO = False" +| "nonazero r = True" + + +lemma flts_single1: + assumes "nonalt r" "nonazero r" + shows "rflts [r] = [r]" + using assms + apply(induct r) + apply(auto) + done + +lemma nonalt0_flts_keeps: + shows "(a \ RZERO) \ (\rs. a \ RALTS rs) \ rflts (a # xs) = a # rflts xs" + apply(case_tac a) + apply simp+ + done + + +lemma nonalt0_fltseq: + shows "\r \ set rs. r \ RZERO \ nonalt r \ rflts rs = rs" + apply(induct rs) + apply simp + apply(case_tac "a = RZERO") + apply fastforce + apply(case_tac "\rs1. a = RALTS rs1") + apply(erule exE) + apply simp+ + using nonalt0_flts_keeps by presburger + + + + +lemma goodalts_nonalt: + shows "good (RALTS rs) \ rflts rs = rs" + apply(induct x == "RALTS rs" arbitrary: rs rule: good.induct) + apply simp + + using good.simps(5) apply blast + apply simp + apply(case_tac "r1 = RZERO") + using good.simps(1) apply force + apply(case_tac "r2 = RZERO") + using good.simps(1) apply force + apply(subgoal_tac "rflts (r1 # r2 # rs) = r1 # r2 # rflts rs") + prefer 2 + apply (metis nonalt.simps(1) rflts_def_idiot) + apply(subgoal_tac "\r \ set rs. r \ RZERO \ nonalt r") + apply(subgoal_tac "rflts rs = rs") + apply presburger + using nonalt0_fltseq apply presburger + using good.simps(1) by blast + + + + + +lemma test: + assumes "good r" + shows "rsimp r = r" + + using assms + apply(induct rule: good.induct) + apply simp + apply simp + apply simp + apply simp + apply simp + apply(subgoal_tac "distinct (r1 # r2 # rs)") + prefer 2 + using good.simps(6) apply blast + apply(subgoal_tac "rflts (r1 # r2 # rs ) = r1 # r2 # rs") + prefer 2 + using goodalts_nonalt apply blast + + apply(subgoal_tac "r1 \ r2") + prefer 2 + apply (meson distinct_length_2_or_more) + apply(subgoal_tac "r1 \ set rs") + apply(subgoal_tac "r2 \ set rs") + apply(subgoal_tac "\r \ set rs. rsimp r = r") + apply(subgoal_tac "map rsimp rs = rs") + apply simp + apply(subgoal_tac "\r \ {r1, r2}. r \ set rs") + apply (metis distinct_not_exist rdistinct_on_distinct) + + apply blast + apply (meson map_idI) + apply (metis good.simps(6) insert_iff list.simps(15)) + + apply (meson distinct.simps(2)) + apply (simp add: distinct_length_2_or_more) + apply simp+ + done + + + +lemma rsimp_idem: + shows "rsimp (rsimp r) = rsimp r" + using test good1 + by force + +corollary rsimp_inner_idem4: + shows "rsimp r = RALTS rs \ rflts rs = rs" + by (metis good1 goodalts_nonalt rrexp.simps(12)) + + +lemma head_one_more_simp: + shows "map rsimp (r # rs) = map rsimp (( rsimp r) # rs)" + by (simp add: rsimp_idem) + + +lemma der_simp_nullability: + shows "rnullable r = rnullable (rsimp r)" + using RL_rnullable RL_rsimp by auto + + +lemma no_alt_short_list_after_simp: + shows "RALTS rs = rsimp r \ rsimp_ALTs rs = RALTS rs" + by (metis bbbbs good1 k0a rrexp.simps(12)) + + +lemma no_further_dB_after_simp: + shows "RALTS rs = rsimp r \ rdistinct rs {} = rs" + apply(subgoal_tac "good (RALTS rs)") + apply(subgoal_tac "distinct rs") + using rdistinct_on_distinct apply blast + apply (metis distinct.simps(1) distinct.simps(2) empty_iff good.simps(6) list.exhaust set_empty2) + using good1 by fastforce + + +lemma idem_after_simp1: + shows "rsimp_ALTs (rdistinct (rflts [rsimp aa]) {}) = rsimp aa" + apply(case_tac "rsimp aa") + apply simp+ + apply (metis no_alt_short_list_after_simp no_further_dB_after_simp) + by simp + +lemma identity_wwo0: + shows "rsimp (rsimp_ALTs (RZERO # rs)) = rsimp (rsimp_ALTs rs)" + by (metis idem_after_simp1 list.exhaust list.simps(8) list.simps(9) rflts.simps(2) rsimp.simps(2) rsimp.simps(3) rsimp_ALTs.simps(1) rsimp_ALTs.simps(2) rsimp_ALTs.simps(3)) + + +lemma distinct_removes_last: + shows "\a \ set as\ + \ rdistinct as rset = rdistinct (as @ [a]) rset" +and "rdistinct (ab # as @ [ab]) rset1 = rdistinct (ab # as) rset1" + apply(induct as arbitrary: rset ab rset1 a) + apply simp + apply simp + apply(case_tac "aa \ rset") + apply(case_tac "a = aa") + apply (metis append_Cons) + apply simp + apply(case_tac "a \ set as") + apply (metis append_Cons rdistinct.simps(2) set_ConsD) + apply(case_tac "a = aa") + prefer 2 + apply simp + apply (metis append_Cons) + apply(case_tac "ab \ rset1") + prefer 2 + apply(subgoal_tac "rdistinct (ab # (a # as) @ [ab]) rset1 = + ab # (rdistinct ((a # as) @ [ab]) (insert ab rset1))") + prefer 2 + apply force + apply(simp only:) + apply(subgoal_tac "rdistinct (ab # a # as) rset1 = ab # (rdistinct (a # as) (insert ab rset1))") + apply(simp only:) + apply(subgoal_tac "rdistinct ((a # as) @ [ab]) (insert ab rset1) = rdistinct (a # as) (insert ab rset1)") + apply blast + apply(case_tac "a \ insert ab rset1") + apply simp + apply (metis insertI1) + apply simp + apply (meson insertI1) + apply simp + apply(subgoal_tac "rdistinct ((a # as) @ [ab]) rset1 = rdistinct (a # as) rset1") + apply simp + by (metis append_Cons insert_iff insert_is_Un rdistinct.simps(2)) + + +lemma distinct_removes_middle: + shows "\a \ set as\ + \ rdistinct (as @ as2) rset = rdistinct (as @ [a] @ as2) rset" +and "rdistinct (ab # as @ [ab] @ as3) rset1 = rdistinct (ab # as @ as3) rset1" + apply(induct as arbitrary: rset rset1 ab as2 as3 a) + apply simp + apply simp + apply(case_tac "a \ rset") + apply simp + apply metis + apply simp + apply (metis insertI1) + apply(case_tac "a = ab") + apply simp + apply(case_tac "ab \ rset") + apply simp + apply presburger + apply (meson insertI1) + apply(case_tac "a \ rset") + apply (metis (no_types, opaque_lifting) Un_insert_left append_Cons insert_iff rdistinct.simps(2) sup_bot_left) + apply(case_tac "ab \ rset") + apply simp + apply (meson insert_iff) + apply simp + by (metis insertI1) + + +lemma distinct_removes_middle3: + shows "\a \ set as\ + \ rdistinct (as @ a #as2) rset = rdistinct (as @ as2) rset" + using distinct_removes_middle(1) by fastforce + + +lemma distinct_removes_list: + shows "\ \r \ set rs. r \ set as\ \ rdistinct (as @ rs) {} = rdistinct as {}" + apply(induct rs) + apply simp+ + apply(subgoal_tac "rdistinct (as @ a # rs) {} = rdistinct (as @ rs) {}") + prefer 2 + apply (metis append_Cons append_Nil distinct_removes_middle(1)) + by presburger + + +lemma spawn_simp_rsimpalts: + shows "rsimp (rsimp_ALTs rs) = rsimp (rsimp_ALTs (map rsimp rs))" + apply(cases rs) + apply simp + apply(case_tac list) + apply simp + apply(subst rsimp_idem[symmetric]) + apply simp + apply(subgoal_tac "rsimp_ALTs rs = RALTS rs") + apply(simp only:) + apply(subgoal_tac "rsimp_ALTs (map rsimp rs) = RALTS (map rsimp rs)") + apply(simp only:) + prefer 2 + apply simp + prefer 2 + using rsimp_ALTs.simps(3) apply presburger + apply auto + apply(subst rsimp_idem)+ + by (metis comp_apply rsimp_idem) + + +lemma simp_singlealt_flatten: + shows "rsimp (RALTS [RALTS rsa]) = rsimp (RALTS (rsa @ []))" + apply(induct rsa) + apply simp + apply simp + by (metis idem_after_simp1 list.simps(9) rsimp.simps(2)) + + +lemma good1_rsimpalts: + shows "rsimp r = RALTS rs \ rsimp_ALTs rs = RALTS rs" + by (metis no_alt_short_list_after_simp) + + + + +lemma good1_flatten: + shows "\ rsimp r = (RALTS rs1)\ + \ rflts (rsimp_ALTs rs1 # map rsimp rsb) = rflts (rs1 @ map rsimp rsb)" + apply(subst good1_rsimpalts) + apply simp+ + apply(subgoal_tac "rflts (rs1 @ map rsimp rsb) = rs1 @ rflts (map rsimp rsb)") + apply simp + using flts_append rsimp_inner_idem4 by presburger + + +lemma flatten_rsimpalts: + shows "rflts (rsimp_ALTs (rdistinct (rflts (map rsimp rsa)) {}) # map rsimp rsb) = + rflts ( (rdistinct (rflts (map rsimp rsa)) {}) @ map rsimp rsb)" + apply(case_tac "map rsimp rsa") + apply simp + apply(case_tac "list") + apply simp + apply(case_tac a) + apply simp+ + apply(rename_tac rs1) + apply (metis good1_flatten map_eq_Cons_D no_further_dB_after_simp) + + apply simp + + apply(subgoal_tac "\r \ set( rflts (map rsimp rsa)). good r") + apply(case_tac "rdistinct (rflts (map rsimp rsa)) {}") + apply simp + apply(case_tac "listb") + apply simp+ + apply (metis Cons_eq_appendI good1_flatten rflts.simps(3) rsimp.simps(2) rsimp_ALTs.simps(3)) + by (metis (mono_tags, lifting) flts3 good1 image_iff list.set_map) + + +lemma last_elem_out: + shows "\x \ set xs; x \ rset \ \ rdistinct (xs @ [x]) rset = rdistinct xs rset @ [x]" + apply(induct xs arbitrary: rset) + apply simp+ + done + + + + +lemma rdistinct_concat_general: + shows "rdistinct (rs1 @ rs2) {} = (rdistinct rs1 {}) @ (rdistinct rs2 (set rs1))" + apply(induct rs1 arbitrary: rs2 rule: rev_induct) + apply simp + apply(drule_tac x = "x # rs2" in meta_spec) + apply simp + apply(case_tac "x \ set xs") + apply simp + + apply (simp add: distinct_removes_middle3 insert_absorb) + apply simp + by (simp add: last_elem_out) + + + + +lemma distinct_once_enough: + shows "rdistinct (rs @ rsa) {} = rdistinct (rdistinct rs {} @ rsa) {}" + apply(subgoal_tac "distinct (rdistinct rs {})") + apply(subgoal_tac +" rdistinct (rdistinct rs {} @ rsa) {} = rdistinct rs {} @ (rdistinct rsa (set rs))") + apply(simp only:) + using rdistinct_concat_general apply blast + apply (simp add: distinct_rdistinct_append rdistinct_set_equality1) + by (simp add: rdistinct_does_the_job) + + +lemma simp_flatten: + shows "rsimp (RALTS ((RALTS rsa) # rsb)) = rsimp (RALTS (rsa @ rsb))" + apply simp + apply(subst flatten_rsimpalts) + apply(simp add: flts_append) + by (metis Diff_empty distinct_once_enough flts_append nonalt0_fltseq nonalt_flts_rd qqq1 rdistinct_set_equality1) + +lemma basic_rsimp_SEQ_property1: + shows "rsimp_SEQ RONE r = r" + by (simp add: idiot) + + + +lemma basic_rsimp_SEQ_property3: + shows "rsimp_SEQ r RZERO = RZERO" + using rsimp_SEQ.elims by blast + + + +fun vsuf :: "char list \ rrexp \ char list list" where +"vsuf [] _ = []" +|"vsuf (c#cs) r1 = (if (rnullable r1) then (vsuf cs (rder c r1)) @ [c # cs] + else (vsuf cs (rder c r1)) + ) " + + + + + + +fun star_update :: "char \ rrexp \ char list list \ char list list" where +"star_update c r [] = []" +|"star_update c r (s # Ss) = (if (rnullable (rders r s)) + then (s@[c]) # [c] # (star_update c r Ss) + else (s@[c]) # (star_update c r Ss) )" + + +fun star_updates :: "char list \ rrexp \ char list list \ char list list" + where +"star_updates [] r Ss = Ss" +| "star_updates (c # cs) r Ss = star_updates cs r (star_update c r Ss)" + +lemma stupdates_append: shows +"star_updates (s @ [c]) r Ss = star_update c r (star_updates s r Ss)" + apply(induct s arbitrary: Ss) + apply simp + apply simp + done + +lemma flts_removes0: + shows " rflts (rs @ [RZERO]) = + rflts rs" + apply(induct rs) + apply simp + by (metis append_Cons rflts.simps(2) rflts.simps(3) rflts_def_idiot) + + +lemma rflts_spills_last: + shows "rflts (rs1 @ [RALTS rs]) = rflts rs1 @ rs" + apply (induct rs1 rule: rflts.induct) + apply(auto) + done + +lemma flts_keeps1: + shows "rflts (rs @ [RONE]) = rflts rs @ [RONE]" + apply (induct rs rule: rflts.induct) + apply(auto) + done + +lemma flts_keeps_others: + shows "\a \ RZERO; \rs1. a = RALTS rs1\ \rflts (rs @ [a]) = rflts rs @ [a]" + apply(induct rs rule: rflts.induct) + apply(auto) + by (meson k0b nonalt.elims(3)) + +lemma spilled_alts_contained: + shows "\a = RALTS rs ; a \ set rs1\ \ \r \ set rs. r \ set (rflts rs1)" + apply(induct rs1) + apply simp + apply(case_tac "a = aa") + apply simp + apply(subgoal_tac " a \ set rs1") + prefer 2 + apply (meson set_ConsD) + apply(case_tac aa) + using rflts.simps(2) apply presburger + apply fastforce + apply fastforce + apply fastforce + apply fastforce + by fastforce + + +lemma distinct_removes_duplicate_flts: + shows " a \ set rsa + \ rdistinct (rflts (map rsimp rsa @ [rsimp a])) {} = + rdistinct (rflts (map rsimp rsa)) {}" + apply(subgoal_tac "rsimp a \ set (map rsimp rsa)") + prefer 2 + apply simp + apply(induct "rsimp a") + apply simp + using flts_removes0 apply presburger + apply(subgoal_tac " rdistinct (rflts (map rsimp rsa @ [rsimp a])) {} = + rdistinct (rflts (map rsimp rsa @ [RONE])) {}") + apply (simp only:) + apply(subst flts_keeps1) + apply (metis distinct_removes_last(1) rflts_def_idiot2 rrexp.simps(20) rrexp.simps(6)) + apply presburger + apply(subgoal_tac " rdistinct (rflts (map rsimp rsa @ [rsimp a])) {} = + rdistinct ((rflts (map rsimp rsa)) @ [RCHAR x]) {}") + apply (simp only:) + prefer 2 + apply (metis flts_keeps_others rrexp.distinct(21) rrexp.distinct(3)) + apply (metis distinct_removes_last(1) rflts_def_idiot2 rrexp.distinct(21) rrexp.distinct(3)) + + apply (metis distinct_removes_last(1) flts_keeps_others rflts_def_idiot2 rrexp.distinct(25) rrexp.distinct(5)) + prefer 2 + apply (metis distinct_removes_last(1) flts_keeps_others flts_removes0 rflts_def_idiot2 rrexp.distinct(29)) + apply(subgoal_tac "rflts (map rsimp rsa @ [rsimp a]) = rflts (map rsimp rsa) @ x") + prefer 2 + apply (simp add: rflts_spills_last) + apply(subgoal_tac "\ r \ set x. r \ set (rflts (map rsimp rsa))") + prefer 2 + apply (metis (mono_tags, lifting) image_iff image_set spilled_alts_contained) + apply (metis rflts_spills_last) + by (metis distinct_removes_list spilled_alts_contained) + + + +(*some basic facts about rsimp*) + +unused_thms + + +end \ No newline at end of file diff -r f9cdc295ccf7 -r f493a20feeb3 thys3/src/Blexer.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys3/src/Blexer.thy Sat Apr 30 00:50:08 2022 +0100 @@ -0,0 +1,454 @@ + +theory Blexer + imports "Lexer" "PDerivs" +begin + +section \Bit-Encodings\ + +datatype bit = Z | S + +fun code :: "val \ bit list" +where + "code Void = []" +| "code (Char c) = []" +| "code (Left v) = Z # (code v)" +| "code (Right v) = S # (code v)" +| "code (Seq v1 v2) = (code v1) @ (code v2)" +| "code (Stars []) = [S]" +| "code (Stars (v # vs)) = (Z # code v) @ code (Stars vs)" + + +fun + Stars_add :: "val \ val \ val" +where + "Stars_add v (Stars vs) = Stars (v # vs)" + +function + decode' :: "bit list \ rexp \ (val * bit list)" +where + "decode' bs ZERO = (undefined, bs)" +| "decode' bs ONE = (Void, bs)" +| "decode' bs (CH d) = (Char d, bs)" +| "decode' [] (ALT r1 r2) = (Void, [])" +| "decode' (Z # bs) (ALT r1 r2) = (let (v, bs') = decode' bs r1 in (Left v, bs'))" +| "decode' (S # bs) (ALT r1 r2) = (let (v, bs') = decode' bs r2 in (Right v, bs'))" +| "decode' bs (SEQ r1 r2) = (let (v1, bs') = decode' bs r1 in + let (v2, bs'') = decode' bs' r2 in (Seq v1 v2, bs''))" +| "decode' [] (STAR r) = (Void, [])" +| "decode' (S # bs) (STAR r) = (Stars [], bs)" +| "decode' (Z # bs) (STAR r) = (let (v, bs') = decode' bs r in + let (vs, bs'') = decode' bs' (STAR r) + in (Stars_add v vs, bs''))" +by pat_completeness auto + +lemma decode'_smaller: + assumes "decode'_dom (bs, r)" + shows "length (snd (decode' bs r)) \ length bs" +using assms +apply(induct bs r) +apply(auto simp add: decode'.psimps split: prod.split) +using dual_order.trans apply blast +by (meson dual_order.trans le_SucI) + +termination "decode'" +apply(relation "inv_image (measure(%cs. size cs) <*lex*> measure(%s. size s)) (%(ds,r). (r,ds))") +apply(auto dest!: decode'_smaller) +by (metis less_Suc_eq_le snd_conv) + +definition + decode :: "bit list \ rexp \ val option" +where + "decode ds r \ (let (v, ds') = decode' ds r + in (if ds' = [] then Some v else None))" + +lemma decode'_code_Stars: + assumes "\v\set vs. \ v : r \ (\x. decode' (code v @ x) r = (v, x)) \ flat v \ []" + shows "decode' (code (Stars vs) @ ds) (STAR r) = (Stars vs, ds)" + using assms + apply(induct vs) + apply(auto) + done + +lemma decode'_code: + assumes "\ v : r" + shows "decode' ((code v) @ ds) r = (v, ds)" +using assms + apply(induct v r arbitrary: ds) + apply(auto) + using decode'_code_Stars by blast + +lemma decode_code: + assumes "\ v : r" + shows "decode (code v) r = Some v" + using assms unfolding decode_def + by (smt append_Nil2 decode'_code old.prod.case) + + +section {* Annotated Regular Expressions *} + +datatype arexp = + AZERO +| AONE "bit list" +| ACHAR "bit list" char +| ASEQ "bit list" arexp arexp +| AALTs "bit list" "arexp list" +| ASTAR "bit list" arexp + +abbreviation + "AALT bs r1 r2 \ AALTs bs [r1, r2]" + +fun asize :: "arexp \ nat" where + "asize AZERO = 1" +| "asize (AONE cs) = 1" +| "asize (ACHAR cs c) = 1" +| "asize (AALTs cs rs) = Suc (sum_list (map asize rs))" +| "asize (ASEQ cs r1 r2) = Suc (asize r1 + asize r2)" +| "asize (ASTAR cs r) = Suc (asize r)" + +fun + erase :: "arexp \ rexp" +where + "erase AZERO = ZERO" +| "erase (AONE _) = ONE" +| "erase (ACHAR _ c) = CH c" +| "erase (AALTs _ []) = ZERO" +| "erase (AALTs _ [r]) = (erase r)" +| "erase (AALTs bs (r#rs)) = ALT (erase r) (erase (AALTs bs rs))" +| "erase (ASEQ _ r1 r2) = SEQ (erase r1) (erase r2)" +| "erase (ASTAR _ r) = STAR (erase r)" + + +fun fuse :: "bit list \ arexp \ arexp" where + "fuse bs AZERO = AZERO" +| "fuse bs (AONE cs) = AONE (bs @ cs)" +| "fuse bs (ACHAR cs c) = ACHAR (bs @ cs) c" +| "fuse bs (AALTs cs rs) = AALTs (bs @ cs) rs" +| "fuse bs (ASEQ cs r1 r2) = ASEQ (bs @ cs) r1 r2" +| "fuse bs (ASTAR cs r) = ASTAR (bs @ cs) r" + +lemma fuse_append: + shows "fuse (bs1 @ bs2) r = fuse bs1 (fuse bs2 r)" + apply(induct r) + apply(auto) + done + + +fun intern :: "rexp \ arexp" where + "intern ZERO = AZERO" +| "intern ONE = AONE []" +| "intern (CH c) = ACHAR [] c" +| "intern (ALT r1 r2) = AALT [] (fuse [Z] (intern r1)) + (fuse [S] (intern r2))" +| "intern (SEQ r1 r2) = ASEQ [] (intern r1) (intern r2)" +| "intern (STAR r) = ASTAR [] (intern r)" + + +fun retrieve :: "arexp \ val \ bit list" where + "retrieve (AONE bs) Void = bs" +| "retrieve (ACHAR bs c) (Char d) = bs" +| "retrieve (AALTs bs [r]) v = bs @ retrieve r v" +| "retrieve (AALTs bs (r#rs)) (Left v) = bs @ retrieve r v" +| "retrieve (AALTs bs (r#rs)) (Right v) = bs @ retrieve (AALTs [] rs) v" +| "retrieve (ASEQ bs r1 r2) (Seq v1 v2) = bs @ retrieve r1 v1 @ retrieve r2 v2" +| "retrieve (ASTAR bs r) (Stars []) = bs @ [S]" +| "retrieve (ASTAR bs r) (Stars (v#vs)) = + bs @ [Z] @ retrieve r v @ retrieve (ASTAR [] r) (Stars vs)" + + + +fun + bnullable :: "arexp \ bool" +where + "bnullable (AZERO) = False" +| "bnullable (AONE bs) = True" +| "bnullable (ACHAR bs c) = False" +| "bnullable (AALTs bs rs) = (\r \ set rs. bnullable r)" +| "bnullable (ASEQ bs r1 r2) = (bnullable r1 \ bnullable r2)" +| "bnullable (ASTAR bs r) = True" + +abbreviation + bnullables :: "arexp list \ bool" +where + "bnullables rs \ (\r \ set rs. bnullable r)" + +fun + bmkeps :: "arexp \ bit list" and + bmkepss :: "arexp list \ bit list" +where + "bmkeps(AONE bs) = bs" +| "bmkeps(ASEQ bs r1 r2) = bs @ (bmkeps r1) @ (bmkeps r2)" +| "bmkeps(AALTs bs rs) = bs @ (bmkepss rs)" +| "bmkeps(ASTAR bs r) = bs @ [S]" +| "bmkepss (r # rs) = (if bnullable(r) then (bmkeps r) else (bmkepss rs))" + +lemma bmkepss1: + assumes "\ bnullables rs1" + shows "bmkepss (rs1 @ rs2) = bmkepss rs2" + using assms + by (induct rs1) (auto) + +lemma bmkepss2: + assumes "bnullables rs1" + shows "bmkepss (rs1 @ rs2) = bmkepss rs1" + using assms + by (induct rs1) (auto) + + +fun + bder :: "char \ arexp \ arexp" +where + "bder c (AZERO) = AZERO" +| "bder c (AONE bs) = AZERO" +| "bder c (ACHAR bs d) = (if c = d then AONE bs else AZERO)" +| "bder c (AALTs bs rs) = AALTs bs (map (bder c) rs)" +| "bder c (ASEQ bs r1 r2) = + (if bnullable r1 + then AALT bs (ASEQ [] (bder c r1) r2) (fuse (bmkeps r1) (bder c r2)) + else ASEQ bs (bder c r1) r2)" +| "bder c (ASTAR bs r) = ASEQ (bs @ [Z]) (bder c r) (ASTAR [] r)" + + +fun + bders :: "arexp \ string \ arexp" +where + "bders r [] = r" +| "bders r (c#s) = bders (bder c r) s" + +lemma bders_append: + "bders c (s1 @ s2) = bders (bders c s1) s2" + apply(induct s1 arbitrary: c s2) + apply(simp_all) + done + +lemma bnullable_correctness: + shows "nullable (erase r) = bnullable r" + apply(induct r rule: erase.induct) + apply(simp_all) + done + +lemma erase_fuse: + shows "erase (fuse bs r) = erase r" + apply(induct r rule: erase.induct) + apply(simp_all) + done + +lemma erase_intern [simp]: + shows "erase (intern r) = r" + apply(induct r) + apply(simp_all add: erase_fuse) + done + +lemma erase_bder [simp]: + shows "erase (bder a r) = der a (erase r)" + apply(induct r rule: erase.induct) + apply(simp_all add: erase_fuse bnullable_correctness) + done + +lemma erase_bders [simp]: + shows "erase (bders r s) = ders s (erase r)" + apply(induct s arbitrary: r ) + apply(simp_all) + done + +lemma bnullable_fuse: + shows "bnullable (fuse bs r) = bnullable r" + apply(induct r arbitrary: bs) + apply(auto) + done + +lemma retrieve_encode_STARS: + assumes "\v\set vs. \ v : r \ code v = retrieve (intern r) v" + shows "code (Stars vs) = retrieve (ASTAR [] (intern r)) (Stars vs)" + using assms + apply(induct vs) + apply(simp_all) + done + +lemma retrieve_fuse2: + assumes "\ v : (erase r)" + shows "retrieve (fuse bs r) v = bs @ retrieve r v" + using assms + apply(induct r arbitrary: v bs) + apply(auto elim: Prf_elims)[4] + apply(case_tac x2a) + apply(simp) + using Prf_elims(1) apply blast + apply(case_tac x2a) + apply(simp) + apply(simp) + apply(case_tac list) + apply(simp) + apply(simp) + apply (smt (verit, best) Prf_elims(3) append_assoc retrieve.simps(4) retrieve.simps(5)) + apply(simp) + using retrieve_encode_STARS + apply(auto elim!: Prf_elims)[1] + apply(case_tac vs) + apply(simp) + apply(simp) + done + +lemma retrieve_fuse: + assumes "\ v : r" + shows "retrieve (fuse bs (intern r)) v = bs @ retrieve (intern r) v" + using assms + by (simp_all add: retrieve_fuse2) + + +lemma retrieve_code: + assumes "\ v : r" + shows "code v = retrieve (intern r) v" + using assms + apply(induct v r ) + apply(simp_all add: retrieve_fuse retrieve_encode_STARS) + done + + +lemma retrieve_AALTs_bnullable1: + assumes "bnullable r" + shows "retrieve (AALTs bs (r # rs)) (mkeps (erase (AALTs bs (r # rs)))) + = bs @ retrieve r (mkeps (erase r))" + using assms + apply(case_tac rs) + apply(auto simp add: bnullable_correctness) + done + +lemma retrieve_AALTs_bnullable2: + assumes "\bnullable r" "bnullables rs" + shows "retrieve (AALTs bs (r # rs)) (mkeps (erase (AALTs bs (r # rs)))) + = retrieve (AALTs bs rs) (mkeps (erase (AALTs bs rs)))" + using assms + apply(induct rs arbitrary: r bs) + apply(auto) + using bnullable_correctness apply blast + apply(case_tac rs) + apply(auto) + using bnullable_correctness apply blast + apply(case_tac rs) + apply(auto) + done + +lemma bmkeps_retrieve_AALTs: + assumes "\r \ set rs. bnullable r \ bmkeps r = retrieve r (mkeps (erase r))" + "bnullables rs" + shows "bs @ bmkepss rs = retrieve (AALTs bs rs) (mkeps (erase (AALTs bs rs)))" + using assms + apply(induct rs arbitrary: bs) + apply(auto) + using retrieve_AALTs_bnullable1 apply presburger + apply (metis retrieve_AALTs_bnullable2) + apply (simp add: retrieve_AALTs_bnullable1) + by (metis retrieve_AALTs_bnullable2) + + +lemma bmkeps_retrieve: + assumes "bnullable r" + shows "bmkeps r = retrieve r (mkeps (erase r))" + using assms + apply(induct r) + apply(auto) + using bmkeps_retrieve_AALTs by auto + +lemma bder_retrieve: + assumes "\ v : der c (erase r)" + shows "retrieve (bder c r) v = retrieve r (injval (erase r) c v)" + using assms + apply(induct r arbitrary: v rule: erase.induct) + using Prf_elims(1) apply auto[1] + using Prf_elims(1) apply auto[1] + apply(auto)[1] + apply (metis Prf_elims(4) injval.simps(1) retrieve.simps(1) retrieve.simps(2)) + using Prf_elims(1) apply blast + (* AALTs case *) + apply(simp) + apply(erule Prf_elims) + apply(simp) + apply(simp) + apply(rename_tac "r\<^sub>1" "r\<^sub>2" rs v) + apply(erule Prf_elims) + apply(simp) + apply(simp) + apply(case_tac rs) + apply(simp) + apply(simp) + using Prf_elims(3) apply fastforce + (* ASEQ case *) + apply(simp) + apply(case_tac "nullable (erase r1)") + apply(simp) + apply(erule Prf_elims) + using Prf_elims(2) bnullable_correctness apply force + apply (simp add: bmkeps_retrieve bnullable_correctness retrieve_fuse2) + apply (simp add: bmkeps_retrieve bnullable_correctness retrieve_fuse2) + using Prf_elims(2) apply force + (* ASTAR case *) + apply(rename_tac bs r v) + apply(simp) + apply(erule Prf_elims) + apply(clarify) + apply(erule Prf_elims) + apply(clarify) + by (simp add: retrieve_fuse2) + + +lemma MAIN_decode: + assumes "\ v : ders s r" + shows "Some (flex r id s v) = decode (retrieve (bders (intern r) s) v) r" + using assms +proof (induct s arbitrary: v rule: rev_induct) + case Nil + have "\ v : ders [] r" by fact + then have "\ v : r" by simp + then have "Some v = decode (retrieve (intern r) v) r" + using decode_code retrieve_code by auto + then show "Some (flex r id [] v) = decode (retrieve (bders (intern r) []) v) r" + by simp +next + case (snoc c s v) + have IH: "\v. \ v : ders s r \ + Some (flex r id s v) = decode (retrieve (bders (intern r) s) v) r" by fact + have asm: "\ v : ders (s @ [c]) r" by fact + then have asm2: "\ injval (ders s r) c v : ders s r" + by (simp add: Prf_injval ders_append) + have "Some (flex r id (s @ [c]) v) = Some (flex r id s (injval (ders s r) c v))" + by (simp add: flex_append) + also have "... = decode (retrieve (bders (intern r) s) (injval (ders s r) c v)) r" + using asm2 IH by simp + also have "... = decode (retrieve (bder c (bders (intern r) s)) v) r" + using asm by (simp_all add: bder_retrieve ders_append) + finally show "Some (flex r id (s @ [c]) v) = + decode (retrieve (bders (intern r) (s @ [c])) v) r" by (simp add: bders_append) +qed + +definition blexer where + "blexer r s \ if bnullable (bders (intern r) s) then + decode (bmkeps (bders (intern r) s)) r else None" + +lemma blexer_correctness: + shows "blexer r s = lexer r s" +proof - + { define bds where "bds \ bders (intern r) s" + define ds where "ds \ ders s r" + assume asm: "nullable ds" + have era: "erase bds = ds" + unfolding ds_def bds_def by simp + have mke: "\ mkeps ds : ds" + using asm by (simp add: mkeps_nullable) + have "decode (bmkeps bds) r = decode (retrieve bds (mkeps ds)) r" + using bmkeps_retrieve + using asm era + using bnullable_correctness by force + also have "... = Some (flex r id s (mkeps ds))" + using mke by (simp_all add: MAIN_decode ds_def bds_def) + finally have "decode (bmkeps bds) r = Some (flex r id s (mkeps ds))" + unfolding bds_def ds_def . + } + then show "blexer r s = lexer r s" + unfolding blexer_def lexer_flex + by (auto simp add: bnullable_correctness[symmetric]) +qed + + +unused_thms + +end diff -r f9cdc295ccf7 -r f493a20feeb3 thys3/src/BlexerSimp.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys3/src/BlexerSimp.thy Sat Apr 30 00:50:08 2022 +0100 @@ -0,0 +1,617 @@ +theory BlexerSimp + imports Blexer +begin + +fun distinctWith :: "'a list \ ('a \ 'a \ bool) \ 'a set \ 'a list" + where + "distinctWith [] eq acc = []" +| "distinctWith (x # xs) eq acc = + (if (\ y \ acc. eq x y) then distinctWith xs eq acc + else x # (distinctWith xs eq ({x} \ acc)))" + + +fun eq1 ("_ ~1 _" [80, 80] 80) where + "AZERO ~1 AZERO = True" +| "(AONE bs1) ~1 (AONE bs2) = True" +| "(ACHAR bs1 c) ~1 (ACHAR bs2 d) = (if c = d then True else False)" +| "(ASEQ bs1 ra1 ra2) ~1 (ASEQ bs2 rb1 rb2) = (ra1 ~1 rb1 \ ra2 ~1 rb2)" +| "(AALTs bs1 []) ~1 (AALTs bs2 []) = True" +| "(AALTs bs1 (r1 # rs1)) ~1 (AALTs bs2 (r2 # rs2)) = (r1 ~1 r2 \ (AALTs bs1 rs1) ~1 (AALTs bs2 rs2))" +| "(ASTAR bs1 r1) ~1 (ASTAR bs2 r2) = r1 ~1 r2" +| "_ ~1 _ = False" + + + +lemma eq1_L: + assumes "x ~1 y" + shows "L (erase x) = L (erase y)" + using assms + apply(induct rule: eq1.induct) + apply(auto elim: eq1.elims) + apply presburger + done + +fun flts :: "arexp list \ arexp list" + where + "flts [] = []" +| "flts (AZERO # rs) = flts rs" +| "flts ((AALTs bs rs1) # rs) = (map (fuse bs) rs1) @ flts rs" +| "flts (r1 # rs) = r1 # flts rs" + + + +fun bsimp_ASEQ :: "bit list \ arexp \ arexp \ arexp" + where + "bsimp_ASEQ _ AZERO _ = AZERO" +| "bsimp_ASEQ _ _ AZERO = AZERO" +| "bsimp_ASEQ bs1 (AONE bs2) r2 = fuse (bs1 @ bs2) r2" +| "bsimp_ASEQ bs1 r1 r2 = ASEQ bs1 r1 r2" + +lemma bsimp_ASEQ0[simp]: + shows "bsimp_ASEQ bs r1 AZERO = AZERO" + by (case_tac r1)(simp_all) + +lemma bsimp_ASEQ1: + assumes "r1 \ AZERO" "r2 \ AZERO" "\bs. r1 = AONE bs" + shows "bsimp_ASEQ bs r1 r2 = ASEQ bs r1 r2" + using assms + apply(induct bs r1 r2 rule: bsimp_ASEQ.induct) + apply(auto) + done + +lemma bsimp_ASEQ2[simp]: + shows "bsimp_ASEQ bs1 (AONE bs2) r2 = fuse (bs1 @ bs2) r2" + by (case_tac r2) (simp_all) + + +fun bsimp_AALTs :: "bit list \ arexp list \ arexp" + where + "bsimp_AALTs _ [] = AZERO" +| "bsimp_AALTs bs1 [r] = fuse bs1 r" +| "bsimp_AALTs bs1 rs = AALTs bs1 rs" + + + + +fun bsimp :: "arexp \ arexp" + where + "bsimp (ASEQ bs1 r1 r2) = bsimp_ASEQ bs1 (bsimp r1) (bsimp r2)" +| "bsimp (AALTs bs1 rs) = bsimp_AALTs bs1 (distinctWith (flts (map bsimp rs)) eq1 {}) " +| "bsimp r = r" + + +fun + bders_simp :: "arexp \ string \ arexp" +where + "bders_simp r [] = r" +| "bders_simp r (c # s) = bders_simp (bsimp (bder c r)) s" + +definition blexer_simp where + "blexer_simp r s \ if bnullable (bders_simp (intern r) s) then + decode (bmkeps (bders_simp (intern r) s)) r else None" + + + +lemma bders_simp_append: + shows "bders_simp r (s1 @ s2) = bders_simp (bders_simp r s1) s2" + apply(induct s1 arbitrary: r s2) + apply(simp_all) + done + +lemma bmkeps_fuse: + assumes "bnullable r" + shows "bmkeps (fuse bs r) = bs @ bmkeps r" + using assms + by (induct r rule: bnullable.induct) (auto) + +lemma bmkepss_fuse: + assumes "bnullables rs" + shows "bmkepss (map (fuse bs) rs) = bs @ bmkepss rs" + using assms + apply(induct rs arbitrary: bs) + apply(auto simp add: bmkeps_fuse bnullable_fuse) + done + +lemma bder_fuse: + shows "bder c (fuse bs a) = fuse bs (bder c a)" + apply(induct a arbitrary: bs c) + apply(simp_all) + done + + + + +inductive + rrewrite:: "arexp \ arexp \ bool" ("_ \ _" [99, 99] 99) +and + srewrite:: "arexp list \ arexp list \ bool" (" _ s\ _" [100, 100] 100) +where + bs1: "ASEQ bs AZERO r2 \ AZERO" +| bs2: "ASEQ bs r1 AZERO \ AZERO" +| bs3: "ASEQ bs1 (AONE bs2) r \ fuse (bs1@bs2) r" +| bs4: "r1 \ r2 \ ASEQ bs r1 r3 \ ASEQ bs r2 r3" +| bs5: "r3 \ r4 \ ASEQ bs r1 r3 \ ASEQ bs r1 r4" +| bs6: "AALTs bs [] \ AZERO" +| bs7: "AALTs bs [r] \ fuse bs r" +| bs10: "rs1 s\ rs2 \ AALTs bs rs1 \ AALTs bs rs2" +| ss1: "[] s\ []" +| ss2: "rs1 s\ rs2 \ (r # rs1) s\ (r # rs2)" +| ss3: "r1 \ r2 \ (r1 # rs) s\ (r2 # rs)" +| ss4: "(AZERO # rs) s\ rs" +| ss5: "(AALTs bs1 rs1 # rsb) s\ ((map (fuse bs1) rs1) @ rsb)" +| ss6: "L (erase a2) \ L (erase a1) \ (rsa@[a1]@rsb@[a2]@rsc) s\ (rsa@[a1]@rsb@rsc)" + + +inductive + rrewrites:: "arexp \ arexp \ bool" ("_ \* _" [100, 100] 100) +where + rs1[intro, simp]:"r \* r" +| rs2[intro]: "\r1 \* r2; r2 \ r3\ \ r1 \* r3" + +inductive + srewrites:: "arexp list \ arexp list \ bool" ("_ s\* _" [100, 100] 100) +where + sss1[intro, simp]:"rs s\* rs" +| sss2[intro]: "\rs1 s\ rs2; rs2 s\* rs3\ \ rs1 s\* rs3" + + +lemma r_in_rstar : "r1 \ r2 \ r1 \* r2" + using rrewrites.intros(1) rrewrites.intros(2) by blast + +lemma rs_in_rstar: + shows "r1 s\ r2 \ r1 s\* r2" + using rrewrites.intros(1) rrewrites.intros(2) by blast + + +lemma rrewrites_trans[trans]: + assumes a1: "r1 \* r2" and a2: "r2 \* r3" + shows "r1 \* r3" + using a2 a1 + apply(induct r2 r3 arbitrary: r1 rule: rrewrites.induct) + apply(auto) + done + +lemma srewrites_trans[trans]: + assumes a1: "r1 s\* r2" and a2: "r2 s\* r3" + shows "r1 s\* r3" + using a1 a2 + apply(induct r1 r2 arbitrary: r3 rule: srewrites.induct) + apply(auto) + done + + + +lemma contextrewrites0: + "rs1 s\* rs2 \ AALTs bs rs1 \* AALTs bs rs2" + apply(induct rs1 rs2 rule: srewrites.inducts) + apply simp + using bs10 r_in_rstar rrewrites_trans by blast + +lemma contextrewrites1: + "r \* r' \ AALTs bs (r # rs) \* AALTs bs (r' # rs)" + apply(induct r r' rule: rrewrites.induct) + apply simp + using bs10 ss3 by blast + +lemma srewrite1: + shows "rs1 s\ rs2 \ (rs @ rs1) s\ (rs @ rs2)" + apply(induct rs) + apply(auto) + using ss2 by auto + +lemma srewrites1: + shows "rs1 s\* rs2 \ (rs @ rs1) s\* (rs @ rs2)" + apply(induct rs1 rs2 rule: srewrites.induct) + apply(auto) + using srewrite1 by blast + +lemma srewrite2: + shows "r1 \ r2 \ True" + and "rs1 s\ rs2 \ (rs1 @ rs) s\* (rs2 @ rs)" + apply(induct rule: rrewrite_srewrite.inducts) + apply(auto) + apply (metis append_Cons append_Nil srewrites1) + apply(meson srewrites.simps ss3) + apply (meson srewrites.simps ss4) + apply (meson srewrites.simps ss5) + by (metis append_Cons append_Nil srewrites.simps ss6) + + +lemma srewrites3: + shows "rs1 s\* rs2 \ (rs1 @ rs) s\* (rs2 @ rs)" + apply(induct rs1 rs2 arbitrary: rs rule: srewrites.induct) + apply(auto) + by (meson srewrite2(2) srewrites_trans) + +(* +lemma srewrites4: + assumes "rs3 s\* rs4" "rs1 s\* rs2" + shows "(rs1 @ rs3) s\* (rs2 @ rs4)" + using assms + apply(induct rs3 rs4 arbitrary: rs1 rs2 rule: srewrites.induct) + apply (simp add: srewrites3) + using srewrite1 by blast +*) + +lemma srewrites6: + assumes "r1 \* r2" + shows "[r1] s\* [r2]" + using assms + apply(induct r1 r2 rule: rrewrites.induct) + apply(auto) + by (meson srewrites.simps srewrites_trans ss3) + +lemma srewrites7: + assumes "rs3 s\* rs4" "r1 \* r2" + shows "(r1 # rs3) s\* (r2 # rs4)" + using assms + by (smt (verit, best) append_Cons append_Nil srewrites1 srewrites3 srewrites6 srewrites_trans) + +lemma ss6_stronger_aux: + shows "(rs1 @ rs2) s\* (rs1 @ distinctWith rs2 eq1 (set rs1))" + apply(induct rs2 arbitrary: rs1) + apply(auto) + prefer 2 + apply(drule_tac x="rs1 @ [a]" in meta_spec) + apply(simp) + apply(drule_tac x="rs1" in meta_spec) + apply(subgoal_tac "(rs1 @ a # rs2) s\* (rs1 @ rs2)") + using srewrites_trans apply blast + apply(subgoal_tac "\rs1a rs1b. rs1 = rs1a @ [x] @ rs1b") + prefer 2 + apply (simp add: split_list) + apply(erule exE)+ + apply(simp) + using eq1_L rs_in_rstar ss6 by force + +lemma ss6_stronger: + shows "rs1 s\* distinctWith rs1 eq1 {}" + by (metis append_Nil list.set(1) ss6_stronger_aux) + + +lemma rewrite_preserves_fuse: + shows "r2 \ r3 \ fuse bs r2 \ fuse bs r3" + and "rs2 s\ rs3 \ map (fuse bs) rs2 s\* map (fuse bs) rs3" +proof(induct rule: rrewrite_srewrite.inducts) + case (bs3 bs1 bs2 r) + then show ?case + by (metis fuse.simps(5) fuse_append rrewrite_srewrite.bs3) +next + case (bs7 bs r) + then show ?case + by (metis fuse.simps(4) fuse_append rrewrite_srewrite.bs7) +next + case (ss2 rs1 rs2 r) + then show ?case + using srewrites7 by force +next + case (ss3 r1 r2 rs) + then show ?case by (simp add: r_in_rstar srewrites7) +next + case (ss5 bs1 rs1 rsb) + then show ?case + apply(simp) + by (metis (mono_tags, lifting) comp_def fuse_append map_eq_conv rrewrite_srewrite.ss5 srewrites.simps) +next + case (ss6 a1 a2 rsa rsb rsc) + then show ?case + apply(simp only: map_append) + by (smt (verit, best) erase_fuse list.simps(8) list.simps(9) rrewrite_srewrite.ss6 srewrites.simps) +qed (auto intro: rrewrite_srewrite.intros) + + +lemma rewrites_fuse: + assumes "r1 \* r2" + shows "fuse bs r1 \* fuse bs r2" +using assms +apply(induction r1 r2 arbitrary: bs rule: rrewrites.induct) +apply(auto intro: rewrite_preserves_fuse rrewrites_trans) +done + + +lemma star_seq: + assumes "r1 \* r2" + shows "ASEQ bs r1 r3 \* ASEQ bs r2 r3" +using assms +apply(induct r1 r2 arbitrary: r3 rule: rrewrites.induct) +apply(auto intro: rrewrite_srewrite.intros) +done + +lemma star_seq2: + assumes "r3 \* r4" + shows "ASEQ bs r1 r3 \* ASEQ bs r1 r4" + using assms +apply(induct r3 r4 arbitrary: r1 rule: rrewrites.induct) +apply(auto intro: rrewrite_srewrite.intros) +done + +lemma continuous_rewrite: + assumes "r1 \* AZERO" + shows "ASEQ bs1 r1 r2 \* AZERO" +using assms bs1 star_seq by blast + +(* +lemma continuous_rewrite2: + assumes "r1 \* AONE bs" + shows "ASEQ bs1 r1 r2 \* (fuse (bs1 @ bs) r2)" + using assms by (meson bs3 rrewrites.simps star_seq) +*) + +lemma bsimp_aalts_simpcases: + shows "AONE bs \* bsimp (AONE bs)" + and "AZERO \* bsimp AZERO" + and "ACHAR bs c \* bsimp (ACHAR bs c)" + by (simp_all) + +lemma bsimp_AALTs_rewrites: + shows "AALTs bs1 rs \* bsimp_AALTs bs1 rs" + by (smt (verit) bs6 bs7 bsimp_AALTs.elims rrewrites.simps) + +lemma trivialbsimp_srewrites: + "\\x. x \ set rs \ x \* f x \ \ rs s\* (map f rs)" + apply(induction rs) + apply simp + apply(simp) + using srewrites7 by auto + + + +lemma fltsfrewrites: "rs s\* flts rs" + apply(induction rs rule: flts.induct) + apply(auto intro: rrewrite_srewrite.intros) + apply (meson srewrites.simps srewrites1 ss5) + using rs1 srewrites7 apply presburger + using srewrites7 apply force + apply (simp add: srewrites7) + by (simp add: srewrites7) + +lemma bnullable0: +shows "r1 \ r2 \ bnullable r1 = bnullable r2" + and "rs1 s\ rs2 \ bnullables rs1 = bnullables rs2" + apply(induct rule: rrewrite_srewrite.inducts) + apply(auto simp add: bnullable_fuse) + apply (meson UnCI bnullable_fuse imageI) + using bnullable_correctness nullable_correctness by blast + + +lemma rewritesnullable: + assumes "r1 \* r2" + shows "bnullable r1 = bnullable r2" +using assms + apply(induction r1 r2 rule: rrewrites.induct) + apply simp + using bnullable0(1) by auto + +lemma rewrite_bmkeps_aux: + shows "r1 \ r2 \ (bnullable r1 \ bnullable r2 \ bmkeps r1 = bmkeps r2)" + and "rs1 s\ rs2 \ (bnullables rs1 \ bnullables rs2 \ bmkepss rs1 = bmkepss rs2)" +proof (induct rule: rrewrite_srewrite.inducts) + case (bs3 bs1 bs2 r) + then show ?case by (simp add: bmkeps_fuse) +next + case (bs7 bs r) + then show ?case by (simp add: bmkeps_fuse) +next + case (ss3 r1 r2 rs) + then show ?case + using bmkepss.simps bnullable0(1) by presburger +next + case (ss5 bs1 rs1 rsb) + then show ?case + by (simp add: bmkepss1 bmkepss2 bmkepss_fuse bnullable_fuse) +next + case (ss6 a1 a2 rsa rsb rsc) + then show ?case + by (smt (verit, best) Nil_is_append_conv bmkepss1 bmkepss2 bnullable_correctness in_set_conv_decomp list.distinct(1) list.set_intros(1) nullable_correctness set_ConsD subsetD) +qed (auto) + +lemma rewrites_bmkeps: + assumes "r1 \* r2" "bnullable r1" + shows "bmkeps r1 = bmkeps r2" + using assms +proof(induction r1 r2 rule: rrewrites.induct) + case (rs1 r) + then show "bmkeps r = bmkeps r" by simp +next + case (rs2 r1 r2 r3) + then have IH: "bmkeps r1 = bmkeps r2" by simp + have a1: "bnullable r1" by fact + have a2: "r1 \* r2" by fact + have a3: "r2 \ r3" by fact + have a4: "bnullable r2" using a1 a2 by (simp add: rewritesnullable) + then have "bmkeps r2 = bmkeps r3" + using a3 bnullable0(1) rewrite_bmkeps_aux(1) by blast + then show "bmkeps r1 = bmkeps r3" using IH by simp +qed + + +lemma rewrites_to_bsimp: + shows "r \* bsimp r" +proof (induction r rule: bsimp.induct) + case (1 bs1 r1 r2) + have IH1: "r1 \* bsimp r1" by fact + have IH2: "r2 \* bsimp r2" by fact + { assume as: "bsimp r1 = AZERO \ bsimp r2 = AZERO" + with IH1 IH2 have "r1 \* AZERO \ r2 \* AZERO" by auto + then have "ASEQ bs1 r1 r2 \* AZERO" + by (metis bs2 continuous_rewrite rrewrites.simps star_seq2) + then have "ASEQ bs1 r1 r2 \* bsimp (ASEQ bs1 r1 r2)" using as by auto + } + moreover + { assume "\bs. bsimp r1 = AONE bs" + then obtain bs where as: "bsimp r1 = AONE bs" by blast + with IH1 have "r1 \* AONE bs" by simp + then have "ASEQ bs1 r1 r2 \* fuse (bs1 @ bs) r2" using bs3 star_seq by blast + with IH2 have "ASEQ bs1 r1 r2 \* fuse (bs1 @ bs) (bsimp r2)" + using rewrites_fuse by (meson rrewrites_trans) + then have "ASEQ bs1 r1 r2 \* bsimp (ASEQ bs1 (AONE bs) r2)" by simp + then have "ASEQ bs1 r1 r2 \* bsimp (ASEQ bs1 r1 r2)" by (simp add: as) + } + moreover + { assume as1: "bsimp r1 \ AZERO" "bsimp r2 \ AZERO" and as2: "(\bs. bsimp r1 = AONE bs)" + then have "bsimp_ASEQ bs1 (bsimp r1) (bsimp r2) = ASEQ bs1 (bsimp r1) (bsimp r2)" + by (simp add: bsimp_ASEQ1) + then have "ASEQ bs1 r1 r2 \* bsimp_ASEQ bs1 (bsimp r1) (bsimp r2)" using as1 as2 IH1 IH2 + by (metis rrewrites_trans star_seq star_seq2) + then have "ASEQ bs1 r1 r2 \* bsimp (ASEQ bs1 r1 r2)" by simp + } + ultimately show "ASEQ bs1 r1 r2 \* bsimp (ASEQ bs1 r1 r2)" by blast +next + case (2 bs1 rs) + have IH: "\x. x \ set rs \ x \* bsimp x" by fact + then have "rs s\* (map bsimp rs)" by (simp add: trivialbsimp_srewrites) + also have "... s\* flts (map bsimp rs)" by (simp add: fltsfrewrites) + also have "... s\* distinctWith (flts (map bsimp rs)) eq1 {}" by (simp add: ss6_stronger) + finally have "AALTs bs1 rs \* AALTs bs1 (distinctWith (flts (map bsimp rs)) eq1 {})" + using contextrewrites0 by auto + also have "... \* bsimp_AALTs bs1 (distinctWith (flts (map bsimp rs)) eq1 {})" + by (simp add: bsimp_AALTs_rewrites) + finally show "AALTs bs1 rs \* bsimp (AALTs bs1 rs)" by simp +qed (simp_all) + + +lemma to_zero_in_alt: + shows "AALT bs (ASEQ [] AZERO r) r2 \ AALT bs AZERO r2" + by (simp add: bs1 bs10 ss3) + + + +lemma bder_fuse_list: + shows "map (bder c \ fuse bs1) rs1 = map (fuse bs1 \ bder c) rs1" + apply(induction rs1) + apply(simp_all add: bder_fuse) + done + +lemma map1: + shows "(map f [a]) = [f a]" + by (simp) + +lemma rewrite_preserves_bder: + shows "r1 \ r2 \ (bder c r1) \* (bder c r2)" + and "rs1 s\ rs2 \ map (bder c) rs1 s\* map (bder c) rs2" +proof(induction rule: rrewrite_srewrite.inducts) + case (bs1 bs r2) + then show ?case + by (simp add: continuous_rewrite) +next + case (bs2 bs r1) + then show ?case + apply(auto) + apply (meson bs6 contextrewrites0 rrewrite_srewrite.bs2 rs2 ss3 ss4 sss1 sss2) + by (simp add: r_in_rstar rrewrite_srewrite.bs2) +next + case (bs3 bs1 bs2 r) + then show ?case + apply(simp) + + by (metis (no_types, lifting) bder_fuse bs10 bs7 fuse_append rrewrites.simps ss4 to_zero_in_alt) +next + case (bs4 r1 r2 bs r3) + have as: "r1 \ r2" by fact + have IH: "bder c r1 \* bder c r2" by fact + from as IH show "bder c (ASEQ bs r1 r3) \* bder c (ASEQ bs r2 r3)" + by (metis bder.simps(5) bnullable0(1) contextrewrites1 rewrite_bmkeps_aux(1) star_seq) +next + case (bs5 r3 r4 bs r1) + have as: "r3 \ r4" by fact + have IH: "bder c r3 \* bder c r4" by fact + from as IH show "bder c (ASEQ bs r1 r3) \* bder c (ASEQ bs r1 r4)" + apply(simp) + apply(auto) + using contextrewrites0 r_in_rstar rewrites_fuse srewrites6 srewrites7 star_seq2 apply presburger + using star_seq2 by blast +next + case (bs6 bs) + then show ?case + using rrewrite_srewrite.bs6 by force +next + case (bs7 bs r) + then show ?case + by (simp add: bder_fuse r_in_rstar rrewrite_srewrite.bs7) +next + case (bs10 rs1 rs2 bs) + then show ?case + using contextrewrites0 by force +next + case ss1 + then show ?case by simp +next + case (ss2 rs1 rs2 r) + then show ?case + by (simp add: srewrites7) +next + case (ss3 r1 r2 rs) + then show ?case + by (simp add: srewrites7) +next + case (ss4 rs) + then show ?case + using rrewrite_srewrite.ss4 by fastforce +next + case (ss5 bs1 rs1 rsb) + then show ?case + apply(simp) + using bder_fuse_list map_map rrewrite_srewrite.ss5 srewrites.simps by blast +next + case (ss6 a1 a2 bs rsa rsb) + then show ?case + apply(simp only: map_append map1) + apply(rule srewrites_trans) + apply(rule rs_in_rstar) + apply(rule_tac rrewrite_srewrite.ss6) + using Der_def der_correctness apply auto[1] + by blast +qed + +lemma rewrites_preserves_bder: + assumes "r1 \* r2" + shows "bder c r1 \* bder c r2" +using assms +apply(induction r1 r2 rule: rrewrites.induct) +apply(simp_all add: rewrite_preserves_bder rrewrites_trans) +done + + +lemma central: + shows "bders r s \* bders_simp r s" +proof(induct s arbitrary: r rule: rev_induct) + case Nil + then show "bders r [] \* bders_simp r []" by simp +next + case (snoc x xs) + have IH: "\r. bders r xs \* bders_simp r xs" by fact + have "bders r (xs @ [x]) = bders (bders r xs) [x]" by (simp add: bders_append) + also have "... \* bders (bders_simp r xs) [x]" using IH + by (simp add: rewrites_preserves_bder) + also have "... \* bders_simp (bders_simp r xs) [x]" using IH + by (simp add: rewrites_to_bsimp) + finally show "bders r (xs @ [x]) \* bders_simp r (xs @ [x])" + by (simp add: bders_simp_append) +qed + +lemma main_aux: + assumes "bnullable (bders r s)" + shows "bmkeps (bders r s) = bmkeps (bders_simp r s)" +proof - + have "bders r s \* bders_simp r s" by (rule central) + then + show "bmkeps (bders r s) = bmkeps (bders_simp r s)" using assms + by (rule rewrites_bmkeps) +qed + + + + +theorem main_blexer_simp: + shows "blexer r s = blexer_simp r s" + unfolding blexer_def blexer_simp_def + by (metis central main_aux rewritesnullable) + +theorem blexersimp_correctness: + shows "lexer r s = blexer_simp r s" + using blexer_correctness main_blexer_simp by simp + + +unused_thms + +end diff -r f9cdc295ccf7 -r f493a20feeb3 thys3/src/ClosedForms.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys3/src/ClosedForms.thy Sat Apr 30 00:50:08 2022 +0100 @@ -0,0 +1,1682 @@ +theory ClosedForms + imports "BasicIdentities" +begin + +lemma flts_middle0: + shows "rflts (rsa @ RZERO # rsb) = rflts (rsa @ rsb)" + apply(induct rsa) + apply simp + by (metis append_Cons rflts.simps(2) rflts.simps(3) rflts_def_idiot) + + + +lemma simp_flatten_aux0: + shows "rsimp (RALTS rs) = rsimp (RALTS (map rsimp rs))" + by (metis append_Nil head_one_more_simp identity_wwo0 list.simps(8) rdistinct.simps(1) rflts.simps(1) rsimp.simps(2) rsimp_ALTs.simps(1) rsimp_ALTs.simps(3) simp_flatten spawn_simp_rsimpalts) + + +inductive + hrewrite:: "rrexp \ rrexp \ bool" ("_ h\ _" [99, 99] 99) +where + "RSEQ RZERO r2 h\ RZERO" +| "RSEQ r1 RZERO h\ RZERO" +| "RSEQ RONE r h\ r" +| "r1 h\ r2 \ RSEQ r1 r3 h\ RSEQ r2 r3" +| "r3 h\ r4 \ RSEQ r1 r3 h\ RSEQ r1 r4" +| "r h\ r' \ (RALTS (rs1 @ [r] @ rs2)) h\ (RALTS (rs1 @ [r'] @ rs2))" +(*context rule for eliminating 0, alts--corresponds to the recursive call flts r::rs = r::(flts rs)*) +| "RALTS (rsa @ [RZERO] @ rsb) h\ RALTS (rsa @ rsb)" +| "RALTS (rsa @ [RALTS rs1] @ rsb) h\ RALTS (rsa @ rs1 @ rsb)" +| "RALTS [] h\ RZERO" +| "RALTS [r] h\ r" +| "a1 = a2 \ RALTS (rsa@[a1]@rsb@[a2]@rsc) h\ RALTS (rsa @ [a1] @ rsb @ rsc)" + +inductive + hrewrites:: "rrexp \ rrexp \ bool" ("_ h\* _" [100, 100] 100) +where + rs1[intro, simp]:"r h\* r" +| rs2[intro]: "\r1 h\* r2; r2 h\ r3\ \ r1 h\* r3" + + +lemma hr_in_rstar : "r1 h\ r2 \ r1 h\* r2" + using hrewrites.intros(1) hrewrites.intros(2) by blast + +lemma hreal_trans[trans]: + assumes a1: "r1 h\* r2" and a2: "r2 h\* r3" + shows "r1 h\* r3" + using a2 a1 + apply(induct r2 r3 arbitrary: r1 rule: hrewrites.induct) + apply(auto) + done + +lemma hrewrites_seq_context: + shows "r1 h\* r2 \ RSEQ r1 r3 h\* RSEQ r2 r3" + apply(induct r1 r2 rule: hrewrites.induct) + apply simp + using hrewrite.intros(4) by blast + +lemma hrewrites_seq_context2: + shows "r1 h\* r2 \ RSEQ r0 r1 h\* RSEQ r0 r2" + apply(induct r1 r2 rule: hrewrites.induct) + apply simp + using hrewrite.intros(5) by blast + + +lemma hrewrites_seq_contexts: + shows "\r1 h\* r2; r3 h\* r4\ \ RSEQ r1 r3 h\* RSEQ r2 r4" + by (meson hreal_trans hrewrites_seq_context hrewrites_seq_context2) + + +lemma simp_removes_duplicate1: + shows " a \ set rsa \ rsimp (RALTS (rsa @ [a])) = rsimp (RALTS (rsa))" +and " rsimp (RALTS (a1 # rsa @ [a1])) = rsimp (RALTS (a1 # rsa))" + apply(induct rsa arbitrary: a1) + apply simp + apply simp + prefer 2 + apply(case_tac "a = aa") + apply simp + apply simp + apply (metis Cons_eq_appendI Cons_eq_map_conv distinct_removes_duplicate_flts list.set_intros(2)) + apply (metis append_Cons append_Nil distinct_removes_duplicate_flts list.set_intros(1) list.simps(8) list.simps(9)) + by (metis (mono_tags, lifting) append_Cons distinct_removes_duplicate_flts list.set_intros(1) list.simps(8) list.simps(9) map_append rsimp.simps(2)) + +lemma simp_removes_duplicate2: + shows "a \ set rsa \ rsimp (RALTS (rsa @ [a] @ rsb)) = rsimp (RALTS (rsa @ rsb))" + apply(induct rsb arbitrary: rsa) + apply simp + using distinct_removes_duplicate_flts apply auto[1] + by (metis append.assoc head_one_more_simp rsimp.simps(2) simp_flatten simp_removes_duplicate1(1)) + +lemma simp_removes_duplicate3: + shows "a \ set rsa \ rsimp (RALTS (rsa @ a # rsb)) = rsimp (RALTS (rsa @ rsb))" + using simp_removes_duplicate2 by auto + +(* +lemma distinct_removes_middle4: + shows "a \ set rsa \ rdistinct (rsa @ [a] @ rsb) rset = rdistinct (rsa @ rsb) rset" + using distinct_removes_middle(1) by fastforce +*) + +(* +lemma distinct_removes_middle_list: + shows "\a \ set x. a \ set rsa \ rdistinct (rsa @ x @ rsb) rset = rdistinct (rsa @ rsb) rset" + apply(induct x) + apply simp + by (simp add: distinct_removes_middle3) +*) + +inductive frewrite:: "rrexp list \ rrexp list \ bool" ("_ \f _" [10, 10] 10) + where + "(RZERO # rs) \f rs" +| "((RALTS rs) # rsa) \f (rs @ rsa)" +| "rs1 \f rs2 \ (r # rs1) \f (r # rs2)" + + +inductive + frewrites:: "rrexp list \ rrexp list \ bool" ("_ \f* _" [10, 10] 10) +where + [intro, simp]:"rs \f* rs" +| [intro]: "\rs1 \f* rs2; rs2 \f rs3\ \ rs1 \f* rs3" + +inductive grewrite:: "rrexp list \ rrexp list \ bool" ("_ \g _" [10, 10] 10) + where + "(RZERO # rs) \g rs" +| "((RALTS rs) # rsa) \g (rs @ rsa)" +| "rs1 \g rs2 \ (r # rs1) \g (r # rs2)" +| "rsa @ [a] @ rsb @ [a] @ rsc \g rsa @ [a] @ rsb @ rsc" + +lemma grewrite_variant1: + shows "a \ set rs1 \ rs1 @ a # rs \g rs1 @ rs" + apply (metis append.assoc append_Cons append_Nil grewrite.intros(4) split_list_first) + done + + +inductive + grewrites:: "rrexp list \ rrexp list \ bool" ("_ \g* _" [10, 10] 10) +where + [intro, simp]:"rs \g* rs" +| [intro]: "\rs1 \g* rs2; rs2 \g rs3\ \ rs1 \g* rs3" + + + +(* +inductive + frewrites2:: "rrexp list \ rrexp list \ bool" ("_ <\f* _" [10, 10] 10) +where + [intro]: "\rs1 \f* rs2; rs2 \f* rs1\ \ rs1 <\f* rs2" +*) + +lemma fr_in_rstar : "r1 \f r2 \ r1 \f* r2" + using frewrites.intros(1) frewrites.intros(2) by blast + +lemma freal_trans[trans]: + assumes a1: "r1 \f* r2" and a2: "r2 \f* r3" + shows "r1 \f* r3" + using a2 a1 + apply(induct r2 r3 arbitrary: r1 rule: frewrites.induct) + apply(auto) + done + + +lemma many_steps_later: "\r1 \f r2; r2 \f* r3 \ \ r1 \f* r3" + by (meson fr_in_rstar freal_trans) + + +lemma gr_in_rstar : "r1 \g r2 \ r1 \g* r2" + using grewrites.intros(1) grewrites.intros(2) by blast + +lemma greal_trans[trans]: + assumes a1: "r1 \g* r2" and a2: "r2 \g* r3" + shows "r1 \g* r3" + using a2 a1 + apply(induct r2 r3 arbitrary: r1 rule: grewrites.induct) + apply(auto) + done + + +lemma gmany_steps_later: "\r1 \g r2; r2 \g* r3 \ \ r1 \g* r3" + by (meson gr_in_rstar greal_trans) + +lemma gstar_rdistinct_general: + shows "rs1 @ rs \g* rs1 @ (rdistinct rs (set rs1))" + apply(induct rs arbitrary: rs1) + apply simp + apply(case_tac " a \ set rs1") + apply simp + apply(subgoal_tac "rs1 @ a # rs \g rs1 @ rs") + using gmany_steps_later apply auto[1] + apply (metis append.assoc append_Cons append_Nil grewrite.intros(4) split_list_first) + apply simp + apply(drule_tac x = "rs1 @ [a]" in meta_spec) + by simp + + +lemma gstar_rdistinct: + shows "rs \g* rdistinct rs {}" + apply(induct rs) + apply simp + by (metis append.left_neutral empty_set gstar_rdistinct_general) + + +lemma grewrite_append: + shows "\ rsa \g rsb \ \ rs @ rsa \g rs @ rsb" + apply(induct rs) + apply simp+ + using grewrite.intros(3) by blast + + + +lemma frewrites_cons: + shows "\ rsa \f* rsb \ \ r # rsa \f* r # rsb" + apply(induct rsa rsb rule: frewrites.induct) + apply simp + using frewrite.intros(3) by blast + + +lemma grewrites_cons: + shows "\ rsa \g* rsb \ \ r # rsa \g* r # rsb" + apply(induct rsa rsb rule: grewrites.induct) + apply simp + using grewrite.intros(3) by blast + + +lemma frewrites_append: + shows " \rsa \f* rsb\ \ (rs @ rsa) \f* (rs @ rsb)" + apply(induct rs) + apply simp + by (simp add: frewrites_cons) + +lemma grewrites_append: + shows " \rsa \g* rsb\ \ (rs @ rsa) \g* (rs @ rsb)" + apply(induct rs) + apply simp + by (simp add: grewrites_cons) + + +lemma grewrites_concat: + shows "\rs1 \g rs2; rsa \g* rsb \ \ (rs1 @ rsa) \g* (rs2 @ rsb)" + apply(induct rs1 rs2 rule: grewrite.induct) + apply(simp) + apply(subgoal_tac "(RZERO # rs @ rsa) \g (rs @ rsa)") + prefer 2 + using grewrite.intros(1) apply blast + apply(subgoal_tac "(rs @ rsa) \g* (rs @ rsb)") + using gmany_steps_later apply blast + apply (simp add: grewrites_append) + apply (metis append.assoc append_Cons grewrite.intros(2) grewrites_append gmany_steps_later) + using grewrites_cons apply auto + apply(subgoal_tac "rsaa @ a # rsba @ a # rsc @ rsa \g* rsaa @ a # rsba @ a # rsc @ rsb") + using grewrite.intros(4) grewrites.intros(2) apply force + using grewrites_append by auto + + +lemma grewritess_concat: + shows "\rsa \g* rsb; rsc \g* rsd \ \ (rsa @ rsc) \g* (rsb @ rsd)" + apply(induct rsa rsb rule: grewrites.induct) + apply(case_tac rs) + apply simp + using grewrites_append apply blast + by (meson greal_trans grewrites.simps grewrites_concat) + +fun alt_set:: "rrexp \ rrexp set" + where + "alt_set (RALTS rs) = set rs \ \ (alt_set ` (set rs))" +| "alt_set r = {r}" + + +lemma grewrite_cases_middle: + shows "rs1 \g rs2 \ +(\rsa rsb rsc. rs1 = (rsa @ [RALTS rsb] @ rsc) \ rs2 = (rsa @ rsb @ rsc)) \ +(\rsa rsc. rs1 = rsa @ [RZERO] @ rsc \ rs2 = rsa @ rsc) \ +(\rsa rsb rsc a. rs1 = rsa @ [a] @ rsb @ [a] @ rsc \ rs2 = rsa @ [a] @ rsb @ rsc)" + apply( induct rs1 rs2 rule: grewrite.induct) + apply simp + apply blast + apply (metis append_Cons append_Nil) + apply (metis append_Cons) + by blast + + +lemma good_singleton: + shows "good a \ nonalt a \ rflts [a] = [a]" + using good.simps(1) k0b by blast + + + + + + + +lemma all_that_same_elem: + shows "\ a \ rset; rdistinct rs {a} = []\ + \ rdistinct (rs @ rsb) rset = rdistinct rsb rset" + apply(induct rs) + apply simp + apply(subgoal_tac "aa = a") + apply simp + by (metis empty_iff insert_iff list.discI rdistinct.simps(2)) + +lemma distinct_early_app1: + shows "rset1 \ rset \ rdistinct rs rset = rdistinct (rdistinct rs rset1) rset" + apply(induct rs arbitrary: rset rset1) + apply simp + apply simp + apply(case_tac "a \ rset1") + apply simp + apply(case_tac "a \ rset") + apply simp+ + + apply blast + apply(case_tac "a \ rset1") + apply simp+ + apply(case_tac "a \ rset") + apply simp + apply (metis insert_subsetI) + apply simp + by (meson insert_mono) + + +lemma distinct_early_app: + shows " rdistinct (rs @ rsb) rset = rdistinct (rdistinct rs {} @ rsb) rset" + apply(induct rsb) + apply simp + using distinct_early_app1 apply blast + by (metis distinct_early_app1 distinct_once_enough empty_subsetI) + + +lemma distinct_eq_interesting1: + shows "a \ rset \ rdistinct (rs @ rsb) rset = rdistinct (rdistinct (a # rs) {} @ rsb) rset" + apply(subgoal_tac "rdistinct (rdistinct (a # rs) {} @ rsb) rset = rdistinct (rdistinct rs {} @ rsb) rset") + apply(simp only:) + using distinct_early_app apply blast + by (metis append_Cons distinct_early_app rdistinct.simps(2)) + + + +lemma good_flatten_aux_aux1: + shows "\ size rs \2; +\r \ set rs. good r \ r \ RZERO \ nonalt r; \r \ set rsb. good r \ r \ RZERO \ nonalt r \ + \ rdistinct (rs @ rsb) rset = + rdistinct (rflts [rsimp_ALTs (rdistinct rs {})] @ rsb) rset" + apply(induct rs arbitrary: rset) + apply simp + apply(case_tac "a \ rset") + apply simp + apply(case_tac "rdistinct rs {a}") + apply simp + apply(subst good_singleton) + apply force + apply simp + apply (meson all_that_same_elem) + apply(subgoal_tac "rflts [rsimp_ALTs (a # rdistinct rs {a})] = a # rdistinct rs {a} ") + prefer 2 + using k0a rsimp_ALTs.simps(3) apply presburger + apply(simp only:) + apply(subgoal_tac "rdistinct (rs @ rsb) rset = rdistinct ((rdistinct (a # rs) {}) @ rsb) rset ") + apply (metis insert_absorb insert_is_Un insert_not_empty rdistinct.simps(2)) + apply (meson distinct_eq_interesting1) + apply simp + apply(case_tac "rdistinct rs {a}") + prefer 2 + apply(subgoal_tac "rsimp_ALTs (a # rdistinct rs {a}) = RALTS (a # rdistinct rs {a})") + apply(simp only:) + apply(subgoal_tac "a # rdistinct (rs @ rsb) (insert a rset) = + rdistinct (rflts [RALTS (a # rdistinct rs {a})] @ rsb) rset") + apply simp + apply (metis append_Cons distinct_early_app empty_iff insert_is_Un k0a rdistinct.simps(2)) + using rsimp_ALTs.simps(3) apply presburger + by (metis Un_insert_left append_Cons distinct_early_app empty_iff good_singleton rdistinct.simps(2) rsimp_ALTs.simps(2) sup_bot_left) + + + + + +lemma good_flatten_aux_aux: + shows "\\a aa lista list. rs = a # list \ list = aa # lista; +\r \ set rs. good r \ r \ RZERO \ nonalt r; \r \ set rsb. good r \ r \ RZERO \ nonalt r \ + \ rdistinct (rs @ rsb) rset = + rdistinct (rflts [rsimp_ALTs (rdistinct rs {})] @ rsb) rset" + apply(erule exE)+ + apply(subgoal_tac "size rs \ 2") + apply (metis good_flatten_aux_aux1) + by (simp add: Suc_leI length_Cons less_add_Suc1) + + + +lemma good_flatten_aux: + shows " \\r\set rs. good r \ r = RZERO; \r\set rsa . good r \ r = RZERO; + \r\set rsb. good r \ r = RZERO; + rsimp (RALTS (rsa @ rs @ rsb)) = rsimp_ALTs (rdistinct (rflts (rsa @ rs @ rsb)) {}); + rsimp (RALTS (rsa @ [RALTS rs] @ rsb)) = + rsimp_ALTs (rdistinct (rflts (rsa @ [rsimp (RALTS rs)] @ rsb)) {}); + map rsimp rsa = rsa; map rsimp rsb = rsb; map rsimp rs = rs; + rdistinct (rflts rsa @ rflts rs @ rflts rsb) {} = + rdistinct (rflts rsa) {} @ rdistinct (rflts rs @ rflts rsb) (set (rflts rsa)); + rdistinct (rflts rsa @ rflts [rsimp (RALTS rs)] @ rflts rsb) {} = + rdistinct (rflts rsa) {} @ rdistinct (rflts [rsimp (RALTS rs)] @ rflts rsb) (set (rflts rsa))\ + \ rdistinct (rflts rs @ rflts rsb) rset = + rdistinct (rflts [rsimp (RALTS rs)] @ rflts rsb) rset" + apply simp + apply(case_tac "rflts rs ") + apply simp + apply(case_tac "list") + apply simp + apply(case_tac "a \ rset") + apply simp + apply (metis append.left_neutral append_Cons equals0D k0b list.set_intros(1) nonalt_flts_rd qqq1 rdistinct.simps(2)) + apply simp + apply (metis Un_insert_left append_Cons append_Nil ex_in_conv flts_single1 insertI1 list.simps(15) nonalt_flts_rd nonazero.elims(3) qqq1 rdistinct.simps(2) sup_bot_left) + apply(subgoal_tac "\r \ set (rflts rs). good r \ r \ RZERO \ nonalt r") + prefer 2 + apply (metis Diff_empty flts3 nonalt_flts_rd qqq1 rdistinct_set_equality1) + apply(subgoal_tac "\r \ set (rflts rsb). good r \ r \ RZERO \ nonalt r") + prefer 2 + apply (metis Diff_empty flts3 good.simps(1) nonalt_flts_rd rdistinct_set_equality1) + by (smt (verit, ccfv_threshold) good_flatten_aux_aux) + + + + +lemma good_flatten_middle: + shows "\\r \ set rs. good r \ r = RZERO; \r \ set rsa. good r \ r = RZERO; \r \ set rsb. good r \ r = RZERO\ \ +rsimp (RALTS (rsa @ rs @ rsb)) = rsimp (RALTS (rsa @ [RALTS rs] @ rsb))" + apply(subgoal_tac "rsimp (RALTS (rsa @ rs @ rsb)) = rsimp_ALTs (rdistinct (rflts (map rsimp rsa @ +map rsimp rs @ map rsimp rsb)) {})") + prefer 2 + apply simp + apply(simp only:) + apply(subgoal_tac "rsimp (RALTS (rsa @ [RALTS rs] @ rsb)) = rsimp_ALTs (rdistinct (rflts (map rsimp rsa @ +[rsimp (RALTS rs)] @ map rsimp rsb)) {})") + prefer 2 + apply simp + apply(simp only:) + apply(subgoal_tac "map rsimp rsa = rsa") + prefer 2 + apply (metis map_idI rsimp.simps(3) test) + apply(simp only:) + apply(subgoal_tac "map rsimp rsb = rsb") + prefer 2 + apply (metis map_idI rsimp.simps(3) test) + apply(simp only:) + apply(subst k00)+ + apply(subgoal_tac "map rsimp rs = rs") + apply(simp only:) + prefer 2 + apply (metis map_idI rsimp.simps(3) test) + apply(subgoal_tac "rdistinct (rflts rsa @ rflts rs @ rflts rsb) {} = +rdistinct (rflts rsa) {} @ rdistinct (rflts rs @ rflts rsb) (set (rflts rsa))") + apply(simp only:) + prefer 2 + using rdistinct_concat_general apply blast + apply(subgoal_tac "rdistinct (rflts rsa @ rflts [rsimp (RALTS rs)] @ rflts rsb) {} = +rdistinct (rflts rsa) {} @ rdistinct (rflts [rsimp (RALTS rs)] @ rflts rsb) (set (rflts rsa))") + apply(simp only:) + prefer 2 + using rdistinct_concat_general apply blast + apply(subgoal_tac "rdistinct (rflts rs @ rflts rsb) (set (rflts rsa)) = + rdistinct (rflts [rsimp (RALTS rs)] @ rflts rsb) (set (rflts rsa))") + apply presburger + using good_flatten_aux by blast + + +lemma simp_flatten3: + shows "rsimp (RALTS (rsa @ [RALTS rs] @ rsb)) = rsimp (RALTS (rsa @ rs @ rsb))" + apply(subgoal_tac "rsimp (RALTS (rsa @ [RALTS rs] @ rsb)) = + rsimp (RALTS (map rsimp rsa @ [rsimp (RALTS rs)] @ map rsimp rsb)) ") + prefer 2 + apply (metis append.left_neutral append_Cons list.simps(9) map_append simp_flatten_aux0) + apply (simp only:) + apply(subgoal_tac "rsimp (RALTS (rsa @ rs @ rsb)) = +rsimp (RALTS (map rsimp rsa @ map rsimp rs @ map rsimp rsb))") + prefer 2 + apply (metis map_append simp_flatten_aux0) + apply(simp only:) + apply(subgoal_tac "rsimp (RALTS (map rsimp rsa @ map rsimp rs @ map rsimp rsb)) = + rsimp (RALTS (map rsimp rsa @ [RALTS (map rsimp rs)] @ map rsimp rsb))") + + apply (metis (no_types, lifting) head_one_more_simp map_append simp_flatten_aux0) + apply(subgoal_tac "\r \ set (map rsimp rsa). good r \ r = RZERO") + apply(subgoal_tac "\r \ set (map rsimp rs). good r \ r = RZERO") + apply(subgoal_tac "\r \ set (map rsimp rsb). good r \ r = RZERO") + + using good_flatten_middle apply presburger + + apply (simp add: good1) + apply (simp add: good1) + apply (simp add: good1) + + done + + + + + +lemma grewrite_equal_rsimp: + shows "rs1 \g rs2 \ rsimp (RALTS rs1) = rsimp (RALTS rs2)" + apply(frule grewrite_cases_middle) + apply(case_tac "(\rsa rsb rsc. rs1 = rsa @ [RALTS rsb] @ rsc \ rs2 = rsa @ rsb @ rsc)") + using simp_flatten3 apply auto[1] + apply(case_tac "(\rsa rsc. rs1 = rsa @ [RZERO] @ rsc \ rs2 = rsa @ rsc)") + apply (metis (mono_tags, opaque_lifting) append_Cons append_Nil list.set_intros(1) list.simps(9) rflts.simps(2) rsimp.simps(2) rsimp.simps(3) simp_removes_duplicate3) + by (smt (verit) append.assoc append_Cons append_Nil in_set_conv_decomp simp_removes_duplicate3) + + +lemma grewrites_equal_rsimp: + shows "rs1 \g* rs2 \ rsimp (RALTS rs1) = rsimp (RALTS rs2)" + apply (induct rs1 rs2 rule: grewrites.induct) + apply simp + using grewrite_equal_rsimp by presburger + + + +lemma grewrites_last: + shows "r # [RALTS rs] \g* r # rs" + by (metis gr_in_rstar grewrite.intros(2) grewrite.intros(3) self_append_conv) + +lemma simp_flatten2: + shows "rsimp (RALTS (r # [RALTS rs])) = rsimp (RALTS (r # rs))" + using grewrites_equal_rsimp grewrites_last by blast + + +lemma frewrites_alt: + shows "rs1 \f* rs2 \ (RALT r1 r2) # rs1 \f* r1 # r2 # rs2" + by (metis Cons_eq_appendI append_self_conv2 frewrite.intros(2) frewrites_cons many_steps_later) + +lemma early_late_der_frewrites: + shows "map (rder x) (rflts rs) \f* rflts (map (rder x) rs)" + apply(induct rs) + apply simp + apply(case_tac a) + apply simp+ + using frewrite.intros(1) many_steps_later apply blast + apply(case_tac "x = x3") + apply simp + using frewrites_cons apply presburger + using frewrite.intros(1) many_steps_later apply fastforce + apply(case_tac "rnullable x41") + apply simp+ + apply (simp add: frewrites_alt) + apply (simp add: frewrites_cons) + apply (simp add: frewrites_append) + by (simp add: frewrites_cons) + + +lemma gstar0: + shows "rsa @ (rdistinct rs (set rsa)) \g* rsa @ (rdistinct rs (insert RZERO (set rsa)))" + apply(induct rs arbitrary: rsa) + apply simp + apply(case_tac "a = RZERO") + apply simp + + using gr_in_rstar grewrite.intros(1) grewrites_append apply presburger + apply(case_tac "a \ set rsa") + apply simp+ + apply(drule_tac x = "rsa @ [a]" in meta_spec) + by simp + +lemma grewrite_rdistinct_aux: + shows "rs @ rdistinct rsa rset \g* rs @ rdistinct rsa (rset \ set rs)" + apply(induct rsa arbitrary: rs rset) + apply simp + apply(case_tac " a \ rset") + apply simp + apply(case_tac "a \ set rs") + apply simp + apply (metis Un_insert_left Un_insert_right gmany_steps_later grewrite_variant1 insert_absorb) + apply simp + apply(drule_tac x = "rs @ [a]" in meta_spec) + by (metis Un_insert_left Un_insert_right append.assoc append.right_neutral append_Cons append_Nil insert_absorb2 list.simps(15) set_append) + + +lemma flts_gstar: + shows "rs \g* rflts rs" + apply(induct rs) + apply simp + apply(case_tac "a = RZERO") + apply simp + using gmany_steps_later grewrite.intros(1) apply blast + apply(case_tac "\rsa. a = RALTS rsa") + apply(erule exE) + apply simp + apply (meson grewrite.intros(2) grewrites.simps grewrites_cons) + by (simp add: grewrites_cons rflts_def_idiot) + +lemma more_distinct1: + shows " \\rsb rset rset2. + rset2 \ set rsb \ rsb @ rdistinct rs rset \g* rsb @ rdistinct rs (rset \ rset2); + rset2 \ set rsb; a \ rset; a \ rset2\ + \ rsb @ a # rdistinct rs (insert a rset) \g* rsb @ rdistinct rs (rset \ rset2)" + apply(subgoal_tac "rsb @ a # rdistinct rs (insert a rset) \g* rsb @ rdistinct rs (insert a rset)") + apply(subgoal_tac "rsb @ rdistinct rs (insert a rset) \g* rsb @ rdistinct rs (rset \ rset2)") + apply (meson greal_trans) + apply (metis Un_iff Un_insert_left insert_absorb) + by (simp add: gr_in_rstar grewrite_variant1 in_mono) + + + + + +lemma frewrite_rd_grewrites_aux: + shows " RALTS rs \ set rsb \ + rsb @ + RALTS rs # + rdistinct rsa + (insert (RALTS rs) + (set rsb)) \g* rflts rsb @ + rdistinct rs (set rsb) @ rdistinct rsa (set rs \ set rsb \ {RALTS rs})" + + apply simp + apply(subgoal_tac "rsb @ + RALTS rs # + rdistinct rsa + (insert (RALTS rs) + (set rsb)) \g* rsb @ + rs @ + rdistinct rsa + (insert (RALTS rs) + (set rsb)) ") + apply(subgoal_tac " rsb @ + rs @ + rdistinct rsa + (insert (RALTS rs) + (set rsb)) \g* + rsb @ + rdistinct rs (set rsb) @ + rdistinct rsa + (insert (RALTS rs) + (set rsb)) ") + apply (smt (verit, ccfv_SIG) Un_insert_left flts_gstar greal_trans grewrite_rdistinct_aux grewritess_concat inf_sup_aci(5) rdistinct_concat_general rdistinct_set_equality set_append) + apply (metis append_assoc grewrites.intros(1) grewritess_concat gstar_rdistinct_general) + by (simp add: gr_in_rstar grewrite.intros(2) grewrites_append) + + + + +lemma list_dlist_union: + shows "set rs \ set rsb \ set (rdistinct rs (set rsb))" + by (metis rdistinct_concat_general rdistinct_set_equality set_append sup_ge2) + +lemma r_finite1: + shows "r = RALTS (r # rs) = False" + apply(induct r) + apply simp+ + apply (metis list.set_intros(1)) + by blast + + + +lemma grewrite_singleton: + shows "[r] \g r # rs \ rs = []" + apply (induct "[r]" "r # rs" rule: grewrite.induct) + apply simp + apply (metis r_finite1) + using grewrite.simps apply blast + by simp + + + +lemma concat_rdistinct_equality1: + shows "rdistinct (rs @ rsa) rset = rdistinct rs rset @ rdistinct rsa (rset \ (set rs))" + apply(induct rs arbitrary: rsa rset) + apply simp + apply(case_tac "a \ rset") + apply simp + apply (simp add: insert_absorb) + by auto + + +lemma grewrites_rev_append: + shows "rs1 \g* rs2 \ rs1 @ [x] \g* rs2 @ [x]" + using grewritess_concat by auto + +lemma grewrites_inclusion: + shows "set rs \ set rs1 \ rs1 @ rs \g* rs1" + apply(induct rs arbitrary: rs1) + apply simp + by (meson gmany_steps_later grewrite_variant1 list.set_intros(1) set_subset_Cons subset_code(1)) + +lemma distinct_keeps_last: + shows "\x \ rset; x \ set xs \ \ rdistinct (xs @ [x]) rset = rdistinct xs rset @ [x]" + by (simp add: concat_rdistinct_equality1) + +lemma grewrites_shape2_aux: + shows " RALTS rs \ set rsb \ + rsb @ + rdistinct (rs @ rsa) + (set rsb) \g* rsb @ + rdistinct rs (set rsb) @ + rdistinct rsa (set rs \ set rsb \ {RALTS rs})" + apply(subgoal_tac " rdistinct (rs @ rsa) (set rsb) = rdistinct rs (set rsb) @ rdistinct rsa (set rs \ set rsb)") + apply (simp only:) + prefer 2 + apply (simp add: Un_commute concat_rdistinct_equality1) + apply(induct rsa arbitrary: rs rsb rule: rev_induct) + apply simp + apply(case_tac "x \ set rs") + apply (simp add: distinct_removes_middle3) + apply(case_tac "x = RALTS rs") + apply simp + apply(case_tac "x \ set rsb") + apply simp + apply (simp add: concat_rdistinct_equality1) + apply (simp add: concat_rdistinct_equality1) + apply simp + apply(drule_tac x = "rs " in meta_spec) + apply(drule_tac x = rsb in meta_spec) + apply simp + apply(subgoal_tac " rsb @ rdistinct rs (set rsb) @ rdistinct xs (set rs \ set rsb) \g* rsb @ rdistinct rs (set rsb) @ rdistinct xs (insert (RALTS rs) (set rs \ set rsb))") + prefer 2 + apply (simp add: concat_rdistinct_equality1) + apply(case_tac "x \ set xs") + apply simp + apply (simp add: distinct_removes_last) + apply(case_tac "x \ set rsb") + apply (smt (verit, ccfv_threshold) Un_iff append.right_neutral concat_rdistinct_equality1 insert_is_Un rdistinct.simps(2)) + apply(subgoal_tac "rsb @ rdistinct rs (set rsb) @ rdistinct (xs @ [x]) (set rs \ set rsb) = rsb @ rdistinct rs (set rsb) @ rdistinct xs (set rs \ set rsb) @ [x]") + apply(simp only:) + apply(case_tac "x = RALTS rs") + apply(subgoal_tac "rsb @ rdistinct rs (set rsb) @ rdistinct xs (set rs \ set rsb) @ [x] \g* rsb @ rdistinct rs (set rsb) @ rdistinct xs (set rs \ set rsb) @ rs") + apply(subgoal_tac "rsb @ rdistinct rs (set rsb) @ rdistinct xs (set rs \ set rsb) @ rs \g* rsb @ rdistinct rs (set rsb) @ rdistinct xs (set rs \ set rsb) ") + apply (smt (verit, ccfv_SIG) Un_insert_left append.right_neutral concat_rdistinct_equality1 greal_trans insert_iff rdistinct.simps(2)) + apply(subgoal_tac "set rs \ set ( rsb @ rdistinct rs (set rsb) @ rdistinct xs (set rs \ set rsb))") + apply (metis append.assoc grewrites_inclusion) + apply (metis Un_upper1 append.assoc dual_order.trans list_dlist_union set_append) + apply (metis append_Nil2 gr_in_rstar grewrite.intros(2) grewrite_append) + apply(subgoal_tac " rsb @ rdistinct rs (set rsb) @ rdistinct (xs @ [x]) (insert (RALTS rs) (set rs \ set rsb)) = rsb @ rdistinct rs (set rsb) @ rdistinct (xs) (insert (RALTS rs) (set rs \ set rsb)) @ [x]") + apply(simp only:) + apply (metis append.assoc grewrites_rev_append) + apply (simp add: insert_absorb) + apply (simp add: distinct_keeps_last)+ + done + +lemma grewrites_shape2: + shows " RALTS rs \ set rsb \ + rsb @ + rdistinct (rs @ rsa) + (set rsb) \g* rflts rsb @ + rdistinct rs (set rsb) @ + rdistinct rsa (set rs \ set rsb \ {RALTS rs})" + apply (meson flts_gstar greal_trans grewrites.simps grewrites_shape2_aux grewritess_concat) + done + +lemma rdistinct_add_acc: + shows "rset2 \ set rsb \ rsb @ rdistinct rs rset \g* rsb @ rdistinct rs (rset \ rset2)" + apply(induct rs arbitrary: rsb rset rset2) + apply simp + apply (case_tac "a \ rset") + apply simp + apply(case_tac "a \ rset2") + apply simp + apply (simp add: more_distinct1) + apply simp + apply(drule_tac x = "rsb @ [a]" in meta_spec) + by (metis Un_insert_left append.assoc append_Cons append_Nil set_append sup.coboundedI1) + + +lemma frewrite_fun1: + shows " RALTS rs \ set rsb \ + rsb @ rdistinct rsa (set rsb) \g* rflts rsb @ rdistinct rsa (set rsb \ set rs)" + apply(subgoal_tac "rsb @ rdistinct rsa (set rsb) \g* rflts rsb @ rdistinct rsa (set rsb)") + apply(subgoal_tac " set rs \ set (rflts rsb)") + prefer 2 + using spilled_alts_contained apply blast + apply(subgoal_tac "rflts rsb @ rdistinct rsa (set rsb) \g* rflts rsb @ rdistinct rsa (set rsb \ set rs)") + using greal_trans apply blast + using rdistinct_add_acc apply presburger + using flts_gstar grewritess_concat by auto + +lemma frewrite_rd_grewrites: + shows "rs1 \f rs2 \ +\rs3. (rs @ (rdistinct rs1 (set rs)) \g* rs3) \ (rs @ (rdistinct rs2 (set rs)) \g* rs3) " + apply(induct rs1 rs2 arbitrary: rs rule: frewrite.induct) + apply(rule_tac x = "rsa @ (rdistinct rs ({RZERO} \ set rsa))" in exI) + apply(rule conjI) + apply(case_tac "RZERO \ set rsa") + apply simp+ + using gstar0 apply fastforce + apply (simp add: gr_in_rstar grewrite.intros(1) grewrites_append) + apply (simp add: gstar0) + prefer 2 + apply(case_tac "r \ set rs") + apply simp + apply(drule_tac x = "rs @ [r]" in meta_spec) + apply(erule exE) + apply(rule_tac x = "rs3" in exI) + apply simp + apply(case_tac "RALTS rs \ set rsb") + apply simp + apply(rule_tac x = "rflts rsb @ rdistinct rsa (set rsb \ set rs)" in exI) + apply(rule conjI) + using frewrite_fun1 apply force + apply (metis frewrite_fun1 rdistinct_concat sup_ge2) + apply(simp) + apply(rule_tac x = + "rflts rsb @ + rdistinct rs (set rsb) @ + rdistinct rsa (set rs \ set rsb \ {RALTS rs})" in exI) + apply(rule conjI) + prefer 2 + using grewrites_shape2 apply force + using frewrite_rd_grewrites_aux by blast + + +lemma frewrite_simpeq2: + shows "rs1 \f rs2 \ rsimp (RALTS (rdistinct rs1 {})) = rsimp (RALTS (rdistinct rs2 {}))" + apply(subgoal_tac "\ rs3. (rdistinct rs1 {} \g* rs3) \ (rdistinct rs2 {} \g* rs3)") + using grewrites_equal_rsimp apply fastforce + by (metis append_self_conv2 frewrite_rd_grewrites list.set(1)) + + + + +(*a more refined notion of h\* is needed, +this lemma fails when rs1 contains some RALTS rs where elements +of rs appear in later parts of rs1, which will be picked up by rs2 +and deduplicated*) +lemma frewrites_simpeq: + shows "rs1 \f* rs2 \ + rsimp (RALTS (rdistinct rs1 {})) = rsimp (RALTS ( rdistinct rs2 {})) " + apply(induct rs1 rs2 rule: frewrites.induct) + apply simp + using frewrite_simpeq2 by presburger + + +lemma frewrite_single_step: + shows " rs2 \f rs3 \ rsimp (RALTS rs2) = rsimp (RALTS rs3)" + apply(induct rs2 rs3 rule: frewrite.induct) + apply simp + using simp_flatten apply blast + by (metis (no_types, opaque_lifting) list.simps(9) rsimp.simps(2) simp_flatten2) + +lemma grewrite_simpalts: + shows " rs2 \g rs3 \ rsimp (rsimp_ALTs rs2) = rsimp (rsimp_ALTs rs3)" + apply(induct rs2 rs3 rule : grewrite.induct) + using identity_wwo0 apply presburger + apply (metis frewrite.intros(1) frewrite_single_step identity_wwo0 rsimp_ALTs.simps(3) simp_flatten) + apply (smt (verit, ccfv_SIG) gmany_steps_later grewrites.intros(1) grewrites_cons grewrites_equal_rsimp identity_wwo0 rsimp_ALTs.simps(3)) + apply simp + apply(subst rsimp_alts_equal) + apply(case_tac "rsa = [] \ rsb = [] \ rsc = []") + apply(subgoal_tac "rsa @ a # rsb @ rsc = [a]") + apply (simp only:) + apply (metis append_Nil frewrite.intros(1) frewrite_single_step identity_wwo0 rsimp_ALTs.simps(3) simp_removes_duplicate1(2)) + apply simp + by (smt (verit, best) append.assoc append_Cons frewrite.intros(1) frewrite_single_step identity_wwo0 in_set_conv_decomp rsimp_ALTs.simps(3) simp_removes_duplicate3) + + +lemma grewrites_simpalts: + shows " rs2 \g* rs3 \ rsimp (rsimp_ALTs rs2) = rsimp (rsimp_ALTs rs3)" + apply(induct rs2 rs3 rule: grewrites.induct) + apply simp + using grewrite_simpalts by presburger + + +lemma simp_der_flts: + shows "rsimp (RALTS (rdistinct (map (rder x) (rflts rs)) {})) = + rsimp (RALTS (rdistinct (rflts (map (rder x) rs)) {}))" + apply(subgoal_tac "map (rder x) (rflts rs) \f* rflts (map (rder x) rs)") + using frewrites_simpeq apply presburger + using early_late_der_frewrites by auto + + +lemma simp_der_pierce_flts_prelim: + shows "rsimp (rsimp_ALTs (rdistinct (map (rder x) (rflts rs)) {})) + = rsimp (rsimp_ALTs (rdistinct (rflts (map (rder x) rs)) {}))" + by (metis append.right_neutral grewrite.intros(2) grewrite_simpalts rsimp_ALTs.simps(2) simp_der_flts) + + +lemma basic_regex_property1: + shows "rnullable r \ rsimp r \ RZERO" + apply(induct r rule: rsimp.induct) + apply(auto) + apply (metis idiot idiot2 rrexp.distinct(5)) + by (metis der_simp_nullability rnullable.simps(1) rnullable.simps(4) rsimp.simps(2)) + + +lemma inside_simp_seq_nullable: + shows +"\r1 r2. + \rsimp (rder x (rsimp r1)) = rsimp (rder x r1); rsimp (rder x (rsimp r2)) = rsimp (rder x r2); + rnullable r1\ + \ rsimp (rder x (rsimp_SEQ (rsimp r1) (rsimp r2))) = + rsimp_ALTs (rdistinct (rflts [rsimp_SEQ (rsimp (rder x r1)) (rsimp r2), rsimp (rder x r2)]) {})" + apply(case_tac "rsimp r1 = RONE") + apply(simp) + apply(subst basic_rsimp_SEQ_property1) + apply (simp add: idem_after_simp1) + apply(case_tac "rsimp r1 = RZERO") + + using basic_regex_property1 apply blast + apply(case_tac "rsimp r2 = RZERO") + + apply (simp add: basic_rsimp_SEQ_property3) + apply(subst idiot2) + apply simp+ + apply(subgoal_tac "rnullable (rsimp r1)") + apply simp + using rsimp_idem apply presburger + using der_simp_nullability by presburger + + + +lemma grewrite_ralts: + shows "rs \g rs' \ RALTS rs h\* RALTS rs'" + by (smt (verit) grewrite_cases_middle hr_in_rstar hrewrite.intros(11) hrewrite.intros(7) hrewrite.intros(8)) + +lemma grewrites_ralts: + shows "rs \g* rs' \ RALTS rs h\* RALTS rs'" + apply(induct rule: grewrites.induct) + apply simp + using grewrite_ralts hreal_trans by blast + + +lemma distinct_grewrites_subgoal1: + shows " + \rs1 \g* [a]; RALTS rs1 h\* a; [a] \g rs3\ \ RALTS rs1 h\* rsimp_ALTs rs3" + apply(subgoal_tac "RALTS rs1 h\* RALTS rs3") + apply (metis hrewrite.intros(10) hrewrite.intros(9) rs2 rsimp_ALTs.cases rsimp_ALTs.simps(1) rsimp_ALTs.simps(2) rsimp_ALTs.simps(3)) + apply(subgoal_tac "rs1 \g* rs3") + using grewrites_ralts apply blast + using grewrites.intros(2) by presburger + +lemma grewrites_ralts_rsimpalts: + shows "rs \g* rs' \ RALTS rs h\* rsimp_ALTs rs' " + apply(induct rs rs' rule: grewrites.induct) + apply(case_tac rs) + using hrewrite.intros(9) apply force + apply(case_tac list) + apply simp + using hr_in_rstar hrewrite.intros(10) rsimp_ALTs.simps(2) apply presburger + apply simp + apply(case_tac rs2) + apply simp + apply (metis grewrite.intros(3) grewrite_singleton rsimp_ALTs.simps(1)) + apply(case_tac list) + apply(simp) + using distinct_grewrites_subgoal1 apply blast + apply simp + apply(case_tac rs3) + apply simp + using grewrites_ralts hrewrite.intros(9) apply blast + by (metis (no_types, opaque_lifting) grewrite_ralts hr_in_rstar hreal_trans hrewrite.intros(10) neq_Nil_conv rsimp_ALTs.simps(2) rsimp_ALTs.simps(3)) + +lemma hrewrites_alts: + shows " r h\* r' \ (RALTS (rs1 @ [r] @ rs2)) h\* (RALTS (rs1 @ [r'] @ rs2))" + apply(induct r r' rule: hrewrites.induct) + apply simp + using hrewrite.intros(6) by blast + +inductive + srewritescf:: "rrexp list \ rrexp list \ bool" (" _ scf\* _" [100, 100] 100) +where + ss1: "[] scf\* []" +| ss2: "\r h\* r'; rs scf\* rs'\ \ (r#rs) scf\* (r'#rs')" + + +lemma srewritescf_alt: "rs1 scf\* rs2 \ (RALTS (rs@rs1)) h\* (RALTS (rs@rs2))" + + apply(induct rs1 rs2 arbitrary: rs rule: srewritescf.induct) + apply(rule rs1) + apply(drule_tac x = "rsa@[r']" in meta_spec) + apply simp + apply(rule hreal_trans) + prefer 2 + apply(assumption) + apply(drule hrewrites_alts) + by auto + + +corollary srewritescf_alt1: + assumes "rs1 scf\* rs2" + shows "RALTS rs1 h\* RALTS rs2" + using assms + by (metis append_Nil srewritescf_alt) + + + + +lemma trivialrsimp_srewrites: + "\\x. x \ set rs \ x h\* f x \ \ rs scf\* (map f rs)" + + apply(induction rs) + apply simp + apply(rule ss1) + by (metis insert_iff list.simps(15) list.simps(9) srewritescf.simps) + +lemma hrewrites_list: + shows +" (\xa. xa \ set x \ xa h\* rsimp xa) \ RALTS x h\* RALTS (map rsimp x)" + apply(induct x) + apply(simp)+ + by (simp add: srewritescf_alt1 ss2 trivialrsimp_srewrites) +(* apply(subgoal_tac "RALTS x h\* RALTS (map rsimp x)")*) + + +lemma hrewrite_simpeq: + shows "r1 h\ r2 \ rsimp r1 = rsimp r2" + apply(induct rule: hrewrite.induct) + apply simp+ + apply (simp add: basic_rsimp_SEQ_property3) + apply (simp add: basic_rsimp_SEQ_property1) + using rsimp.simps(1) apply presburger + apply simp+ + using flts_middle0 apply force + + + using simp_flatten3 apply presburger + + apply simp+ + apply (simp add: idem_after_simp1) + using grewrite.intros(4) grewrite_equal_rsimp by presburger + +lemma hrewrites_simpeq: + shows "r1 h\* r2 \ rsimp r1 = rsimp r2" + apply(induct rule: hrewrites.induct) + apply simp + apply(subgoal_tac "rsimp r2 = rsimp r3") + apply auto[1] + using hrewrite_simpeq by presburger + + + +lemma simp_hrewrites: + shows "r1 h\* rsimp r1" + apply(induct r1) + apply simp+ + apply(case_tac "rsimp r11 = RONE") + apply simp + apply(subst basic_rsimp_SEQ_property1) + apply(subgoal_tac "RSEQ r11 r12 h\* RSEQ RONE r12") + using hreal_trans hrewrite.intros(3) apply blast + using hrewrites_seq_context apply presburger + apply(case_tac "rsimp r11 = RZERO") + apply simp + using hrewrite.intros(1) hrewrites_seq_context apply blast + apply(case_tac "rsimp r12 = RZERO") + apply simp + apply(subst basic_rsimp_SEQ_property3) + apply (meson hrewrite.intros(2) hrewrites.simps hrewrites_seq_context2) + apply(subst idiot2) + apply simp+ + using hrewrites_seq_contexts apply presburger + apply simp + apply(subgoal_tac "RALTS x h\* RALTS (map rsimp x)") + apply(subgoal_tac "RALTS (map rsimp x) h\* rsimp_ALTs (rdistinct (rflts (map rsimp x)) {}) ") + using hreal_trans apply blast + apply (meson flts_gstar greal_trans grewrites_ralts_rsimpalts gstar_rdistinct) + + apply (simp add: grewrites_ralts hrewrites_list) + by simp + +lemma interleave_aux1: + shows " RALT (RSEQ RZERO r1) r h\* r" + apply(subgoal_tac "RSEQ RZERO r1 h\* RZERO") + apply(subgoal_tac "RALT (RSEQ RZERO r1) r h\* RALT RZERO r") + apply (meson grewrite.intros(1) grewrite_ralts hreal_trans hrewrite.intros(10) hrewrites.simps) + using rs1 srewritescf_alt1 ss1 ss2 apply presburger + by (simp add: hr_in_rstar hrewrite.intros(1)) + + + +lemma rnullable_hrewrite: + shows "r1 h\ r2 \ rnullable r1 = rnullable r2" + apply(induct rule: hrewrite.induct) + apply simp+ + apply blast + apply simp+ + done + + +lemma interleave1: + shows "r h\ r' \ rder c r h\* rder c r'" + apply(induct r r' rule: hrewrite.induct) + apply (simp add: hr_in_rstar hrewrite.intros(1)) + apply (metis (no_types, lifting) basic_rsimp_SEQ_property3 list.simps(8) list.simps(9) rder.simps(1) rder.simps(5) rdistinct.simps(1) rflts.simps(1) rflts.simps(2) rsimp.simps(1) rsimp.simps(2) rsimp.simps(3) rsimp_ALTs.simps(1) simp_hrewrites) + apply simp + apply(subst interleave_aux1) + apply simp + apply(case_tac "rnullable r1") + apply simp + + apply (simp add: hrewrites_seq_context rnullable_hrewrite srewritescf_alt1 ss1 ss2) + + apply (simp add: hrewrites_seq_context rnullable_hrewrite) + apply(case_tac "rnullable r1") + apply simp + + using hr_in_rstar hrewrites_seq_context2 srewritescf_alt1 ss1 ss2 apply presburger + apply simp + using hr_in_rstar hrewrites_seq_context2 apply blast + apply simp + + using hrewrites_alts apply auto[1] + apply simp + using grewrite.intros(1) grewrite_append grewrite_ralts apply auto[1] + apply simp + apply (simp add: grewrite.intros(2) grewrite_append grewrite_ralts) + apply (simp add: hr_in_rstar hrewrite.intros(9)) + apply (simp add: hr_in_rstar hrewrite.intros(10)) + apply simp + using hrewrite.intros(11) by auto + +lemma interleave_star1: + shows "r h\* r' \ rder c r h\* rder c r'" + apply(induct rule : hrewrites.induct) + apply simp + by (meson hreal_trans interleave1) + + + +lemma inside_simp_removal: + shows " rsimp (rder x (rsimp r)) = rsimp (rder x r)" + apply(induct r) + apply simp+ + apply(case_tac "rnullable r1") + apply simp + + using inside_simp_seq_nullable apply blast + apply simp + apply (smt (verit, del_insts) idiot2 basic_rsimp_SEQ_property3 der_simp_nullability rder.simps(1) rder.simps(5) rnullable.simps(2) rsimp.simps(1) rsimp_SEQ.simps(1) rsimp_idem) + apply(subgoal_tac "rder x (RALTS xa) h\* rder x (rsimp (RALTS xa))") + using hrewrites_simpeq apply presburger + using interleave_star1 simp_hrewrites apply presburger + by simp + + + + +lemma rders_simp_same_simpders: + shows "s \ [] \ rders_simp r s = rsimp (rders r s)" + apply(induct s rule: rev_induct) + apply simp + apply(case_tac "xs = []") + apply simp + apply(simp add: rders_append rders_simp_append) + using inside_simp_removal by blast + + + + +lemma distinct_der: + shows "rsimp (rsimp_ALTs (map (rder x) (rdistinct rs {}))) = + rsimp (rsimp_ALTs (rdistinct (map (rder x) rs) {}))" + by (metis grewrites_simpalts gstar_rdistinct inside_simp_removal rder_rsimp_ALTs_commute) + + + + + +lemma rders_simp_lambda: + shows " rsimp \ rder x \ (\r. rders_simp r xs) = (\r. rders_simp r (xs @ [x]))" + using rders_simp_append by auto + +lemma rders_simp_nonempty_simped: + shows "xs \ [] \ rsimp \ (\r. rders_simp r xs) = (\r. rders_simp r xs)" + using rders_simp_same_simpders rsimp_idem by auto + +lemma repeated_altssimp: + shows "\r \ set rs. rsimp r = r \ rsimp (rsimp_ALTs (rdistinct (rflts rs) {})) = + rsimp_ALTs (rdistinct (rflts rs) {})" + by (metis map_idI rsimp.simps(2) rsimp_idem) + + + +lemma alts_closed_form: + shows "rsimp (rders_simp (RALTS rs) s) = rsimp (RALTS (map (\r. rders_simp r s) rs))" + apply(induct s rule: rev_induct) + apply simp + apply simp + apply(subst rders_simp_append) + apply(subgoal_tac " rsimp (rders_simp (rders_simp (RALTS rs) xs) [x]) = + rsimp(rders_simp (rsimp_ALTs (rdistinct (rflts (map (rsimp \ (\r. rders_simp r xs)) rs)) {})) [x])") + prefer 2 + apply (metis inside_simp_removal rders_simp_one_char) + apply(simp only: ) + apply(subst rders_simp_one_char) + apply(subst rsimp_idem) + apply(subgoal_tac "rsimp (rder x (rsimp_ALTs (rdistinct (rflts (map (rsimp \ (\r. rders_simp r xs)) rs)) {}))) = + rsimp ((rsimp_ALTs (map (rder x) (rdistinct (rflts (map (rsimp \ (\r. rders_simp r xs)) rs)) {})))) ") + prefer 2 + using rder_rsimp_ALTs_commute apply presburger + apply(simp only:) + apply(subgoal_tac "rsimp (rsimp_ALTs (map (rder x) (rdistinct (rflts (map (rsimp \ (\r. rders_simp r xs)) rs)) {}))) += rsimp (rsimp_ALTs (rdistinct (map (rder x) (rflts (map (rsimp \ (\r. rders_simp r xs)) rs))) {}))") + prefer 2 + + using distinct_der apply presburger + apply(simp only:) + apply(subgoal_tac " rsimp (rsimp_ALTs (rdistinct (map (rder x) (rflts (map (rsimp \ (\r. rders_simp r xs)) rs))) {})) = + rsimp (rsimp_ALTs (rdistinct ( (rflts (map (rder x) (map (rsimp \ (\r. rders_simp r xs)) rs)))) {}))") + apply(simp only:) + apply(subgoal_tac " rsimp (rsimp_ALTs (rdistinct (rflts (map (rder x) (map (rsimp \ (\r. rders_simp r xs)) rs))) {})) = + rsimp (rsimp_ALTs (rdistinct (rflts ( (map (rsimp \ (rder x) \ (\r. rders_simp r xs)) rs))) {}))") + apply(simp only:) + apply(subst rders_simp_lambda) + apply(subst rders_simp_nonempty_simped) + apply simp + apply(subgoal_tac "\r \ set (map (\r. rders_simp r (xs @ [x])) rs). rsimp r = r") + prefer 2 + apply (simp add: rders_simp_same_simpders rsimp_idem) + apply(subst repeated_altssimp) + apply simp + apply fastforce + apply (metis inside_simp_removal list.map_comp rder.simps(4) rsimp.simps(2) rsimp_idem) + using simp_der_pierce_flts_prelim by blast + + +lemma alts_closed_form_variant: + shows "s \ [] \ rders_simp (RALTS rs) s = rsimp (RALTS (map (\r. rders_simp r s) rs))" + by (metis alts_closed_form comp_apply rders_simp_nonempty_simped) + + +lemma rsimp_seq_equal1: + shows "rsimp_SEQ (rsimp r1) (rsimp r2) = rsimp_ALTs (rdistinct (rflts [rsimp_SEQ (rsimp r1) (rsimp r2)]) {})" + by (metis idem_after_simp1 rsimp.simps(1)) + + +fun sflat_aux :: "rrexp \ rrexp list " where + "sflat_aux (RALTS (r # rs)) = sflat_aux r @ rs" +| "sflat_aux (RALTS []) = []" +| "sflat_aux r = [r]" + + +fun sflat :: "rrexp \ rrexp" where + "sflat (RALTS (r # [])) = r" +| "sflat (RALTS (r # rs)) = RALTS (sflat_aux r @ rs)" +| "sflat r = r" + +inductive created_by_seq:: "rrexp \ bool" where + "created_by_seq (RSEQ r1 r2) " +| "created_by_seq r1 \ created_by_seq (RALT r1 r2)" + +lemma seq_ders_shape1: + shows "\r1 r2. \r3 r4. (rders (RSEQ r1 r2) s) = RSEQ r3 r4 \ (rders (RSEQ r1 r2) s) = RALT r3 r4" + apply(induct s rule: rev_induct) + apply auto[1] + apply(rule allI)+ + apply(subst rders_append)+ + apply(subgoal_tac " \r3 r4. rders (RSEQ r1 r2) xs = RSEQ r3 r4 \ rders (RSEQ r1 r2) xs = RALT r3 r4 ") + apply(erule exE)+ + apply(erule disjE) + apply simp+ + done + +lemma created_by_seq_der: + shows "created_by_seq r \ created_by_seq (rder c r)" + apply(induct r) + apply simp+ + + using created_by_seq.cases apply blast + + apply (meson created_by_seq.cases rrexp.distinct(19) rrexp.distinct(21)) + apply (metis created_by_seq.simps rder.simps(5)) + apply (smt (verit, ccfv_threshold) created_by_seq.simps list.set_intros(1) list.simps(8) list.simps(9) rder.simps(4) rrexp.distinct(25) rrexp.inject(3)) + using created_by_seq.intros(1) by force + +lemma createdbyseq_left_creatable: + shows "created_by_seq (RALT r1 r2) \ created_by_seq r1" + using created_by_seq.cases by blast + + + +lemma recursively_derseq: + shows " created_by_seq (rders (RSEQ r1 r2) s)" + apply(induct s rule: rev_induct) + apply simp + using created_by_seq.intros(1) apply force + apply(subgoal_tac "created_by_seq (rders (RSEQ r1 r2) (xs @ [x]))") + apply blast + apply(subst rders_append) + apply(subgoal_tac "\r3 r4. rders (RSEQ r1 r2) xs = RSEQ r3 r4 \ + rders (RSEQ r1 r2) xs = RALT r3 r4") + prefer 2 + using seq_ders_shape1 apply presburger + apply(erule exE)+ + apply(erule disjE) + apply(subgoal_tac "created_by_seq (rders (RSEQ r3 r4) [x])") + apply presburger + apply simp + using created_by_seq.intros(1) created_by_seq.intros(2) apply presburger + apply simp + apply(subgoal_tac "created_by_seq r3") + prefer 2 + using createdbyseq_left_creatable apply blast + using created_by_seq.intros(2) created_by_seq_der by blast + + +lemma recursively_derseq1: + shows "r = rders (RSEQ r1 r2) s \ created_by_seq r" + using recursively_derseq by blast + + +lemma sfau_head: + shows " created_by_seq r \ \ra rb rs. sflat_aux r = RSEQ ra rb # rs" + apply(induction r rule: created_by_seq.induct) + apply simp + by fastforce + + +lemma vsuf_prop1: + shows "vsuf (xs @ [x]) r = (if (rnullable (rders r xs)) + then [x] # (map (\s. s @ [x]) (vsuf xs r) ) + else (map (\s. s @ [x]) (vsuf xs r)) ) + " + apply(induct xs arbitrary: r) + apply simp + apply(case_tac "rnullable r") + apply simp + apply simp + done + +fun breakHead :: "rrexp list \ rrexp list" where + "breakHead [] = [] " +| "breakHead (RALT r1 r2 # rs) = r1 # r2 # rs" +| "breakHead (r # rs) = r # rs" + + +lemma sfau_idem_der: + shows "created_by_seq r \ sflat_aux (rder c r) = breakHead (map (rder c) (sflat_aux r))" + apply(induct rule: created_by_seq.induct) + apply simp+ + using sfau_head by fastforce + +lemma vsuf_compose1: + shows " \ rnullable (rders r1 xs) + \ map (rder x \ rders r2) (vsuf xs r1) = map (rders r2) (vsuf (xs @ [x]) r1)" + apply(subst vsuf_prop1) + apply simp + by (simp add: rders_append) + + + + +lemma seq_sfau0: + shows "sflat_aux (rders (RSEQ r1 r2) s) = (RSEQ (rders r1 s) r2) # + (map (rders r2) (vsuf s r1)) " + apply(induct s rule: rev_induct) + apply simp + apply(subst rders_append)+ + apply(subgoal_tac "created_by_seq (rders (RSEQ r1 r2) xs)") + prefer 2 + using recursively_derseq1 apply blast + apply simp + apply(subst sfau_idem_der) + + apply blast + apply(case_tac "rnullable (rders r1 xs)") + apply simp + apply(subst vsuf_prop1) + apply simp + apply (simp add: rders_append) + apply simp + using vsuf_compose1 by blast + + + + + + + + + +thm sflat.elims + + + + + +lemma sflat_rsimpeq: + shows "created_by_seq r1 \ sflat_aux r1 = rs \ rsimp r1 = rsimp (RALTS rs)" + apply(induct r1 arbitrary: rs rule: created_by_seq.induct) + apply simp + using rsimp_seq_equal1 apply force + by (metis head_one_more_simp rsimp.simps(2) sflat_aux.simps(1) simp_flatten) + + + +lemma seq_closed_form_general: + shows "rsimp (rders (RSEQ r1 r2) s) = +rsimp ( (RALTS ( (RSEQ (rders r1 s) r2 # (map (rders r2) (vsuf s r1))))))" + apply(case_tac "s \ []") + apply(subgoal_tac "created_by_seq (rders (RSEQ r1 r2) s)") + apply(subgoal_tac "sflat_aux (rders (RSEQ r1 r2) s) = RSEQ (rders r1 s) r2 # (map (rders r2) (vsuf s r1))") + using sflat_rsimpeq apply blast + apply (simp add: seq_sfau0) + using recursively_derseq1 apply blast + apply simp + by (metis idem_after_simp1 rsimp.simps(1)) + +lemma seq_closed_form_aux1a: + shows "rsimp (RALTS (RSEQ (rders r1 s) r2 # rs)) = + rsimp (RALTS (RSEQ (rders_simp r1 s) r2 # rs))" + by (metis head_one_more_simp rders.simps(1) rders_simp.simps(1) rders_simp_same_simpders rsimp.simps(1) rsimp_idem simp_flatten_aux0) + + +lemma seq_closed_form_aux1: + shows "rsimp (RALTS (RSEQ (rders r1 s) r2 # (map (rders r2) (vsuf s r1)))) = + rsimp (RALTS (RSEQ (rders_simp r1 s) r2 # (map (rders r2) (vsuf s r1))))" + by (smt (verit, best) list.simps(9) rders.simps(1) rders_simp.simps(1) rders_simp_same_simpders rsimp.simps(1) rsimp.simps(2) rsimp_idem) + +lemma add_simp_to_rest: + shows "rsimp (RALTS (r # rs)) = rsimp (RALTS (r # map rsimp rs))" + by (metis append_Nil2 grewrite.intros(2) grewrite_simpalts head_one_more_simp list.simps(9) rsimp_ALTs.simps(2) spawn_simp_rsimpalts) + +lemma rsimp_compose_der2: + shows "\s \ set ss. s \ [] \ map rsimp (map (rders r) ss) = map (\s. (rders_simp r s)) ss" + by (simp add: rders_simp_same_simpders) + +lemma vsuf_nonempty: + shows "\s \ set ( vsuf s1 r). s \ []" + apply(induct s1 arbitrary: r) + apply simp + apply simp + done + + + +lemma seq_closed_form_aux2: + shows "s \ [] \ rsimp ( (RALTS ( (RSEQ (rders_simp r1 s) r2 # (map (rders r2) (vsuf s r1)))))) = + rsimp ( (RALTS ( (RSEQ (rders_simp r1 s) r2 # (map (rders_simp r2) (vsuf s r1))))))" + + by (metis add_simp_to_rest rsimp_compose_der2 vsuf_nonempty) + + +lemma seq_closed_form: + shows "rsimp (rders_simp (RSEQ r1 r2) s) = + rsimp (RALTS ((RSEQ (rders_simp r1 s) r2) # (map (rders_simp r2) (vsuf s r1))))" +proof (cases s) + case Nil + then show ?thesis + by (simp add: rsimp_seq_equal1[symmetric]) +next + case (Cons a list) + have "rsimp (rders_simp (RSEQ r1 r2) s) = rsimp (rsimp (rders (RSEQ r1 r2) s))" + using local.Cons by (subst rders_simp_same_simpders)(simp_all) + also have "... = rsimp (rders (RSEQ r1 r2) s)" + by (simp add: rsimp_idem) + also have "... = rsimp (RALTS (RSEQ (rders r1 s) r2 # map (rders r2) (vsuf s r1)))" + using seq_closed_form_general by blast + also have "... = rsimp (RALTS (RSEQ (rders_simp r1 s) r2 # map (rders r2) (vsuf s r1)))" + by (simp only: seq_closed_form_aux1) + also have "... = rsimp (RALTS (RSEQ (rders_simp r1 s) r2 # map (rders_simp r2) (vsuf s r1)))" + using local.Cons by (subst seq_closed_form_aux2)(simp_all) + finally show ?thesis . +qed + +lemma q: "s \ [] \ rders_simp (RSEQ r1 r2) s = rsimp (rders_simp (RSEQ r1 r2) s)" + using rders_simp_same_simpders rsimp_idem by presburger + + +lemma seq_closed_form_variant: + assumes "s \ []" + shows "rders_simp (RSEQ r1 r2) s = + rsimp (RALTS (RSEQ (rders_simp r1 s) r2 # (map (rders_simp r2) (vsuf s r1))))" + using assms q seq_closed_form by force + + +fun hflat_aux :: "rrexp \ rrexp list" where + "hflat_aux (RALT r1 r2) = hflat_aux r1 @ hflat_aux r2" +| "hflat_aux r = [r]" + + +fun hflat :: "rrexp \ rrexp" where + "hflat (RALT r1 r2) = RALTS ((hflat_aux r1) @ (hflat_aux r2))" +| "hflat r = r" + +inductive created_by_star :: "rrexp \ bool" where + "created_by_star (RSEQ ra (RSTAR rb))" +| "\created_by_star r1; created_by_star r2\ \ created_by_star (RALT r1 r2)" + +fun hElem :: "rrexp \ rrexp list" where + "hElem (RALT r1 r2) = (hElem r1 ) @ (hElem r2)" +| "hElem r = [r]" + + + + +lemma cbs_ders_cbs: + shows "created_by_star r \ created_by_star (rder c r)" + apply(induct r rule: created_by_star.induct) + apply simp + using created_by_star.intros(1) created_by_star.intros(2) apply auto[1] + by (metis (mono_tags, lifting) created_by_star.simps list.simps(8) list.simps(9) rder.simps(4)) + +lemma star_ders_cbs: + shows "created_by_star (rders (RSEQ r1 (RSTAR r2)) s)" + apply(induct s rule: rev_induct) + apply simp + apply (simp add: created_by_star.intros(1)) + apply(subst rders_append) + apply simp + using cbs_ders_cbs by auto + +(* +lemma created_by_star_cases: + shows "created_by_star r \ \ra rb. (r = RALT ra rb \ created_by_star ra \ created_by_star rb) \ r = RSEQ ra rb " + by (meson created_by_star.cases) +*) + + +lemma hfau_pushin: + shows "created_by_star r \ hflat_aux (rder c r) = concat (map hElem (map (rder c) (hflat_aux r)))" + apply(induct r rule: created_by_star.induct) + apply simp + apply(subgoal_tac "created_by_star (rder c r1)") + prefer 2 + apply(subgoal_tac "created_by_star (rder c r2)") + using cbs_ders_cbs apply blast + using cbs_ders_cbs apply auto[1] + apply simp + done + +lemma stupdate_induct1: + shows " concat (map (hElem \ (rder x \ (\s1. RSEQ (rders r0 s1) (RSTAR r0)))) Ss) = + map (\s1. RSEQ (rders r0 s1) (RSTAR r0)) (star_update x r0 Ss)" + apply(induct Ss) + apply simp+ + by (simp add: rders_append) + + + +lemma stupdates_join_general: + shows "concat (map hElem (map (rder x) (map (\s1. RSEQ (rders r0 s1) (RSTAR r0)) (star_updates xs r0 Ss)))) = + map (\s1. RSEQ (rders r0 s1) (RSTAR r0)) (star_updates (xs @ [x]) r0 Ss)" + apply(induct xs arbitrary: Ss) + apply (simp) + prefer 2 + apply auto[1] + using stupdate_induct1 by blast + +lemma star_hfau_induct: + shows "hflat_aux (rders (RSEQ (rder c r0) (RSTAR r0)) s) = + map (\s1. RSEQ (rders r0 s1) (RSTAR r0)) (star_updates s r0 [[c]])" + apply(induct s rule: rev_induct) + apply simp + apply(subst rders_append)+ + apply simp + apply(subst stupdates_append) + apply(subgoal_tac "created_by_star (rders (RSEQ (rder c r0) (RSTAR r0)) xs)") + prefer 2 + apply (simp add: star_ders_cbs) + apply(subst hfau_pushin) + apply simp + apply(subgoal_tac "concat (map hElem (map (rder x) (hflat_aux (rders (RSEQ (rder c r0) (RSTAR r0)) xs)))) = + concat (map hElem (map (rder x) ( map (\s1. RSEQ (rders r0 s1) (RSTAR r0)) (star_updates xs r0 [[c]])))) ") + apply(simp only:) + prefer 2 + apply presburger + apply(subst stupdates_append[symmetric]) + using stupdates_join_general by blast + +lemma starders_hfau_also1: + shows "hflat_aux (rders (RSTAR r) (c # xs)) = map (\s1. RSEQ (rders r s1) (RSTAR r)) (star_updates xs r [[c]])" + using star_hfau_induct by force + +lemma hflat_aux_grewrites: + shows "a # rs \g* hflat_aux a @ rs" + apply(induct a arbitrary: rs) + apply simp+ + apply(case_tac x) + apply simp + apply(case_tac list) + + apply (metis append.right_neutral append_Cons append_eq_append_conv2 grewrites.simps hflat_aux.simps(7) same_append_eq) + apply(case_tac lista) + apply simp + apply (metis (no_types, lifting) append_Cons append_eq_append_conv2 gmany_steps_later greal_trans grewrite.intros(2) grewrites_append self_append_conv) + apply simp + by simp + + + + +lemma cbs_hfau_rsimpeq1: + shows "rsimp (RALT a b) = rsimp (RALTS ((hflat_aux a) @ (hflat_aux b)))" + apply(subgoal_tac "[a, b] \g* hflat_aux a @ hflat_aux b") + using grewrites_equal_rsimp apply presburger + by (metis append.right_neutral greal_trans grewrites_cons hflat_aux_grewrites) + + +lemma hfau_rsimpeq2: + shows "created_by_star r \ rsimp r = rsimp ( (RALTS (hflat_aux r)))" + apply(induct r) + apply simp+ + + apply (metis rsimp_seq_equal1) + prefer 2 + apply simp + apply(case_tac x) + apply simp + apply(case_tac "list") + apply simp + + apply (metis idem_after_simp1) + apply(case_tac "lista") + prefer 2 + apply (metis hflat_aux.simps(8) idem_after_simp1 list.simps(8) list.simps(9) rsimp.simps(2)) + apply(subgoal_tac "rsimp (RALT a aa) = rsimp (RALTS (hflat_aux (RALT a aa)))") + apply simp + apply(subgoal_tac "rsimp (RALT a aa) = rsimp (RALTS (hflat_aux a @ hflat_aux aa))") + using hflat_aux.simps(1) apply presburger + apply simp + using cbs_hfau_rsimpeq1 by fastforce + +lemma star_closed_form1: + shows "rsimp (rders (RSTAR r0) (c#s)) = +rsimp ( ( RALTS ( (map (\s1. RSEQ (rders r0 s1) (RSTAR r0) ) (star_updates s r0 [[c]]) ) )))" + using hfau_rsimpeq2 rder.simps(6) rders.simps(2) star_ders_cbs starders_hfau_also1 by presburger + +lemma star_closed_form2: + shows "rsimp (rders_simp (RSTAR r0) (c#s)) = +rsimp ( ( RALTS ( (map (\s1. RSEQ (rders r0 s1) (RSTAR r0) ) (star_updates s r0 [[c]]) ) )))" + by (metis list.distinct(1) rders_simp_same_simpders rsimp_idem star_closed_form1) + +lemma star_closed_form3: + shows "rsimp (rders_simp (RSTAR r0) (c#s)) = (rders_simp (RSTAR r0) (c#s))" + by (metis list.distinct(1) rders_simp_same_simpders star_closed_form1 star_closed_form2) + +lemma star_closed_form4: + shows " (rders_simp (RSTAR r0) (c#s)) = +rsimp ( ( RALTS ( (map (\s1. RSEQ (rders r0 s1) (RSTAR r0) ) (star_updates s r0 [[c]]) ) )))" + using star_closed_form2 star_closed_form3 by presburger + +lemma star_closed_form5: + shows " rsimp ( ( RALTS ( (map (\s1. RSEQ (rders r0 s1) (RSTAR r0) ) Ss )))) = + rsimp ( ( RALTS ( (map (\s1. rsimp (RSEQ (rders r0 s1) (RSTAR r0)) ) Ss ))))" + by (metis (mono_tags, lifting) list.map_comp map_eq_conv o_apply rsimp.simps(2) rsimp_idem) + +lemma star_closed_form6_hrewrites: + shows " + (map (\s1. (RSEQ (rsimp (rders r0 s1)) (RSTAR r0)) ) Ss ) + scf\* +(map (\s1. rsimp (RSEQ (rders r0 s1) (RSTAR r0)) ) Ss )" + apply(induct Ss) + apply simp + apply (simp add: ss1) + by (metis (no_types, lifting) list.simps(9) rsimp.simps(1) rsimp_idem simp_hrewrites ss2) + +lemma star_closed_form6: + shows " rsimp ( ( RALTS ( (map (\s1. rsimp (RSEQ (rders r0 s1) (RSTAR r0)) ) Ss )))) = + rsimp ( ( RALTS ( (map (\s1. (RSEQ (rsimp (rders r0 s1)) (RSTAR r0)) ) Ss ))))" + apply(subgoal_tac " map (\s1. (RSEQ (rsimp (rders r0 s1)) (RSTAR r0)) ) Ss scf\* + map (\s1. rsimp (RSEQ (rders r0 s1) (RSTAR r0)) ) Ss ") + using hrewrites_simpeq srewritescf_alt1 apply fastforce + using star_closed_form6_hrewrites by blast + +lemma stupdate_nonempty: + shows "\s \ set Ss. s \ [] \ \s \ set (star_update c r Ss). s \ []" + apply(induct Ss) + apply simp + apply(case_tac "rnullable (rders r a)") + apply simp+ + done + + +lemma stupdates_nonempty: + shows "\s \ set Ss. s\ [] \ \s \ set (star_updates s r Ss). s \ []" + apply(induct s arbitrary: Ss) + apply simp + apply simp + using stupdate_nonempty by presburger + + +lemma star_closed_form8: + shows +"rsimp ( ( RALTS ( (map (\s1. RSEQ (rsimp (rders r0 s1)) (RSTAR r0) ) (star_updates s r0 [[c]]) ) ))) = + rsimp ( ( RALTS ( (map (\s1. RSEQ ( (rders_simp r0 s1)) (RSTAR r0) ) (star_updates s r0 [[c]]) ) )))" + by (smt (verit, ccfv_SIG) list.simps(8) map_eq_conv rders__onechar rders_simp_same_simpders set_ConsD stupdates_nonempty) + + +lemma star_closed_form: + shows "rders_simp (RSTAR r0) (c#s) = +rsimp ( RALTS ( (map (\s1. RSEQ (rders_simp r0 s1) (RSTAR r0) ) (star_updates s r0 [[c]]) ) ))" + apply(induct s) + apply simp + apply (metis idem_after_simp1 rsimp.simps(1) rsimp.simps(6) rsimp_idem) + using star_closed_form4 star_closed_form5 star_closed_form6 star_closed_form8 by presburger + + +unused_thms + +end \ No newline at end of file diff -r f9cdc295ccf7 -r f493a20feeb3 thys3/src/ClosedFormsBounds.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys3/src/ClosedFormsBounds.thy Sat Apr 30 00:50:08 2022 +0100 @@ -0,0 +1,448 @@ + +theory ClosedFormsBounds + imports "GeneralRegexBound" "ClosedForms" +begin +lemma alts_ders_lambda_shape_ders: + shows "\r \ set (map (\r. rders_simp r ( s)) rs ). \r1 \ set rs. r = rders_simp r1 s" + by (simp add: image_iff) + +lemma rlist_bound: + assumes "\r \ set rs. rsize r \ N" + shows "rsizes rs \ N * (length rs)" + using assms + apply(induct rs) + apply simp + by simp + +lemma alts_closed_form_bounded: + assumes "\r \ set rs. \s. rsize (rders_simp r s) \ N" + shows "rsize (rders_simp (RALTS rs) s) \ max (Suc (N * (length rs))) (rsize (RALTS rs))" +proof (cases s) + case Nil + then show "rsize (rders_simp (RALTS rs) s) \ max (Suc (N * length rs)) (rsize (RALTS rs))" + by simp +next + case (Cons a s) + + from assms have "\r \ set (map (\r. rders_simp r (a # s)) rs ). rsize r \ N" + by (metis alts_ders_lambda_shape_ders) + then have a: "rsizes (map (\r. rders_simp r (a # s)) rs ) \ N * (length rs)" + by (metis length_map rlist_bound) + + have "rsize (rders_simp (RALTS rs) (a # s)) + = rsize (rsimp (RALTS (map (\r. rders_simp r (a # s)) rs)))" + by (metis alts_closed_form_variant list.distinct(1)) + also have "... \ rsize (RALTS (map (\r. rders_simp r (a # s)) rs))" + using rsimp_mono by blast + also have "... = Suc (rsizes (map (\r. rders_simp r (a # s)) rs))" + by simp + also have "... \ Suc (N * (length rs))" + using a by blast + finally have "rsize (rders_simp (RALTS rs) (a # s)) \ max (Suc (N * length rs)) (rsize (RALTS rs))" + by auto + then show ?thesis using local.Cons by simp +qed + +lemma alts_simp_ineq_unfold: + shows "rsize (rsimp (RALTS rs)) \ Suc (rsizes (rdistinct (rflts (map rsimp rs)) {}))" + using rsimp_aalts_smaller by auto + + +lemma rdistinct_mono_list: + shows "rsizes (rdistinct (x5 @ rs) rset) \ rsizes x5 + rsizes (rdistinct rs ((set x5 ) \ rset))" + apply(induct x5 arbitrary: rs rset) + apply simp + apply(case_tac "a \ rset") + apply simp + apply (simp add: add.assoc insert_absorb trans_le_add2) + apply simp + by (metis Un_insert_right) + + +lemma flts_size_reduction_alts: + assumes a: "\noalts_set alts_set corr_set. + (\r\noalts_set. \xs. r \ RALTS xs) \ + (\a\alts_set. \xs. a = RALTS xs \ set xs \ corr_set) \ + Suc (rsizes (rdistinct (rflts rs) (noalts_set \ corr_set))) + \ Suc (rsizes (rdistinct rs (insert RZERO (noalts_set \ alts_set))))" + and b: "\r\noalts_set. \xs. r \ RALTS xs" + and c: "\a\alts_set. \xs. a = RALTS xs \ set xs \ corr_set" + and d: "a = RALTS x5" + shows "rsizes (rdistinct (rflts (a # rs)) (noalts_set \ corr_set)) + \ rsizes (rdistinct (a # rs) (insert RZERO (noalts_set \ alts_set)))" + + apply(case_tac "a \ alts_set") + using a b c d + apply simp + apply(subgoal_tac "set x5 \ corr_set") + apply(subst rdistinct_concat) + apply auto[1] + apply presburger + apply fastforce + using a b c d + apply (subgoal_tac "a \ noalts_set") + prefer 2 + apply blast + apply simp + apply(subgoal_tac "rsizes (rdistinct (x5 @ rflts rs) (noalts_set \ corr_set)) + \ rsizes x5 + rsizes (rdistinct (rflts rs) ((set x5) \ (noalts_set \ corr_set)))") + prefer 2 + using rdistinct_mono_list apply presburger + apply(subgoal_tac "insert (RALTS x5) (noalts_set \ alts_set) = noalts_set \ (insert (RALTS x5) alts_set)") + apply(simp only:) + apply(subgoal_tac "rsizes x5 + rsizes (rdistinct (rflts rs) (noalts_set \ (corr_set \ (set x5)))) \ + rsizes x5 + rsizes (rdistinct rs (insert RZERO (noalts_set \ insert (RALTS x5) alts_set)))") + + apply (simp add: Un_left_commute inf_sup_aci(5)) + apply(subgoal_tac "rsizes (rdistinct (rflts rs) (noalts_set \ (corr_set \ set x5))) \ + rsizes (rdistinct rs (insert RZERO (noalts_set \ insert (RALTS x5) alts_set)))") + apply linarith + apply(subgoal_tac "\r \ insert (RALTS x5) alts_set. \xs1.( r = RALTS xs1 \ set xs1 \ corr_set \ set x5)") + apply presburger + apply (meson insert_iff sup.cobounded2 sup.coboundedI1) + by blast + + +lemma flts_vs_nflts1: + assumes "\r \ noalts_set. \xs. r \ RALTS xs" + and "\a \ alts_set. (\xs. a = RALTS xs \ set xs \ corr_set)" + shows "rsizes (rdistinct (rflts rs) (noalts_set \ corr_set)) + \ rsizes (rdistinct rs (insert RZERO (noalts_set \ alts_set)))" + using assms + apply(induct rs arbitrary: noalts_set alts_set corr_set) + apply simp + apply(case_tac a) + apply(case_tac "RZERO \ noalts_set") + apply simp + apply(subgoal_tac "RZERO \ alts_set") + apply simp + apply fastforce + apply(case_tac "RONE \ noalts_set") + apply simp + apply(subgoal_tac "RONE \ alts_set") + prefer 2 + apply fastforce + apply(case_tac "RONE \ corr_set") + apply(subgoal_tac "rflts (a # rs) = RONE # rflts rs") + apply(simp only:) + apply(subgoal_tac "rdistinct (RONE # rflts rs) (noalts_set \ corr_set) = + rdistinct (rflts rs) (noalts_set \ corr_set)") + apply(simp only:) + apply(subgoal_tac "rdistinct (RONE # rs) (insert RZERO (noalts_set \ alts_set)) = + RONE # (rdistinct rs (insert RONE (insert RZERO (noalts_set \ alts_set)))) ") + apply(simp only:) + apply(subgoal_tac "rdistinct (rflts rs) (noalts_set \ corr_set) = + rdistinct (rflts rs) (insert RONE (noalts_set \ corr_set))") + apply (simp only:) + apply(subgoal_tac "insert RONE (noalts_set \ corr_set) = (insert RONE noalts_set) \ corr_set") + apply(simp only:) + apply(subgoal_tac "insert RONE (insert RZERO (noalts_set \ alts_set)) = + insert RZERO ((insert RONE noalts_set) \ alts_set)") + apply(simp only:) + apply(subgoal_tac "rsizes (rdistinct rs (insert RZERO (insert RONE noalts_set \ alts_set))) + \ rsizes (RONE # rdistinct rs (insert RZERO (insert RONE noalts_set \ alts_set)))") + apply (smt (verit, best) dual_order.trans insert_iff rrexp.distinct(15)) + apply (metis (no_types, opaque_lifting) le_add_same_cancel2 list.simps(9) sum_list.Cons zero_le) + apply fastforce + apply fastforce + apply (metis Un_iff insert_absorb) + apply (metis UnE insertE insert_is_Un rdistinct.simps(2) rrexp.distinct(1)) + apply (meson UnCI rdistinct.simps(2)) + using rflts.simps(4) apply presburger + apply simp + apply(subgoal_tac "insert RONE (noalts_set \ corr_set) = (insert RONE noalts_set) \ corr_set") + apply(simp only:) + apply (metis Un_insert_left insertE rrexp.distinct(15)) + apply fastforce + apply(case_tac "a \ noalts_set") + apply simp + apply(subgoal_tac "a \ alts_set") + prefer 2 + apply blast + apply(case_tac "a \ corr_set") + apply(subgoal_tac "noalts_set \ corr_set = insert a ( noalts_set \ corr_set)") + prefer 2 + apply fastforce + apply(simp only:) + apply(subgoal_tac "rsizes (rdistinct (a # rs) (insert RZERO ((insert a noalts_set) \ alts_set))) \ + rsizes (rdistinct (a # rs) (insert RZERO (noalts_set \ alts_set)))") + + apply(subgoal_tac "rsizes (rdistinct (rflts (a # rs)) ((insert a noalts_set) \ corr_set)) \ + rsizes (rdistinct (a # rs) (insert RZERO ((insert a noalts_set) \ alts_set)))") + apply fastforce + apply simp + apply(subgoal_tac "(insert a (noalts_set \ alts_set)) = (insert a noalts_set) \ alts_set") + apply(simp only:) + apply(subgoal_tac "noalts_set \ corr_set = (insert a noalts_set) \ corr_set") + apply(simp only:) + apply (metis insertE rrexp.distinct(21)) + apply blast + + apply fastforce + apply force + apply simp + apply (metis Un_insert_left insert_iff rrexp.distinct(21)) + apply(case_tac "a \ noalts_set") + apply simp + apply(subgoal_tac "a \ alts_set") + prefer 2 + apply blast + apply(case_tac "a \ corr_set") + apply(subgoal_tac "noalts_set \ corr_set = insert a ( noalts_set \ corr_set)") + prefer 2 + apply fastforce + apply(simp only:) + apply(subgoal_tac "rsizes (rdistinct (a # rs) (insert RZERO ((insert a noalts_set) \ alts_set))) \ + rsizes (rdistinct (a # rs) (insert RZERO (noalts_set \ alts_set)))") + + apply(subgoal_tac "rsizes (rdistinct (rflts (a # rs)) ((insert a noalts_set) \ corr_set)) \ + rsizes (rdistinct (a # rs) (insert RZERO ((insert a noalts_set) \ alts_set)))") + apply fastforce + apply simp + apply(subgoal_tac "(insert a (noalts_set \ alts_set)) = (insert a noalts_set) \ alts_set") + apply(simp only:) + apply(subgoal_tac "noalts_set \ corr_set = (insert a noalts_set) \ corr_set") + apply(simp only:) + + + apply (metis insertE rrexp.distinct(25)) + apply blast + apply fastforce + apply force + apply simp + + apply (metis Un_insert_left insertE rrexp.distinct(25)) + + using Suc_le_mono flts_size_reduction_alts apply presburger + apply(case_tac "a \ noalts_set") + apply simp + apply(subgoal_tac "a \ alts_set") + prefer 2 + apply blast + apply(case_tac "a \ corr_set") + apply(subgoal_tac "noalts_set \ corr_set = insert a ( noalts_set \ corr_set)") + prefer 2 + apply fastforce + apply(simp only:) + apply(subgoal_tac "rsizes (rdistinct (a # rs) (insert RZERO ((insert a noalts_set) \ alts_set))) \ + rsizes (rdistinct (a # rs) (insert RZERO (noalts_set \ alts_set)))") + + apply(subgoal_tac "rsizes (rdistinct (rflts (a # rs)) ((insert a noalts_set) \ corr_set)) \ + rsizes (rdistinct (a # rs) (insert RZERO ((insert a noalts_set) \ alts_set)))") + apply fastforce + apply simp + apply(subgoal_tac "(insert a (noalts_set \ alts_set)) = (insert a noalts_set) \ alts_set") + apply(simp only:) + apply(subgoal_tac "noalts_set \ corr_set = (insert a noalts_set) \ corr_set") + apply(simp only:) + apply (metis insertE rrexp.distinct(29)) + + apply blast + + apply fastforce + apply force + apply simp + apply (metis Un_insert_left insert_iff rrexp.distinct(29)) + done + + +lemma flts_vs_nflts: + assumes "\r \ noalts_set. \xs. r \ RALTS xs" + and "\a \ alts_set. (\xs. a = RALTS xs \ set xs \ corr_set)" + shows "rsizes (rdistinct (rflts rs) (noalts_set \ corr_set)) + \ rsizes (rdistinct rs (insert RZERO (noalts_set \ alts_set)))" + by (simp add: assms flts_vs_nflts1) + +lemma distinct_simp_ineq_general: + assumes "rsimp ` no_simp = has_simp" "finite no_simp" + shows "rsizes (rdistinct (map rsimp rs) has_simp) \ rsizes (rdistinct rs no_simp)" + using assms + apply(induct rs no_simp arbitrary: has_simp rule: rdistinct.induct) + apply simp + apply(auto) + using add_le_mono rsimp_mono by presburger + +lemma larger_acc_smaller_distinct_res0: + assumes "ss \ SS" + shows "rsizes (rdistinct rs SS) \ rsizes (rdistinct rs ss)" + using assms + apply(induct rs arbitrary: ss SS) + apply simp + by (metis distinct_early_app1 rdistinct_smaller) + +lemma without_flts_ineq: + shows "rsizes (rdistinct (rflts rs) {}) \ rsizes (rdistinct rs {})" +proof - + have "rsizes (rdistinct (rflts rs) {}) \ rsizes (rdistinct rs (insert RZERO {}))" + by (metis empty_iff flts_vs_nflts sup_bot_left) + also have "... \ rsizes (rdistinct rs {})" + by (simp add: larger_acc_smaller_distinct_res0) + finally show ?thesis + by blast +qed + + +lemma distinct_simp_ineq: + shows "rsizes (rdistinct (map rsimp rs) {}) \ rsizes (rdistinct rs {})" + using distinct_simp_ineq_general by blast + + +lemma alts_simp_control: + shows "rsize (rsimp (RALTS rs)) \ Suc (rsizes (rdistinct rs {}))" +proof - + have "rsize (rsimp (RALTS rs)) \ Suc (rsizes (rdistinct (rflts (map rsimp rs)) {}))" + using alts_simp_ineq_unfold by auto + moreover have "\ \ Suc (rsizes (rdistinct (map rsimp rs) {}))" + using without_flts_ineq by blast + ultimately show "rsize (rsimp (RALTS rs)) \ Suc (rsizes (rdistinct rs {}))" + by (meson Suc_le_mono distinct_simp_ineq le_trans) +qed + + +lemma larger_acc_smaller_distinct_res: + shows "rsizes (rdistinct rs (insert a ss)) \ rsizes (rdistinct rs ss)" + by (simp add: larger_acc_smaller_distinct_res0 subset_insertI) + +lemma triangle_inequality_distinct: + shows "rsizes (rdistinct (a # rs) ss) \ rsize a + rsizes (rdistinct rs ss)" + apply(case_tac "a \ ss") + apply simp + by (simp add: larger_acc_smaller_distinct_res) + + +lemma distinct_list_size_len_bounded: + assumes "\r \ set rs. rsize r \ N" "length rs \ lrs" + shows "rsizes rs \ lrs * N " + using assms + by (metis rlist_bound dual_order.trans mult.commute mult_le_mono1) + + + +lemma rdistinct_same_set: + shows "r \ set rs \ r \ set (rdistinct rs {})" + apply(induct rs) + apply simp + by (metis rdistinct_set_equality) + +(* distinct_list_rexp_up_to_certain_size_bouded_by_set_enumerating_up_to_that_size *) +lemma distinct_list_rexp_upto: + assumes "\r\ set rs. (rsize r) \ N" + shows "rsizes (rdistinct rs {}) \ (card (sizeNregex N)) * N" + + apply(subgoal_tac "distinct (rdistinct rs {})") + prefer 2 + using rdistinct_does_the_job apply blast + apply(subgoal_tac "length (rdistinct rs {}) \ card (sizeNregex N)") + apply(rule distinct_list_size_len_bounded) + using assms + apply (meson rdistinct_same_set) + apply blast + apply(subgoal_tac "\r \ set (rdistinct rs {}). rsize r \ N") + prefer 2 + using assms + apply (meson rdistinct_same_set) + apply(subgoal_tac "length (rdistinct rs {}) = card (set (rdistinct rs {}))") + prefer 2 + apply (simp add: distinct_card) + apply(simp) + by (metis card_mono finite_size_n mem_Collect_eq sizeNregex_def subsetI) + + +lemma star_control_bounded: + assumes "\s. rsize (rders_simp r s) \ N" + shows "rsizes (rdistinct (map (\s1. RSEQ (rders_simp r s1) (RSTAR r)) (star_updates s r [[c]])) {}) + \ (card (sizeNregex (Suc (N + rsize (RSTAR r))))) * (Suc (N + rsize (RSTAR r)))" + by (smt (verit) add_Suc_shift add_mono_thms_linordered_semiring(3) assms distinct_list_rexp_upto image_iff list.set_map plus_nat.simps(2) rsize.simps(5)) + + +lemma star_closed_form_bounded: + assumes "\s. rsize (rders_simp r s) \ N" + shows "rsize (rders_simp (RSTAR r) s) \ + max ((Suc (card (sizeNregex (Suc (N + rsize (RSTAR r))))) * (Suc (N + rsize (RSTAR r))))) (rsize (RSTAR r))" +proof(cases s) + case Nil + then show "rsize (rders_simp (RSTAR r) s) + \ max (Suc (card (sizeNregex (Suc (N + rsize (RSTAR r))))) * Suc (N + rsize (RSTAR r))) (rsize (RSTAR r))" + by simp +next + case (Cons a list) + then have "rsize (rders_simp (RSTAR r) s) = + rsize (rsimp (RALTS ((map (\s1. RSEQ (rders_simp r s1) (RSTAR r)) (star_updates list r [[a]])))))" + using star_closed_form by fastforce + also have "... \ Suc (rsizes (rdistinct (map (\s1. RSEQ (rders_simp r s1) (RSTAR r)) (star_updates list r [[a]])) {}))" + using alts_simp_control by blast + also have "... \ Suc (card (sizeNregex (Suc (N + rsize (RSTAR r))))) * (Suc (N + rsize (RSTAR r)))" + using star_control_bounded[OF assms] by (metis add_mono le_add1 mult_Suc plus_1_eq_Suc) + also have "... \ max (Suc (card (sizeNregex (Suc (N + rsize (RSTAR r))))) * Suc (N + rsize (RSTAR r))) (rsize (RSTAR r))" + by simp + finally show ?thesis by simp +qed + + +lemma seq_estimate_bounded: + assumes "\s. rsize (rders_simp r1 s) \ N1" + and "\s. rsize (rders_simp r2 s) \ N2" + shows + "rsizes (rdistinct (RSEQ (rders_simp r1 s) r2 # map (rders_simp r2) (vsuf s r1)) {}) + \ (Suc (N1 + (rsize r2)) + (N2 * card (sizeNregex N2)))" +proof - + have a: "rsizes (rdistinct (map (rders_simp r2) (vsuf s r1)) {}) \ N2 * card (sizeNregex N2)" + by (metis assms(2) distinct_list_rexp_upto ex_map_conv mult.commute) + + have "rsizes (rdistinct (RSEQ (rders_simp r1 s) r2 # map (rders_simp r2) (vsuf s r1)) {}) \ + rsize (RSEQ (rders_simp r1 s) r2) + rsizes (rdistinct (map (rders_simp r2) (vsuf s r1)) {})" + using triangle_inequality_distinct by blast + also have "... \ rsize (RSEQ (rders_simp r1 s) r2) + N2 * card (sizeNregex N2)" + by (simp add: a) + also have "... \ Suc (N1 + (rsize r2) + N2 * card (sizeNregex N2))" + by (simp add: assms(1)) + finally show ?thesis + by force +qed + + +lemma seq_closed_form_bounded2: + assumes "\s. rsize (rders_simp r1 s) \ N1" + and "\s. rsize (rders_simp r2 s) \ N2" +shows "rsize (rders_simp (RSEQ r1 r2) s) + \ max (2 + N1 + (rsize r2) + (N2 * card (sizeNregex N2))) (rsize (RSEQ r1 r2))" +proof(cases s) + case Nil + then show "rsize (rders_simp (RSEQ r1 r2) s) + \ max (2 + N1 + (rsize r2) + (N2 * card (sizeNregex N2))) (rsize (RSEQ r1 r2))" + by simp +next + case (Cons a list) + then have "rsize (rders_simp (RSEQ r1 r2) s) = + rsize (rsimp (RALTS ((RSEQ (rders_simp r1 s) r2) # (map (rders_simp r2) (vsuf s r1)))))" + using seq_closed_form_variant by (metis list.distinct(1)) + also have "... \ Suc (rsizes (rdistinct (RSEQ (rders_simp r1 s) r2 # map (rders_simp r2) (vsuf s r1)) {}))" + using alts_simp_control by blast + also have "... \ 2 + N1 + (rsize r2) + (N2 * card (sizeNregex N2))" + using seq_estimate_bounded[OF assms] by auto + ultimately show "rsize (rders_simp (RSEQ r1 r2) s) + \ max (2 + N1 + (rsize r2) + N2 * card (sizeNregex N2)) (rsize (RSEQ r1 r2))" + by auto +qed + + +lemma rders_simp_bounded: + shows "\N. \s. rsize (rders_simp r s) \ N" + apply(induct r) + apply(rule_tac x = "Suc 0 " in exI) + using three_easy_cases0 apply force + using three_easy_cases1 apply blast + using three_easy_casesC apply blast + apply(erule exE)+ + apply(rule exI) + apply(rule allI) + apply(rule seq_closed_form_bounded2) + apply(assumption) + apply(assumption) + apply (metis alts_closed_form_bounded size_list_estimation') + using star_closed_form_bounded by blast + + +unused_thms + +end diff -r f9cdc295ccf7 -r f493a20feeb3 thys3/src/FBound.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys3/src/FBound.thy Sat Apr 30 00:50:08 2022 +0100 @@ -0,0 +1,180 @@ + +theory FBound + imports "BlexerSimp" "ClosedFormsBounds" +begin + +fun distinctBy :: "'a list \ ('a \ 'b) \ 'b set \ 'a list" + where + "distinctBy [] f acc = []" +| "distinctBy (x#xs) f acc = + (if (f x) \ acc then distinctBy xs f acc + else x # (distinctBy xs f ({f x} \ acc)))" + +fun rerase :: "arexp \ rrexp" +where + "rerase AZERO = RZERO" +| "rerase (AONE _) = RONE" +| "rerase (ACHAR _ c) = RCHAR c" +| "rerase (AALTs bs rs) = RALTS (map rerase rs)" +| "rerase (ASEQ _ r1 r2) = RSEQ (rerase r1) (rerase r2)" +| "rerase (ASTAR _ r) = RSTAR (rerase r)" + +lemma eq1_rerase: + shows "x ~1 y \ (rerase x) = (rerase y)" + apply(induct x y rule: eq1.induct) + apply(auto) + done + + +lemma distinctBy_distinctWith: + shows "distinctBy xs f (f ` acc) = distinctWith xs (\x y. f x = f y) acc" + apply(induct xs arbitrary: acc) + apply(auto) + by (metis image_insert) + +lemma distinctBy_distinctWith2: + shows "distinctBy xs rerase {} = distinctWith xs eq1 {}" + apply(subst distinctBy_distinctWith[of _ _ "{}", simplified]) + using eq1_rerase by presburger + +lemma asize_rsize: + shows "rsize (rerase r) = asize r" + apply(induct r rule: rerase.induct) + apply(auto) + apply (metis (mono_tags, lifting) comp_apply map_eq_conv) + done + +lemma rerase_fuse: + shows "rerase (fuse bs r) = rerase r" + apply(induct r) + apply simp+ + done + +lemma rerase_bsimp_ASEQ: + shows "rerase (bsimp_ASEQ x1 a1 a2) = rsimp_SEQ (rerase a1) (rerase a2)" + apply(induct x1 a1 a2 rule: bsimp_ASEQ.induct) + apply(auto) + done + +lemma rerase_bsimp_AALTs: + shows "rerase (bsimp_AALTs bs rs) = rsimp_ALTs (map rerase rs)" + apply(induct bs rs rule: bsimp_AALTs.induct) + apply(auto simp add: rerase_fuse) + done + +fun anonalt :: "arexp \ bool" + where + "anonalt (AALTs bs2 rs) = False" +| "anonalt r = True" + + +fun agood :: "arexp \ bool" where + "agood AZERO = False" +| "agood (AONE cs) = True" +| "agood (ACHAR cs c) = True" +| "agood (AALTs cs []) = False" +| "agood (AALTs cs [r]) = False" +| "agood (AALTs cs (r1#r2#rs)) = (distinct (map rerase (r1 # r2 # rs)) \(\r' \ set (r1#r2#rs). agood r' \ anonalt r'))" +| "agood (ASEQ _ AZERO _) = False" +| "agood (ASEQ _ (AONE _) _) = False" +| "agood (ASEQ _ _ AZERO) = False" +| "agood (ASEQ cs r1 r2) = (agood r1 \ agood r2)" +| "agood (ASTAR cs r) = True" + + +fun anonnested :: "arexp \ bool" + where + "anonnested (AALTs bs2 []) = True" +| "anonnested (AALTs bs2 ((AALTs bs1 rs1) # rs2)) = False" +| "anonnested (AALTs bs2 (r # rs2)) = anonnested (AALTs bs2 rs2)" +| "anonnested r = True" + + +lemma asize0: + shows "0 < asize r" + apply(induct r) + apply(auto) + done + +lemma rnullable: + shows "rnullable (rerase r) = bnullable r" + apply(induct r rule: rerase.induct) + apply(auto) + done + +lemma rder_bder_rerase: + shows "rder c (rerase r ) = rerase (bder c r)" + apply (induct r) + apply (auto) + using rerase_fuse apply presburger + using rnullable apply blast + using rnullable by blast + +lemma rerase_map_bsimp: + assumes "\ r. r \ set rs \ rerase (bsimp r) = (rsimp \ rerase) r" + shows "map rerase (map bsimp rs) = map (rsimp \ rerase) rs" + using assms + apply(induct rs) + by simp_all + + +lemma rerase_flts: + shows "map rerase (flts rs) = rflts (map rerase rs)" + apply(induct rs rule: flts.induct) + apply(auto simp add: rerase_fuse) + done + +lemma rerase_dB: + shows "map rerase (distinctBy rs rerase acc) = rdistinct (map rerase rs) acc" + apply(induct rs arbitrary: acc) + apply simp+ + done + +lemma rerase_earlier_later_same: + assumes " \r. r \ set rs \ rerase (bsimp r) = rsimp (rerase r)" + shows " (map rerase (distinctBy (flts (map bsimp rs)) rerase {})) = + (rdistinct (rflts (map (rsimp \ rerase) rs)) {})" + apply(subst rerase_dB) + apply(subst rerase_flts) + apply(subst rerase_map_bsimp) + apply auto + using assms + apply simp + done + +lemma bsimp_rerase: + shows "rerase (bsimp a) = rsimp (rerase a)" + apply(induct a rule: bsimp.induct) + apply(auto) + using rerase_bsimp_ASEQ apply presburger + using distinctBy_distinctWith2 rerase_bsimp_AALTs rerase_earlier_later_same by fastforce + +lemma rders_simp_size: + shows "rders_simp (rerase r) s = rerase (bders_simp r s)" + apply(induct s rule: rev_induct) + apply simp + by (simp add: bders_simp_append rder_bder_rerase rders_simp_append bsimp_rerase) + + +corollary aders_simp_finiteness: + assumes "\N. \s. rsize (rders_simp (rerase r) s) \ N" + shows " \N. \s. asize (bders_simp r s) \ N" +proof - + from assms obtain N where "\s. rsize (rders_simp (rerase r) s) \ N" + by blast + then have "\s. rsize (rerase (bders_simp r s)) \ N" + by (simp add: rders_simp_size) + then have "\s. asize (bders_simp r s) \ N" + by (simp add: asize_rsize) + then show "\N. \s. asize (bders_simp r s) \ N" by blast +qed + +theorem annotated_size_bound: + shows "\N. \s. asize (bders_simp r s) \ N" + apply(insert aders_simp_finiteness) + by (simp add: rders_simp_bounded) + + +unused_thms + +end diff -r f9cdc295ccf7 -r f493a20feeb3 thys3/src/GeneralRegexBound.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys3/src/GeneralRegexBound.thy Sat Apr 30 00:50:08 2022 +0100 @@ -0,0 +1,212 @@ +theory GeneralRegexBound + imports "BasicIdentities" +begin + +lemma size_geq1: + shows "rsize r \ 1" + by (induct r) auto + +definition RSEQ_set where + "RSEQ_set A n \ {RSEQ r1 r2 | r1 r2. r1 \ A \ r2 \ A \ rsize r1 + rsize r2 \ n}" + +definition RSEQ_set_cartesian where + "RSEQ_set_cartesian A = {RSEQ r1 r2 | r1 r2. r1 \ A \ r2 \ A}" + +definition RALT_set where + "RALT_set A n \ {RALTS rs | rs. set rs \ A \ rsizes rs \ n}" + +definition RALTs_set where + "RALTs_set A n \ {RALTS rs | rs. \r \ set rs. r \ A \ rsizes rs \ n}" + +definition + "sizeNregex N \ {r. rsize r \ N}" + + +lemma sizenregex_induct1: + "sizeNregex (Suc n) = (({RZERO, RONE} \ {RCHAR c| c. True}) + \ (RSTAR ` sizeNregex n) + \ (RSEQ_set (sizeNregex n) n) + \ (RALTs_set (sizeNregex n) n))" + apply(auto) + apply(case_tac x) + apply(auto simp add: RSEQ_set_def) + using sizeNregex_def apply force + using sizeNregex_def apply auto[1] + apply (simp add: sizeNregex_def) + apply (simp add: sizeNregex_def) + apply (simp add: RALTs_set_def) + apply (metis imageI list.set_map member_le_sum_list order_trans) + apply (simp add: sizeNregex_def) + apply (simp add: sizeNregex_def) + apply (simp add: sizeNregex_def) + using sizeNregex_def apply force + apply (simp add: sizeNregex_def) + apply (simp add: sizeNregex_def) + apply (simp add: RALTs_set_def) + apply(simp add: sizeNregex_def) + apply(auto) + using ex_in_conv by fastforce + +lemma s4: + "RSEQ_set A n \ RSEQ_set_cartesian A" + using RSEQ_set_cartesian_def RSEQ_set_def by fastforce + +lemma s5: + assumes "finite A" + shows "finite (RSEQ_set_cartesian A)" + using assms + apply(subgoal_tac "RSEQ_set_cartesian A = (\(x1, x2). RSEQ x1 x2) ` (A \ A)") + apply simp + unfolding RSEQ_set_cartesian_def + apply(auto) + done + + +definition RALTs_set_length + where + "RALTs_set_length A n l \ {RALTS rs | rs. \r \ set rs. r \ A \ rsizes rs \ n \ length rs \ l}" + + +definition RALTs_set_length2 + where + "RALTs_set_length2 A l \ {RALTS rs | rs. \r \ set rs. r \ A \ length rs \ l}" + +definition set_length2 + where + "set_length2 A l \ {rs. \r \ set rs. r \ A \ length rs \ l}" + + +lemma r000: + shows "RALTs_set_length A n l \ RALTs_set_length2 A l" + apply(auto simp add: RALTs_set_length2_def RALTs_set_length_def) + done + + +lemma r02: + shows "set_length2 A 0 \ {[]}" + apply(auto simp add: set_length2_def) + apply(case_tac x) + apply(auto) + done + +lemma r03: + shows "set_length2 A (Suc n) \ + {[]} \ (\(h, t). h # t) ` (A \ (set_length2 A n))" + apply(auto simp add: set_length2_def) + apply(case_tac x) + apply(auto) + done + +lemma r1: + assumes "finite A" + shows "finite (set_length2 A n)" + using assms + apply(induct n) + apply(rule finite_subset) + apply(rule r02) + apply(simp) + apply(rule finite_subset) + apply(rule r03) + apply(simp) + done + +lemma size_sum_more_than_len: + shows "rsizes rs \ length rs" + apply(induct rs) + apply simp + apply simp + apply(subgoal_tac "rsize a \ 1") + apply linarith + using size_geq1 by auto + + +lemma sum_list_len: + shows "rsizes rs \ n \ length rs \ n" + by (meson order.trans size_sum_more_than_len) + + +lemma t2: + shows "RALTs_set A n \ RALTs_set_length A n n" + unfolding RALTs_set_length_def RALTs_set_def + apply(auto) + using sum_list_len by blast + +lemma s8_aux: + assumes "finite A" + shows "finite (RALTs_set_length A n n)" +proof - + have "finite A" by fact + then have "finite (set_length2 A n)" + by (simp add: r1) + moreover have "(RALTS ` (set_length2 A n)) = RALTs_set_length2 A n" + unfolding RALTs_set_length2_def set_length2_def + by (auto) + ultimately have "finite (RALTs_set_length2 A n)" + by (metis finite_imageI) + then show ?thesis + by (metis infinite_super r000) +qed + +lemma char_finite: + shows "finite {RCHAR c |c. True}" + apply simp + apply(subgoal_tac "finite (RCHAR ` (UNIV::char set))") + prefer 2 + apply simp + by (simp add: full_SetCompr_eq) + + +lemma finite_size_n: + shows "finite (sizeNregex n)" + apply(induct n) + apply(simp add: sizeNregex_def) + apply (metis (mono_tags, lifting) not_finite_existsD not_one_le_zero size_geq1) + apply(subst sizenregex_induct1) + apply(simp only: finite_Un) + apply(rule conjI)+ + apply(simp) + + using char_finite apply blast + apply(simp) + apply(rule finite_subset) + apply(rule s4) + apply(rule s5) + apply(simp) + apply(rule finite_subset) + apply(rule t2) + apply(rule s8_aux) + apply(simp) + done + +lemma three_easy_cases0: + shows "rsize (rders_simp RZERO s) \ Suc 0" + apply(induct s) + apply simp + apply simp + done + + +lemma three_easy_cases1: + shows "rsize (rders_simp RONE s) \ Suc 0" + apply(induct s) + apply simp + apply simp + using three_easy_cases0 by auto + + +lemma three_easy_casesC: + shows "rsize (rders_simp (RCHAR c) s) \ Suc 0" + apply(induct s) + apply simp + apply simp + apply(case_tac " a = c") + using three_easy_cases1 apply blast + apply simp + using three_easy_cases0 by force + + +unused_thms + + +end + diff -r f9cdc295ccf7 -r f493a20feeb3 thys3/src/Lexer.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys3/src/Lexer.thy Sat Apr 30 00:50:08 2022 +0100 @@ -0,0 +1,417 @@ + +theory Lexer + imports PosixSpec +begin + +section {* The Lexer Functions by Sulzmann and Lu (without simplification) *} + +fun + mkeps :: "rexp \ val" +where + "mkeps(ONE) = Void" +| "mkeps(SEQ r1 r2) = Seq (mkeps r1) (mkeps r2)" +| "mkeps(ALT r1 r2) = (if nullable(r1) then Left (mkeps r1) else Right (mkeps r2))" +| "mkeps(STAR r) = Stars []" + +fun injval :: "rexp \ char \ val \ val" +where + "injval (CH d) c Void = Char d" +| "injval (ALT r1 r2) c (Left v1) = Left(injval r1 c v1)" +| "injval (ALT r1 r2) c (Right v2) = Right(injval r2 c v2)" +| "injval (SEQ r1 r2) c (Seq v1 v2) = Seq (injval r1 c v1) v2" +| "injval (SEQ r1 r2) c (Left (Seq v1 v2)) = Seq (injval r1 c v1) v2" +| "injval (SEQ r1 r2) c (Right v2) = Seq (mkeps r1) (injval r2 c v2)" +| "injval (STAR r) c (Seq v (Stars vs)) = Stars ((injval r c v) # vs)" + +fun + lexer :: "rexp \ string \ val option" +where + "lexer r [] = (if nullable r then Some(mkeps r) else None)" +| "lexer r (c#s) = (case (lexer (der c r) s) of + None \ None + | Some(v) \ Some(injval r c v))" + + + +section {* Mkeps, Injval Properties *} + +lemma mkeps_nullable: + assumes "nullable(r)" + shows "\ mkeps r : r" +using assms +by (induct rule: nullable.induct) + (auto intro: Prf.intros) + +lemma mkeps_flat: + assumes "nullable(r)" + shows "flat (mkeps r) = []" +using assms +by (induct rule: nullable.induct) (auto) + +lemma Prf_injval_flat: + assumes "\ v : der c r" + shows "flat (injval r c v) = c # (flat v)" +using assms +apply(induct c r arbitrary: v rule: der.induct) +apply(auto elim!: Prf_elims intro: mkeps_flat split: if_splits) +done + +lemma Prf_injval: + assumes "\ v : der c r" + shows "\ (injval r c v) : r" +using assms +apply(induct r arbitrary: c v rule: rexp.induct) +apply(auto intro!: Prf.intros mkeps_nullable elim!: Prf_elims split: if_splits) +apply(simp add: Prf_injval_flat) +done + + + +text {* + Mkeps and injval produce, or preserve, Posix values. +*} + +lemma Posix_mkeps: + assumes "nullable r" + shows "[] \ r \ mkeps r" +using assms +apply(induct r rule: nullable.induct) +apply(auto intro: Posix.intros simp add: nullable_correctness Sequ_def) +apply(subst append.simps(1)[symmetric]) +apply(rule Posix.intros) +apply(auto) +done + +lemma Posix_injval: + assumes "s \ (der c r) \ v" + shows "(c # s) \ r \ (injval r c v)" +using assms +proof(induct r arbitrary: s v rule: rexp.induct) + case ZERO + have "s \ der c ZERO \ v" by fact + then have "s \ ZERO \ v" by simp + then have "False" by cases + then show "(c # s) \ ZERO \ (injval ZERO c v)" by simp +next + case ONE + have "s \ der c ONE \ v" by fact + then have "s \ ZERO \ v" by simp + then have "False" by cases + then show "(c # s) \ ONE \ (injval ONE c v)" by simp +next + case (CH d) + consider (eq) "c = d" | (ineq) "c \ d" by blast + then show "(c # s) \ (CH d) \ (injval (CH d) c v)" + proof (cases) + case eq + have "s \ der c (CH d) \ v" by fact + then have "s \ ONE \ v" using eq by simp + then have eqs: "s = [] \ v = Void" by cases simp + show "(c # s) \ CH d \ injval (CH d) c v" using eq eqs + by (auto intro: Posix.intros) + next + case ineq + have "s \ der c (CH d) \ v" by fact + then have "s \ ZERO \ v" using ineq by simp + then have "False" by cases + then show "(c # s) \ CH d \ injval (CH d) c v" by simp + qed +next + case (ALT r1 r2) + have IH1: "\s v. s \ der c r1 \ v \ (c # s) \ r1 \ injval r1 c v" by fact + have IH2: "\s v. s \ der c r2 \ v \ (c # s) \ r2 \ injval r2 c v" by fact + have "s \ der c (ALT r1 r2) \ v" by fact + then have "s \ ALT (der c r1) (der c r2) \ v" by simp + then consider (left) v' where "v = Left v'" "s \ der c r1 \ v'" + | (right) v' where "v = Right v'" "s \ L (der c r1)" "s \ der c r2 \ v'" + by cases auto + then show "(c # s) \ ALT r1 r2 \ injval (ALT r1 r2) c v" + proof (cases) + case left + have "s \ der c r1 \ v'" by fact + then have "(c # s) \ r1 \ injval r1 c v'" using IH1 by simp + then have "(c # s) \ ALT r1 r2 \ injval (ALT r1 r2) c (Left v')" by (auto intro: Posix.intros) + then show "(c # s) \ ALT r1 r2 \ injval (ALT r1 r2) c v" using left by simp + next + case right + have "s \ L (der c r1)" by fact + then have "c # s \ L r1" by (simp add: der_correctness Der_def) + moreover + have "s \ der c r2 \ v'" by fact + then have "(c # s) \ r2 \ injval r2 c v'" using IH2 by simp + ultimately have "(c # s) \ ALT r1 r2 \ injval (ALT r1 r2) c (Right v')" + by (auto intro: Posix.intros) + then show "(c # s) \ ALT r1 r2 \ injval (ALT r1 r2) c v" using right by simp + qed +next + case (SEQ r1 r2) + have IH1: "\s v. s \ der c r1 \ v \ (c # s) \ r1 \ injval r1 c v" by fact + have IH2: "\s v. s \ der c r2 \ v \ (c # s) \ r2 \ injval r2 c v" by fact + have "s \ der c (SEQ r1 r2) \ v" by fact + then consider + (left_nullable) v1 v2 s1 s2 where + "v = Left (Seq v1 v2)" "s = s1 @ s2" + "s1 \ der c r1 \ v1" "s2 \ r2 \ v2" "nullable r1" + "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L (der c r1) \ s\<^sub>4 \ L r2)" + | (right_nullable) v1 s1 s2 where + "v = Right v1" "s = s1 @ s2" + "s \ der c r2 \ v1" "nullable r1" "s1 @ s2 \ L (SEQ (der c r1) r2)" + | (not_nullable) v1 v2 s1 s2 where + "v = Seq v1 v2" "s = s1 @ s2" + "s1 \ der c r1 \ v1" "s2 \ r2 \ v2" "\nullable r1" + "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L (der c r1) \ s\<^sub>4 \ L r2)" + by (force split: if_splits elim!: Posix_elims simp add: Sequ_def der_correctness Der_def) + then show "(c # s) \ SEQ r1 r2 \ injval (SEQ r1 r2) c v" + proof (cases) + case left_nullable + have "s1 \ der c r1 \ v1" by fact + then have "(c # s1) \ r1 \ injval r1 c v1" using IH1 by simp + moreover + have "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L (der c r1) \ s\<^sub>4 \ L r2)" by fact + then have "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (c # s1) @ s\<^sub>3 \ L r1 \ s\<^sub>4 \ L r2)" by (simp add: der_correctness Der_def) + ultimately have "((c # s1) @ s2) \ SEQ r1 r2 \ Seq (injval r1 c v1) v2" using left_nullable by (rule_tac Posix.intros) + then show "(c # s) \ SEQ r1 r2 \ injval (SEQ r1 r2) c v" using left_nullable by simp + next + case right_nullable + have "nullable r1" by fact + then have "[] \ r1 \ (mkeps r1)" by (rule Posix_mkeps) + moreover + have "s \ der c r2 \ v1" by fact + then have "(c # s) \ r2 \ (injval r2 c v1)" using IH2 by simp + moreover + have "s1 @ s2 \ L (SEQ (der c r1) r2)" by fact + then have "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = c # s \ [] @ s\<^sub>3 \ L r1 \ s\<^sub>4 \ L r2)" using right_nullable + by(auto simp add: der_correctness Der_def append_eq_Cons_conv Sequ_def) + ultimately have "([] @ (c # s)) \ SEQ r1 r2 \ Seq (mkeps r1) (injval r2 c v1)" + by(rule Posix.intros) + then show "(c # s) \ SEQ r1 r2 \ injval (SEQ r1 r2) c v" using right_nullable by simp + next + case not_nullable + have "s1 \ der c r1 \ v1" by fact + then have "(c # s1) \ r1 \ injval r1 c v1" using IH1 by simp + moreover + have "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L (der c r1) \ s\<^sub>4 \ L r2)" by fact + then have "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (c # s1) @ s\<^sub>3 \ L r1 \ s\<^sub>4 \ L r2)" by (simp add: der_correctness Der_def) + ultimately have "((c # s1) @ s2) \ SEQ r1 r2 \ Seq (injval r1 c v1) v2" using not_nullable + by (rule_tac Posix.intros) (simp_all) + then show "(c # s) \ SEQ r1 r2 \ injval (SEQ r1 r2) c v" using not_nullable by simp + qed +next + case (STAR r) + have IH: "\s v. s \ der c r \ v \ (c # s) \ r \ injval r c v" by fact + have "s \ der c (STAR r) \ v" by fact + then consider + (cons) v1 vs s1 s2 where + "v = Seq v1 (Stars vs)" "s = s1 @ s2" + "s1 \ der c r \ v1" "s2 \ (STAR r) \ (Stars vs)" + "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L (der c r) \ s\<^sub>4 \ L (STAR r))" + apply(auto elim!: Posix_elims(1-5) simp add: der_correctness Der_def intro: Posix.intros) + apply(rotate_tac 3) + apply(erule_tac Posix_elims(6)) + apply (simp add: Posix.intros(6)) + using Posix.intros(7) by blast + then show "(c # s) \ STAR r \ injval (STAR r) c v" + proof (cases) + case cons + have "s1 \ der c r \ v1" by fact + then have "(c # s1) \ r \ injval r c v1" using IH by simp + moreover + have "s2 \ STAR r \ Stars vs" by fact + moreover + have "(c # s1) \ r \ injval r c v1" by fact + then have "flat (injval r c v1) = (c # s1)" by (rule Posix1) + then have "flat (injval r c v1) \ []" by simp + moreover + have "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L (der c r) \ s\<^sub>4 \ L (STAR r))" by fact + then have "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (c # s1) @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (STAR r))" + by (simp add: der_correctness Der_def) + ultimately + have "((c # s1) @ s2) \ STAR r \ Stars (injval r c v1 # vs)" by (rule Posix.intros) + then show "(c # s) \ STAR r \ injval (STAR r) c v" using cons by(simp) + qed +qed + + +section {* Lexer Correctness *} + + +lemma lexer_correct_None: + shows "s \ L r \ lexer r s = None" + apply(induct s arbitrary: r) + apply(simp) + apply(simp add: nullable_correctness) + apply(simp) + apply(drule_tac x="der a r" in meta_spec) + apply(auto) + apply(auto simp add: der_correctness Der_def) +done + +lemma lexer_correct_Some: + shows "s \ L r \ (\v. lexer r s = Some(v) \ s \ r \ v)" + apply(induct s arbitrary : r) + apply(simp only: lexer.simps) + apply(simp) + apply(simp add: nullable_correctness Posix_mkeps) + apply(drule_tac x="der a r" in meta_spec) + apply(simp (no_asm_use) add: der_correctness Der_def del: lexer.simps) + apply(simp del: lexer.simps) + apply(simp only: lexer.simps) + apply(case_tac "lexer (der a r) s = None") + apply(auto)[1] + apply(simp) + apply(erule exE) + apply(simp) + apply(rule iffI) + apply(simp add: Posix_injval) + apply(simp add: Posix1(1)) +done + +lemma lexer_correctness: + shows "(lexer r s = Some v) \ s \ r \ v" + and "(lexer r s = None) \ \(\v. s \ r \ v)" +using Posix1(1) Posix_determ lexer_correct_None lexer_correct_Some apply fastforce +using Posix1(1) lexer_correct_None lexer_correct_Some by blast + + +subsection {* A slight reformulation of the lexer algorithm using stacked functions*} + +fun flex :: "rexp \ (val \ val) => string \ (val \ val)" + where + "flex r f [] = f" +| "flex r f (c#s) = flex (der c r) (\v. f (injval r c v)) s" + +lemma flex_fun_apply: + shows "g (flex r f s v) = flex r (g o f) s v" + apply(induct s arbitrary: g f r v) + apply(simp_all add: comp_def) + by meson + +lemma flex_fun_apply2: + shows "g (flex r id s v) = flex r g s v" + by (simp add: flex_fun_apply) + + +lemma flex_append: + shows "flex r f (s1 @ s2) = flex (ders s1 r) (flex r f s1) s2" + apply(induct s1 arbitrary: s2 r f) + apply(simp_all) + done + +lemma lexer_flex: + shows "lexer r s = (if nullable (ders s r) + then Some(flex r id s (mkeps (ders s r))) else None)" + apply(induct s arbitrary: r) + apply(simp_all add: flex_fun_apply) + done + +lemma Posix_flex: + assumes "s2 \ (ders s1 r) \ v" + shows "(s1 @ s2) \ r \ flex r id s1 v" + using assms + apply(induct s1 arbitrary: r v s2) + apply(simp) + apply(simp) + apply(drule_tac x="der a r" in meta_spec) + apply(drule_tac x="v" in meta_spec) + apply(drule_tac x="s2" in meta_spec) + apply(simp) + using Posix_injval + apply(drule_tac Posix_injval) + apply(subst (asm) (5) flex_fun_apply) + apply(simp) + done + +lemma injval_inj: + assumes "\ a : (der c r)" "\ v : (der c r)" "injval r c a = injval r c v" + shows "a = v" + using assms + apply(induct r arbitrary: a c v) + apply(auto) + using Prf_elims(1) apply blast + using Prf_elims(1) apply blast + apply(case_tac "c = x") + apply(auto) + using Prf_elims(4) apply auto[1] + using Prf_elims(1) apply blast + prefer 2 + apply (smt Prf_elims(3) injval.simps(2) injval.simps(3) val.distinct(25) val.inject(3) val.inject(4)) + apply(case_tac "nullable r1") + apply(auto) + apply(erule Prf_elims) + apply(erule Prf_elims) + apply(erule Prf_elims) + apply(erule Prf_elims) + apply(auto) + apply (metis Prf_injval_flat list.distinct(1) mkeps_flat) + apply(erule Prf_elims) + apply(erule Prf_elims) + apply(auto) + using Prf_injval_flat mkeps_flat apply fastforce + apply(erule Prf_elims) + apply(erule Prf_elims) + apply(auto) + apply(erule Prf_elims) + apply(erule Prf_elims) + apply(auto) + apply (smt Prf_elims(6) injval.simps(7) list.inject val.inject(5)) + by (smt Prf_elims(6) injval.simps(7) list.inject val.inject(5)) + + + +lemma uu: + assumes "(c # s) \ r \ injval r c v" "\ v : (der c r)" + shows "s \ der c r \ v" + using assms + apply - + apply(subgoal_tac "lexer r (c # s) = Some (injval r c v)") + prefer 2 + using lexer_correctness(1) apply blast + apply(simp add: ) + apply(case_tac "lexer (der c r) s") + apply(simp) + apply(simp) + apply(case_tac "s \ der c r \ a") + prefer 2 + apply (simp add: lexer_correctness(1)) + apply(subgoal_tac "\ a : (der c r)") + prefer 2 + using Posix_Prf apply blast + using injval_inj by blast + + +lemma Posix_flex2: + assumes "(s1 @ s2) \ r \ flex r id s1 v" "\ v : ders s1 r" + shows "s2 \ (ders s1 r) \ v" + using assms + apply(induct s1 arbitrary: r v s2 rule: rev_induct) + apply(simp) + apply(simp) + apply(drule_tac x="r" in meta_spec) + apply(drule_tac x="injval (ders xs r) x v" in meta_spec) + apply(drule_tac x="x#s2" in meta_spec) + apply(simp add: flex_append ders_append) + using Prf_injval uu by blast + +lemma Posix_flex3: + assumes "s1 \ r \ flex r id s1 v" "\ v : ders s1 r" + shows "[] \ (ders s1 r) \ v" + using assms + by (simp add: Posix_flex2) + +lemma flex_injval: + shows "flex (der a r) (injval r a) s v = injval r a (flex (der a r) id s v)" + by (simp add: flex_fun_apply) + +lemma Prf_flex: + assumes "\ v : ders s r" + shows "\ flex r id s v : r" + using assms + apply(induct s arbitrary: v r) + apply(simp) + apply(simp) + by (simp add: Prf_injval flex_injval) + + +unused_thms + +end \ No newline at end of file diff -r f9cdc295ccf7 -r f493a20feeb3 thys3/src/LexerSimp.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys3/src/LexerSimp.thy Sat Apr 30 00:50:08 2022 +0100 @@ -0,0 +1,246 @@ +theory LexerSimp + imports "Lexer" +begin + +section {* Lexer including some simplifications *} + + +fun F_RIGHT where + "F_RIGHT f v = Right (f v)" + +fun F_LEFT where + "F_LEFT f v = Left (f v)" + +fun F_ALT where + "F_ALT f\<^sub>1 f\<^sub>2 (Right v) = Right (f\<^sub>2 v)" +| "F_ALT f\<^sub>1 f\<^sub>2 (Left v) = Left (f\<^sub>1 v)" +| "F_ALT f1 f2 v = v" + + +fun F_SEQ1 where + "F_SEQ1 f\<^sub>1 f\<^sub>2 v = Seq (f\<^sub>1 Void) (f\<^sub>2 v)" + +fun F_SEQ2 where + "F_SEQ2 f\<^sub>1 f\<^sub>2 v = Seq (f\<^sub>1 v) (f\<^sub>2 Void)" + +fun F_SEQ where + "F_SEQ f\<^sub>1 f\<^sub>2 (Seq v\<^sub>1 v\<^sub>2) = Seq (f\<^sub>1 v\<^sub>1) (f\<^sub>2 v\<^sub>2)" +| "F_SEQ f1 f2 v = v" + +fun simp_ALT where + "simp_ALT (ZERO, f\<^sub>1) (r\<^sub>2, f\<^sub>2) = (r\<^sub>2, F_RIGHT f\<^sub>2)" +| "simp_ALT (r\<^sub>1, f\<^sub>1) (ZERO, f\<^sub>2) = (r\<^sub>1, F_LEFT f\<^sub>1)" +| "simp_ALT (r\<^sub>1, f\<^sub>1) (r\<^sub>2, f\<^sub>2) = (ALT r\<^sub>1 r\<^sub>2, F_ALT f\<^sub>1 f\<^sub>2)" + + +fun simp_SEQ where + "simp_SEQ (ONE, f\<^sub>1) (r\<^sub>2, f\<^sub>2) = (r\<^sub>2, F_SEQ1 f\<^sub>1 f\<^sub>2)" +| "simp_SEQ (r\<^sub>1, f\<^sub>1) (ONE, f\<^sub>2) = (r\<^sub>1, F_SEQ2 f\<^sub>1 f\<^sub>2)" +| "simp_SEQ (ZERO, f\<^sub>1) (r\<^sub>2, f\<^sub>2) = (ZERO, undefined)" +| "simp_SEQ (r\<^sub>1, f\<^sub>1) (ZERO, f\<^sub>2) = (ZERO, undefined)" +| "simp_SEQ (r\<^sub>1, f\<^sub>1) (r\<^sub>2, f\<^sub>2) = (SEQ r\<^sub>1 r\<^sub>2, F_SEQ f\<^sub>1 f\<^sub>2)" + +lemma simp_SEQ_simps[simp]: + "simp_SEQ p1 p2 = (if (fst p1 = ONE) then (fst p2, F_SEQ1 (snd p1) (snd p2)) + else (if (fst p2 = ONE) then (fst p1, F_SEQ2 (snd p1) (snd p2)) + else (if (fst p1 = ZERO) then (ZERO, undefined) + else (if (fst p2 = ZERO) then (ZERO, undefined) + else (SEQ (fst p1) (fst p2), F_SEQ (snd p1) (snd p2))))))" +by (induct p1 p2 rule: simp_SEQ.induct) (auto) + +lemma simp_ALT_simps[simp]: + "simp_ALT p1 p2 = (if (fst p1 = ZERO) then (fst p2, F_RIGHT (snd p2)) + else (if (fst p2 = ZERO) then (fst p1, F_LEFT (snd p1)) + else (ALT (fst p1) (fst p2), F_ALT (snd p1) (snd p2))))" +by (induct p1 p2 rule: simp_ALT.induct) (auto) + +fun + simp :: "rexp \ rexp * (val \ val)" +where + "simp (ALT r1 r2) = simp_ALT (simp r1) (simp r2)" +| "simp (SEQ r1 r2) = simp_SEQ (simp r1) (simp r2)" +| "simp r = (r, id)" + +fun + slexer :: "rexp \ string \ val option" +where + "slexer r [] = (if nullable r then Some(mkeps r) else None)" +| "slexer r (c#s) = (let (rs, fr) = simp (der c r) in + (case (slexer rs s) of + None \ None + | Some(v) \ Some(injval r c (fr v))))" + + +lemma slexer_better_simp: + "slexer r (c#s) = (case (slexer (fst (simp (der c r))) s) of + None \ None + | Some(v) \ Some(injval r c ((snd (simp (der c r))) v)))" +by (auto split: prod.split option.split) + + +lemma L_fst_simp: + shows "L(r) = L(fst (simp r))" +by (induct r) (auto) + +lemma Posix_simp: + assumes "s \ (fst (simp r)) \ v" + shows "s \ r \ ((snd (simp r)) v)" +using assms +proof(induct r arbitrary: s v rule: rexp.induct) + case (ALT r1 r2 s v) + have IH1: "\s v. s \ fst (simp r1) \ v \ s \ r1 \ snd (simp r1) v" by fact + have IH2: "\s v. s \ fst (simp r2) \ v \ s \ r2 \ snd (simp r2) v" by fact + have as: "s \ fst (simp (ALT r1 r2)) \ v" by fact + consider (ZERO_ZERO) "fst (simp r1) = ZERO" "fst (simp r2) = ZERO" + | (ZERO_NZERO) "fst (simp r1) = ZERO" "fst (simp r2) \ ZERO" + | (NZERO_ZERO) "fst (simp r1) \ ZERO" "fst (simp r2) = ZERO" + | (NZERO_NZERO) "fst (simp r1) \ ZERO" "fst (simp r2) \ ZERO" by auto + then show "s \ ALT r1 r2 \ snd (simp (ALT r1 r2)) v" + proof(cases) + case (ZERO_ZERO) + with as have "s \ ZERO \ v" by simp + then show "s \ ALT r1 r2 \ snd (simp (ALT r1 r2)) v" by (rule Posix_elims(1)) + next + case (ZERO_NZERO) + with as have "s \ fst (simp r2) \ v" by simp + with IH2 have "s \ r2 \ snd (simp r2) v" by simp + moreover + from ZERO_NZERO have "fst (simp r1) = ZERO" by simp + then have "L (fst (simp r1)) = {}" by simp + then have "L r1 = {}" using L_fst_simp by simp + then have "s \ L r1" by simp + ultimately have "s \ ALT r1 r2 \ Right (snd (simp r2) v)" by (rule Posix_ALT2) + then show "s \ ALT r1 r2 \ snd (simp (ALT r1 r2)) v" + using ZERO_NZERO by simp + next + case (NZERO_ZERO) + with as have "s \ fst (simp r1) \ v" by simp + with IH1 have "s \ r1 \ snd (simp r1) v" by simp + then have "s \ ALT r1 r2 \ Left (snd (simp r1) v)" by (rule Posix_ALT1) + then show "s \ ALT r1 r2 \ snd (simp (ALT r1 r2)) v" using NZERO_ZERO by simp + next + case (NZERO_NZERO) + with as have "s \ ALT (fst (simp r1)) (fst (simp r2)) \ v" by simp + then consider (Left) v1 where "v = Left v1" "s \ (fst (simp r1)) \ v1" + | (Right) v2 where "v = Right v2" "s \ (fst (simp r2)) \ v2" "s \ L (fst (simp r1))" + by (erule_tac Posix_elims(4)) + then show "s \ ALT r1 r2 \ snd (simp (ALT r1 r2)) v" + proof(cases) + case (Left) + then have "v = Left v1" "s \ r1 \ (snd (simp r1) v1)" using IH1 by simp_all + then show "s \ ALT r1 r2 \ snd (simp (ALT r1 r2)) v" using NZERO_NZERO + by (simp_all add: Posix_ALT1) + next + case (Right) + then have "v = Right v2" "s \ r2 \ (snd (simp r2) v2)" "s \ L r1" using IH2 L_fst_simp by simp_all + then show "s \ ALT r1 r2 \ snd (simp (ALT r1 r2)) v" using NZERO_NZERO + by (simp_all add: Posix_ALT2) + qed + qed +next + case (SEQ r1 r2 s v) + have IH1: "\s v. s \ fst (simp r1) \ v \ s \ r1 \ snd (simp r1) v" by fact + have IH2: "\s v. s \ fst (simp r2) \ v \ s \ r2 \ snd (simp r2) v" by fact + have as: "s \ fst (simp (SEQ r1 r2)) \ v" by fact + consider (ONE_ONE) "fst (simp r1) = ONE" "fst (simp r2) = ONE" + | (ONE_NONE) "fst (simp r1) = ONE" "fst (simp r2) \ ONE" + | (NONE_ONE) "fst (simp r1) \ ONE" "fst (simp r2) = ONE" + | (NONE_NONE) "fst (simp r1) \ ONE" "fst (simp r2) \ ONE" + by auto + then show "s \ SEQ r1 r2 \ snd (simp (SEQ r1 r2)) v" + proof(cases) + case (ONE_ONE) + with as have b: "s \ ONE \ v" by simp + from b have "s \ r1 \ snd (simp r1) v" using IH1 ONE_ONE by simp + moreover + from b have c: "s = []" "v = Void" using Posix_elims(2) by auto + moreover + have "[] \ ONE \ Void" by (simp add: Posix_ONE) + then have "[] \ fst (simp r2) \ Void" using ONE_ONE by simp + then have "[] \ r2 \ snd (simp r2) Void" using IH2 by simp + ultimately have "([] @ []) \ SEQ r1 r2 \ Seq (snd (simp r1) Void) (snd (simp r2) Void)" + using Posix_SEQ by blast + then show "s \ SEQ r1 r2 \ snd (simp (SEQ r1 r2)) v" using c ONE_ONE by simp + next + case (ONE_NONE) + with as have b: "s \ fst (simp r2) \ v" by simp + from b have "s \ r2 \ snd (simp r2) v" using IH2 ONE_NONE by simp + moreover + have "[] \ ONE \ Void" by (simp add: Posix_ONE) + then have "[] \ fst (simp r1) \ Void" using ONE_NONE by simp + then have "[] \ r1 \ snd (simp r1) Void" using IH1 by simp + moreover + from ONE_NONE(1) have "L (fst (simp r1)) = {[]}" by simp + then have "L r1 = {[]}" by (simp add: L_fst_simp[symmetric]) + ultimately have "([] @ s) \ SEQ r1 r2 \ Seq (snd (simp r1) Void) (snd (simp r2) v)" + by(rule_tac Posix_SEQ) auto + then show "s \ SEQ r1 r2 \ snd (simp (SEQ r1 r2)) v" using ONE_NONE by simp + next + case (NONE_ONE) + with as have "s \ fst (simp r1) \ v" by simp + with IH1 have "s \ r1 \ snd (simp r1) v" by simp + moreover + have "[] \ ONE \ Void" by (simp add: Posix_ONE) + then have "[] \ fst (simp r2) \ Void" using NONE_ONE by simp + then have "[] \ r2 \ snd (simp r2) Void" using IH2 by simp + ultimately have "(s @ []) \ SEQ r1 r2 \ Seq (snd (simp r1) v) (snd (simp r2) Void)" + by(rule_tac Posix_SEQ) auto + then show "s \ SEQ r1 r2 \ snd (simp (SEQ r1 r2)) v" using NONE_ONE by simp + next + case (NONE_NONE) + from as have 00: "fst (simp r1) \ ZERO" "fst (simp r2) \ ZERO" + apply(auto) + apply(smt Posix_elims(1) fst_conv) + by (smt NONE_NONE(2) Posix_elims(1) fstI) + with NONE_NONE as have "s \ SEQ (fst (simp r1)) (fst (simp r2)) \ v" by simp + then obtain s1 s2 v1 v2 where eqs: "s = s1 @ s2" "v = Seq v1 v2" + "s1 \ (fst (simp r1)) \ v1" "s2 \ (fst (simp r2)) \ v2" + "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L r1 \ s\<^sub>4 \ L r2)" + by (erule_tac Posix_elims(5)) (auto simp add: L_fst_simp[symmetric]) + then have "s1 \ r1 \ (snd (simp r1) v1)" "s2 \ r2 \ (snd (simp r2) v2)" + using IH1 IH2 by auto + then show "s \ SEQ r1 r2 \ snd (simp (SEQ r1 r2)) v" using eqs NONE_NONE 00 + by(auto intro: Posix_SEQ) + qed +qed (simp_all) + + +lemma slexer_correctness: + shows "slexer r s = lexer r s" +proof(induct s arbitrary: r) + case Nil + show "slexer r [] = lexer r []" by simp +next + case (Cons c s r) + have IH: "\r. slexer r s = lexer r s" by fact + show "slexer r (c # s) = lexer r (c # s)" + proof (cases "s \ L (der c r)") + case True + assume a1: "s \ L (der c r)" + then obtain v1 where a2: "lexer (der c r) s = Some v1" "s \ der c r \ v1" + using lexer_correct_Some by auto + from a1 have "s \ L (fst (simp (der c r)))" using L_fst_simp[symmetric] by simp + then obtain v2 where a3: "lexer (fst (simp (der c r))) s = Some v2" "s \ (fst (simp (der c r))) \ v2" + using lexer_correct_Some by auto + then have a4: "slexer (fst (simp (der c r))) s = Some v2" using IH by simp + from a3(2) have "s \ der c r \ (snd (simp (der c r))) v2" using Posix_simp by simp + with a2(2) have "v1 = (snd (simp (der c r))) v2" using Posix_determ by simp + with a2(1) a4 show "slexer r (c # s) = lexer r (c # s)" by (auto split: prod.split) + next + case False + assume b1: "s \ L (der c r)" + then have "lexer (der c r) s = None" using lexer_correct_None by simp + moreover + from b1 have "s \ L (fst (simp (der c r)))" using L_fst_simp[symmetric] by simp + then have "lexer (fst (simp (der c r))) s = None" using lexer_correct_None by simp + then have "slexer (fst (simp (der c r))) s = None" using IH by simp + ultimately show "slexer r (c # s) = lexer r (c # s)" + by (simp del: slexer.simps add: slexer_better_simp) + qed + qed + + +unused_thms + + +end \ No newline at end of file diff -r f9cdc295ccf7 -r f493a20feeb3 thys3/src/PDerivs.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys3/src/PDerivs.thy Sat Apr 30 00:50:08 2022 +0100 @@ -0,0 +1,603 @@ + +theory PDerivs + imports PosixSpec +begin + + + +abbreviation + "SEQs rs r \ (\r' \ rs. {SEQ r' r})" + +lemma SEQs_eq_image: + "SEQs rs r = (\r'. SEQ r' r) ` rs" + by auto + +fun + pder :: "char \ rexp \ rexp set" +where + "pder c ZERO = {}" +| "pder c ONE = {}" +| "pder c (CH d) = (if c = d then {ONE} else {})" +| "pder c (ALT r1 r2) = (pder c r1) \ (pder c r2)" +| "pder c (SEQ r1 r2) = + (if nullable r1 then SEQs (pder c r1) r2 \ pder c r2 else SEQs (pder c r1) r2)" +| "pder c (STAR r) = SEQs (pder c r) (STAR r)" + +fun + pders :: "char list \ rexp \ rexp set" +where + "pders [] r = {r}" +| "pders (c # s) r = \ (pders s ` pder c r)" + +abbreviation + pder_set :: "char \ rexp set \ rexp set" +where + "pder_set c rs \ \ (pder c ` rs)" + +abbreviation + pders_set :: "char list \ rexp set \ rexp set" +where + "pders_set s rs \ \ (pders s ` rs)" + +lemma pders_append: + "pders (s1 @ s2) r = \ (pders s2 ` pders s1 r)" +by (induct s1 arbitrary: r) (simp_all) + +lemma pders_snoc: + shows "pders (s @ [c]) r = pder_set c (pders s r)" +by (simp add: pders_append) + +lemma pders_simps [simp]: + shows "pders s ZERO = (if s = [] then {ZERO} else {})" + and "pders s ONE = (if s = [] then {ONE} else {})" + and "pders s (ALT r1 r2) = (if s = [] then {ALT r1 r2} else (pders s r1) \ (pders s r2))" +by (induct s) (simp_all) + +lemma pders_CHAR: + shows "pders s (CH c) \ {CH c, ONE}" +by (induct s) (simp_all) + +subsection \Relating left-quotients and partial derivatives\ + +lemma Sequ_UNION_distrib: +shows "A ;; \(M ` I) = \((\i. A ;; M i) ` I)" +and "\(M ` I) ;; A = \((\i. M i ;; A) ` I)" +by (auto simp add: Sequ_def) + + +lemma Der_pder: + shows "Der c (L r) = \ (L ` pder c r)" +by (induct r) (simp_all add: nullable_correctness Sequ_UNION_distrib) + +lemma Ders_pders: + shows "Ders s (L r) = \ (L ` pders s r)" +proof (induct s arbitrary: r) + case (Cons c s) + have ih: "\r. Ders s (L r) = \ (L ` pders s r)" by fact + have "Ders (c # s) (L r) = Ders s (Der c (L r))" by (simp add: Ders_def Der_def) + also have "\ = Ders s (\ (L ` pder c r))" by (simp add: Der_pder) + also have "\ = (\A\(L ` (pder c r)). (Ders s A))" + by (auto simp add: Ders_def) + also have "\ = \ (L ` (pders_set s (pder c r)))" + using ih by auto + also have "\ = \ (L ` (pders (c # s) r))" by simp + finally show "Ders (c # s) (L r) = \ (L ` pders (c # s) r)" . +qed (simp add: Ders_def) + +subsection \Relating derivatives and partial derivatives\ + +lemma der_pder: + shows "\ (L ` (pder c r)) = L (der c r)" +unfolding der_correctness Der_pder by simp + +lemma ders_pders: + shows "\ (L ` (pders s r)) = L (ders s r)" +unfolding der_correctness ders_correctness Ders_pders by simp + + +subsection \Finiteness property of partial derivatives\ + +definition + pders_Set :: "string set \ rexp \ rexp set" +where + "pders_Set A r \ \x \ A. pders x r" + +lemma pders_Set_subsetI: + assumes "\s. s \ A \ pders s r \ C" + shows "pders_Set A r \ C" +using assms unfolding pders_Set_def by (rule UN_least) + +lemma pders_Set_union: + shows "pders_Set (A \ B) r = (pders_Set A r \ pders_Set B r)" +by (simp add: pders_Set_def) + +lemma pders_Set_subset: + shows "A \ B \ pders_Set A r \ pders_Set B r" +by (auto simp add: pders_Set_def) + +definition + "UNIV1 \ UNIV - {[]}" + +lemma pders_Set_ZERO [simp]: + shows "pders_Set UNIV1 ZERO = {}" +unfolding UNIV1_def pders_Set_def by auto + +lemma pders_Set_ONE [simp]: + shows "pders_Set UNIV1 ONE = {}" +unfolding UNIV1_def pders_Set_def by (auto split: if_splits) + +lemma pders_Set_CHAR [simp]: + shows "pders_Set UNIV1 (CH c) = {ONE}" +unfolding UNIV1_def pders_Set_def +apply(auto) +apply(frule rev_subsetD) +apply(rule pders_CHAR) +apply(simp) +apply(case_tac xa) +apply(auto split: if_splits) +done + +lemma pders_Set_ALT [simp]: + shows "pders_Set UNIV1 (ALT r1 r2) = pders_Set UNIV1 r1 \ pders_Set UNIV1 r2" +unfolding UNIV1_def pders_Set_def by auto + + +text \Non-empty suffixes of a string (needed for the cases of @{const SEQ} and @{const STAR} below)\ + +definition + "PSuf s \ {v. v \ [] \ (\u. u @ v = s)}" + +lemma PSuf_snoc: + shows "PSuf (s @ [c]) = (PSuf s) ;; {[c]} \ {[c]}" +unfolding PSuf_def Sequ_def +by (auto simp add: append_eq_append_conv2 append_eq_Cons_conv) + +lemma PSuf_Union: + shows "(\v \ PSuf s ;; {[c]}. f v) = (\v \ PSuf s. f (v @ [c]))" +by (auto simp add: Sequ_def) + +lemma pders_Set_snoc: + shows "pders_Set (PSuf s ;; {[c]}) r = (pder_set c (pders_Set (PSuf s) r))" +unfolding pders_Set_def +by (simp add: PSuf_Union pders_snoc) + +lemma pders_SEQ: + shows "pders s (SEQ r1 r2) \ SEQs (pders s r1) r2 \ (pders_Set (PSuf s) r2)" +proof (induct s rule: rev_induct) + case (snoc c s) + have ih: "pders s (SEQ r1 r2) \ SEQs (pders s r1) r2 \ (pders_Set (PSuf s) r2)" + by fact + have "pders (s @ [c]) (SEQ r1 r2) = pder_set c (pders s (SEQ r1 r2))" + by (simp add: pders_snoc) + also have "\ \ pder_set c (SEQs (pders s r1) r2 \ (pders_Set (PSuf s) r2))" + using ih by fastforce + also have "\ = pder_set c (SEQs (pders s r1) r2) \ pder_set c (pders_Set (PSuf s) r2)" + by (simp) + also have "\ = pder_set c (SEQs (pders s r1) r2) \ pders_Set (PSuf s ;; {[c]}) r2" + by (simp add: pders_Set_snoc) + also + have "\ \ pder_set c (SEQs (pders s r1) r2) \ pder c r2 \ pders_Set (PSuf s ;; {[c]}) r2" + by auto + also + have "\ \ SEQs (pder_set c (pders s r1)) r2 \ pder c r2 \ pders_Set (PSuf s ;; {[c]}) r2" + by (auto simp add: if_splits) + also have "\ = SEQs (pders (s @ [c]) r1) r2 \ pder c r2 \ pders_Set (PSuf s ;; {[c]}) r2" + by (simp add: pders_snoc) + also have "\ \ SEQs (pders (s @ [c]) r1) r2 \ pders_Set (PSuf (s @ [c])) r2" + unfolding pders_Set_def by (auto simp add: PSuf_snoc) + finally show ?case . +qed (simp) + +lemma pders_Set_SEQ_aux1: + assumes a: "s \ UNIV1" + shows "pders_Set (PSuf s) r \ pders_Set UNIV1 r" +using a unfolding UNIV1_def PSuf_def pders_Set_def by auto + +lemma pders_Set_SEQ_aux2: + assumes a: "s \ UNIV1" + shows "SEQs (pders s r1) r2 \ SEQs (pders_Set UNIV1 r1) r2" +using a unfolding pders_Set_def by auto + +lemma pders_Set_SEQ: + shows "pders_Set UNIV1 (SEQ r1 r2) \ SEQs (pders_Set UNIV1 r1) r2 \ pders_Set UNIV1 r2" +apply(rule pders_Set_subsetI) +apply(rule subset_trans) +apply(rule pders_SEQ) +using pders_Set_SEQ_aux1 pders_Set_SEQ_aux2 +apply auto +apply blast +done + +lemma pders_STAR: + assumes a: "s \ []" + shows "pders s (STAR r) \ SEQs (pders_Set (PSuf s) r) (STAR r)" +using a +proof (induct s rule: rev_induct) + case (snoc c s) + have ih: "s \ [] \ pders s (STAR r) \ SEQs (pders_Set (PSuf s) r) (STAR r)" by fact + { assume asm: "s \ []" + have "pders (s @ [c]) (STAR r) = pder_set c (pders s (STAR r))" by (simp add: pders_snoc) + also have "\ \ pder_set c (SEQs (pders_Set (PSuf s) r) (STAR r))" + using ih[OF asm] by fast + also have "\ \ SEQs (pder_set c (pders_Set (PSuf s) r)) (STAR r) \ pder c (STAR r)" + by (auto split: if_splits) + also have "\ \ SEQs (pders_Set (PSuf (s @ [c])) r) (STAR r) \ (SEQs (pder c r) (STAR r))" + by (simp only: PSuf_snoc pders_Set_snoc pders_Set_union) + (auto simp add: pders_Set_def) + also have "\ = SEQs (pders_Set (PSuf (s @ [c])) r) (STAR r)" + by (auto simp add: PSuf_snoc PSuf_Union pders_snoc pders_Set_def) + finally have ?case . + } + moreover + { assume asm: "s = []" + then have ?case by (auto simp add: pders_Set_def pders_snoc PSuf_def) + } + ultimately show ?case by blast +qed (simp) + +lemma pders_Set_STAR: + shows "pders_Set UNIV1 (STAR r) \ SEQs (pders_Set UNIV1 r) (STAR r)" +apply(rule pders_Set_subsetI) +apply(rule subset_trans) +apply(rule pders_STAR) +apply(simp add: UNIV1_def) +apply(simp add: UNIV1_def PSuf_def) +apply(auto simp add: pders_Set_def) +done + +lemma finite_SEQs [simp]: + assumes a: "finite A" + shows "finite (SEQs A r)" +using a by auto + + +lemma finite_pders_Set_UNIV1: + shows "finite (pders_Set UNIV1 r)" +apply(induct r) +apply(simp_all add: + finite_subset[OF pders_Set_SEQ] + finite_subset[OF pders_Set_STAR]) +done + +lemma pders_Set_UNIV: + shows "pders_Set UNIV r = pders [] r \ pders_Set UNIV1 r" +unfolding UNIV1_def pders_Set_def +by blast + +lemma finite_pders_Set_UNIV: + shows "finite (pders_Set UNIV r)" +unfolding pders_Set_UNIV +by (simp add: finite_pders_Set_UNIV1) + +lemma finite_pders_set: + shows "finite (pders_Set A r)" +by (metis finite_pders_Set_UNIV pders_Set_subset rev_finite_subset subset_UNIV) + + +text \The following relationship between the alphabetic width of regular expressions +(called \awidth\ below) and the number of partial derivatives was proved +by Antimirov~\cite{Antimirov95} and formalized by Max Haslbeck.\ + +fun awidth :: "rexp \ nat" where +"awidth ZERO = 0" | +"awidth ONE = 0" | +"awidth (CH a) = 1" | +"awidth (ALT r1 r2) = awidth r1 + awidth r2" | +"awidth (SEQ r1 r2) = awidth r1 + awidth r2" | +"awidth (STAR r1) = awidth r1" + +lemma card_SEQs_pders_Set_le: + shows "card (SEQs (pders_Set A r) s) \ card (pders_Set A r)" + using finite_pders_set + unfolding SEQs_eq_image + by (rule card_image_le) + +lemma card_pders_set_UNIV1_le_awidth: + shows "card (pders_Set UNIV1 r) \ awidth r" +proof (induction r) + case (ALT r1 r2) + have "card (pders_Set UNIV1 (ALT r1 r2)) = card (pders_Set UNIV1 r1 \ pders_Set UNIV1 r2)" by simp + also have "\ \ card (pders_Set UNIV1 r1) + card (pders_Set UNIV1 r2)" + by(simp add: card_Un_le) + also have "\ \ awidth (ALT r1 r2)" using ALT.IH by simp + finally show ?case . +next + case (SEQ r1 r2) + have "card (pders_Set UNIV1 (SEQ r1 r2)) \ card (SEQs (pders_Set UNIV1 r1) r2 \ pders_Set UNIV1 r2)" + by (simp add: card_mono finite_pders_set pders_Set_SEQ) + also have "\ \ card (SEQs (pders_Set UNIV1 r1) r2) + card (pders_Set UNIV1 r2)" + by (simp add: card_Un_le) + also have "\ \ card (pders_Set UNIV1 r1) + card (pders_Set UNIV1 r2)" + by (simp add: card_SEQs_pders_Set_le) + also have "\ \ awidth (SEQ r1 r2)" using SEQ.IH by simp + finally show ?case . +next + case (STAR r) + have "card (pders_Set UNIV1 (STAR r)) \ card (SEQs (pders_Set UNIV1 r) (STAR r))" + by (simp add: card_mono finite_pders_set pders_Set_STAR) + also have "\ \ card (pders_Set UNIV1 r)" by (rule card_SEQs_pders_Set_le) + also have "\ \ awidth (STAR r)" by (simp add: STAR.IH) + finally show ?case . +qed (auto) + +text\Antimirov's Theorem 3.4:\ + +theorem card_pders_set_UNIV_le_awidth: + shows "card (pders_Set UNIV r) \ awidth r + 1" +proof - + have "card (insert r (pders_Set UNIV1 r)) \ Suc (card (pders_Set UNIV1 r))" + by(auto simp: card_insert_if[OF finite_pders_Set_UNIV1]) + also have "\ \ Suc (awidth r)" by(simp add: card_pders_set_UNIV1_le_awidth) + finally show ?thesis by(simp add: pders_Set_UNIV) +qed + +text\Antimirov's Corollary 3.5:\ +(*W stands for word set*) +corollary card_pders_set_le_awidth: + shows "card (pders_Set W r) \ awidth r + 1" +proof - + have "card (pders_Set W r) \ card (pders_Set UNIV r)" + by (simp add: card_mono finite_pders_set pders_Set_subset) + also have "... \ awidth r + 1" + by (rule card_pders_set_UNIV_le_awidth) + finally show "card (pders_Set W r) \ awidth r + 1" by simp +qed + +(* other result by antimirov *) + +lemma card_pders_awidth: + shows "card (pders s r) \ awidth r + 1" +proof - + have "pders s r \ pders_Set UNIV r" + using pders_Set_def by auto + then have "card (pders s r) \ card (pders_Set UNIV r)" + by (simp add: card_mono finite_pders_set) + then show "card (pders s r) \ awidth r + 1" + using card_pders_set_le_awidth order_trans by blast +qed + + + + + +fun subs :: "rexp \ rexp set" where +"subs ZERO = {ZERO}" | +"subs ONE = {ONE}" | +"subs (CH a) = {CH a, ONE}" | +"subs (ALT r1 r2) = (subs r1 \ subs r2 \ {ALT r1 r2})" | +"subs (SEQ r1 r2) = (subs r1 \ subs r2 \ {SEQ r1 r2} \ SEQs (subs r1) r2)" | +"subs (STAR r1) = (subs r1 \ {STAR r1} \ SEQs (subs r1) (STAR r1))" + +lemma subs_finite: + shows "finite (subs r)" + apply(induct r) + apply(simp_all) + done + + + +lemma pders_subs: + shows "pders s r \ subs r" + apply(induct r arbitrary: s) + apply(simp) + apply(simp) + apply(simp add: pders_CHAR) +(* SEQ case *) + apply(simp) + apply(rule subset_trans) + apply(rule pders_SEQ) + defer +(* ALT case *) + apply(simp) + apply(rule impI) + apply(rule conjI) + apply blast + apply blast +(* STAR case *) + apply(case_tac s) + apply(simp) + apply(rule subset_trans) + thm pders_STAR + apply(rule pders_STAR) + apply(simp) + apply(auto simp add: pders_Set_def)[1] + apply(simp) + apply(rule conjI) + apply blast +apply(auto simp add: pders_Set_def)[1] + done + +fun size2 :: "rexp \ nat" where + "size2 ZERO = 1" | + "size2 ONE = 1" | + "size2 (CH c) = 1" | + "size2 (ALT r1 r2) = Suc (size2 r1 + size2 r2)" | + "size2 (SEQ r1 r2) = Suc (size2 r1 + size2 r2)" | + "size2 (STAR r1) = Suc (size2 r1)" + + +lemma size_rexp: + fixes r :: rexp + shows "1 \ size2 r" + apply(induct r) + apply(simp) + apply(simp_all) + done + +lemma subs_size2: + shows "\r1 \ subs r. size2 r1 \ Suc (size2 r * size2 r)" + apply(induct r) + apply(simp) + apply(simp) + apply(simp) +(* SEQ case *) + apply(simp) + apply(auto)[1] + apply (smt Suc_n_not_le_n add.commute distrib_left le_Suc_eq left_add_mult_distrib nat_le_linear trans_le_add1) + apply (smt Suc_le_mono Suc_n_not_le_n le_trans nat_le_linear power2_eq_square power2_sum semiring_normalization_rules(23) trans_le_add2) + apply (smt Groups.add_ac(3) Suc_n_not_le_n distrib_left le_Suc_eq left_add_mult_distrib nat_le_linear trans_le_add1) +(* ALT case *) + apply(simp) + apply(auto)[1] + apply (smt Groups.add_ac(2) Suc_le_mono Suc_n_not_le_n le_add2 linear order_trans power2_eq_square power2_sum) + apply (smt Groups.add_ac(2) Suc_le_mono Suc_n_not_le_n left_add_mult_distrib linear mult.commute order.trans trans_le_add1) +(* STAR case *) + apply(auto)[1] + apply(drule_tac x="r'" in bspec) + apply(simp) + apply(rule le_trans) + apply(assumption) + apply(simp) + using size_rexp + apply(simp) + done + +lemma awidth_size: + shows "awidth r \ size2 r" + apply(induct r) + apply(simp_all) + done + +lemma Sum1: + fixes A B :: "nat set" + assumes "A \ B" "finite A" "finite B" + shows "\A \ \B" + using assms + by (simp add: sum_mono2) + +lemma Sum2: + fixes A :: "rexp set" + and f g :: "rexp \ nat" + assumes "finite A" "\x \ A. f x \ g x" + shows "sum f A \ sum g A" + using assms + apply(induct A) + apply(auto) + done + + + + + +lemma pders_max_size: + shows "(sum size2 (pders s r)) \ (Suc (size2 r)) ^ 3" +proof - + have "(sum size2 (pders s r)) \ sum (\_. Suc (size2 r * size2 r)) (pders s r)" + apply(rule_tac Sum2) + apply (meson pders_subs rev_finite_subset subs_finite) + using pders_subs subs_size2 by blast + also have "... \ (Suc (size2 r * size2 r)) * (sum (\_. 1) (pders s r))" + by simp + also have "... \ (Suc (size2 r * size2 r)) * card (pders s r)" + by simp + also have "... \ (Suc (size2 r * size2 r)) * (Suc (awidth r))" + using Suc_eq_plus1 card_pders_awidth mult_le_mono2 by presburger + also have "... \ (Suc (size2 r * size2 r)) * (Suc (size2 r))" + using Suc_le_mono awidth_size mult_le_mono2 by presburger + also have "... \ (Suc (size2 r)) ^ 3" + by (smt One_nat_def Suc_1 Suc_mult_le_cancel1 Suc_n_not_le_n antisym_conv le_Suc_eq mult.commute nat_le_linear numeral_3_eq_3 power2_eq_square power2_le_imp_le power_Suc size_rexp) + finally show ?thesis . +qed + +lemma pders_Set_max_size: + shows "(sum size2 (pders_Set A r)) \ (Suc (size2 r)) ^ 3" +proof - + have "(sum size2 (pders_Set A r)) \ sum (\_. Suc (size2 r * size2 r)) (pders_Set A r)" + apply(rule_tac Sum2) + apply (simp add: finite_pders_set) + by (meson pders_Set_subsetI pders_subs subs_size2 subsetD) + also have "... \ (Suc (size2 r * size2 r)) * (sum (\_. 1) (pders_Set A r))" + by simp + also have "... \ (Suc (size2 r * size2 r)) * card (pders_Set A r)" + by simp + also have "... \ (Suc (size2 r * size2 r)) * (Suc (awidth r))" + using Suc_eq_plus1 card_pders_set_le_awidth mult_le_mono2 by presburger + also have "... \ (Suc (size2 r * size2 r)) * (Suc (size2 r))" + using Suc_le_mono awidth_size mult_le_mono2 by presburger + also have "... \ (Suc (size2 r)) ^ 3" + by (smt One_nat_def Suc_1 Suc_mult_le_cancel1 Suc_n_not_le_n antisym_conv le_Suc_eq mult.commute nat_le_linear numeral_3_eq_3 power2_eq_square power2_le_imp_le power_Suc size_rexp) + finally show ?thesis . +qed + +fun height :: "rexp \ nat" where + "height ZERO = 1" | + "height ONE = 1" | + "height (CH c) = 1" | + "height (ALT r1 r2) = Suc (max (height r1) (height r2))" | + "height (SEQ r1 r2) = Suc (max (height r1) (height r2))" | + "height (STAR r1) = Suc (height r1)" + +lemma height_size2: + shows "height r \ size2 r" + apply(induct r) + apply(simp_all) + done + +lemma height_rexp: + fixes r :: rexp + shows "1 \ height r" + apply(induct r) + apply(simp_all) + done + +lemma subs_height: + shows "\r1 \ subs r. height r1 \ Suc (height r)" + apply(induct r) + apply(auto)+ + done + +fun lin_concat :: "(char \ rexp) \ rexp \ (char \ rexp)" (infixl "[.]" 91) + where +"(c, ZERO) [.] t = (c, ZERO)" +| "(c, ONE) [.] t = (c, t)" +| "(c, p) [.] t = (c, SEQ p t)" + + +fun circle_concat :: "(char \ rexp ) set \ rexp \ (char \ rexp) set" ( infixl "\" 90) + where +"l \ ZERO = {}" +| "l \ ONE = l" +| "l \ t = ( (\p. p [.] t) ` l ) " + + + +fun linear_form :: "rexp \( char \ rexp ) set" + where + "linear_form ZERO = {}" +| "linear_form ONE = {}" +| "linear_form (CH c) = {(c, ONE)}" +| "linear_form (ALT r1 r2) = (linear_form) r1 \ (linear_form r2)" +| "linear_form (SEQ r1 r2) = (if (nullable r1) then (linear_form r1) \ r2 \ linear_form r2 else (linear_form r1) \ r2) " +| "linear_form (STAR r ) = (linear_form r) \ (STAR r)" + + +value "linear_form (SEQ (STAR (CH x)) (STAR (ALT (SEQ (CH x) (CH x)) (CH y) )) )" + + +value "linear_form (STAR (ALT (SEQ (CH x) (CH x)) (CH y) )) " + +fun pdero :: "char \ rexp \ rexp set" + where +"pdero c t = \ ((\(d, p). if d = c then {p} else {}) ` (linear_form t) )" + +fun pderso :: "char list \ rexp \ rexp set" + where + "pderso [] r = {r}" +| "pderso (c # s) r = \ ( pderso s ` (pdero c r) )" + +lemma pdero_result: + shows "pdero c (STAR (ALT (CH c) (SEQ (CH c) (CH c)))) = {SEQ (CH c)(STAR (ALT (CH c) (SEQ (CH c) (CH c)))),(STAR (ALT (CH c) (SEQ (CH c) (CH c))))}" + apply(simp) + by auto + +fun concatLen :: "rexp \ nat" where +"concatLen ZERO = 0" | +"concatLen ONE = 0" | +"concatLen (CH c) = 0" | +"concatLen (SEQ v1 v2) = Suc (max (concatLen v1) (concatLen v2))" | +" concatLen (ALT v1 v2) = max (concatLen v1) (concatLen v2)" | +" concatLen (STAR v) = Suc (concatLen v)" + + + +end \ No newline at end of file diff -r f9cdc295ccf7 -r f493a20feeb3 thys3/src/Positions.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys3/src/Positions.thy Sat Apr 30 00:50:08 2022 +0100 @@ -0,0 +1,773 @@ + +theory Positions + imports PosixSpec Lexer +begin + +chapter \An alternative definition for POSIX values\ + +section \Positions in Values\ + +fun + at :: "val \ nat list \ val" +where + "at v [] = v" +| "at (Left v) (0#ps)= at v ps" +| "at (Right v) (Suc 0#ps)= at v ps" +| "at (Seq v1 v2) (0#ps)= at v1 ps" +| "at (Seq v1 v2) (Suc 0#ps)= at v2 ps" +| "at (Stars vs) (n#ps)= at (nth vs n) ps" + + + +fun Pos :: "val \ (nat list) set" +where + "Pos (Void) = {[]}" +| "Pos (Char c) = {[]}" +| "Pos (Left v) = {[]} \ {0#ps | ps. ps \ Pos v}" +| "Pos (Right v) = {[]} \ {1#ps | ps. ps \ Pos v}" +| "Pos (Seq v1 v2) = {[]} \ {0#ps | ps. ps \ Pos v1} \ {1#ps | ps. ps \ Pos v2}" +| "Pos (Stars []) = {[]}" +| "Pos (Stars (v#vs)) = {[]} \ {0#ps | ps. ps \ Pos v} \ {Suc n#ps | n ps. n#ps \ Pos (Stars vs)}" + + +lemma Pos_stars: + "Pos (Stars vs) = {[]} \ (\n < length vs. {n#ps | ps. ps \ Pos (vs ! n)})" +apply(induct vs) +apply(auto simp add: insert_ident less_Suc_eq_0_disj) +done + +lemma Pos_empty: + shows "[] \ Pos v" +by (induct v rule: Pos.induct)(auto) + + +abbreviation + "intlen vs \ int (length vs)" + + +definition pflat_len :: "val \ nat list => int" +where + "pflat_len v p \ (if p \ Pos v then intlen (flat (at v p)) else -1)" + +lemma pflat_len_simps: + shows "pflat_len (Seq v1 v2) (0#p) = pflat_len v1 p" + and "pflat_len (Seq v1 v2) (Suc 0#p) = pflat_len v2 p" + and "pflat_len (Left v) (0#p) = pflat_len v p" + and "pflat_len (Left v) (Suc 0#p) = -1" + and "pflat_len (Right v) (Suc 0#p) = pflat_len v p" + and "pflat_len (Right v) (0#p) = -1" + and "pflat_len (Stars (v#vs)) (Suc n#p) = pflat_len (Stars vs) (n#p)" + and "pflat_len (Stars (v#vs)) (0#p) = pflat_len v p" + and "pflat_len v [] = intlen (flat v)" +by (auto simp add: pflat_len_def Pos_empty) + +lemma pflat_len_Stars_simps: + assumes "n < length vs" + shows "pflat_len (Stars vs) (n#p) = pflat_len (vs!n) p" +using assms +apply(induct vs arbitrary: n p) +apply(auto simp add: less_Suc_eq_0_disj pflat_len_simps) +done + +lemma pflat_len_outside: + assumes "p \ Pos v1" + shows "pflat_len v1 p = -1 " +using assms by (simp add: pflat_len_def) + + + +section \Orderings\ + + +definition prefix_list:: "'a list \ 'a list \ bool" ("_ \pre _" [60,59] 60) +where + "ps1 \pre ps2 \ \ps'. ps1 @ps' = ps2" + +definition sprefix_list:: "'a list \ 'a list \ bool" ("_ \spre _" [60,59] 60) +where + "ps1 \spre ps2 \ ps1 \pre ps2 \ ps1 \ ps2" + +inductive lex_list :: "nat list \ nat list \ bool" ("_ \lex _" [60,59] 60) +where + "[] \lex (p#ps)" +| "ps1 \lex ps2 \ (p#ps1) \lex (p#ps2)" +| "p1 < p2 \ (p1#ps1) \lex (p2#ps2)" + +lemma lex_irrfl: + fixes ps1 ps2 :: "nat list" + assumes "ps1 \lex ps2" + shows "ps1 \ ps2" +using assms +by(induct rule: lex_list.induct)(auto) + +lemma lex_simps [simp]: + fixes xs ys :: "nat list" + shows "[] \lex ys \ ys \ []" + and "xs \lex [] \ False" + and "(x # xs) \lex (y # ys) \ (x < y \ (x = y \ xs \lex ys))" +by (auto simp add: neq_Nil_conv elim: lex_list.cases intro: lex_list.intros) + +lemma lex_trans: + fixes ps1 ps2 ps3 :: "nat list" + assumes "ps1 \lex ps2" "ps2 \lex ps3" + shows "ps1 \lex ps3" +using assms +by (induct arbitrary: ps3 rule: lex_list.induct) + (auto elim: lex_list.cases) + + +lemma lex_trichotomous: + fixes p q :: "nat list" + shows "p = q \ p \lex q \ q \lex p" +apply(induct p arbitrary: q) +apply(auto elim: lex_list.cases) +apply(case_tac q) +apply(auto) +done + + + + +section \POSIX Ordering of Values According to Okui \& Suzuki\ + + +definition PosOrd:: "val \ nat list \ val \ bool" ("_ \val _ _" [60, 60, 59] 60) +where + "v1 \val p v2 \ pflat_len v1 p > pflat_len v2 p \ + (\q \ Pos v1 \ Pos v2. q \lex p \ pflat_len v1 q = pflat_len v2 q)" + +lemma PosOrd_def2: + shows "v1 \val p v2 \ + pflat_len v1 p > pflat_len v2 p \ + (\q \ Pos v1. q \lex p \ pflat_len v1 q = pflat_len v2 q) \ + (\q \ Pos v2. q \lex p \ pflat_len v1 q = pflat_len v2 q)" +unfolding PosOrd_def +apply(auto) +done + + +definition PosOrd_ex:: "val \ val \ bool" ("_ :\val _" [60, 59] 60) +where + "v1 :\val v2 \ \p. v1 \val p v2" + +definition PosOrd_ex_eq:: "val \ val \ bool" ("_ :\val _" [60, 59] 60) +where + "v1 :\val v2 \ v1 :\val v2 \ v1 = v2" + + +lemma PosOrd_trans: + assumes "v1 :\val v2" "v2 :\val v3" + shows "v1 :\val v3" +proof - + from assms obtain p p' + where as: "v1 \val p v2" "v2 \val p' v3" unfolding PosOrd_ex_def by blast + then have pos: "p \ Pos v1" "p' \ Pos v2" unfolding PosOrd_def pflat_len_def + by (smt not_int_zless_negative)+ + have "p = p' \ p \lex p' \ p' \lex p" + by (rule lex_trichotomous) + moreover + { assume "p = p'" + with as have "v1 \val p v3" unfolding PosOrd_def pflat_len_def + by (smt Un_iff) + then have " v1 :\val v3" unfolding PosOrd_ex_def by blast + } + moreover + { assume "p \lex p'" + with as have "v1 \val p v3" unfolding PosOrd_def pflat_len_def + by (smt Un_iff lex_trans) + then have " v1 :\val v3" unfolding PosOrd_ex_def by blast + } + moreover + { assume "p' \lex p" + with as have "v1 \val p' v3" unfolding PosOrd_def + by (smt Un_iff lex_trans pflat_len_def) + then have "v1 :\val v3" unfolding PosOrd_ex_def by blast + } + ultimately show "v1 :\val v3" by blast +qed + +lemma PosOrd_irrefl: + assumes "v :\val v" + shows "False" +using assms unfolding PosOrd_ex_def PosOrd_def +by auto + +lemma PosOrd_assym: + assumes "v1 :\val v2" + shows "\(v2 :\val v1)" +using assms +using PosOrd_irrefl PosOrd_trans by blast + +(* + :\val and :\val are partial orders. +*) + +lemma PosOrd_ordering: + shows "ordering (\v1 v2. v1 :\val v2) (\ v1 v2. v1 :\val v2)" +unfolding ordering_def PosOrd_ex_eq_def +apply(auto) +using PosOrd_trans partial_preordering_def apply blast +using PosOrd_assym ordering_axioms_def by blast + +lemma PosOrd_order: + shows "class.order (\v1 v2. v1 :\val v2) (\ v1 v2. v1 :\val v2)" +using PosOrd_ordering +apply(simp add: class.order_def class.preorder_def class.order_axioms_def) + by (metis (full_types) PosOrd_ex_eq_def PosOrd_irrefl PosOrd_trans) + + +lemma PosOrd_ex_eq2: + shows "v1 :\val v2 \ (v1 :\val v2 \ v1 \ v2)" + using PosOrd_ordering + using PosOrd_ex_eq_def PosOrd_irrefl by blast + +lemma PosOrdeq_trans: + assumes "v1 :\val v2" "v2 :\val v3" + shows "v1 :\val v3" +using assms PosOrd_ordering + unfolding ordering_def + by (metis partial_preordering.trans) + +lemma PosOrdeq_antisym: + assumes "v1 :\val v2" "v2 :\val v1" + shows "v1 = v2" +using assms PosOrd_ordering + unfolding ordering_def + by (simp add: ordering_axioms_def) + +lemma PosOrdeq_refl: + shows "v :\val v" +unfolding PosOrd_ex_eq_def +by auto + + +lemma PosOrd_shorterE: + assumes "v1 :\val v2" + shows "length (flat v2) \ length (flat v1)" +using assms unfolding PosOrd_ex_def PosOrd_def +apply(auto) +apply(case_tac p) +apply(simp add: pflat_len_simps) +apply(drule_tac x="[]" in bspec) +apply(simp add: Pos_empty) +apply(simp add: pflat_len_simps) +done + +lemma PosOrd_shorterI: + assumes "length (flat v2) < length (flat v1)" + shows "v1 :\val v2" +unfolding PosOrd_ex_def PosOrd_def pflat_len_def +using assms Pos_empty by force + +lemma PosOrd_spreI: + assumes "flat v' \spre flat v" + shows "v :\val v'" +using assms +apply(rule_tac PosOrd_shorterI) +unfolding prefix_list_def sprefix_list_def +by (metis append_Nil2 append_eq_conv_conj drop_all le_less_linear) + +lemma pflat_len_inside: + assumes "pflat_len v2 p < pflat_len v1 p" + shows "p \ Pos v1" +using assms +unfolding pflat_len_def +by (auto split: if_splits) + + +lemma PosOrd_Left_Right: + assumes "flat v1 = flat v2" + shows "Left v1 :\val Right v2" +unfolding PosOrd_ex_def +apply(rule_tac x="[0]" in exI) +apply(auto simp add: PosOrd_def pflat_len_simps assms) +done + +lemma PosOrd_LeftE: + assumes "Left v1 :\val Left v2" "flat v1 = flat v2" + shows "v1 :\val v2" +using assms +unfolding PosOrd_ex_def PosOrd_def2 +apply(auto simp add: pflat_len_simps) +apply(frule pflat_len_inside) +apply(auto simp add: pflat_len_simps) +by (metis lex_simps(3) pflat_len_simps(3)) + +lemma PosOrd_LeftI: + assumes "v1 :\val v2" "flat v1 = flat v2" + shows "Left v1 :\val Left v2" +using assms +unfolding PosOrd_ex_def PosOrd_def2 +apply(auto simp add: pflat_len_simps) +by (metis less_numeral_extra(3) lex_simps(3) pflat_len_simps(3)) + +lemma PosOrd_Left_eq: + assumes "flat v1 = flat v2" + shows "Left v1 :\val Left v2 \ v1 :\val v2" +using assms PosOrd_LeftE PosOrd_LeftI +by blast + + +lemma PosOrd_RightE: + assumes "Right v1 :\val Right v2" "flat v1 = flat v2" + shows "v1 :\val v2" +using assms +unfolding PosOrd_ex_def PosOrd_def2 +apply(auto simp add: pflat_len_simps) +apply(frule pflat_len_inside) +apply(auto simp add: pflat_len_simps) +by (metis lex_simps(3) pflat_len_simps(5)) + +lemma PosOrd_RightI: + assumes "v1 :\val v2" "flat v1 = flat v2" + shows "Right v1 :\val Right v2" +using assms +unfolding PosOrd_ex_def PosOrd_def2 +apply(auto simp add: pflat_len_simps) +by (metis lex_simps(3) nat_neq_iff pflat_len_simps(5)) + + +lemma PosOrd_Right_eq: + assumes "flat v1 = flat v2" + shows "Right v1 :\val Right v2 \ v1 :\val v2" +using assms PosOrd_RightE PosOrd_RightI +by blast + + +lemma PosOrd_SeqI1: + assumes "v1 :\val w1" "flat (Seq v1 v2) = flat (Seq w1 w2)" + shows "Seq v1 v2 :\val Seq w1 w2" +using assms(1) +apply(subst (asm) PosOrd_ex_def) +apply(subst (asm) PosOrd_def) +apply(clarify) +apply(subst PosOrd_ex_def) +apply(rule_tac x="0#p" in exI) +apply(subst PosOrd_def) +apply(rule conjI) +apply(simp add: pflat_len_simps) +apply(rule ballI) +apply(rule impI) +apply(simp only: Pos.simps) +apply(auto)[1] +apply(simp add: pflat_len_simps) +apply(auto simp add: pflat_len_simps) +using assms(2) +apply(simp) +apply(metis length_append of_nat_add) +done + +lemma PosOrd_SeqI2: + assumes "v2 :\val w2" "flat v2 = flat w2" + shows "Seq v v2 :\val Seq v w2" +using assms(1) +apply(subst (asm) PosOrd_ex_def) +apply(subst (asm) PosOrd_def) +apply(clarify) +apply(subst PosOrd_ex_def) +apply(rule_tac x="Suc 0#p" in exI) +apply(subst PosOrd_def) +apply(rule conjI) +apply(simp add: pflat_len_simps) +apply(rule ballI) +apply(rule impI) +apply(simp only: Pos.simps) +apply(auto)[1] +apply(simp add: pflat_len_simps) +using assms(2) +apply(simp) +apply(auto simp add: pflat_len_simps) +done + +lemma PosOrd_Seq_eq: + assumes "flat v2 = flat w2" + shows "(Seq v v2) :\val (Seq v w2) \ v2 :\val w2" +using assms +apply(auto) +prefer 2 +apply(simp add: PosOrd_SeqI2) +apply(simp add: PosOrd_ex_def) +apply(auto) +apply(case_tac p) +apply(simp add: PosOrd_def pflat_len_simps) +apply(case_tac a) +apply(simp add: PosOrd_def pflat_len_simps) +apply(clarify) +apply(case_tac nat) +prefer 2 +apply(simp add: PosOrd_def pflat_len_simps pflat_len_outside) +apply(rule_tac x="list" in exI) +apply(auto simp add: PosOrd_def2 pflat_len_simps) +apply(smt Collect_disj_eq lex_list.intros(2) mem_Collect_eq pflat_len_simps(2)) +apply(smt Collect_disj_eq lex_list.intros(2) mem_Collect_eq pflat_len_simps(2)) +done + + + +lemma PosOrd_StarsI: + assumes "v1 :\val v2" "flats (v1#vs1) = flats (v2#vs2)" + shows "Stars (v1#vs1) :\val Stars (v2#vs2)" +using assms(1) +apply(subst (asm) PosOrd_ex_def) +apply(subst (asm) PosOrd_def) +apply(clarify) +apply(subst PosOrd_ex_def) +apply(subst PosOrd_def) +apply(rule_tac x="0#p" in exI) +apply(simp add: pflat_len_Stars_simps pflat_len_simps) +using assms(2) +apply(simp add: pflat_len_simps) +apply(auto simp add: pflat_len_Stars_simps pflat_len_simps) +by (metis length_append of_nat_add) + +lemma PosOrd_StarsI2: + assumes "Stars vs1 :\val Stars vs2" "flats vs1 = flats vs2" + shows "Stars (v#vs1) :\val Stars (v#vs2)" +using assms(1) +apply(subst (asm) PosOrd_ex_def) +apply(subst (asm) PosOrd_def) +apply(clarify) +apply(subst PosOrd_ex_def) +apply(subst PosOrd_def) +apply(case_tac p) +apply(simp add: pflat_len_simps) +apply(rule_tac x="Suc a#list" in exI) +apply(auto simp add: pflat_len_Stars_simps pflat_len_simps assms(2)) +done + +lemma PosOrd_Stars_appendI: + assumes "Stars vs1 :\val Stars vs2" "flat (Stars vs1) = flat (Stars vs2)" + shows "Stars (vs @ vs1) :\val Stars (vs @ vs2)" +using assms +apply(induct vs) +apply(simp) +apply(simp add: PosOrd_StarsI2) +done + +lemma PosOrd_StarsE2: + assumes "Stars (v # vs1) :\val Stars (v # vs2)" + shows "Stars vs1 :\val Stars vs2" +using assms +apply(subst (asm) PosOrd_ex_def) +apply(erule exE) +apply(case_tac p) +apply(simp) +apply(simp add: PosOrd_def pflat_len_simps) +apply(subst PosOrd_ex_def) +apply(rule_tac x="[]" in exI) +apply(simp add: PosOrd_def pflat_len_simps Pos_empty) +apply(simp) +apply(case_tac a) +apply(clarify) +apply(auto simp add: pflat_len_simps PosOrd_def pflat_len_def split: if_splits)[1] +apply(clarify) +apply(simp add: PosOrd_ex_def) +apply(rule_tac x="nat#list" in exI) +apply(auto simp add: PosOrd_def pflat_len_simps)[1] +apply(case_tac q) +apply(simp add: PosOrd_def pflat_len_simps) +apply(clarify) +apply(drule_tac x="Suc a # lista" in bspec) +apply(simp) +apply(auto simp add: PosOrd_def pflat_len_simps)[1] +apply(case_tac q) +apply(simp add: PosOrd_def pflat_len_simps) +apply(clarify) +apply(drule_tac x="Suc a # lista" in bspec) +apply(simp) +apply(auto simp add: PosOrd_def pflat_len_simps)[1] +done + +lemma PosOrd_Stars_appendE: + assumes "Stars (vs @ vs1) :\val Stars (vs @ vs2)" + shows "Stars vs1 :\val Stars vs2" +using assms +apply(induct vs) +apply(simp) +apply(simp add: PosOrd_StarsE2) +done + +lemma PosOrd_Stars_append_eq: + assumes "flats vs1 = flats vs2" + shows "Stars (vs @ vs1) :\val Stars (vs @ vs2) \ Stars vs1 :\val Stars vs2" +using assms +apply(rule_tac iffI) +apply(erule PosOrd_Stars_appendE) +apply(rule PosOrd_Stars_appendI) +apply(auto) +done + +lemma PosOrd_almost_trichotomous: + shows "v1 :\val v2 \ v2 :\val v1 \ (length (flat v1) = length (flat v2))" +apply(auto simp add: PosOrd_ex_def) +apply(auto simp add: PosOrd_def) +apply(rule_tac x="[]" in exI) +apply(auto simp add: Pos_empty pflat_len_simps) +apply(drule_tac x="[]" in spec) +apply(auto simp add: Pos_empty pflat_len_simps) +done + + + +section \The Posix Value is smaller than any other Value\ + + +lemma Posix_PosOrd: + assumes "s \ r \ v1" "v2 \ LV r s" + shows "v1 :\val v2" +using assms +proof (induct arbitrary: v2 rule: Posix.induct) + case (Posix_ONE v) + have "v \ LV ONE []" by fact + then have "v = Void" + by (simp add: LV_simps) + then show "Void :\val v" + by (simp add: PosOrd_ex_eq_def) +next + case (Posix_CH c v) + have "v \ LV (CH c) [c]" by fact + then have "v = Char c" + by (simp add: LV_simps) + then show "Char c :\val v" + by (simp add: PosOrd_ex_eq_def) +next + case (Posix_ALT1 s r1 v r2 v2) + have as1: "s \ r1 \ v" by fact + have IH: "\v2. v2 \ LV r1 s \ v :\val v2" by fact + have "v2 \ LV (ALT r1 r2) s" by fact + then have "\ v2 : ALT r1 r2" "flat v2 = s" + by(auto simp add: LV_def prefix_list_def) + then consider + (Left) v3 where "v2 = Left v3" "\ v3 : r1" "flat v3 = s" + | (Right) v3 where "v2 = Right v3" "\ v3 : r2" "flat v3 = s" + by (auto elim: Prf.cases) + then show "Left v :\val v2" + proof(cases) + case (Left v3) + have "v3 \ LV r1 s" using Left(2,3) + by (auto simp add: LV_def prefix_list_def) + with IH have "v :\val v3" by simp + moreover + have "flat v3 = flat v" using as1 Left(3) + by (simp add: Posix1(2)) + ultimately have "Left v :\val Left v3" + by (simp add: PosOrd_ex_eq_def PosOrd_Left_eq) + then show "Left v :\val v2" unfolding Left . + next + case (Right v3) + have "flat v3 = flat v" using as1 Right(3) + by (simp add: Posix1(2)) + then have "Left v :\val Right v3" + unfolding PosOrd_ex_eq_def + by (simp add: PosOrd_Left_Right) + then show "Left v :\val v2" unfolding Right . + qed +next + case (Posix_ALT2 s r2 v r1 v2) + have as1: "s \ r2 \ v" by fact + have as2: "s \ L r1" by fact + have IH: "\v2. v2 \ LV r2 s \ v :\val v2" by fact + have "v2 \ LV (ALT r1 r2) s" by fact + then have "\ v2 : ALT r1 r2" "flat v2 = s" + by(auto simp add: LV_def prefix_list_def) + then consider + (Left) v3 where "v2 = Left v3" "\ v3 : r1" "flat v3 = s" + | (Right) v3 where "v2 = Right v3" "\ v3 : r2" "flat v3 = s" + by (auto elim: Prf.cases) + then show "Right v :\val v2" + proof (cases) + case (Right v3) + have "v3 \ LV r2 s" using Right(2,3) + by (auto simp add: LV_def prefix_list_def) + with IH have "v :\val v3" by simp + moreover + have "flat v3 = flat v" using as1 Right(3) + by (simp add: Posix1(2)) + ultimately have "Right v :\val Right v3" + by (auto simp add: PosOrd_ex_eq_def PosOrd_RightI) + then show "Right v :\val v2" unfolding Right . + next + case (Left v3) + have "v3 \ LV r1 s" using Left(2,3) as2 + by (auto simp add: LV_def prefix_list_def) + then have "flat v3 = flat v \ \ v3 : r1" using as1 Left(3) + by (simp add: Posix1(2) LV_def) + then have "False" using as1 as2 Left + by (auto simp add: Posix1(2) L_flat_Prf1) + then show "Right v :\val v2" by simp + qed +next + case (Posix_SEQ s1 r1 v1 s2 r2 v2 v3) + have "s1 \ r1 \ v1" "s2 \ r2 \ v2" by fact+ + then have as1: "s1 = flat v1" "s2 = flat v2" by (simp_all add: Posix1(2)) + have IH1: "\v3. v3 \ LV r1 s1 \ v1 :\val v3" by fact + have IH2: "\v3. v3 \ LV r2 s2 \ v2 :\val v3" by fact + have cond: "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L r1 \ s\<^sub>4 \ L r2)" by fact + have "v3 \ LV (SEQ r1 r2) (s1 @ s2)" by fact + then obtain v3a v3b where eqs: + "v3 = Seq v3a v3b" "\ v3a : r1" "\ v3b : r2" + "flat v3a @ flat v3b = s1 @ s2" + by (force simp add: prefix_list_def LV_def elim: Prf.cases) + with cond have "flat v3a \pre s1" unfolding prefix_list_def + by (smt L_flat_Prf1 append_eq_append_conv2 append_self_conv) + then have "flat v3a \spre s1 \ (flat v3a = s1 \ flat v3b = s2)" using eqs + by (simp add: sprefix_list_def append_eq_conv_conj) + then have q2: "v1 :\val v3a \ (flat v3a = s1 \ flat v3b = s2)" + using PosOrd_spreI as1(1) eqs by blast + then have "v1 :\val v3a \ (v3a \ LV r1 s1 \ v3b \ LV r2 s2)" using eqs(2,3) + by (auto simp add: LV_def) + then have "v1 :\val v3a \ (v1 :\val v3a \ v2 :\val v3b)" using IH1 IH2 by blast + then have "Seq v1 v2 :\val Seq v3a v3b" using eqs q2 as1 + unfolding PosOrd_ex_eq_def by (auto simp add: PosOrd_SeqI1 PosOrd_Seq_eq) + then show "Seq v1 v2 :\val v3" unfolding eqs by blast +next + case (Posix_STAR1 s1 r v s2 vs v3) + have "s1 \ r \ v" "s2 \ STAR r \ Stars vs" by fact+ + then have as1: "s1 = flat v" "s2 = flat (Stars vs)" by (auto dest: Posix1(2)) + have IH1: "\v3. v3 \ LV r s1 \ v :\val v3" by fact + have IH2: "\v3. v3 \ LV (STAR r) s2 \ Stars vs :\val v3" by fact + have cond: "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (STAR r))" by fact + have cond2: "flat v \ []" by fact + have "v3 \ LV (STAR r) (s1 @ s2)" by fact + then consider + (NonEmpty) v3a vs3 where "v3 = Stars (v3a # vs3)" + "\ v3a : r" "\ Stars vs3 : STAR r" + "flat (Stars (v3a # vs3)) = s1 @ s2" + | (Empty) "v3 = Stars []" + unfolding LV_def + apply(auto) + apply(erule Prf.cases) + apply(auto) + apply(case_tac vs) + apply(auto intro: Prf.intros) + done + then show "Stars (v # vs) :\val v3" + proof (cases) + case (NonEmpty v3a vs3) + have "flat (Stars (v3a # vs3)) = s1 @ s2" using NonEmpty(4) . + with cond have "flat v3a \pre s1" using NonEmpty(2,3) + unfolding prefix_list_def + by (smt L_flat_Prf1 append_Nil2 append_eq_append_conv2 flat.simps(7)) + then have "flat v3a \spre s1 \ (flat v3a = s1 \ flat (Stars vs3) = s2)" using NonEmpty(4) + by (simp add: sprefix_list_def append_eq_conv_conj) + then have q2: "v :\val v3a \ (flat v3a = s1 \ flat (Stars vs3) = s2)" + using PosOrd_spreI as1(1) NonEmpty(4) by blast + then have "v :\val v3a \ (v3a \ LV r s1 \ Stars vs3 \ LV (STAR r) s2)" + using NonEmpty(2,3) by (auto simp add: LV_def) + then have "v :\val v3a \ (v :\val v3a \ Stars vs :\val Stars vs3)" using IH1 IH2 by blast + then have "v :\val v3a \ (v = v3a \ Stars vs :\val Stars vs3)" + unfolding PosOrd_ex_eq_def by auto + then have "Stars (v # vs) :\val Stars (v3a # vs3)" using NonEmpty(4) q2 as1 + unfolding PosOrd_ex_eq_def + using PosOrd_StarsI PosOrd_StarsI2 by auto + then show "Stars (v # vs) :\val v3" unfolding NonEmpty by blast + next + case Empty + have "v3 = Stars []" by fact + then show "Stars (v # vs) :\val v3" + unfolding PosOrd_ex_eq_def using cond2 + by (simp add: PosOrd_shorterI) + qed +next + case (Posix_STAR2 r v2) + have "v2 \ LV (STAR r) []" by fact + then have "v2 = Stars []" + unfolding LV_def by (auto elim: Prf.cases) + then show "Stars [] :\val v2" + by (simp add: PosOrd_ex_eq_def) +qed + + +lemma Posix_PosOrd_reverse: + assumes "s \ r \ v1" + shows "\(\v2 \ LV r s. v2 :\val v1)" +using assms +by (metis Posix_PosOrd less_irrefl PosOrd_def + PosOrd_ex_eq_def PosOrd_ex_def PosOrd_trans) + +lemma PosOrd_Posix: + assumes "v1 \ LV r s" "\v\<^sub>2 \ LV r s. \ v\<^sub>2 :\val v1" + shows "s \ r \ v1" +proof - + have "s \ L r" using assms(1) unfolding LV_def + using L_flat_Prf1 by blast + then obtain vposix where vp: "s \ r \ vposix" + using lexer_correct_Some by blast + with assms(1) have "vposix :\val v1" by (simp add: Posix_PosOrd) + then have "vposix = v1 \ vposix :\val v1" unfolding PosOrd_ex_eq2 by auto + moreover + { assume "vposix :\val v1" + moreover + have "vposix \ LV r s" using vp + using Posix_LV by blast + ultimately have "False" using assms(2) by blast + } + ultimately show "s \ r \ v1" using vp by blast +qed + +lemma Least_existence: + assumes "LV r s \ {}" + shows " \vmin \ LV r s. \v \ LV r s. vmin :\val v" +proof - + from assms + obtain vposix where "s \ r \ vposix" + unfolding LV_def + using L_flat_Prf1 lexer_correct_Some by blast + then have "\v \ LV r s. vposix :\val v" + by (simp add: Posix_PosOrd) + then show "\vmin \ LV r s. \v \ LV r s. vmin :\val v" + using Posix_LV \s \ r \ vposix\ by blast +qed + +lemma Least_existence1: + assumes "LV r s \ {}" + shows " \!vmin \ LV r s. \v \ LV r s. vmin :\val v" +using Least_existence[OF assms] assms +using PosOrdeq_antisym by blast + +lemma Least_existence2: + assumes "LV r s \ {}" + shows " \!vmin \ LV r s. lexer r s = Some vmin \ (\v \ LV r s. vmin :\val v)" +using Least_existence[OF assms] assms +using PosOrdeq_antisym + using PosOrd_Posix PosOrd_ex_eq2 lexer_correctness(1) by auto + + +lemma Least_existence1_pre: + assumes "LV r s \ {}" + shows " \!vmin \ LV r s. \v \ (LV r s \ {v'. flat v' \spre s}). vmin :\val v" +using Least_existence[OF assms] assms +apply - +apply(erule bexE) +apply(rule_tac a="vmin" in ex1I) +apply(auto)[1] +apply (metis PosOrd_Posix PosOrd_ex_eq2 PosOrd_spreI PosOrdeq_antisym Posix1(2)) +apply(auto)[1] +apply(simp add: PosOrdeq_antisym) +done + +lemma + shows "partial_order_on UNIV {(v1, v2). v1 :\val v2}" +apply(simp add: partial_order_on_def) +apply(simp add: preorder_on_def refl_on_def) +apply(simp add: PosOrdeq_refl) +apply(auto) +apply(rule transI) +apply(auto intro: PosOrdeq_trans)[1] +apply(rule antisymI) +apply(simp add: PosOrdeq_antisym) +done + +lemma + "wf {(v1, v2). v1 :\val v2 \ v1 \ LV r s \ v2 \ LV r s}" +apply(rule finite_acyclic_wf) +prefer 2 +apply(simp add: acyclic_def) +apply(induct_tac rule: trancl.induct) +apply(auto)[1] +oops + + +unused_thms + +end \ No newline at end of file diff -r f9cdc295ccf7 -r f493a20feeb3 thys3/src/PosixSpec.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys3/src/PosixSpec.thy Sat Apr 30 00:50:08 2022 +0100 @@ -0,0 +1,380 @@ + +theory PosixSpec + imports RegLangs +begin + +section \"Plain" Values\ + +datatype val = + Void +| Char char +| Seq val val +| Right val +| Left val +| Stars "val list" + + +section \The string behind a value\ + +fun + flat :: "val \ string" +where + "flat (Void) = []" +| "flat (Char c) = [c]" +| "flat (Left v) = flat v" +| "flat (Right v) = flat v" +| "flat (Seq v1 v2) = (flat v1) @ (flat v2)" +| "flat (Stars []) = []" +| "flat (Stars (v#vs)) = (flat v) @ (flat (Stars vs))" + +abbreviation + "flats vs \ concat (map flat vs)" + +lemma flat_Stars [simp]: + "flat (Stars vs) = flats vs" +by (induct vs) (auto) + + +section \Lexical Values\ + +inductive + Prf :: "val \ rexp \ bool" ("\ _ : _" [100, 100] 100) +where + "\\ v1 : r1; \ v2 : r2\ \ \ Seq v1 v2 : SEQ r1 r2" +| "\ v1 : r1 \ \ Left v1 : ALT r1 r2" +| "\ v2 : r2 \ \ Right v2 : ALT r1 r2" +| "\ Void : ONE" +| "\ Char c : CH c" +| "\v \ set vs. \ v : r \ flat v \ [] \ \ Stars vs : STAR r" + +inductive_cases Prf_elims: + "\ v : ZERO" + "\ v : SEQ r1 r2" + "\ v : ALT r1 r2" + "\ v : ONE" + "\ v : CH c" + "\ vs : STAR r" + +lemma Prf_Stars_appendE: + assumes "\ Stars (vs1 @ vs2) : STAR r" + shows "\ Stars vs1 : STAR r \ \ Stars vs2 : STAR r" +using assms +by (auto intro: Prf.intros elim!: Prf_elims) + + +lemma flats_Prf_value: + assumes "\s\set ss. \v. s = flat v \ \ v : r" + shows "\vs. flats vs = concat ss \ (\v\set vs. \ v : r \ flat v \ [])" +using assms +apply(induct ss) +apply(auto) +apply(rule_tac x="[]" in exI) +apply(simp) +apply(case_tac "flat v = []") +apply(rule_tac x="vs" in exI) +apply(simp) +apply(rule_tac x="v#vs" in exI) +apply(simp) +done + + +lemma L_flat_Prf1: + assumes "\ v : r" + shows "flat v \ L r" +using assms +by (induct) (auto simp add: Sequ_def Star_concat) + +lemma L_flat_Prf2: + assumes "s \ L r" + shows "\v. \ v : r \ flat v = s" +using assms +proof(induct r arbitrary: s) + case (STAR r s) + have IH: "\s. s \ L r \ \v. \ v : r \ flat v = s" by fact + have "s \ L (STAR r)" by fact + then obtain ss where "concat ss = s" "\s \ set ss. s \ L r \ s \ []" + using Star_split by auto + then obtain vs where "flats vs = s" "\v\set vs. \ v : r \ flat v \ []" + using IH flats_Prf_value by metis + then show "\v. \ v : STAR r \ flat v = s" + using Prf.intros(6) flat_Stars by blast +next + case (SEQ r1 r2 s) + then show "\v. \ v : SEQ r1 r2 \ flat v = s" + unfolding Sequ_def L.simps by (fastforce intro: Prf.intros) +next + case (ALT r1 r2 s) + then show "\v. \ v : ALT r1 r2 \ flat v = s" + unfolding L.simps by (fastforce intro: Prf.intros) +qed (auto intro: Prf.intros) + + +lemma L_flat_Prf: + shows "L(r) = {flat v | v. \ v : r}" +using L_flat_Prf1 L_flat_Prf2 by blast + + + +section \Sets of Lexical Values\ + +text \ + Shows that lexical values are finite for a given regex and string. +\ + +definition + LV :: "rexp \ string \ val set" +where "LV r s \ {v. \ v : r \ flat v = s}" + +lemma LV_simps: + shows "LV ZERO s = {}" + and "LV ONE s = (if s = [] then {Void} else {})" + and "LV (CH c) s = (if s = [c] then {Char c} else {})" + and "LV (ALT r1 r2) s = Left ` LV r1 s \ Right ` LV r2 s" +unfolding LV_def +by (auto intro: Prf.intros elim: Prf.cases) + + +abbreviation + "Prefixes s \ {s'. prefix s' s}" + +abbreviation + "Suffixes s \ {s'. suffix s' s}" + +abbreviation + "SSuffixes s \ {s'. strict_suffix s' s}" + +lemma Suffixes_cons [simp]: + shows "Suffixes (c # s) = Suffixes s \ {c # s}" +by (auto simp add: suffix_def Cons_eq_append_conv) + + +lemma finite_Suffixes: + shows "finite (Suffixes s)" +by (induct s) (simp_all) + +lemma finite_SSuffixes: + shows "finite (SSuffixes s)" +proof - + have "SSuffixes s \ Suffixes s" + unfolding strict_suffix_def suffix_def by auto + then show "finite (SSuffixes s)" + using finite_Suffixes finite_subset by blast +qed + +lemma finite_Prefixes: + shows "finite (Prefixes s)" +proof - + have "finite (Suffixes (rev s))" + by (rule finite_Suffixes) + then have "finite (rev ` Suffixes (rev s))" by simp + moreover + have "rev ` (Suffixes (rev s)) = Prefixes s" + unfolding suffix_def prefix_def image_def + by (auto)(metis rev_append rev_rev_ident)+ + ultimately show "finite (Prefixes s)" by simp +qed + +lemma LV_STAR_finite: + assumes "\s. finite (LV r s)" + shows "finite (LV (STAR r) s)" +proof(induct s rule: length_induct) + fix s::"char list" + assume "\s'. length s' < length s \ finite (LV (STAR r) s')" + then have IH: "\s' \ SSuffixes s. finite (LV (STAR r) s')" + by (force simp add: strict_suffix_def suffix_def) + define f where "f \ \(v, vs). Stars (v # vs)" + define S1 where "S1 \ \s' \ Prefixes s. LV r s'" + define S2 where "S2 \ \s2 \ SSuffixes s. Stars -` (LV (STAR r) s2)" + have "finite S1" using assms + unfolding S1_def by (simp_all add: finite_Prefixes) + moreover + with IH have "finite S2" unfolding S2_def + by (auto simp add: finite_SSuffixes inj_on_def finite_vimageI) + ultimately + have "finite ({Stars []} \ f ` (S1 \ S2))" by simp + moreover + have "LV (STAR r) s \ {Stars []} \ f ` (S1 \ S2)" + unfolding S1_def S2_def f_def + unfolding LV_def image_def prefix_def strict_suffix_def + apply(auto) + apply(case_tac x) + apply(auto elim: Prf_elims) + apply(erule Prf_elims) + apply(auto) + apply(case_tac vs) + apply(auto intro: Prf.intros) + apply(rule exI) + apply(rule conjI) + apply(rule_tac x="flat a" in exI) + apply(rule conjI) + apply(rule_tac x="flats list" in exI) + apply(simp) + apply(blast) + apply(simp add: suffix_def) + using Prf.intros(6) by blast + ultimately + show "finite (LV (STAR r) s)" by (simp add: finite_subset) +qed + + +lemma LV_finite: + shows "finite (LV r s)" +proof(induct r arbitrary: s) + case (ZERO s) + show "finite (LV ZERO s)" by (simp add: LV_simps) +next + case (ONE s) + show "finite (LV ONE s)" by (simp add: LV_simps) +next + case (CH c s) + show "finite (LV (CH c) s)" by (simp add: LV_simps) +next + case (ALT r1 r2 s) + then show "finite (LV (ALT r1 r2) s)" by (simp add: LV_simps) +next + case (SEQ r1 r2 s) + define f where "f \ \(v1, v2). Seq v1 v2" + define S1 where "S1 \ \s' \ Prefixes s. LV r1 s'" + define S2 where "S2 \ \s' \ Suffixes s. LV r2 s'" + have IHs: "\s. finite (LV r1 s)" "\s. finite (LV r2 s)" by fact+ + then have "finite S1" "finite S2" unfolding S1_def S2_def + by (simp_all add: finite_Prefixes finite_Suffixes) + moreover + have "LV (SEQ r1 r2) s \ f ` (S1 \ S2)" + unfolding f_def S1_def S2_def + unfolding LV_def image_def prefix_def suffix_def + apply (auto elim!: Prf_elims) + by (metis (mono_tags, lifting) mem_Collect_eq) + ultimately + show "finite (LV (SEQ r1 r2) s)" + by (simp add: finite_subset) +next + case (STAR r s) + then show "finite (LV (STAR r) s)" by (simp add: LV_STAR_finite) +qed + + + +section \Our inductive POSIX Definition\ + +inductive + Posix :: "string \ rexp \ val \ bool" ("_ \ _ \ _" [100, 100, 100] 100) +where + Posix_ONE: "[] \ ONE \ Void" +| Posix_CH: "[c] \ (CH c) \ (Char c)" +| Posix_ALT1: "s \ r1 \ v \ s \ (ALT r1 r2) \ (Left v)" +| Posix_ALT2: "\s \ r2 \ v; s \ L(r1)\ \ s \ (ALT r1 r2) \ (Right v)" +| Posix_SEQ: "\s1 \ r1 \ v1; s2 \ r2 \ v2; + \(\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (s1 @ s\<^sub>3) \ L r1 \ s\<^sub>4 \ L r2)\ \ + (s1 @ s2) \ (SEQ r1 r2) \ (Seq v1 v2)" +| Posix_STAR1: "\s1 \ r \ v; s2 \ STAR r \ Stars vs; flat v \ []; + \(\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ (s1 @ s\<^sub>3) \ L r \ s\<^sub>4 \ L (STAR r))\ + \ (s1 @ s2) \ STAR r \ Stars (v # vs)" +| Posix_STAR2: "[] \ STAR r \ Stars []" + +inductive_cases Posix_elims: + "s \ ZERO \ v" + "s \ ONE \ v" + "s \ CH c \ v" + "s \ ALT r1 r2 \ v" + "s \ SEQ r1 r2 \ v" + "s \ STAR r \ v" + +lemma Posix1: + assumes "s \ r \ v" + shows "s \ L r" "flat v = s" +using assms + by(induct s r v rule: Posix.induct) + (auto simp add: Sequ_def) + +text \ + For a give value and string, our Posix definition + determines a unique value. +\ + +lemma Posix_determ: + assumes "s \ r \ v1" "s \ r \ v2" + shows "v1 = v2" +using assms +proof (induct s r v1 arbitrary: v2 rule: Posix.induct) + case (Posix_ONE v2) + have "[] \ ONE \ v2" by fact + then show "Void = v2" by cases auto +next + case (Posix_CH c v2) + have "[c] \ CH c \ v2" by fact + then show "Char c = v2" by cases auto +next + case (Posix_ALT1 s r1 v r2 v2) + have "s \ ALT r1 r2 \ v2" by fact + moreover + have "s \ r1 \ v" by fact + then have "s \ L r1" by (simp add: Posix1) + ultimately obtain v' where eq: "v2 = Left v'" "s \ r1 \ v'" by cases auto + moreover + have IH: "\v2. s \ r1 \ v2 \ v = v2" by fact + ultimately have "v = v'" by simp + then show "Left v = v2" using eq by simp +next + case (Posix_ALT2 s r2 v r1 v2) + have "s \ ALT r1 r2 \ v2" by fact + moreover + have "s \ L r1" by fact + ultimately obtain v' where eq: "v2 = Right v'" "s \ r2 \ v'" + by cases (auto simp add: Posix1) + moreover + have IH: "\v2. s \ r2 \ v2 \ v = v2" by fact + ultimately have "v = v'" by simp + then show "Right v = v2" using eq by simp +next + case (Posix_SEQ s1 r1 v1 s2 r2 v2 v') + have "(s1 @ s2) \ SEQ r1 r2 \ v'" + "s1 \ r1 \ v1" "s2 \ r2 \ v2" + "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L r1 \ s\<^sub>4 \ L r2)" by fact+ + then obtain v1' v2' where "v' = Seq v1' v2'" "s1 \ r1 \ v1'" "s2 \ r2 \ v2'" + apply(cases) apply (auto simp add: append_eq_append_conv2) + using Posix1(1) by fastforce+ + moreover + have IHs: "\v1'. s1 \ r1 \ v1' \ v1 = v1'" + "\v2'. s2 \ r2 \ v2' \ v2 = v2'" by fact+ + ultimately show "Seq v1 v2 = v'" by simp +next + case (Posix_STAR1 s1 r v s2 vs v2) + have "(s1 @ s2) \ STAR r \ v2" + "s1 \ r \ v" "s2 \ STAR r \ Stars vs" "flat v \ []" + "\ (\s\<^sub>3 s\<^sub>4. s\<^sub>3 \ [] \ s\<^sub>3 @ s\<^sub>4 = s2 \ s1 @ s\<^sub>3 \ L r \ s\<^sub>4 \ L (STAR r))" by fact+ + then obtain v' vs' where "v2 = Stars (v' # vs')" "s1 \ r \ v'" "s2 \ (STAR r) \ (Stars vs')" + apply(cases) apply (auto simp add: append_eq_append_conv2) + using Posix1(1) apply fastforce + apply (metis Posix1(1) Posix_STAR1.hyps(6) append_Nil append_Nil2) + using Posix1(2) by blast + moreover + have IHs: "\v2. s1 \ r \ v2 \ v = v2" + "\v2. s2 \ STAR r \ v2 \ Stars vs = v2" by fact+ + ultimately show "Stars (v # vs) = v2" by auto +next + case (Posix_STAR2 r v2) + have "[] \ STAR r \ v2" by fact + then show "Stars [] = v2" by cases (auto simp add: Posix1) +qed + + +text \ + Our POSIX values are lexical values. +\ + +lemma Posix_LV: + assumes "s \ r \ v" + shows "v \ LV r s" + using assms unfolding LV_def + apply(induct rule: Posix.induct) + apply(auto simp add: intro!: Prf.intros elim!: Prf_elims) + done + +lemma Posix_Prf: + assumes "s \ r \ v" + shows "\ v : r" + using assms Posix_LV LV_def + by simp + +end diff -r f9cdc295ccf7 -r f493a20feeb3 thys3/src/RegLangs.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/thys3/src/RegLangs.thy Sat Apr 30 00:50:08 2022 +0100 @@ -0,0 +1,236 @@ +theory RegLangs + imports Main "HOL-Library.Sublist" +begin + +section \Sequential Composition of Languages\ + +definition + Sequ :: "string set \ string set \ string set" ("_ ;; _" [100,100] 100) +where + "A ;; B = {s1 @ s2 | s1 s2. s1 \ A \ s2 \ B}" + +text \Two Simple Properties about Sequential Composition\ + +lemma Sequ_empty_string [simp]: + shows "A ;; {[]} = A" + and "{[]} ;; A = A" +by (simp_all add: Sequ_def) + +lemma Sequ_empty [simp]: + shows "A ;; {} = {}" + and "{} ;; A = {}" + by (simp_all add: Sequ_def) + + +section \Semantic Derivative (Left Quotient) of Languages\ + +definition + Der :: "char \ string set \ string set" +where + "Der c A \ {s. c # s \ A}" + +definition + Ders :: "string \ string set \ string set" +where + "Ders s A \ {s'. s @ s' \ A}" + +lemma Der_null [simp]: + shows "Der c {} = {}" +unfolding Der_def +by auto + +lemma Der_empty [simp]: + shows "Der c {[]} = {}" +unfolding Der_def +by auto + +lemma Der_char [simp]: + shows "Der c {[d]} = (if c = d then {[]} else {})" +unfolding Der_def +by auto + +lemma Der_union [simp]: + shows "Der c (A \ B) = Der c A \ Der c B" +unfolding Der_def +by auto + +lemma Der_Sequ [simp]: + shows "Der c (A ;; B) = (Der c A) ;; B \ (if [] \ A then Der c B else {})" +unfolding Der_def Sequ_def +by (auto simp add: Cons_eq_append_conv) + + +section \Kleene Star for Languages\ + +inductive_set + Star :: "string set \ string set" ("_\" [101] 102) + for A :: "string set" +where + start[intro]: "[] \ A\" +| step[intro]: "\s1 \ A; s2 \ A\\ \ s1 @ s2 \ A\" + +(* Arden's lemma *) + +lemma Star_cases: + shows "A\ = {[]} \ A ;; A\" +unfolding Sequ_def +by (auto) (metis Star.simps) + +lemma Star_decomp: + assumes "c # x \ A\" + shows "\s1 s2. x = s1 @ s2 \ c # s1 \ A \ s2 \ A\" +using assms +by (induct x\"c # x" rule: Star.induct) + (auto simp add: append_eq_Cons_conv) + +lemma Star_Der_Sequ: + shows "Der c (A\) \ (Der c A) ;; A\" +unfolding Der_def Sequ_def +by(auto simp add: Star_decomp) + + +lemma Der_star[simp]: + shows "Der c (A\) = (Der c A) ;; A\" +proof - + have "Der c (A\) = Der c ({[]} \ A ;; A\)" + by (simp only: Star_cases[symmetric]) + also have "... = Der c (A ;; A\)" + by (simp only: Der_union Der_empty) (simp) + also have "... = (Der c A) ;; A\ \ (if [] \ A then Der c (A\) else {})" + by simp + also have "... = (Der c A) ;; A\" + using Star_Der_Sequ by auto + finally show "Der c (A\) = (Der c A) ;; A\" . +qed + +lemma Star_concat: + assumes "\s \ set ss. s \ A" + shows "concat ss \ A\" +using assms by (induct ss) (auto) + +lemma Star_split: + assumes "s \ A\" + shows "\ss. concat ss = s \ (\s \ set ss. s \ A \ s \ [])" +using assms + apply(induct rule: Star.induct) + using concat.simps(1) apply fastforce + apply(clarify) + by (metis append_Nil concat.simps(2) set_ConsD) + + + +section \Regular Expressions\ + +datatype rexp = + ZERO +| ONE +| CH char +| SEQ rexp rexp +| ALT rexp rexp +| STAR rexp + +section \Semantics of Regular Expressions\ + +fun + L :: "rexp \ string set" +where + "L (ZERO) = {}" +| "L (ONE) = {[]}" +| "L (CH c) = {[c]}" +| "L (SEQ r1 r2) = (L r1) ;; (L r2)" +| "L (ALT r1 r2) = (L r1) \ (L r2)" +| "L (STAR r) = (L r)\" + + +section \Nullable, Derivatives\ + +fun + nullable :: "rexp \ bool" +where + "nullable (ZERO) = False" +| "nullable (ONE) = True" +| "nullable (CH c) = False" +| "nullable (ALT r1 r2) = (nullable r1 \ nullable r2)" +| "nullable (SEQ r1 r2) = (nullable r1 \ nullable r2)" +| "nullable (STAR r) = True" + + +fun + der :: "char \ rexp \ rexp" +where + "der c (ZERO) = ZERO" +| "der c (ONE) = ZERO" +| "der c (CH d) = (if c = d then ONE else ZERO)" +| "der c (ALT r1 r2) = ALT (der c r1) (der c r2)" +| "der c (SEQ r1 r2) = + (if nullable r1 + then ALT (SEQ (der c r1) r2) (der c r2) + else SEQ (der c r1) r2)" +| "der c (STAR r) = SEQ (der c r) (STAR r)" + +fun + ders :: "string \ rexp \ rexp" +where + "ders [] r = r" +| "ders (c # s) r = ders s (der c r)" + + +lemma nullable_correctness: + shows "nullable r \ [] \ (L r)" +by (induct r) (auto simp add: Sequ_def) + +lemma der_correctness: + shows "L (der c r) = Der c (L r)" +by (induct r) (simp_all add: nullable_correctness) + +lemma ders_correctness: + shows "L (ders s r) = Ders s (L r)" + by (induct s arbitrary: r) + (simp_all add: Ders_def der_correctness Der_def) + +lemma ders_append: + shows "ders (s1 @ s2) r = ders s2 (ders s1 r)" + by (induct s1 arbitrary: s2 r) (auto) + +lemma ders_snoc: + shows "ders (s @ [c]) r = der c (ders s r)" + by (simp add: ders_append) + + +(* +datatype ctxt = + SeqC rexp bool + | AltCL rexp + | AltCH rexp + | StarC rexp + +function + down :: "char \ rexp \ ctxt list \ rexp * ctxt list" +and up :: "char \ rexp \ ctxt list \ rexp * ctxt list" +where + "down c (SEQ r1 r2) ctxts = + (if (nullable r1) then down c r1 (SeqC r2 True # ctxts) + else down c r1 (SeqC r2 False # ctxts))" +| "down c (CH d) ctxts = + (if c = d then up c ONE ctxts else up c ZERO ctxts)" +| "down c ONE ctxts = up c ZERO ctxts" +| "down c ZERO ctxts = up c ZERO ctxts" +| "down c (ALT r1 r2) ctxts = down c r1 (AltCH r2 # ctxts)" +| "down c (STAR r1) ctxts = down c r1 (StarC r1 # ctxts)" +| "up c r [] = (r, [])" +| "up c r (SeqC r2 False # ctxts) = up c (SEQ r r2) ctxts" +| "up c r (SeqC r2 True # ctxts) = down c r2 (AltCL (SEQ r r2) # ctxts)" +| "up c r (AltCL r1 # ctxts) = up c (ALT r1 r) ctxts" +| "up c r (AltCH r2 # ctxts) = down c r2 (AltCL r # ctxts)" +| "up c r (StarC r1 # ctxts) = up c (SEQ r (STAR r1)) ctxts" + apply(pat_completeness) + apply(auto) + done + +termination + sorry + +*) + + +end \ No newline at end of file