| author | webertj | 
| Mon, 15 Nov 2004 17:04:11 +0100 | |
| changeset 15287 | 55b7f7920622 | 
| parent 14379 | ea10a8c3e9cf | 
| child 15481 | fc075ae929e4 | 
| permissions | -rw-r--r-- | 
| 10217 | 1 | % | 
| 2 | \begin{isabellebody}%
 | |
| 3 | \def\isabellecontext{AB}%
 | |
| 11866 | 4 | \isamarkupfalse% | 
| 10225 | 5 | % | 
| 10878 | 6 | \isamarkupsection{Case Study: A Context Free Grammar%
 | 
| 10395 | 7 | } | 
| 11866 | 8 | \isamarkuptrue% | 
| 10236 | 9 | % | 
| 10 | \begin{isamarkuptext}%
 | |
| 10242 | 11 | \label{sec:CFG}
 | 
| 11494 | 12 | \index{grammars!defining inductively|(}%
 | 
| 10236 | 13 | Grammars are nothing but shorthands for inductive definitions of nonterminals | 
| 14 | which represent sets of strings. For example, the production | |
| 15 | $A \to B c$ is short for | |
| 16 | \[ w \in B \Longrightarrow wc \in A \] | |
| 10878 | 17 | This section demonstrates this idea with an example | 
| 18 | due to Hopcroft and Ullman, a grammar for generating all words with an | |
| 19 | equal number of $a$'s and~$b$'s: | |
| 10236 | 20 | \begin{eqnarray}
 | 
| 21 | S &\to& \epsilon \mid b A \mid a B \nonumber\\ | |
| 22 | A &\to& a S \mid b A A \nonumber\\ | |
| 23 | B &\to& b S \mid a B B \nonumber | |
| 24 | \end{eqnarray}
 | |
| 10878 | 25 | At the end we say a few words about the relationship between | 
| 26 | the original proof \cite[p.\ts81]{HopcroftUllman} and our formal version.
 | |
| 10236 | 27 | |
| 10299 | 28 | We start by fixing the alphabet, which consists only of \isa{a}'s
 | 
| 10878 | 29 | and~\isa{b}'s:%
 | 
| 10236 | 30 | \end{isamarkuptext}%
 | 
| 11866 | 31 | \isamarkuptrue% | 
| 32 | \isacommand{datatype}\ alfa\ {\isacharequal}\ a\ {\isacharbar}\ b\isamarkupfalse%
 | |
| 33 | % | |
| 10236 | 34 | \begin{isamarkuptext}%
 | 
| 35 | \noindent | |
| 10299 | 36 | For convenience we include the following easy lemmas as simplification rules:% | 
| 10236 | 37 | \end{isamarkuptext}%
 | 
| 11866 | 38 | \isamarkuptrue% | 
| 10236 | 39 | \isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}x\ {\isasymnoteq}\ a{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharequal}\ b{\isacharparenright}\ {\isasymand}\ {\isacharparenleft}x\ {\isasymnoteq}\ b{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}x\ {\isacharequal}\ a{\isacharparenright}{\isachardoublequote}\isanewline
 | 
| 11866 | 40 | \isamarkupfalse% | 
| 41 | \isacommand{by}\ {\isacharparenleft}case{\isacharunderscore}tac\ x{\isacharcomma}\ auto{\isacharparenright}\isamarkupfalse%
 | |
| 42 | % | |
| 10236 | 43 | \begin{isamarkuptext}%
 | 
| 44 | \noindent | |
| 45 | Words over this alphabet are of type \isa{alfa\ list}, and
 | |
| 10878 | 46 | the three nonterminals are declared as sets of such words:% | 
| 10236 | 47 | \end{isamarkuptext}%
 | 
| 11866 | 48 | \isamarkuptrue% | 
| 10217 | 49 | \isacommand{consts}\ S\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}alfa\ list\ set{\isachardoublequote}\isanewline
 | 
| 50 | \ \ \ \ \ \ \ A\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}alfa\ list\ set{\isachardoublequote}\isanewline
 | |
| 11866 | 51 | \ \ \ \ \ \ \ B\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}alfa\ list\ set{\isachardoublequote}\isamarkupfalse%
 | 
| 52 | % | |
| 10236 | 53 | \begin{isamarkuptext}%
 | 
| 54 | \noindent | |
| 10878 | 55 | The productions above are recast as a \emph{mutual} inductive
 | 
| 10242 | 56 | definition\index{inductive definition!simultaneous}
 | 
| 10878 | 57 | of \isa{S}, \isa{A} and~\isa{B}:%
 | 
| 10236 | 58 | \end{isamarkuptext}%
 | 
| 11866 | 59 | \isamarkuptrue% | 
| 10217 | 60 | \isacommand{inductive}\ S\ A\ B\isanewline
 | 
| 61 | \isakeyword{intros}\isanewline
 | |
