| author | wenzelm | 
| Mon, 06 Apr 2015 23:14:05 +0200 | |
| changeset 59940 | 087d81f5213e | 
| parent 58882 | 6e2010ab8bd9 | 
| child 61541 | 846c72206207 | 
| permissions | -rw-r--r-- | 
| 33026 | 1 | (* Title: HOL/Isar_Examples/Basic_Logic.thy | 
| 6444 
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
 wenzelm parents: diff
changeset | 2 | Author: Markus Wenzel, TU Muenchen | 
| 
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
 wenzelm parents: diff
changeset | 3 | |
| 
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
 wenzelm parents: diff
changeset | 4 | Basic propositional and quantifier reasoning. | 
| 
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
 wenzelm parents: diff
changeset | 5 | *) | 
| 
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
 wenzelm parents: diff
changeset | 6 | |
| 58882 | 7 | section \<open>Basic logical reasoning\<close> | 
| 7748 | 8 | |
| 31758 | 9 | theory Basic_Logic | 
| 10 | imports Main | |
| 11 | begin | |
| 6444 
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
 wenzelm parents: diff
changeset | 12 | |
| 7761 | 13 | |
| 58614 | 14 | subsection \<open>Pure backward reasoning\<close> | 
| 7740 | 15 | |
| 58614 | 16 | text \<open>In order to get a first idea of how Isabelle/Isar proof | 
| 37671 | 17 |   documents may look like, we consider the propositions @{text I},
 | 
| 18 |   @{text K}, and @{text S}.  The following (rather explicit) proofs
 | |
| 58614 | 19 | should require little extra explanations.\<close> | 
| 7001 | 20 | |
| 55640 | 21 | lemma I: "A \<longrightarrow> A" | 
| 10007 | 22 | proof | 
| 23 | assume A | |
| 23393 | 24 | show A by fact | 
| 10007 | 25 | qed | 
| 6444 
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
 wenzelm parents: diff
changeset | 26 | |
| 55640 | 27 | lemma K: "A \<longrightarrow> B \<longrightarrow> A" | 
| 10007 | 28 | proof | 
| 29 | assume A | |
| 55640 | 30 | show "B \<longrightarrow> A" | 
| 10007 | 31 | proof | 
| 23393 | 32 | show A by fact | 
| 10007 | 33 | qed | 
| 34 | qed | |
| 6444 
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
 wenzelm parents: diff
changeset | 35 | |
| 55640 | 36 | lemma S: "(A \<longrightarrow> B \<longrightarrow> C) \<longrightarrow> (A \<longrightarrow> B) \<longrightarrow> A \<longrightarrow> C" | 
| 10007 | 37 | proof | 
| 55640 | 38 | assume "A \<longrightarrow> B \<longrightarrow> C" | 
| 39 | show "(A \<longrightarrow> B) \<longrightarrow> A \<longrightarrow> C" | |
| 10007 | 40 | proof | 
| 55640 | 41 | assume "A \<longrightarrow> B" | 
| 42 | show "A \<longrightarrow> C" | |
| 10007 | 43 | proof | 
| 44 | assume A | |
| 45 | show C | |
| 46 | proof (rule mp) | |
| 55640 | 47 | show "B \<longrightarrow> C" by (rule mp) fact+ | 
| 23393 | 48 | show B by (rule mp) fact+ | 
| 10007 | 49 | qed | 
| 50 | qed | |
| 51 | qed | |
| 52 | qed | |
| 6444 
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
 wenzelm parents: diff
