| author | wenzelm | 
| Tue, 17 Sep 2024 18:49:46 +0200 | |
| changeset 80898 | 71756d95b7df | 
| parent 76987 | 4c275405faae | 
| permissions | -rw-r--r-- | 
| 17914 | 1 | (*<*)theory AB imports Main begin(*>*) | 
| 10225 | 2 | |
| 67406 | 3 | section\<open>Case Study: A Context Free Grammar\<close> | 
| 10217 | 4 | |
| 67406 | 5 | text\<open>\label{sec:CFG}
 | 
| 11494 | 6 | \index{grammars!defining inductively|(}%
 | 
| 10236 | 7 | Grammars are nothing but shorthands for inductive definitions of nonterminals | 
| 8 | which represent sets of strings. For example, the production | |
| 9 | $A \to B c$ is short for | |
| 10 | \[ w \in B \Longrightarrow wc \in A \] | |
| 10884 | 11 | This section demonstrates this idea with an example | 
| 12 | due to Hopcroft and Ullman, a grammar for generating all words with an | |
| 13 | equal number of $a$'s and~$b$'s: | |
| 10236 | 14 | \begin{eqnarray}
 | 
| 15 | S &\to& \epsilon \mid b A \mid a B \nonumber\\ | |
| 16 | A &\to& a S \mid b A A \nonumber\\ | |
| 17 | B &\to& b S \mid a B B \nonumber | |
| 18 | \end{eqnarray}
 | |
| 10884 | 19 | At the end we say a few words about the relationship between | 
| 76987 | 20 | the original proof \<^cite>\<open>\<open>p.\ts81\<close> in HopcroftUllman\<close> and our formal version. | 
| 10236 | 21 | |
| 69597 | 22 | We start by fixing the alphabet, which consists only of \<^term>\<open>a\<close>'s | 
| 23 | and~\<^term>\<open>b\<close>'s: | |
| 67406 | 24 | \<close> | 
| 10236 | 25 | |
| 11705 | 26 | datatype alfa = a | b | 
| 10217 | 27 | |
| 67406 | 28 | text\<open>\noindent | 
| 10287 | 29 | For convenience we include the following easy lemmas as simplification rules: | 
| 67406 | 30 | \<close> | 
| 10236 | 31 | |
| 11705 | 32 | lemma [simp]: "(x \<noteq> a) = (x = b) \<and> (x \<noteq> b) = (x = a)" | 
| 10884 | 33 | by (case_tac x, auto) | 
| 10217 | 34 | |
| 67406 | 35 | text\<open>\noindent | 
| 69597 | 36 | Words over this alphabet are of type \<^typ>\<open>alfa list\<close>, and | 
| 23733 | 37 | the three nonterminals are declared as sets of such words. | 
| 10884 | 38 | The productions above are recast as a \emph{mutual} inductive
 | 
| 10242 | 39 | definition\index{inductive definition!simultaneous}
 | 
| 69597 | 40 | of \<^term>\<open>S\<close>, \<^term>\<open>A\<close> and~\<^term>\<open>B\<close>: | 
| 67406 | 41 | \<close> | 
| 10236 | 42 | |
| 23733 | 43 | inductive_set | 
| 25330 | 44 | S :: "alfa list set" and | 
| 45 | A :: "alfa list set" and | |
| 46 | B :: "alfa list set" | |
| 23733 | 47 | where | 
| 10236 | 48 | "[] \<in> S" | 
| 23733 | 49 | | "w \<in> A \<Longrightarrow> b#w \<in> S" | 
| 50 | | "w \<in> B \<Longrightarrow> a#w \<in> S" | |
| 10217 | 51 | |
| 23733 | 52 | | "w \<in> S \<Longrightarrow> a#w \<in> A" | 
| 53 | | "\<lbrakk> v\<in>A; w\<in>A \<rbrakk> \<Longrightarrow> b#v@w \<in> A" | |
| 10236 | 54 | |
| 23733 | 55 | | "w \<in> S \<Longrightarrow> b#w \<in> B" | 
| 56 | | "\<lbrakk> v \<in> B; w \<in> B \<rbrakk> \<Longrightarrow> a#v@w \<in> B" | |
| 10217 | 57 | |
| 67406 | 58 | text\<open>\noindent | 
| 69597 | 59 | First we show that all words in \<^term>\<open>S\<close> contain the same number of \<^term>\<open>a\<close>'s and \<^term>\<open>b\<close>'s. Since the definition of \<^term>\<open>S\<close> is by mutual | 
| 10884 | 60 | induction, so is the proof: we show at the same time that all words in | 
| 69597 | 61 | \<^term>\<open>A\<close> contain one more \<^term>\<open>a\<close> than \<^term>\<open>b\<close> and all words in \<^term>\<open>B\<close> contain one more \<^term>\<open>b\<close> than \<^term>\<open>a\<close>. | 
| 67406 | 62 | \<close> | 
| 10217 | 63 | |
| 10236 | 64 | lemma correctness: | 
| 68386 | 65 | "(w \<in> S \<longrightarrow> size[x\<leftarrow>w. x=a] = size[x\<leftarrow>w. x=b]) \<and> | 
| 66 | (w \<in> A \<longrightarrow> size[x\<leftarrow>w. x=a] = size[x\<leftarrow>w. x=b] + 1) \<and> | |
| 67 | (w \<in> B \<longrightarrow> size[x\<leftarrow>w. x=b] = size[x\<leftarrow>w. x=a] + 1)" | |
| 10217 | 68 | |
| 67406 | 69 | txt\<open>\noindent | 
| 69597 | 70 | These propositions are expressed with the help of the predefined \<^term>\<open>filter\<close> function on lists, which has the convenient syntax \<open>[x\<leftarrow>xs. P | 
| 71 | x]\<close>, the list of all elements \<^term>\<open>x\<close> in \<^term>\<open>xs\<close> such that \<^prop>\<open>P x\<close> | |
| 69505 | 72 | holds. Remember that on lists \<open>size\<close> and \<open>length\<close> are synonymous. | 
| 10236 | 73 | |
| 74 | The proof itself is by rule induction and afterwards automatic: | |
| 67406 | 75 | \<close> | 
| 10236 | 76 | |
| 10884 | 77 | by (rule S_A_B.induct, auto) | 
| 10217 | 78 | |
| 67406 | 79 | text\<open>\noindent | 
| 10236 | 80 | This may seem surprising at first, and is indeed an indication of the power | 
| 81 | of inductive definitions. But it is also quite straightforward. For example, | |
| 82 | consider the production $A \to b A A$: if $v,w \in A$ and the elements of $A$ | |
| 10884 | 83 | contain one more $a$ than~$b$'s, then $bvw$ must again contain one more $a$ | 
| 84 | than~$b$'s. | |
| 10236 | 85 | |
| 86 | As usual, the correctness of syntactic descriptions is easy, but completeness | |
| 69597 | 87 | is hard: does \<^term>\<open>S\<close> contain \emph{all} words with an equal number of
 | 
| 88 | \<^term>\<open>a\<close>'s and \<^term>\<open>b\<close>'s? It turns out that this proof requires the | |
| 89 | following lemma: every string with two more \<^term>\<open>a\<close>'s than \<^term>\<open>b\<close>'s can be cut somewhere such that each half has one more \<^term>\<open>a\<close> than | |
| 90 | \<^term>\<open>b\<close>. This is best seen by imagining counting the difference between the | |
| 91 | number of \<^term>\<open>a\<close>'s and \<^term>\<open>b\<close>'s starting at the left end of the | |
| 10283 | 92 | word. We start with 0 and end (at the right end) with 2. Since each move to the | 
| 10236 | 93 | right increases or decreases the difference by 1, we must have passed through | 
| 94 | 1 on our way from 0 to 2. Formally, we appeal to the following discrete | |
| 95 | intermediate value theorem @{thm[source]nat0_intermed_int_val}
 | |
| 16412 | 96 | @{thm[display,margin=60]nat0_intermed_int_val[no_vars]}
 | 
| 69597 | 97 | where \<^term>\<open>f\<close> is of type \<^typ>\<open>nat \<Rightarrow> int\<close>, \<^typ>\<open>int\<close> are the integers, | 
| 69505 | 98 | \<open>\<bar>.\<bar>\<close> is the absolute value function\footnote{See
 | 
| 11308 | 99 | Table~\ref{tab:ascii} in the Appendix for the correct \textsc{ascii}
 | 
| 69597 | 100 | syntax.}, and \<^term>\<open>1::int\<close> is the integer 1 (see \S\ref{sec:numbers}).
 | 
| 10217 | 101 | |
| 11147 | 102 | First we show that our specific function, the difference between the | 
| 69597 | 103 | numbers of \<^term>\<open>a\<close>'s and \<^term>\<open>b\<close>'s, does indeed only change by 1 in every | 
| 104 | move to the right. At this point we also start generalizing from \<^term>\<open>a\<close>'s | |
| 105 | and \<^term>\<open>b\<close>'s to an arbitrary property \<^term>\<open>P\<close>. Otherwise we would have | |
| 10236 | 106 | to prove the desired lemma twice, once as stated above and once with the | 
| 69597 | 107 | roles of \<^term>\<open>a\<close>'s and \<^term>\<open>b\<close>'s interchanged. | 
| 67406 | 108 | \<close> | 
| 10236 | 109 | |
| 110 | lemma step1: "\<forall>i < size w. | |
| 68386 | 111 | \<bar>(int(size[x\<leftarrow>take (i+1) w. P x])-int(size[x\<leftarrow>take (i+1) w. \<not>P x])) | 
| 112 | - (int(size[x\<leftarrow>take i w. P x])-int(size[x\<leftarrow>take i w. \<not>P x]))\<bar> \<le> 1" | |
| 10236 | 113 | |
| 67406 | 114 | txt\<open>\noindent | 
| 10236 | 115 | The lemma is a bit hard to read because of the coercion function | 
| 69597 | 116 | \<open>int :: nat \<Rightarrow> int\<close>. It is required because \<^term>\<open>size\<close> returns | 
| 117 | a natural number, but subtraction on type~\<^typ>\<open>nat\<close> will do the wrong thing. | |
| 118 | Function \<^term>\<open>take\<close> is predefined and \<^term>\<open>take i xs\<close> is the prefix of | |
| 119 | length \<^term>\<open>i\<close> of \<^term>\<open>xs\<close>; below we also need \<^term>\<open>drop i xs\<close>, which | |
| 120 | is what remains after that prefix has been dropped from \<^term>\<open>xs\<close>. | |
| 10236 | 121 | |
| 69597 | 122 | The proof is by induction on \<^term>\<open>w\<close>, with a trivial base case, and a not | 
| 10236 | 123 | so trivial induction step. Since it is essentially just arithmetic, we do not | 
| 124 | discuss it. | |
| 67406 | 125 | \<close> | 
| 10236 | 126 | |
| 12332 | 127 | apply(induct_tac w) | 
| 16585 | 128 | apply(auto simp add: abs_if take_Cons split: nat.split) | 
| 129 | done | |
| 10217 | 130 | |
| 67406 | 131 | text\<open> | 
| 11494 | 132 | 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: | 
| 67406 | 133 | \<close> | 
| 10217 | 134 | |
| 135 | lemma part1: | |
| 68386 | 136 | "size[x\<leftarrow>w. P x] = size[x\<leftarrow>w. \<not>P x]+2 \<Longrightarrow> | 
| 137 | \<exists>i\<le>size w. size[x\<leftarrow>take i w. P x] = size[x\<leftarrow>take i w. \<not>P x]+1" | |
| 10236 | 138 | |
| 67406 | 139 | txt\<open>\noindent | 
| 69505 | 140 | This is proved by \<open>force\<close> with the help of the intermediate value theorem, | 
| 10608 | 141 | instantiated appropriately and with its first premise disposed of by lemma | 
| 142 | @{thm[source]step1}:
 | |
| 67406 | 143 | \<close> | 
| 10236 | 144 | |
| 11870 
181bd2050cf4
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11705diff
changeset | 145 | apply(insert nat0_intermed_int_val[OF step1, of "P" "w" "1"]) | 
| 11705 | 146 | by force | 
| 10236 | 147 | |
| 67406 | 148 | text\<open>\noindent | 
| 10236 | 149 | |
| 69597 | 150 | Lemma @{thm[source]part1} tells us only about the prefix \<^term>\<open>take i w\<close>.
 | 
| 151 | An easy lemma deals with the suffix \<^term>\<open>drop i w\<close>: | |
| 67406 | 152 | \<close> | 
| 10236 | 153 | |
| 10217 | 154 | |
| 155 | lemma part2: | |
| 68386 | 156 | "\<lbrakk>size[x\<leftarrow>take i w @ drop i w. P x] = | 
| 157 | size[x\<leftarrow>take i w @ drop i w. \<not>P x]+2; | |
| 158 | size[x\<leftarrow>take i w. P x] = size[x\<leftarrow>take i w. \<not>P x]+1\<rbrakk> | |
| 159 | \<Longrightarrow> size[x\<leftarrow>drop i w. P x] = size[x\<leftarrow>drop i w. \<not>P x]+1" | |
| 12815 | 160 | by(simp del: append_take_drop_id) | 
| 10217 | 161 | |
| 67406 | 162 | text\<open>\noindent | 
| 11257 | 163 | In the proof we have disabled the normally useful lemma | 
| 10884 | 164 | \begin{isabelle}
 | 
| 165 | @{thm append_take_drop_id[no_vars]}
 | |
| 166 | \rulename{append_take_drop_id}
 | |
| 167 | \end{isabelle}
 | |
| 11257 | 168 | to allow the simplifier to apply the following lemma instead: | 
| 169 | @{text[display]"[x\<in>xs@ys. P x] = [x\<in>xs. P x] @ [x\<in>ys. P x]"}
 | |
| 10236 | 170 | |
| 171 | To dispose of trivial cases automatically, the rules of the inductive | |
| 172 | definition are declared simplification rules: | |
| 67406 | 173 | \<close> | 
| 10236 | 174 | |
| 11705 | 175 | declare S_A_B.intros[simp] | 
| 10236 | 176 | |
| 67406 | 177 | text\<open>\noindent | 
| 10236 | 178 | This could have been done earlier but was not necessary so far. | 
| 179 | ||
| 180 | The completeness theorem tells us that if a word has the same number of | |
| 69597 | 181 | \<^term>\<open>a\<close>'s and \<^term>\<open>b\<close>'s, then it is in \<^term>\<open>S\<close>, and similarly | 
| 182 | for \<^term>\<open>A\<close> and \<^term>\<open>B\<close>: | |
| 67406 | 183 | \<close> | 
| 10236 | 184 | |
| 185 | theorem completeness: | |
| 68386 | 186 | "(size[x\<leftarrow>w. x=a] = size[x\<leftarrow>w. x=b] \<longrightarrow> w \<in> S) \<and> | 
| 187 | (size[x\<leftarrow>w. x=a] = size[x\<leftarrow>w. x=b] + 1 \<longrightarrow> w \<in> A) \<and> | |
| 188 | (size[x\<leftarrow>w. x=b] = size[x\<leftarrow>w. x=a] + 1 \<longrightarrow> w \<in> B)" | |
| 10236 | 189 | |
| 67406 | 190 | txt\<open>\noindent | 
| 69597 | 191 | The proof is by induction on \<^term>\<open>w\<close>. Structural induction would fail here | 
| 10236 | 192 | because, as we can see from the grammar, we need to make bigger steps than | 
| 193 | merely appending a single letter at the front. Hence we induct on the length | |
| 69597 | 194 | of \<^term>\<open>w\<close>, using the induction rule @{thm[source]length_induct}:
 | 
| 67406 | 195 | \<close> | 
| 10217 | 196 | |
| 11705 | 197 | apply(induct_tac w rule: length_induct) | 
| 27167 | 198 | apply(rename_tac w) | 
| 10236 | 199 | |
| 67406 | 200 | txt\<open>\noindent | 
| 69505 | 201 | The \<open>rule\<close> parameter tells \<open>induct_tac\<close> explicitly which induction | 
| 10236 | 202 | rule to use. For details see \S\ref{sec:complete-ind} below.
 | 
| 203 | In this case the result is that we may assume the lemma already | |
| 69597 | 204 | holds for all words shorter than \<^term>\<open>w\<close>. Because the induction step renames | 
| 69505 | 205 | the induction variable we rename it back to \<open>w\<close>. | 
| 10236 | 206 | |
| 69597 | 207 | The proof continues with a case distinction on \<^term>\<open>w\<close>, | 
| 208 | on whether \<^term>\<open>w\<close> is empty or not. | |
| 67406 | 209 | \<close> | 
| 10236 | 210 | |
| 11705 | 211 | apply(case_tac w) | 
| 212 | apply(simp_all) | |
| 10236 | 213 | (*<*)apply(rename_tac x v)(*>*) | 
| 214 | ||
| 67406 | 215 | txt\<open>\noindent | 
| 11257 | 216 | Simplification disposes of the base case and leaves only a conjunction | 
| 217 | of two step cases to be proved: | |
| 69597 | 218 | if \<^prop>\<open>w = a#v\<close> and @{prop[display]"size[x\<in>v. x=a] = size[x\<in>v. x=b]+2"} then
 | 
| 219 | \<^prop>\<open>b#v \<in> A\<close>, and similarly for \<^prop>\<open>w = b#v\<close>. | |
| 10236 | 220 | We only consider the first case in detail. | 
| 221 | ||
| 11257 | 222 | After breaking the conjunction up into two cases, we can apply | 
| 69597 | 223 | @{thm[source]part1} to the assumption that \<^term>\<open>w\<close> contains two more \<^term>\<open>a\<close>'s than \<^term>\<open>b\<close>'s.
 | 
| 67406 | 224 | \<close> | 
| 10236 | 225 | |
| 11257 | 226 | apply(rule conjI) | 
| 227 | apply(clarify) | |
| 228 | apply(frule part1[of "\<lambda>x. x=a", simplified]) | |
| 229 | apply(clarify) | |
| 67406 | 230 | txt\<open>\noindent | 
| 69597 | 231 | This yields an index \<^prop>\<open>i \<le> length v\<close> such that | 
| 68386 | 232 | @{prop[display]"length [x\<leftarrow>take i v . x = a] = length [x\<leftarrow>take i v . x = b] + 1"}
 | 
| 11147 | 233 | With the help of @{thm[source]part2} it follows that
 | 
| 68386 | 234 | @{prop[display]"length [x\<leftarrow>drop i v . x = a] = length [x\<leftarrow>drop i v . x = b] + 1"}
 | 
| 67406 | 235 | \<close> | 
| 10236 | 236 | |
| 11257 | 237 | apply(drule part2[of "\<lambda>x. x=a", simplified]) | 
| 238 | apply(assumption) | |
| 10236 | 239 | |
| 67406 | 240 | txt\<open>\noindent | 
| 69597 | 241 | Now it is time to decompose \<^term>\<open>v\<close> in the conclusion \<^prop>\<open>b#v \<in> A\<close> | 
| 242 | into \<^term>\<open>take i v @ drop i v\<close>, | |
| 67406 | 243 | \<close> | 
| 11257 | 244 | |
| 245 | apply(rule_tac n1=i and t=v in subst[OF append_take_drop_id]) | |
| 246 | ||
| 67406 | 247 | txt\<open>\noindent | 
| 69597 | 248 | (the variables \<^term>\<open>n1\<close> and \<^term>\<open>t\<close> are the result of composing the | 
| 11257 | 249 | theorems @{thm[source]subst} and @{thm[source]append_take_drop_id})
 | 
| 10236 | 250 | after which the appropriate rule of the grammar reduces the goal | 
| 69597 | 251 | to the two subgoals \<^prop>\<open>take i v \<in> A\<close> and \<^prop>\<open>drop i v \<in> A\<close>: | 
| 67406 | 252 | \<close> | 
| 10236 | 253 | |
| 11257 | 254 | apply(rule S_A_B.intros) | 
| 10236 | 255 | |
| 67406 | 256 | txt\<open> | 
| 69597 | 257 | Both subgoals follow from the induction hypothesis because both \<^term>\<open>take i | 
| 258 | v\<close> and \<^term>\<open>drop i v\<close> are shorter than \<^term>\<open>w\<close>: | |
| 67406 | 259 | \<close> | 
| 10236 | 260 | |
| 11257 | 261 | apply(force simp add: min_less_iff_disj) | 
| 63648 | 262 | apply(force split: nat_diff_split) | 
| 10236 | 263 | |
| 67406 | 264 | txt\<open> | 
| 69597 | 265 | The case \<^prop>\<open>w = b#v\<close> is proved analogously: | 
| 67406 | 266 | \<close> | 
| 10236 | 267 | |
| 11257 | 268 | apply(clarify) | 
| 269 | apply(frule part1[of "\<lambda>x. x=b", simplified]) | |
| 270 | apply(clarify) | |
| 271 | apply(drule part2[of "\<lambda>x. x=b", simplified]) | |
| 272 | apply(assumption) | |
| 273 | apply(rule_tac n1=i and t=v in subst[OF append_take_drop_id]) | |
| 274 | apply(rule S_A_B.intros) | |
| 12815 | 275 | apply(force simp add: min_less_iff_disj) | 
| 63648 | 276 | by(force simp add: min_less_iff_disj split: nat_diff_split) | 
| 10217 | 277 | |
| 67406 | 278 | text\<open> | 
| 10884 | 279 | We conclude this section with a comparison of our proof with | 
| 11494 | 280 | Hopcroft\index{Hopcroft, J. E.} and Ullman's\index{Ullman, J. D.}
 | 
| 76987 | 281 | \<^cite>\<open>\<open>p.\ts81\<close> in HopcroftUllman\<close>. | 
| 11494 | 282 | For a start, the textbook | 
| 11257 | 283 | grammar, for no good reason, excludes the empty word, thus complicating | 
| 284 | matters just a little bit: they have 8 instead of our 7 productions. | |
| 10236 | 285 | |
| 11147 | 286 | More importantly, the proof itself is different: rather than | 
| 287 | separating the two directions, they perform one induction on the | |
| 288 | length of a word. This deprives them of the beauty of rule induction, | |
| 289 | and in the easy direction (correctness) their reasoning is more | |
| 69505 | 290 | detailed than our \<open>auto\<close>. For the hard part (completeness), they | 
| 291 | consider just one of the cases that our \<open>simp_all\<close> disposes of | |
| 11147 | 292 | automatically. Then they conclude the proof by saying about the | 
| 293 | remaining cases: ``We do this in a manner similar to our method of | |
| 294 | proof for part (1); this part is left to the reader''. But this is | |
| 295 | precisely the part that requires the intermediate value theorem and | |
| 296 | thus is not at all similar to the other cases (which are automatic in | |
| 297 | Isabelle). The authors are at least cavalier about this point and may | |
| 298 | even have overlooked the slight difficulty lurking in the omitted | |
| 11494 | 299 | cases. Such errors are found in many pen-and-paper proofs when they | 
| 300 | are scrutinized formally.% | |
| 301 | \index{grammars!defining inductively|)}
 | |
| 67406 | 302 | \<close> | 
| 10236 | 303 | |
| 10225 | 304 | (*<*)end(*>*) |