| 10236 | 62 | \ \ {\isachardoublequote}{\isacharbrackleft}{\isacharbrackright}\ {\isasymin}\ S{\isachardoublequote}\isanewline
 | 
| 63 | \ \ {\isachardoublequote}w\ {\isasymin}\ A\ {\isasymLongrightarrow}\ b{\isacharhash}w\ {\isasymin}\ S{\isachardoublequote}\isanewline
 | |
| 64 | \ \ {\isachardoublequote}w\ {\isasymin}\ B\ {\isasymLongrightarrow}\ a{\isacharhash}w\ {\isasymin}\ S{\isachardoublequote}\isanewline
 | |
| 10217 | 65 | \isanewline | 
| 10236 | 66 | \ \ {\isachardoublequote}w\ {\isasymin}\ S\ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ a{\isacharhash}w\ \ \ {\isasymin}\ A{\isachardoublequote}\isanewline
 | 
| 67 | \ \ {\isachardoublequote}{\isasymlbrakk}\ v{\isasymin}A{\isacharsemicolon}\ w{\isasymin}A\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ b{\isacharhash}v{\isacharat}w\ {\isasymin}\ A{\isachardoublequote}\isanewline
 | |
| 10217 | 68 | \isanewline | 
| 10236 | 69 | \ \ {\isachardoublequote}w\ {\isasymin}\ S\ \ \ \ \ \ \ \ \ \ \ \ {\isasymLongrightarrow}\ b{\isacharhash}w\ \ \ {\isasymin}\ B{\isachardoublequote}\isanewline
 | 
| 11866 | 70 | \ \ {\isachardoublequote}{\isasymlbrakk}\ v\ {\isasymin}\ B{\isacharsemicolon}\ w\ {\isasymin}\ B\ {\isasymrbrakk}\ {\isasymLongrightarrow}\ a{\isacharhash}v{\isacharat}w\ {\isasymin}\ B{\isachardoublequote}\isamarkupfalse%
 | 
| 71 | % | |
| 10236 | 72 | \begin{isamarkuptext}%
 | 
| 73 | \noindent | |
| 10878 | 74 | First we show that all words in \isa{S} contain the same number of \isa{a}'s and \isa{b}'s. Since the definition of \isa{S} is by mutual
 | 
| 75 | induction, so is the proof: we show at the same time that all words in | |
| 10236 | 76 | \isa{A} contain one more \isa{a} than \isa{b} and all words in \isa{B} contains one more \isa{b} than \isa{a}.%
 | 
| 77 | \end{isamarkuptext}%
 | |
| 11866 | 78 | \isamarkuptrue% | 
| 10236 | 79 | \isacommand{lemma}\ correctness{\isacharcolon}\isanewline
 | 
| 80 | \ \ {\isachardoublequote}{\isacharparenleft}w\ {\isasymin}\ S\ {\isasymlongrightarrow}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}a{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}b{\isacharbrackright}{\isacharparenright}\ \ \ \ \ {\isasymand}\isanewline
 | |
| 10237 | 81 | \ \ \ {\isacharparenleft}w\ {\isasymin}\ A\ {\isasymlongrightarrow}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}a{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}b{\isacharbrackright}\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}\ {\isasymand}\isanewline
 | 
| 11866 | 82 | \ \ \ {\isacharparenleft}w\ {\isasymin}\ B\ {\isasymlongrightarrow}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}b{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}a{\isacharbrackright}\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}{\isachardoublequote}\isamarkupfalse%
 | 
| 83 | % | |
| 10236 | 84 | \begin{isamarkuptxt}%
 | 
| 85 | \noindent | |
| 10283 | 86 | These propositions are expressed with the help of the predefined \isa{filter} function on lists, which has the convenient syntax \isa{{\isacharbrackleft}x{\isasymin}xs{\isachardot}\ P\ x{\isacharbrackright}}, the list of all elements \isa{x} in \isa{xs} such that \isa{P\ x}
 | 
| 10878 | 87 | holds. Remember that on lists \isa{size} and \isa{length} are synonymous.
 | 
| 10236 | 88 | |
| 89 | The proof itself is by rule induction and afterwards automatic:% | |
| 90 | \end{isamarkuptxt}%
 | |
| 11866 | 91 | \isamarkuptrue% | 
| 92 | \isacommand{by}\ {\isacharparenleft}rule\ S{\isacharunderscore}A{\isacharunderscore}B{\isachardot}induct{\isacharcomma}\ auto{\isacharparenright}\isamarkupfalse%
 | |
| 93 | % | |
| 10236 | 94 | \begin{isamarkuptext}%
 | 
| 95 | \noindent | |
| 96 | This may seem surprising at first, and is indeed an indication of the power | |
| 97 | of inductive definitions. But it is also quite straightforward. For example, | |
| 98 | consider the production $A \to b A A$: if $v,w \in A$ and the elements of $A$ | |
| 10878 | 99 | contain one more $a$ than~$b$'s, then $bvw$ must again contain one more $a$ | 
| 100 | than~$b$'s. | |
| 10236 | 101 | |
| 102 | As usual, the correctness of syntactic descriptions is easy, but completeness | |
| 103 | is hard: does \isa{S} contain \emph{all} words with an equal number of
 | |