changeset | 53 | |
| 58614 | 54 | text \<open>Isar provides several ways to fine-tune the reasoning, | 
| 37671 | 55 | avoiding excessive detail. Several abbreviated language elements | 
| 56 | are available, enabling the writer to express proofs in a more | |
| 57 | concise way, even without referring to any automated proof tools | |
| 58 | yet. | |
| 7761 | 59 | |
| 18193 | 60 | First of all, proof by assumption may be abbreviated as a single | 
| 58614 | 61 | dot.\<close> | 
| 7820 | 62 | |
| 55640 | 63 | lemma "A \<longrightarrow> A" | 
| 10007 | 64 | proof | 
| 65 | assume A | |
| 23393 | 66 | show A by fact+ | 
| 10007 | 67 | qed | 
| 7820 | 68 | |
| 58614 | 69 | text \<open>In fact, concluding any (sub-)proof already involves solving | 
| 37671 | 70 |   any remaining goals by assumption\footnote{This is not a completely
 | 
| 18193 | 71 | trivial operation, as proof by assumption may involve full | 
| 72 | higher-order unification.}. Thus we may skip the rather vacuous | |
| 58614 | 73 | body of the above proof as well.\<close> | 
| 7820 | 74 | |
| 55640 | 75 | lemma "A \<longrightarrow> A" | 
| 10007 | 76 | proof | 
| 77 | qed | |
| 7820 | 78 | |
| 58614 | 79 | text \<open>Note that the \isacommand{proof} command refers to the @{text
 | 
| 37671 | 80 | rule} method (without arguments) by default. Thus it implicitly | 
| 81 | applies a single rule, as determined from the syntactic form of the | |
| 82 |   statements involved.  The \isacommand{by} command abbreviates any
 | |
| 58614 | 83 | proof with empty body, so the proof may be further pruned.\<close> | 
| 7820 | 84 | |
| 55640 | 85 | lemma "A \<longrightarrow> A" | 
| 10007 | 86 | by rule | 
| 7820 | 87 | |
| 58614 | 88 | text \<open>Proof by a single rule may be abbreviated as double-dot.\<close> | 
| 7820 | 89 | |
| 55640 | 90 | lemma "A \<longrightarrow> A" .. | 
| 7820 | 91 | |
| 58614 | 92 | text \<open>Thus we have arrived at an adequate representation of the | 
| 37671 | 93 | proof of a tautology that holds by a single standard | 
| 94 |   rule.\footnote{Apparently, the rule here is implication
 | |
| 58614 | 95 | introduction.}\<close> | 
| 7820 | 96 | |
| 58614 | 97 | text \<open>Let us also reconsider @{text K}.  Its statement is composed
 | 
| 37671 | 98 | of iterated connectives. Basic decomposition is by a single rule at | 
| 99 | a time, which is why our first version above was by nesting two | |
| 18193 | 100 | proofs. | 
| 7820 | 101 | |
| 18193 | 102 |   The @{text intro} proof method repeatedly decomposes a goal's
 | 
