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