| 104 | \isa{a}'s and \isa{b}'s? It turns out that this proof requires the
 | |
| 10878 | 105 | following lemma: every string with two more \isa{a}'s than \isa{b}'s can be cut somewhere such that each half has one more \isa{a} than
 | 
| 10236 | 106 | \isa{b}. This is best seen by imagining counting the difference between the
 | 
| 10283 | 107 | number of \isa{a}'s and \isa{b}'s starting at the left end of the
 | 
| 108 | word. We start with 0 and end (at the right end) with 2. Since each move to the | |
| 10236 | 109 | right increases or decreases the difference by 1, we must have passed through | 
| 110 | 1 on our way from 0 to 2. Formally, we appeal to the following discrete | |
| 111 | intermediate value theorem \isa{nat{\isadigit{0}}{\isacharunderscore}intermed{\isacharunderscore}int{\isacharunderscore}val}
 | |
| 112 | \begin{isabelle}%
 | |
| 14379 | 113 | \ \ \ \ \ {\isasymlbrakk}{\isasymforall}i{\isacharless}n{\isachardot}\ {\isasymbar}f\ {\isacharparenleft}i\ {\isacharplus}\ {\isadigit{1}}{\isacharparenright}\ {\isacharminus}\ f\ i{\isasymbar}\ {\isasymle}\ {\isadigit{1}}{\isacharsemicolon}\ f\ {\isadigit{0}}\ {\isasymle}\ k{\isacharsemicolon}\ k\ {\isasymle}\ f\ n{\isasymrbrakk}\ {\isasymLongrightarrow}\ {\isasymexists}i{\isasymle}n{\isachardot}\ f\ i\ {\isacharequal}\ k%
 | 
| 10236 | 114 | \end{isabelle}
 | 
| 115 | where \isa{f} is of type \isa{nat\ {\isasymRightarrow}\ int}, \isa{int} are the integers,
 | |
| 11308 | 116 | \isa{{\isasymbar}{\isachardot}{\isasymbar}} is the absolute value function\footnote{See
 | 
| 117 | Table~\ref{tab:ascii} in the Appendix for the correct \textsc{ascii}
 | |
| 11708 | 118 | syntax.}, and \isa{{\isadigit{1}}} is the integer 1 (see \S\ref{sec:numbers}).
 | 
| 10236 | 119 | |
| 11147 | 120 | First we show that our specific function, the difference between the | 
| 10236 | 121 | numbers of \isa{a}'s and \isa{b}'s, does indeed only change by 1 in every
 | 
| 122 | move to the right. At this point we also start generalizing from \isa{a}'s
 | |
| 123 | and \isa{b}'s to an arbitrary property \isa{P}. Otherwise we would have
 | |
| 124 | to prove the desired lemma twice, once as stated above and once with the | |
| 125 | roles of \isa{a}'s and \isa{b}'s interchanged.%
 | |
| 126 | \end{isamarkuptext}%
 | |
| 11866 | 127 | \isamarkuptrue% | 
| 10236 | 128 | \isacommand{lemma}\ step{\isadigit{1}}{\isacharcolon}\ {\isachardoublequote}{\isasymforall}i\ {\isacharless}\ size\ w{\isachardot}\isanewline
 | 
| 10608 | 129 | \ \ {\isasymbar}{\isacharparenleft}int{\isacharparenleft}size{\isacharbrackleft}x{\isasymin}take\ {\isacharparenleft}i{\isacharplus}{\isadigit{1}}{\isacharparenright}\ w{\isachardot}\ P\ x{\isacharbrackright}{\isacharparenright}{\isacharminus}int{\isacharparenleft}size{\isacharbrackleft}x{\isasymin}take\ {\isacharparenleft}i{\isacharplus}{\isadigit{1}}{\isacharparenright}\ w{\isachardot}\ {\isasymnot}P\ x{\isacharbrackright}{\isacharparenright}{\isacharparenright}\isanewline
 | 
| 11870 
181bd2050cf4
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11866diff
changeset | 130 | \ \ \ {\isacharminus}\ {\isacharparenleft}int{\isacharparenleft}size{\isacharbrackleft}x{\isasymin}take\ i\ w{\isachardot}\ P\ x{\isacharbrackright}{\isacharparenright}{\isacharminus}int{\isacharparenleft}size{\isacharbrackleft}x{\isasymin}take\ i\ w{\isachardot}\ {\isasymnot}P\ x{\isacharbrackright}{\isacharparenright}{\isacharparenright}{\isasymbar}\ {\isasymle}\ {\isadigit{1}}{\isachardoublequote}\isamarkupfalse%
 | 
| 11866 | 131 | % | 
| 10236 | 132 | \begin{isamarkuptxt}%
 | 