| 103 |   conclusion.\footnote{The dual method is @{text elim}, acting on a
 | |
| 58614 | 104 | goal's premises.}\<close> | 
| 7820 | 105 | |
| 55640 | 106 | lemma "A \<longrightarrow> B \<longrightarrow> A" | 
| 12387 | 107 | proof (intro impI) | 
| 10007 | 108 | assume A | 
| 23393 | 109 | show A by fact | 
| 10007 | 110 | qed | 
| 7820 | 111 | |
| 58614 | 112 | text \<open>Again, the body may be collapsed.\<close> | 
| 7820 | 113 | |
| 55640 | 114 | lemma "A \<longrightarrow> B \<longrightarrow> A" | 
| 12387 | 115 | by (intro impI) | 
| 7820 | 116 | |
| 58614 | 117 | text \<open>Just like @{text rule}, the @{text intro} and @{text elim}
 | 
| 37671 | 118 | proof methods pick standard structural rules, in case no explicit | 
| 18193 | 119 | arguments are given. While implicit rules are usually just fine for | 
| 120 | single rule application, this may go too far with iteration. Thus | |
| 121 |   in practice, @{text intro} and @{text elim} would be typically
 | |
| 122 | restricted to certain structures by giving a few rules only, e.g.\ | |
| 123 |   \isacommand{proof}~@{text "(intro impI allI)"} to strip implications
 | |
| 124 | and universal quantifiers. | |
| 7820 | 125 | |
| 18193 | 126 | Such well-tuned iterated decomposition of certain structures is the | 
| 127 |   prime application of @{text intro} and @{text elim}.  In contrast,
 | |
| 128 | terminal steps that solve a goal completely are usually performed by | |
| 129 |   actual automated proof methods (such as \isacommand{by}~@{text
 | |
| 58614 | 130 | blast}.\<close> | 
| 7820 | 131 | |
| 132 | ||
| 58614 | 133 | subsection \<open>Variations of backward vs.\ forward reasoning\<close> | 
| 7820 | 134 | |
| 58614 | 135 | text \<open>Certainly, any proof may be performed in backward-style only. | 
| 37671 | 136 | On the other hand, small steps of reasoning are often more naturally | 
| 18193 | 137 | expressed in forward-style. Isar supports both backward and forward | 
| 138 | reasoning as a first-class concept. In order to demonstrate the | |
| 139 |   difference, we consider several proofs of @{text "A \<and> B \<longrightarrow> B \<and> A"}.
 | |
| 7820 | 140 | |
| 58614 | 141 | The first version is purely backward.\<close> | 
| 7001 | 142 | |
| 55640 | 143 | lemma "A \<and> B \<longrightarrow> B \<and> A" | 
| 10007 | 144 | proof | 
| 55640 | 145 | assume "A \<and> B" | 
| 146 | show "B \<and> A" | |
| 10007 | 147 | proof | 
| 23393 | 148 | show B by (rule conjunct2) fact | 
| 149 | show A by (rule conjunct1) fact | |
| 10007 | 150 | qed | 
| 151 | qed | |
| 6444 
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
 wenzelm parents: diff
changeset | 152 | |
| 58614 | 153 | text \<open>Above, the @{text "conjunct_1/2"} projection rules had to be
 | 
| 37671 | 154 |   named explicitly, since the goals @{text B} and @{text A} did not
 | 
| 155 | provide any structural clue. This may be avoided using | |
| 156 |   \isacommand{from} to focus on the @{text "A \<and> B"} assumption as the
 | |
| 157 | current facts, enabling the use of double-dot proofs. Note that | |
| 158 |   \isacommand{from} already does forward-chaining, involving the
 | |
| 58614 | 159 |   @{text conjE} rule here.\<close>
 | 
| 6444 
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
 wenzelm parents: diff
changeset | 160 | |
| 55640 | 161 | lemma "A \<and> B \<longrightarrow> B \<and> A" | 
| 10007 | 162 | proof | 
| 55640 | 163 | assume "A \<and> B" | 
| 164 | show "B \<and> A" | |
| 10007 | 165 | proof | 
| 58614 | 166 | from \<open>A \<and> B\<close> show B .. | 
| 167 | from \<open>A \<and> B\<close> show A .. | |
| 10007 | 168 | qed | 
| 169 | qed | |
| 7604 | 170 | |
| 58614 | 171 | text \<open>In the next version, we move the forward step one level | 
| 37671 | 172 | upwards. Forward-chaining from the most recent facts is indicated | 
| 173 |   by the \isacommand{then} command.  Thus the proof of @{text "B \<and> A"}
 | |
| 174 |   from @{text "A \<and> B"} actually becomes an elimination, rather than an
 | |
| 18193 | 175 | introduction. The resulting proof structure directly corresponds to | 
| 176 |   that of the @{text conjE} rule, including the repeated goal
 | |
| 58614 | 177 |   proposition that is abbreviated as @{text ?thesis} below.\<close>
 | 
| 7820 | 178 | |
| 55640 | 179 | lemma "A \<and> B \<longrightarrow> B \<and> A" | 
| 10007 | 180 | proof | 
| 55640 | 181 | assume "A \<and> B" | 
| 182 | then show "B \<and> A" | |
| 58614 | 183 |   proof                    -- \<open>rule @{text conjE} of @{text "A \<and> B"}\<close>
 | 
| 23373 | 184 | assume B A | 
| 58614 | 185 |     then show ?thesis ..   -- \<open>rule @{text conjI} of @{text "B \<and> A"}\<close>
 | 
| 10007 | 186 | qed | 
| 187 | qed | |
| 7820 | 188 | |
| 58614 | 189 | text \<open>In the subsequent version we flatten the structure of the main | 
| 37671 | 190 | body by doing forward reasoning all the time. Only the outermost | 
| 58614 | 191 | decomposition step is left as backward.\<close> | 
| 7820 | 192 | |
| 55640 | 193 | lemma "A \<and> B \<longrightarrow> B \<and> A" | 
| 10007 | 194 | proof | 
| 55640 | 195 | assume "A \<and> B" | 
| 58614 | 196 | from \<open>A \<and> B\<close> have A .. | 
| 197 | from \<open>A \<and> B\<close> have B .. | |
| 198 | from \<open>B\<close> \<open>A\<close> show "B \<and> A" .. | |
| 10007 | 199 | qed | 
| 6444 
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
 wenzelm parents: diff
changeset | 200 | |
| 58614 | 201 | text \<open>We can still push forward-reasoning a bit further, even at the | 
| 37671 | 202 | risk of getting ridiculous. Note that we force the initial proof | 
| 58614 | 203 | step to do nothing here, by referring to the ``-'' proof method.\<close> | 
| 7820 | 204 | |
| 55640 | 205 | lemma "A \<and> B \<longrightarrow> B \<and> A" | 
| 10007 | 206 | proof - | 
| 207 |   {
 | |
| 55640 | 208 | assume "A \<and> B" | 
| 58614 | 209 | from \<open>A \<and> B\<close> have A .. | 
| 210 | from \<open>A \<and> B\<close> have B .. | |
| 211 | from \<open>B\<close> \<open>A\<close> have "B \<and> A" .. | |
| 10007 | 212 | } | 
| 58614 | 213 |   then show ?thesis ..         -- \<open>rule @{text impI}\<close>
 | 
| 10007 | 214 | qed | 
| 7820 | 215 | |
| 58614 | 216 | text \<open>\medskip With these examples we have shifted through a whole | 
| 37671 | 217 | range from purely backward to purely forward reasoning. Apparently, | 
| 218 | in the extreme ends we get slightly ill-structured proofs, which | |
| 219 | also require much explicit naming of either rules (backward) or | |
| 220 | local facts (forward). | |
| 7820 | 221 | |
| 18193 | 222 | The general lesson learned here is that good proof style would | 
| 223 |   achieve just the \emph{right} balance of top-down backward
 | |
| 224 | decomposition, and bottom-up forward composition. In general, there | |
| 225 | is no single best way to arrange some pieces of formal reasoning, of | |
| 226 | course. Depending on the actual applications, the intended audience | |
| 227 | etc., rules (and methods) on the one hand vs.\ facts on the other | |
| 228 | hand have to be emphasized in an appropriate way. This requires the | |
| 58614 | 229 | proof writer to develop good taste, and some practice, of course.\<close> | 
| 7820 | 230 | |
| 58614 | 231 | text \<open>For our example the most appropriate way of reasoning is | 
| 37671 | 232 | probably the middle one, with conjunction introduction done after | 
| 58614 | 233 | elimination.\<close> | 
| 7820 | 234 | |
| 55640 | 235 | lemma "A \<and> B \<longrightarrow> B \<and> A" | 
| 10007 | 236 | proof | 
| 55640 | 237 | assume "A \<and> B" | 
| 238 | then show "B \<and> A" | |
| 10007 | 239 | proof | 
| 23373 | 240 | assume B A | 
| 241 | then show ?thesis .. | |
| 10007 | 242 | qed | 
| 243 | qed | |
| 7820 | 244 | |
| 245 | ||
| 6444 
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
 wenzelm parents: diff
changeset | 246 | |
| 58614 | 247 | subsection \<open>A few examples from ``Introduction to Isabelle''\<close> | 
| 7001 | 248 | |
| 58614 | 249 | text \<open>We rephrase some of the basic reasoning examples of | 
| 250 |   @{cite "isabelle-intro"}, using HOL rather than FOL.\<close>
 | |
| 37671 | 251 | |
| 7820 | 252 | |
| 58614 | 253 | subsubsection \<open>A propositional proof\<close> | 
| 7833 | 254 | |
| 58614 | 255 | text \<open>We consider the proposition @{text "P \<or> P \<longrightarrow> P"}.  The proof
 | 
| 37671 | 256 |   below involves forward-chaining from @{text "P \<or> P"}, followed by an
 | 
| 58614 | 257 |   explicit case-analysis on the two \emph{identical} cases.\<close>
 | 
| 6444 
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
 wenzelm parents: diff
changeset | 258 | |
| 55640 | 259 | lemma "P \<or> P \<longrightarrow> P" | 
| 10007 | 260 | proof | 
| 55640 | 261 | assume "P \<or> P" | 
| 23373 | 262 | then show P | 
| 58614 | 263 | proof -- \<open> | 
| 18193 | 264 |     rule @{text disjE}: \smash{$\infer{C}{A \disj B & \infer*{C}{[A]} & \infer*{C}{[B]}}$}
 | 
| 58614 | 265 | \<close> | 
| 23393 | 266 | assume P show P by fact | 
| 10007 | 267 | next | 
| 23393 | 268 | assume P show P by fact | 
| 10007 | 269 | qed | 
| 270 | qed | |
| 7833 | 271 | |
| 58614 | 272 | text \<open>Case splits are \emph{not} hardwired into the Isar language as
 | 
| 37671 | 273 |   a special feature.  The \isacommand{next} command used to separate
 | 
| 274 | the cases above is just a short form of managing block structure. | |
| 7833 | 275 | |
| 18193 | 276 | \medskip In general, applying proof methods may split up a goal into | 
| 277 | separate ``cases'', i.e.\ new subgoals with individual local | |
| 278 | assumptions. The corresponding proof text typically mimics this by | |
| 279 | establishing results in appropriate contexts, separated by blocks. | |
| 7833 | 280 | |
| 18193 | 281 | In order to avoid too much explicit parentheses, the Isar system | 
| 282 | implicitly opens an additional block for any new goal, the | |
| 283 |   \isacommand{next} statement then closes one block level, opening a
 | |
| 284 | new one. The resulting behavior is what one would expect from | |
| 285 | separating cases, only that it is more flexible. E.g.\ an induction | |
| 286 | base case (which does not introduce local assumptions) would | |
| 287 |   \emph{not} require \isacommand{next} to separate the subsequent step
 | |
| 288 | case. | |
| 7833 | 289 | |
| 18193 | 290 | \medskip In our example the situation is even simpler, since the two | 
| 291 | cases actually coincide. Consequently the proof may be rephrased as | |
| 58614 | 292 | follows.\<close> | 
| 7833 | 293 | |
| 55656 | 294 | lemma "P \<or> P \<longrightarrow> P" | 
| 10007 | 295 | proof | 
| 55640 | 296 | assume "P \<or> P" | 
| 23373 | 297 | then show P | 
| 10007 | 298 | proof | 
| 299 | assume P | |
| 23393 | 300 | show P by fact | 
| 301 | show P by fact | |
| 10007 | 302 | qed | 
| 303 | qed | |
| 6444 
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
 wenzelm parents: diff
changeset | 304 | |
| 58614 | 305 | text \<open>Again, the rather vacuous body of the proof may be collapsed. | 
| 37671 | 306 | Thus the case analysis degenerates into two assumption steps, which | 
| 307 | are implicitly performed when concluding the single rule step of the | |
| 58614 | 308 | double-dot proof as follows.\<close> | 
| 7833 | 309 | |
| 55656 | 310 | lemma "P \<or> P \<longrightarrow> P" | 
| 10007 | 311 | proof | 
| 55640 | 312 | assume "P \<or> P" | 
| 23373 | 313 | then show P .. | 
| 10007 | 314 | qed | 
| 6444 
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
 wenzelm parents: diff
changeset | 315 | |
| 
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
 wenzelm parents: diff
changeset | 316 | |
| 58614 | 317 | subsubsection \<open>A quantifier proof\<close> | 
| 7833 | 318 | |
| 58614 | 319 | text \<open>To illustrate quantifier reasoning, let us prove @{text
 | 
| 37671 | 320 | "(\<exists>x. P (f x)) \<longrightarrow> (\<exists>y. P y)"}. Informally, this holds because any | 
| 321 |   @{text a} with @{text "P (f a)"} may be taken as a witness for the
 | |
| 322 | second existential statement. | |
| 6444 
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
 wenzelm parents: diff
changeset | 323 | |
| 18193 | 324 | The first proof is rather verbose, exhibiting quite a lot of | 
| 325 | (redundant) detail. It gives explicit rules, even with some | |
| 326 | instantiation. Furthermore, we encounter two new language elements: | |
| 327 |   the \isacommand{fix} command augments the context by some new
 | |
| 328 |   ``arbitrary, but fixed'' element; the \isacommand{is} annotation
 | |
| 58614 | 329 | binds term abbreviations by higher-order pattern matching.\<close> | 
| 7833 | 330 | |
| 55640 | 331 | lemma "(\<exists>x. P (f x)) \<longrightarrow> (\<exists>y. P y)" | 
| 10007 | 332 | proof | 
| 55640 | 333 | assume "\<exists>x. P (f x)" | 
| 334 | then show "\<exists>y. P y" | |
| 58614 | 335 | proof (rule exE) -- \<open> | 
| 37671 | 336 |     rule @{text exE}: \smash{$\infer{B}{\ex x A(x) & \infer*{B}{[A(x)]_x}}$}
 | 
| 58614 | 337 | \<close> | 
| 10007 | 338 | fix a | 
| 339 | assume "P (f a)" (is "P ?witness") | |
| 23373 | 340 | then show ?thesis by (rule exI [of P ?witness]) | 
| 10007 | 341 | qed | 
| 342 | qed | |
| 6444 
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
 wenzelm parents: diff
changeset | 343 | |
| 58614 | 344 | text \<open>While explicit rule instantiation may occasionally improve | 
| 18193 | 345 | readability of certain aspects of reasoning, it is usually quite | 
| 346 | redundant. Above, the basic proof outline gives already enough | |
| 347 | structural clues for the system to infer both the rules and their | |
| 348 | instances (by higher-order unification). Thus we may as well prune | |
| 58614 | 349 | the text as follows.\<close> | 
| 7833 | 350 | |
| 55640 | 351 | lemma "(\<exists>x. P (f x)) \<longrightarrow> (\<exists>y. P y)" | 
| 10007 | 352 | proof | 
| 55640 | 353 | assume "\<exists>x. P (f x)" | 
| 354 | then show "\<exists>y. P y" | |
| 10007 | 355 | proof | 
| 356 | fix a | |
| 357 | assume "P (f a)" | |
| 23373 | 358 | then show ?thesis .. | 
| 10007 | 359 | qed | 
| 360 | qed | |
| 6444 
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
 wenzelm parents: diff
changeset | 361 | |
| 58614 | 362 | text \<open>Explicit @{text \<exists>}-elimination as seen above can become quite
 | 
| 18193 | 363 | cumbersome in practice. The derived Isar language element | 
| 364 |   ``\isakeyword{obtain}'' provides a more handsome way to do
 | |
| 58614 | 365 | generalized existence reasoning.\<close> | 
| 9477 | 366 | |
| 55640 | 367 | lemma "(\<exists>x. P (f x)) \<longrightarrow> (\<exists>y. P y)" | 
| 10007 | 368 | proof | 
| 55640 | 369 | assume "\<exists>x. P (f x)" | 
| 10636 | 370 | then obtain a where "P (f a)" .. | 
| 55640 | 371 | then show "\<exists>y. P y" .. | 
| 10007 | 372 | qed | 
| 9477 | 373 | |
| 58614 | 374 | text \<open>Technically, \isakeyword{obtain} is similar to
 | 
| 37671 | 375 |   \isakeyword{fix} and \isakeyword{assume} together with a soundness
 | 
| 376 | proof of the elimination involved. Thus it behaves similar to any | |
| 377 | other forward proof element. Also note that due to the nature of | |
| 378 | general existence reasoning involved here, any result exported from | |
| 379 |   the context of an \isakeyword{obtain} statement may \emph{not} refer
 | |
| 58614 | 380 | to the parameters introduced there.\<close> | 
| 9477 | 381 | |
| 6444 
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
 wenzelm parents: diff
changeset | 382 | |
| 58614 | 383 | subsubsection \<open>Deriving rules in Isabelle\<close> | 
| 7001 | 384 | |
| 58614 | 385 | text \<open>We derive the conjunction elimination rule from the | 
| 37671 | 386 | corresponding projections. The proof is quite straight-forward, | 
| 387 | since Isabelle/Isar supports non-atomic goals and assumptions fully | |
| 58614 | 388 | transparently.\<close> | 
| 7001 | 389 | |
| 55640 | 390 | theorem conjE: "A \<and> B \<Longrightarrow> (A \<Longrightarrow> B \<Longrightarrow> C) \<Longrightarrow> C" | 
| 10007 | 391 | proof - | 
| 55640 | 392 | assume "A \<and> B" | 
| 393 | assume r: "A \<Longrightarrow> B \<Longrightarrow> C" | |
| 10007 | 394 | show C | 
| 395 | proof (rule r) | |
| 23393 | 396 | show A by (rule conjunct1) fact | 
| 397 | show B by (rule conjunct2) fact | |
| 10007 | 398 | qed | 
| 399 | qed | |
| 7001 | 400 | |
| 10007 | 401 | end |