| 133 | \noindent | |
| 134 | The lemma is a bit hard to read because of the coercion function | |
| 11147 | 135 | \isa{int\ {\isacharcolon}{\isacharcolon}\ nat\ {\isasymRightarrow}\ int}. It is required because \isa{size} returns
 | 
| 10878 | 136 | a natural number, but subtraction on type~\isa{nat} will do the wrong thing.
 | 
| 10236 | 137 | Function \isa{take} is predefined and \isa{take\ i\ xs} is the prefix of
 | 
| 10878 | 138 | length \isa{i} of \isa{xs}; below we also need \isa{drop\ i\ xs}, which
 | 
| 10236 | 139 | is what remains after that prefix has been dropped from \isa{xs}.
 | 
| 140 | ||
| 141 | The proof is by induction on \isa{w}, with a trivial base case, and a not
 | |
| 142 | so trivial induction step. Since it is essentially just arithmetic, we do not | |
| 143 | discuss it.% | |
| 144 | \end{isamarkuptxt}%
 | |
| 11866 | 145 | \isamarkuptrue% | 
| 12332 | 146 | \isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ w{\isacharparenright}\isanewline
 | 
| 11866 | 147 | \ \isamarkupfalse% | 
| 148 | \isacommand{apply}{\isacharparenleft}simp{\isacharparenright}\isanewline
 | |
| 149 | \isamarkupfalse% | |
| 150 | \isacommand{by}{\isacharparenleft}force\ simp\ add{\isacharcolon}\ zabs{\isacharunderscore}def\ take{\isacharunderscore}Cons\ split{\isacharcolon}\ nat{\isachardot}split\ if{\isacharunderscore}splits{\isacharparenright}\isamarkupfalse%
 | |
| 151 | % | |
| 10236 | 152 | \begin{isamarkuptext}%
 | 
| 11494 | 153 | Finally we come to the above-mentioned lemma about cutting in half a word with two more elements of one sort than of the other sort:% | 
| 10236 | 154 | \end{isamarkuptext}%
 | 
| 11866 | 155 | \isamarkuptrue% | 
| 10236 | 156 | \isacommand{lemma}\ part{\isadigit{1}}{\isacharcolon}\isanewline
 | 
| 157 | \ {\isachardoublequote}size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ P\ x{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ {\isasymnot}P\ x{\isacharbrackright}{\isacharplus}{\isadigit{2}}\ {\isasymLongrightarrow}\isanewline
 | |
| 11866 | 158 | \ \ {\isasymexists}i{\isasymle}size\ w{\isachardot}\ size{\isacharbrackleft}x{\isasymin}take\ i\ w{\isachardot}\ P\ x{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}take\ i\ w{\isachardot}\ {\isasymnot}P\ x{\isacharbrackright}{\isacharplus}{\isadigit{1}}{\isachardoublequote}\isamarkupfalse%
 | 
| 159 | % | |
| 10236 | 160 | \begin{isamarkuptxt}%
 | 
| 161 | \noindent | |
| 10878 | 162 | This is proved by \isa{force} with the help of the intermediate value theorem,
 | 
| 10608 | 163 | instantiated appropriately and with its first premise disposed of by lemma | 
| 164 | \isa{step{\isadigit{1}}}:%
 | |
| 10236 | 165 | \end{isamarkuptxt}%
 | 
| 11866 | 166 | \isamarkuptrue% | 
| 11870 
181bd2050cf4
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11866diff
changeset | 167 | \isacommand{apply}{\isacharparenleft}insert\ nat{\isadigit{0}}{\isacharunderscore}intermed{\isacharunderscore}int{\isacharunderscore}val{\isacharbrackleft}OF\ step{\isadigit{1}}{\isacharcomma}\ of\ {\isachardoublequote}P{\isachardoublequote}\ {\isachardoublequote}w{\isachardoublequote}\ {\isachardoublequote}{\isadigit{1}}{\isachardoublequote}{\isacharbrackright}{\isacharparenright}\isanewline
 | 
| 11866 | 168 | \isamarkupfalse% | 
| 169 | \isacommand{by}\ force\isamarkupfalse%
 | |
| 170 | % | |
| 10236 | 171 | \begin{isamarkuptext}%
 | 
| 172 | \noindent | |
| 173 | ||
| 174 | Lemma \isa{part{\isadigit{1}}} tells us only about the prefix \isa{take\ i\ w}.
 | |
| 10878 | 175 | An easy lemma deals with the suffix \isa{drop\ i\ w}:%
 | 
| 10236 | 176 | \end{isamarkuptext}%
 | 
| 11866 | 177 | \isamarkuptrue% | 
| 10236 | 178 | \isacommand{lemma}\ part{\isadigit{2}}{\isacharcolon}\isanewline
 | 
| 179 | \ \ {\isachardoublequote}{\isasymlbrakk}size{\isacharbrackleft}x{\isasymin}take\ i\ w\ {\isacharat}\ drop\ i\ w{\isachardot}\ P\ x{\isacharbrackright}\ {\isacharequal}\isanewline
 | |
| 180 | \ \ \ \ size{\isacharbrackleft}x{\isasymin}take\ i\ w\ {\isacharat}\ drop\ i\ w{\isachardot}\ {\isasymnot}P\ x{\isacharbrackright}{\isacharplus}{\isadigit{2}}{\isacharsemicolon}\isanewline
 | |
| 181 | \ \ \ \ size{\isacharbrackleft}x{\isasymin}take\ i\ w{\isachardot}\ P\ x{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}take\ i\ w{\isachardot}\ {\isasymnot}P\ x{\isacharbrackright}{\isacharplus}{\isadigit{1}}{\isasymrbrakk}\isanewline
 | |
| 182 | \ \ \ {\isasymLongrightarrow}\ size{\isacharbrackleft}x{\isasymin}drop\ i\ w{\isachardot}\ P\ x{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}drop\ i\ w{\isachardot}\ {\isasymnot}P\ x{\isacharbrackright}{\isacharplus}{\isadigit{1}}{\isachardoublequote}\isanewline
 | |
| 11866 | 183 | \isamarkupfalse% | 
| 12815 | 184 | \isacommand{by}{\isacharparenleft}simp\ del{\isacharcolon}\ append{\isacharunderscore}take{\isacharunderscore}drop{\isacharunderscore}id{\isacharparenright}\isamarkupfalse%
 | 
| 11866 | 185 | % | 
| 10236 | 186 | \begin{isamarkuptext}%
 | 
| 187 | \noindent | |
| 11257 | 188 | In the proof we have disabled the normally useful lemma | 
| 10878 | 189 | \begin{isabelle}
 | 
| 190 | \isa{take\ n\ xs\ {\isacharat}\ drop\ n\ xs\ {\isacharequal}\ xs}
 | |
| 191 | \rulename{append_take_drop_id}
 | |
| 192 | \end{isabelle}
 | |
| 11257 | 193 | to allow the simplifier to apply the following lemma instead: | 
| 194 | \begin{isabelle}%
 | |
| 195 | \ \ \ \ \ {\isacharbrackleft}x{\isasymin}xs{\isacharat}ys{\isachardot}\ P\ x{\isacharbrackright}\ {\isacharequal}\ {\isacharbrackleft}x{\isasymin}xs{\isachardot}\ P\ x{\isacharbrackright}\ {\isacharat}\ {\isacharbrackleft}x{\isasymin}ys{\isachardot}\ P\ x{\isacharbrackright}%
 | |
| 196 | \end{isabelle}
 | |
| 10236 | 197 | |
| 198 | To dispose of trivial cases automatically, the rules of the inductive | |
| 199 | definition are declared simplification rules:% | |
| 200 | \end{isamarkuptext}%
 | |
| 11866 | 201 | \isamarkuptrue% | 
| 202 | \isacommand{declare}\ S{\isacharunderscore}A{\isacharunderscore}B{\isachardot}intros{\isacharbrackleft}simp{\isacharbrackright}\isamarkupfalse%
 | |
| 203 | % | |
| 10236 | 204 | \begin{isamarkuptext}%
 | 
| 205 | \noindent | |
| 206 | This could have been done earlier but was not necessary so far. | |
| 207 | ||
| 208 | The completeness theorem tells us that if a word has the same number of | |
| 10878 | 209 | \isa{a}'s and \isa{b}'s, then it is in \isa{S}, and similarly 
 | 
| 210 | for \isa{A} and \isa{B}:%
 | |
| 10236 | 211 | \end{isamarkuptext}%
 | 
| 11866 | 212 | \isamarkuptrue% | 
| 10236 | 213 | \isacommand{theorem}\ completeness{\isacharcolon}\isanewline
 | 
| 214 | \ \ {\isachardoublequote}{\isacharparenleft}size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}a{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}b{\isacharbrackright}\ \ \ \ \ {\isasymlongrightarrow}\ w\ {\isasymin}\ S{\isacharparenright}\ {\isasymand}\isanewline
 | |
| 10237 | 215 | \ \ \ {\isacharparenleft}size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}a{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}b{\isacharbrackright}\ {\isacharplus}\ {\isadigit{1}}\ {\isasymlongrightarrow}\ w\ {\isasymin}\ A{\isacharparenright}\ {\isasymand}\isanewline
 | 
| 11866 | 216 | \ \ \ {\isacharparenleft}size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}b{\isacharbrackright}\ {\isacharequal}\ size{\isacharbrackleft}x{\isasymin}w{\isachardot}\ x{\isacharequal}a{\isacharbrackright}\ {\isacharplus}\ {\isadigit{1}}\ {\isasymlongrightarrow}\ w\ {\isasymin}\ B{\isacharparenright}{\isachardoublequote}\isamarkupfalse%
 | 
| 217 | % | |
| 10236 | 218 | \begin{isamarkuptxt}%
 | 
| 219 | \noindent | |
| 220 | The proof is by induction on \isa{w}. Structural induction would fail here
 | |
| 221 | because, as we can see from the grammar, we need to make bigger steps than | |
| 222 | merely appending a single letter at the front. Hence we induct on the length | |
| 223 | of \isa{w}, using the induction rule \isa{length{\isacharunderscore}induct}:%
 | |
| 224 | \end{isamarkuptxt}%
 | |
| 11866 | 225 | \isamarkuptrue% | 
| 13791 | 226 | \isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ w\ rule{\isacharcolon}\ length{\isacharunderscore}induct{\isacharparenright}\isamarkupfalse%
 | 
| 11866 | 227 | \isamarkupfalse% | 
| 228 | % | |
| 10236 | 229 | \begin{isamarkuptxt}%
 | 
| 230 | \noindent | |
| 231 | The \isa{rule} parameter tells \isa{induct{\isacharunderscore}tac} explicitly which induction
 | |
| 232 | rule to use. For details see \S\ref{sec:complete-ind} below.
 | |
| 233 | In this case the result is that we may assume the lemma already | |
| 234 | holds for all words shorter than \isa{w}.
 | |
| 235 | ||
| 236 | The proof continues with a case distinction on \isa{w},
 | |
| 11494 | 237 | on whether \isa{w} is empty or not.%
 | 
| 10236 | 238 | \end{isamarkuptxt}%
 | 
| 11866 | 239 | \isamarkuptrue% | 
| 10236 | 240 | \isacommand{apply}{\isacharparenleft}case{\isacharunderscore}tac\ w{\isacharparenright}\isanewline
 | 
| 11866 | 241 | \ \isamarkupfalse% | 
| 13791 | 242 | \isacommand{apply}{\isacharparenleft}simp{\isacharunderscore}all{\isacharparenright}\isamarkupfalse%
 | 
| 11866 | 243 | \isamarkupfalse% | 
| 244 | % | |
| 10236 | 245 | \begin{isamarkuptxt}%
 | 
| 246 | \noindent | |
| 11257 | 247 | Simplification disposes of the base case and leaves only a conjunction | 
| 248 | of two step cases to be proved: | |
| 10878 | 249 | if \isa{w\ {\isacharequal}\ a\ {\isacharhash}\ v} and \begin{isabelle}%
 | 
| 250 | \ \ \ \ \ length\ {\isacharbrackleft}x{\isasymin}v\ {\isachardot}\ x\ {\isacharequal}\ a{\isacharbrackright}\ {\isacharequal}\ length\ {\isacharbrackleft}x{\isasymin}v\ {\isachardot}\ x\ {\isacharequal}\ b{\isacharbrackright}\ {\isacharplus}\ {\isadigit{2}}%
 | |
| 251 | \end{isabelle} then
 | |
| 10236 | 252 | \isa{b\ {\isacharhash}\ v\ {\isasymin}\ A}, and similarly for \isa{w\ {\isacharequal}\ b\ {\isacharhash}\ v}.
 | 
| 253 | We only consider the first case in detail. | |
| 254 | ||
| 11257 | 255 | After breaking the conjunction up into two cases, we can apply | 
| 10236 | 256 | \isa{part{\isadigit{1}}} to the assumption that \isa{w} contains two more \isa{a}'s than \isa{b}'s.%
 | 
| 257 | \end{isamarkuptxt}%
 | |
| 11866 | 258 | \isamarkuptrue% | 
| 10217 | 259 | \isacommand{apply}{\isacharparenleft}rule\ conjI{\isacharparenright}\isanewline
 | 
| 11866 | 260 | \ \isamarkupfalse% | 
| 261 | \isacommand{apply}{\isacharparenleft}clarify{\isacharparenright}\isanewline
 | |
| 262 | \ \isamarkupfalse% | |
| 263 | \isacommand{apply}{\isacharparenleft}frule\ part{\isadigit{1}}{\isacharbrackleft}of\ {\isachardoublequote}{\isasymlambda}x{\isachardot}\ x{\isacharequal}a{\isachardoublequote}{\isacharcomma}\ simplified{\isacharbrackright}{\isacharparenright}\isanewline
 | |
| 264 | \ \isamarkupfalse% | |
| 265 | \isacommand{apply}{\isacharparenleft}clarify{\isacharparenright}\isamarkupfalse%
 | |
| 266 | % | |
| 10236 | 267 | \begin{isamarkuptxt}%
 | 
| 268 | \noindent | |
| 269 | This yields an index \isa{i\ {\isasymle}\ length\ v} such that
 | |
| 10878 | 270 | \begin{isabelle}%
 | 
| 271 | \ \ \ \ \ length\ {\isacharbrackleft}x{\isasymin}take\ i\ v\ {\isachardot}\ x\ {\isacharequal}\ a{\isacharbrackright}\ {\isacharequal}\ length\ {\isacharbrackleft}x{\isasymin}take\ i\ v\ {\isachardot}\ x\ {\isacharequal}\ b{\isacharbrackright}\ {\isacharplus}\ {\isadigit{1}}%
 | |
| 272 | \end{isabelle}
 | |
| 11147 | 273 | With the help of \isa{part{\isadigit{2}}} it follows that
 | 
| 10878 | 274 | \begin{isabelle}%
 | 
| 275 | \ \ \ \ \ length\ {\isacharbrackleft}x{\isasymin}drop\ i\ v\ {\isachardot}\ x\ {\isacharequal}\ a{\isacharbrackright}\ {\isacharequal}\ length\ {\isacharbrackleft}x{\isasymin}drop\ i\ v\ {\isachardot}\ x\ {\isacharequal}\ b{\isacharbrackright}\ {\isacharplus}\ {\isadigit{1}}%
 | |
| 276 | \end{isabelle}%
 | |
| 10236 | 277 | \end{isamarkuptxt}%
 | 
| 11866 | 278 | \ \isamarkuptrue% | 
| 279 | \isacommand{apply}{\isacharparenleft}drule\ part{\isadigit{2}}{\isacharbrackleft}of\ {\isachardoublequote}{\isasymlambda}x{\isachardot}\ x{\isacharequal}a{\isachardoublequote}{\isacharcomma}\ simplified{\isacharbrackright}{\isacharparenright}\isanewline
 | |
| 280 | \ \ \isamarkupfalse% | |
| 281 | \isacommand{apply}{\isacharparenleft}assumption{\isacharparenright}\isamarkupfalse%
 | |
| 282 | % | |
| 10236 | 283 | \begin{isamarkuptxt}%
 | 
| 284 | \noindent | |
| 285 | Now it is time to decompose \isa{v} in the conclusion \isa{b\ {\isacharhash}\ v\ {\isasymin}\ A}
 | |
| 11257 | 286 | into \isa{take\ i\ v\ {\isacharat}\ drop\ i\ v},%
 | 
| 287 | \end{isamarkuptxt}%
 | |
| 11866 | 288 | \ \isamarkuptrue% | 
| 289 | \isacommand{apply}{\isacharparenleft}rule{\isacharunderscore}tac\ n{\isadigit{1}}{\isacharequal}i\ \isakeyword{and}\ t{\isacharequal}v\ \isakeyword{in}\ subst{\isacharbrackleft}OF\ append{\isacharunderscore}take{\isacharunderscore}drop{\isacharunderscore}id{\isacharbrackright}{\isacharparenright}\isamarkupfalse%
 | |
| 290 | % | |
| 11257 | 291 | \begin{isamarkuptxt}%
 | 
| 292 | \noindent | |
| 293 | (the variables \isa{n{\isadigit{1}}} and \isa{t} are the result of composing the
 | |
| 294 | theorems \isa{subst} and \isa{append{\isacharunderscore}take{\isacharunderscore}drop{\isacharunderscore}id})
 | |
| 10236 | 295 | after which the appropriate rule of the grammar reduces the goal | 
| 296 | to the two subgoals \isa{take\ i\ v\ {\isasymin}\ A} and \isa{drop\ i\ v\ {\isasymin}\ A}:%
 | |
| 297 | \end{isamarkuptxt}%
 | |
| 11866 | 298 | \ \isamarkuptrue% | 
| 299 | \isacommand{apply}{\isacharparenleft}rule\ S{\isacharunderscore}A{\isacharunderscore}B{\isachardot}intros{\isacharparenright}\isamarkupfalse%
 | |
| 300 | % | |
| 10236 | 301 | \begin{isamarkuptxt}%
 | 
| 302 | Both subgoals follow from the induction hypothesis because both \isa{take\ i\ v} and \isa{drop\ i\ v} are shorter than \isa{w}:%
 | |
| 303 | \end{isamarkuptxt}%
 | |
| 11866 | 304 | \ \ \isamarkuptrue% | 
| 305 | \isacommand{apply}{\isacharparenleft}force\ simp\ add{\isacharcolon}\ min{\isacharunderscore}less{\isacharunderscore}iff{\isacharunderscore}disj{\isacharparenright}\isanewline
 | |
| 306 | \ \isamarkupfalse% | |
| 307 | \isacommand{apply}{\isacharparenleft}force\ split\ add{\isacharcolon}\ nat{\isacharunderscore}diff{\isacharunderscore}split{\isacharparenright}\isamarkupfalse%
 | |
| 308 | % | |
| 10236 | 309 | \begin{isamarkuptxt}%
 | 
| 10878 | 310 | The case \isa{w\ {\isacharequal}\ b\ {\isacharhash}\ v} is proved analogously:%
 | 
| 10236 | 311 | \end{isamarkuptxt}%
 | 
| 11866 | 312 | \isamarkuptrue% | 
| 10217 | 313 | \isacommand{apply}{\isacharparenleft}clarify{\isacharparenright}\isanewline
 | 
| 11866 | 314 | \isamarkupfalse% | 
| 10236 | 315 | \isacommand{apply}{\isacharparenleft}frule\ part{\isadigit{1}}{\isacharbrackleft}of\ {\isachardoublequote}{\isasymlambda}x{\isachardot}\ x{\isacharequal}b{\isachardoublequote}{\isacharcomma}\ simplified{\isacharbrackright}{\isacharparenright}\isanewline
 | 
| 11866 | 316 | \isamarkupfalse% | 
| 11257 | 317 | \isacommand{apply}{\isacharparenleft}clarify{\isacharparenright}\isanewline
 | 
| 11866 | 318 | \isamarkupfalse% | 
| 10236 | 319 | \isacommand{apply}{\isacharparenleft}drule\ part{\isadigit{2}}{\isacharbrackleft}of\ {\isachardoublequote}{\isasymlambda}x{\isachardot}\ x{\isacharequal}b{\isachardoublequote}{\isacharcomma}\ simplified{\isacharbrackright}{\isacharparenright}\isanewline
 | 
| 11866 | 320 | \ \isamarkupfalse% | 
| 321 | \isacommand{apply}{\isacharparenleft}assumption{\isacharparenright}\isanewline
 | |
| 322 | \isamarkupfalse% | |
| 10236 | 323 | \isacommand{apply}{\isacharparenleft}rule{\isacharunderscore}tac\ n{\isadigit{1}}{\isacharequal}i\ \isakeyword{and}\ t{\isacharequal}v\ \isakeyword{in}\ subst{\isacharbrackleft}OF\ append{\isacharunderscore}take{\isacharunderscore}drop{\isacharunderscore}id{\isacharbrackright}{\isacharparenright}\isanewline
 | 
| 11866 | 324 | \isamarkupfalse% | 
| 10217 | 325 | \isacommand{apply}{\isacharparenleft}rule\ S{\isacharunderscore}A{\isacharunderscore}B{\isachardot}intros{\isacharparenright}\isanewline
 | 
| 11866 | 326 | \ \isamarkupfalse% | 
| 12815 | 327 | \isacommand{apply}{\isacharparenleft}force\ simp\ add{\isacharcolon}\ min{\isacharunderscore}less{\isacharunderscore}iff{\isacharunderscore}disj{\isacharparenright}\isanewline
 | 
| 11866 | 328 | \isamarkupfalse% | 
| 12815 | 329 | \isacommand{by}{\isacharparenleft}force\ simp\ add{\isacharcolon}\ min{\isacharunderscore}less{\isacharunderscore}iff{\isacharunderscore}disj\ split\ add{\isacharcolon}\ nat{\isacharunderscore}diff{\isacharunderscore}split{\isacharparenright}\isamarkupfalse%
 | 
| 11866 | 330 | % | 
| 10236 | 331 | \begin{isamarkuptext}%
 | 
| 10878 | 332 | We conclude this section with a comparison of our proof with | 
| 11494 | 333 | Hopcroft\index{Hopcroft, J. E.} and Ullman's\index{Ullman, J. D.}
 | 
| 334 | \cite[p.\ts81]{HopcroftUllman}.
 | |
| 335 | For a start, the textbook | |
| 11257 | 336 | grammar, for no good reason, excludes the empty word, thus complicating | 
| 337 | matters just a little bit: they have 8 instead of our 7 productions. | |
| 10236 | 338 | |
| 11158 | 339 | More importantly, the proof itself is different: rather than | 
| 340 | separating the two directions, they perform one induction on the | |
| 341 | length of a word. This deprives them of the beauty of rule induction, | |
| 342 | and in the easy direction (correctness) their reasoning is more | |
| 343 | detailed than our \isa{auto}. For the hard part (completeness), they
 | |
| 344 | consider just one of the cases that our \isa{simp{\isacharunderscore}all} disposes of
 | |
| 345 | automatically. Then they conclude the proof by saying about the | |
| 346 | remaining cases: ``We do this in a manner similar to our method of | |
| 347 | proof for part (1); this part is left to the reader''. But this is | |
| 348 | precisely the part that requires the intermediate value theorem and | |
| 349 | thus is not at all similar to the other cases (which are automatic in | |
| 350 | Isabelle). The authors are at least cavalier about this point and may | |
| 351 | even have overlooked the slight difficulty lurking in the omitted | |
| 11494 | 352 | cases. Such errors are found in many pen-and-paper proofs when they | 
| 353 | are scrutinized formally.% | |
| 354 | \index{grammars!defining inductively|)}%
 | |
| 10236 | 355 | \end{isamarkuptext}%
 | 
| 11866 | 356 | \isamarkuptrue% | 
| 357 | \isamarkupfalse% | |
| 10217 | 358 | \end{isabellebody}%
 | 
| 359 | %%% Local Variables: | |
| 360 | %%% mode: latex | |
| 361 | %%% TeX-master: "root" | |
| 362 | %%% End: |