author | blanchet |
Fri, 25 Oct 2019 15:23:14 +0200 | |
changeset 70938 | 6d84c3c333d5 |
parent 69597 | ff784d5a5bfb |
child 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
61656 | 1 |
(*:maxLineLen=78:*) |
2 |
||
42917 | 3 |
theory Synopsis |
63531 | 4 |
imports Main Base |
42917 | 5 |
begin |
6 |
||
58618 | 7 |
chapter \<open>Synopsis\<close> |
42917 | 8 |
|
58618 | 9 |
section \<open>Notepad\<close> |
42917 | 10 |
|
58618 | 11 |
text \<open> |
42917 | 12 |
An Isar proof body serves as mathematical notepad to compose logical |
42918 | 13 |
content, consisting of types, terms, facts. |
58618 | 14 |
\<close> |
42917 | 15 |
|
16 |
||
58618 | 17 |
subsection \<open>Types and terms\<close> |
42918 | 18 |
|
19 |
notepad |
|
20 |
begin |
|
58618 | 21 |
txt \<open>Locally fixed entities:\<close> |
61580 | 22 |
fix x \<comment> \<open>local constant, without any type information yet\<close> |
23 |
fix x :: 'a \<comment> \<open>variant with explicit type-constraint for subsequent use\<close> |
|
42918 | 24 |
|
25 |
fix a b |
|
61580 | 26 |
assume "a = b" \<comment> \<open>type assignment at first occurrence in concrete term\<close> |
42918 | 27 |
|
58618 | 28 |
txt \<open>Definitions (non-polymorphic):\<close> |
63039 | 29 |
define x :: 'a where "x = t" |
42918 | 30 |
|
58618 | 31 |
txt \<open>Abbreviations (polymorphic):\<close> |
42918 | 32 |
let ?f = "\<lambda>x. x" |
33 |
term "?f ?f" |
|
34 |
||
58618 | 35 |
txt \<open>Notation:\<close> |
42918 | 36 |
write x ("***") |
37 |
end |
|
38 |
||
39 |
||
58618 | 40 |
subsection \<open>Facts\<close> |
42917 | 41 |
|
58618 | 42 |
text \<open> |
42917 | 43 |
A fact is a simultaneous list of theorems. |
58618 | 44 |
\<close> |
42917 | 45 |
|
46 |
||
58618 | 47 |
subsubsection \<open>Producing facts\<close> |
42917 | 48 |
|
49 |
notepad |
|
50 |
begin |
|
51 |
||
58618 | 52 |
txt \<open>Via assumption (``lambda''):\<close> |
42917 | 53 |
assume a: A |
54 |
||
58618 | 55 |
txt \<open>Via proof (``let''):\<close> |
62271 | 56 |
have b: B \<proof> |
42917 | 57 |
|
58618 | 58 |
txt \<open>Via abbreviation (``let''):\<close> |
42917 | 59 |
note c = a b |
60 |
||
61 |
end |
|
62 |
||
63 |
||
58618 | 64 |
subsubsection \<open>Referencing facts\<close> |
42917 | 65 |
|
66 |
notepad |
|
67 |
begin |
|
58618 | 68 |
txt \<open>Via explicit name:\<close> |
42917 | 69 |
assume a: A |
70 |
note a |
|
71 |
||
58618 | 72 |
txt \<open>Via implicit name:\<close> |
42917 | 73 |
assume A |
74 |
note this |
|
75 |
||
58618 | 76 |
txt \<open>Via literal proposition (unification with results from the proof text):\<close> |
42917 | 77 |
assume A |
58618 | 78 |
note \<open>A\<close> |
42917 | 79 |
|
80 |
assume "\<And>x. B x" |
|
58618 | 81 |
note \<open>B a\<close> |
82 |
note \<open>B b\<close> |
|
42917 | 83 |
end |
84 |
||
85 |
||
58618 | 86 |
subsubsection \<open>Manipulating facts\<close> |
42917 | 87 |
|
88 |
notepad |
|
89 |
begin |
|
58618 | 90 |
txt \<open>Instantiation:\<close> |
42917 | 91 |
assume a: "\<And>x. B x" |
92 |
note a |
|
93 |
note a [of b] |
|
94 |
note a [where x = b] |
|
95 |
||
58618 | 96 |
txt \<open>Backchaining:\<close> |
42917 | 97 |
assume 1: A |
98 |
assume 2: "A \<Longrightarrow> C" |
|
99 |
note 2 [OF 1] |
|
100 |
note 1 [THEN 2] |
|
101 |
||
58618 | 102 |
txt \<open>Symmetric results:\<close> |
42917 | 103 |
assume "x = y" |
104 |
note this [symmetric] |
|
105 |
||
106 |
assume "x \<noteq> y" |
|
107 |
note this [symmetric] |
|
108 |
||
58618 | 109 |
txt \<open>Adhoc-simplification (take care!):\<close> |
42917 | 110 |
assume "P ([] @ xs)" |
111 |
note this [simplified] |
|
112 |
end |
|
113 |
||
114 |
||
58618 | 115 |
subsubsection \<open>Projections\<close> |
42917 | 116 |
|
58618 | 117 |
text \<open> |
42917 | 118 |
Isar facts consist of multiple theorems. There is notation to project |
119 |
interval ranges. |
|
58618 | 120 |
\<close> |
42917 | 121 |
|
122 |
notepad |
|
123 |
begin |
|
124 |
assume stuff: A B C D |
|
125 |
note stuff(1) |
|
126 |
note stuff(2-3) |
|
127 |
note stuff(2-) |
|
128 |
end |
|
129 |
||
130 |
||
58618 | 131 |
subsubsection \<open>Naming conventions\<close> |
42917 | 132 |
|
58618 | 133 |
text \<open> |
61421 | 134 |
\<^item> Lower-case identifiers are usually preferred. |
42917 | 135 |
|
61421 | 136 |
\<^item> Facts can be named after the main term within the proposition. |
42917 | 137 |
|
61477 | 138 |
\<^item> Facts should \<^emph>\<open>not\<close> be named after the command that |
42917 | 139 |
introduced them (@{command "assume"}, @{command "have"}). This is |
140 |
misleading and hard to maintain. |
|
141 |
||
61421 | 142 |
\<^item> Natural numbers can be used as ``meaningless'' names (more |
61493 | 143 |
appropriate than \<open>a1\<close>, \<open>a2\<close> etc.) |
42917 | 144 |
|
61493 | 145 |
\<^item> Symbolic identifiers are supported (e.g. \<open>*\<close>, \<open>**\<close>, \<open>***\<close>). |
58618 | 146 |
\<close> |
42917 | 147 |
|
148 |
||
58618 | 149 |
subsection \<open>Block structure\<close> |
42917 | 150 |
|
58618 | 151 |
text \<open> |
42917 | 152 |
The formal notepad is block structured. The fact produced by the last |
153 |
entry of a block is exported into the outer context. |
|
58618 | 154 |
\<close> |
42917 | 155 |
|
156 |
notepad |
|
157 |
begin |
|
158 |
{ |
|
62271 | 159 |
have a: A \<proof> |
160 |
have b: B \<proof> |
|
42917 | 161 |
note a b |
162 |
} |
|
163 |
note this |
|
58618 | 164 |
note \<open>A\<close> |
165 |
note \<open>B\<close> |
|
42917 | 166 |
end |
167 |
||
58618 | 168 |
text \<open>Explicit blocks as well as implicit blocks of nested goal |
42917 | 169 |
statements (e.g.\ @{command have}) automatically introduce one extra |
170 |
pair of parentheses in reserve. The @{command next} command allows |
|
58618 | 171 |
to ``jump'' between these sub-blocks.\<close> |
42917 | 172 |
|
173 |
notepad |
|
174 |
begin |
|
175 |
||
176 |
{ |
|
62271 | 177 |
have a: A \<proof> |
42917 | 178 |
next |
179 |
have b: B |
|
180 |
proof - |
|
62271 | 181 |
show B \<proof> |
42917 | 182 |
next |
62271 | 183 |
have c: C \<proof> |
42917 | 184 |
next |
62271 | 185 |
have d: D \<proof> |
42917 | 186 |
qed |
187 |
} |
|
188 |
||
58618 | 189 |
txt \<open>Alternative version with explicit parentheses everywhere:\<close> |
42917 | 190 |
|
191 |
{ |
|
192 |
{ |
|
62271 | 193 |
have a: A \<proof> |
42917 | 194 |
} |
195 |
{ |
|
196 |
have b: B |
|
197 |
proof - |
|
198 |
{ |
|
62271 | 199 |
show B \<proof> |
42917 | 200 |
} |
201 |
{ |
|
62271 | 202 |
have c: C \<proof> |
42917 | 203 |
} |
204 |
{ |
|
62271 | 205 |
have d: D \<proof> |
42917 | 206 |
} |
207 |
qed |
|
208 |
} |
|
209 |
} |
|
210 |
||
211 |
end |
|
212 |
||
42919 | 213 |
|
58618 | 214 |
section \<open>Calculational reasoning \label{sec:calculations-synopsis}\<close> |
42919 | 215 |
|
58618 | 216 |
text \<open> |
63680 | 217 |
For example, see \<^file>\<open>~~/src/HOL/Isar_Examples/Group.thy\<close>. |
58618 | 218 |
\<close> |
42919 | 219 |
|
220 |
||
58618 | 221 |
subsection \<open>Special names in Isar proofs\<close> |
42919 | 222 |
|
58618 | 223 |
text \<open> |
61493 | 224 |
\<^item> term \<open>?thesis\<close> --- the main conclusion of the |
42919 | 225 |
innermost pending claim |
226 |
||
61493 | 227 |
\<^item> term \<open>\<dots>\<close> --- the argument of the last explicitly |
61421 | 228 |
stated result (for infix application this is the right-hand side) |
42919 | 229 |
|
61493 | 230 |
\<^item> fact \<open>this\<close> --- the last result produced in the text |
58618 | 231 |
\<close> |
42919 | 232 |
|
233 |
notepad |
|
234 |
begin |
|
235 |
have "x = y" |
|
236 |
proof - |
|
237 |
term ?thesis |
|
62271 | 238 |
show ?thesis \<proof> |
61580 | 239 |
term ?thesis \<comment> \<open>static!\<close> |
42919 | 240 |
qed |
241 |
term "\<dots>" |
|
242 |
thm this |
|
243 |
end |
|
244 |
||
58618 | 245 |
text \<open>Calculational reasoning maintains the special fact called |
61493 | 246 |
``\<open>calculation\<close>'' in the background. Certain language |
247 |
elements combine primary \<open>this\<close> with secondary \<open>calculation\<close>.\<close> |
|
42919 | 248 |
|
249 |
||
58618 | 250 |
subsection \<open>Transitive chains\<close> |
42919 | 251 |
|
61493 | 252 |
text \<open>The Idea is to combine \<open>this\<close> and \<open>calculation\<close> |
253 |
via typical \<open>trans\<close> rules (see also @{command |
|
58618 | 254 |
print_trans_rules}):\<close> |
42919 | 255 |
|
256 |
thm trans |
|
257 |
thm less_trans |
|
258 |
thm less_le_trans |
|
259 |
||
260 |
notepad |
|
261 |
begin |
|
58618 | 262 |
txt \<open>Plain bottom-up calculation:\<close> |
62271 | 263 |
have "a = b" \<proof> |
42919 | 264 |
also |
62271 | 265 |
have "b = c" \<proof> |
42919 | 266 |
also |
62271 | 267 |
have "c = d" \<proof> |
42919 | 268 |
finally |
269 |
have "a = d" . |
|
270 |
||
61493 | 271 |
txt \<open>Variant using the \<open>\<dots>\<close> abbreviation:\<close> |
62271 | 272 |
have "a = b" \<proof> |
42919 | 273 |
also |
62271 | 274 |
have "\<dots> = c" \<proof> |
42919 | 275 |
also |
62271 | 276 |
have "\<dots> = d" \<proof> |
42919 | 277 |
finally |
278 |
have "a = d" . |
|
279 |
||
58618 | 280 |
txt \<open>Top-down version with explicit claim at the head:\<close> |
42919 | 281 |
have "a = d" |
282 |
proof - |
|
62271 | 283 |
have "a = b" \<proof> |
42919 | 284 |
also |
62271 | 285 |
have "\<dots> = c" \<proof> |
42919 | 286 |
also |
62271 | 287 |
have "\<dots> = d" \<proof> |
42919 | 288 |
finally |
289 |
show ?thesis . |
|
290 |
qed |
|
291 |
next |
|
58618 | 292 |
txt \<open>Mixed inequalities (require suitable base type):\<close> |
42919 | 293 |
fix a b c d :: nat |
294 |
||
62271 | 295 |
have "a < b" \<proof> |
42919 | 296 |
also |
62271 | 297 |
have "b \<le> c" \<proof> |
42919 | 298 |
also |
62271 | 299 |
have "c = d" \<proof> |
42919 | 300 |
finally |
301 |
have "a < d" . |
|
302 |
end |
|
303 |
||
304 |
||
58618 | 305 |
subsubsection \<open>Notes\<close> |
42919 | 306 |
|
58618 | 307 |
text \<open> |
61493 | 308 |
\<^item> The notion of \<open>trans\<close> rule is very general due to the |
42919 | 309 |
flexibility of Isabelle/Pure rule composition. |
310 |
||
61421 | 311 |
\<^item> User applications may declare their own rules, with some care |
42919 | 312 |
about the operational details of higher-order unification. |
58618 | 313 |
\<close> |
42919 | 314 |
|
315 |
||
62273 | 316 |
subsection \<open>Degenerate calculations\<close> |
42919 | 317 |
|
62273 | 318 |
text \<open>The Idea is to append \<open>this\<close> to \<open>calculation\<close>, without rule composition. |
319 |
This is occasionally useful to avoid naming intermediate facts.\<close> |
|
42919 | 320 |
|
321 |
notepad |
|
322 |
begin |
|
58618 | 323 |
txt \<open>A vacuous proof:\<close> |
62271 | 324 |
have A \<proof> |
42919 | 325 |
moreover |
62271 | 326 |
have B \<proof> |
42919 | 327 |
moreover |
62271 | 328 |
have C \<proof> |
42919 | 329 |
ultimately |
330 |
have A and B and C . |
|
331 |
next |
|
58618 | 332 |
txt \<open>Slightly more content (trivial bigstep reasoning):\<close> |
62271 | 333 |
have A \<proof> |
42919 | 334 |
moreover |
62271 | 335 |
have B \<proof> |
42919 | 336 |
moreover |
62271 | 337 |
have C \<proof> |
42919 | 338 |
ultimately |
339 |
have "A \<and> B \<and> C" by blast |
|
340 |
end |
|
341 |
||
62273 | 342 |
text \<open>Note that For multi-branch case splitting, it is better to use @{command |
343 |
consider}.\<close> |
|
344 |
||
42920 | 345 |
|
58618 | 346 |
section \<open>Induction\<close> |
42921 | 347 |
|
58618 | 348 |
subsection \<open>Induction as Natural Deduction\<close> |
42921 | 349 |
|
58618 | 350 |
text \<open>In principle, induction is just a special case of Natural |
42921 | 351 |
Deduction (see also \secref{sec:natural-deduction-synopsis}). For |
58618 | 352 |
example:\<close> |
42921 | 353 |
|
354 |
thm nat.induct |
|
355 |
print_statement nat.induct |
|
356 |
||
357 |
notepad |
|
358 |
begin |
|
359 |
fix n :: nat |
|
360 |
have "P n" |
|
61580 | 361 |
proof (rule nat.induct) \<comment> \<open>fragile rule application!\<close> |
62271 | 362 |
show "P 0" \<proof> |
42921 | 363 |
next |
364 |
fix n :: nat |
|
365 |
assume "P n" |
|
62271 | 366 |
show "P (Suc n)" \<proof> |
42921 | 367 |
qed |
368 |
end |
|
369 |
||
58618 | 370 |
text \<open> |
42921 | 371 |
In practice, much more proof infrastructure is required. |
372 |
||
373 |
The proof method @{method induct} provides: |
|
374 |
||
61421 | 375 |
\<^item> implicit rule selection and robust instantiation |
42921 | 376 |
|
61421 | 377 |
\<^item> context elements via symbolic case names |
42921 | 378 |
|
61421 | 379 |
\<^item> support for rule-structured induction statements, with local |
380 |
parameters, premises, etc. |
|
58618 | 381 |
\<close> |
42921 | 382 |
|
383 |
notepad |
|
384 |
begin |
|
385 |
fix n :: nat |
|
386 |
have "P n" |
|
387 |
proof (induct n) |
|
388 |
case 0 |
|
62271 | 389 |
show ?case \<proof> |
42921 | 390 |
next |
391 |
case (Suc n) |
|
62271 | 392 |
from Suc.hyps show ?case \<proof> |
42921 | 393 |
qed |
394 |
end |
|
395 |
||
396 |
||
58618 | 397 |
subsubsection \<open>Example\<close> |
42921 | 398 |
|
58618 | 399 |
text \<open> |
42921 | 400 |
The subsequent example combines the following proof patterns: |
401 |
||
61421 | 402 |
\<^item> outermost induction (over the datatype structure of natural |
42921 | 403 |
numbers), to decompose the proof problem in top-down manner |
404 |
||
61421 | 405 |
\<^item> calculational reasoning (\secref{sec:calculations-synopsis}) |
42921 | 406 |
to compose the result in each case |
407 |
||
61421 | 408 |
\<^item> solving local claims within the calculation by simplification |
58618 | 409 |
\<close> |
42921 | 410 |
|
411 |
lemma |
|
412 |
fixes n :: nat |
|
413 |
shows "(\<Sum>i=0..n. i) = n * (n + 1) div 2" |
|
414 |
proof (induct n) |
|
415 |
case 0 |
|
416 |
have "(\<Sum>i=0..0. i) = (0::nat)" by simp |
|
417 |
also have "\<dots> = 0 * (0 + 1) div 2" by simp |
|
418 |
finally show ?case . |
|
419 |
next |
|
420 |
case (Suc n) |
|
421 |
have "(\<Sum>i=0..Suc n. i) = (\<Sum>i=0..n. i) + (n + 1)" by simp |
|
422 |
also have "\<dots> = n * (n + 1) div 2 + (n + 1)" by (simp add: Suc.hyps) |
|
423 |
also have "\<dots> = (n * (n + 1) + 2 * (n + 1)) div 2" by simp |
|
424 |
also have "\<dots> = (Suc n * (Suc n + 1)) div 2" by simp |
|
425 |
finally show ?case . |
|
426 |
qed |
|
427 |
||
58618 | 428 |
text \<open>This demonstrates how induction proofs can be done without |
429 |
having to consider the raw Natural Deduction structure.\<close> |
|
42921 | 430 |
|
431 |
||
58618 | 432 |
subsection \<open>Induction with local parameters and premises\<close> |
42921 | 433 |
|
58618 | 434 |
text \<open>Idea: Pure rule statements are passed through the induction |
42921 | 435 |
rule. This achieves convenient proof patterns, thanks to some |
436 |
internal trickery in the @{method induct} method. |
|
437 |
||
61493 | 438 |
Important: Using compact HOL formulae with \<open>\<forall>/\<longrightarrow>\<close> is a |
42921 | 439 |
well-known anti-pattern! It would produce useless formal noise. |
58618 | 440 |
\<close> |
42921 | 441 |
|
442 |
notepad |
|
443 |
begin |
|
444 |
fix n :: nat |
|
445 |
fix P :: "nat \<Rightarrow> bool" |
|
446 |
fix Q :: "'a \<Rightarrow> nat \<Rightarrow> bool" |
|
447 |
||
448 |
have "P n" |
|
449 |
proof (induct n) |
|
450 |
case 0 |
|
62271 | 451 |
show "P 0" \<proof> |
42921 | 452 |
next |
453 |
case (Suc n) |
|
62271 | 454 |
from \<open>P n\<close> show "P (Suc n)" \<proof> |
42921 | 455 |
qed |
456 |
||
457 |
have "A n \<Longrightarrow> P n" |
|
458 |
proof (induct n) |
|
459 |
case 0 |
|
62271 | 460 |
from \<open>A 0\<close> show "P 0" \<proof> |
42921 | 461 |
next |
462 |
case (Suc n) |
|
58618 | 463 |
from \<open>A n \<Longrightarrow> P n\<close> |
62271 | 464 |
and \<open>A (Suc n)\<close> show "P (Suc n)" \<proof> |
42921 | 465 |
qed |
466 |
||
467 |
have "\<And>x. Q x n" |
|
468 |
proof (induct n) |
|
469 |
case 0 |
|
62271 | 470 |
show "Q x 0" \<proof> |
42921 | 471 |
next |
472 |
case (Suc n) |
|
62271 | 473 |
from \<open>\<And>x. Q x n\<close> show "Q x (Suc n)" \<proof> |
58618 | 474 |
txt \<open>Local quantification admits arbitrary instances:\<close> |
475 |
note \<open>Q a n\<close> and \<open>Q b n\<close> |
|
42921 | 476 |
qed |
477 |
end |
|
478 |
||
479 |
||
58618 | 480 |
subsection \<open>Implicit induction context\<close> |
42921 | 481 |
|
58618 | 482 |
text \<open>The @{method induct} method can isolate local parameters and |
42921 | 483 |
premises directly from the given statement. This is convenient in |
484 |
practical applications, but requires some understanding of what is |
|
58618 | 485 |
going on internally (as explained above).\<close> |
42921 | 486 |
|
487 |
notepad |
|
488 |
begin |
|
489 |
fix n :: nat |
|
490 |
fix Q :: "'a \<Rightarrow> nat \<Rightarrow> bool" |
|
491 |
||
492 |
fix x :: 'a |
|
493 |
assume "A x n" |
|
494 |
then have "Q x n" |
|
495 |
proof (induct n arbitrary: x) |
|
496 |
case 0 |
|
62271 | 497 |
from \<open>A x 0\<close> show "Q x 0" \<proof> |
42921 | 498 |
next |
499 |
case (Suc n) |
|
61580 | 500 |
from \<open>\<And>x. A x n \<Longrightarrow> Q x n\<close> \<comment> \<open>arbitrary instances can be produced here\<close> |
62271 | 501 |
and \<open>A x (Suc n)\<close> show "Q x (Suc n)" \<proof> |
42921 | 502 |
qed |
503 |
end |
|
504 |
||
505 |
||
58618 | 506 |
subsection \<open>Advanced induction with term definitions\<close> |
42921 | 507 |
|
58618 | 508 |
text \<open>Induction over subexpressions of a certain shape are delicate |
42921 | 509 |
to formalize. The Isar @{method induct} method provides |
510 |
infrastructure for this. |
|
511 |
||
512 |
Idea: sub-expressions of the problem are turned into a defined |
|
513 |
induction variable; often accompanied with fixing of auxiliary |
|
58618 | 514 |
parameters in the original expression.\<close> |
42921 | 515 |
|
516 |
notepad |
|
517 |
begin |
|
518 |
fix a :: "'a \<Rightarrow> nat" |
|
519 |
fix A :: "nat \<Rightarrow> bool" |
|
520 |
||
521 |
assume "A (a x)" |
|
522 |
then have "P (a x)" |
|
523 |
proof (induct "a x" arbitrary: x) |
|
524 |
case 0 |
|
58618 | 525 |
note prem = \<open>A (a x)\<close> |
526 |
and defn = \<open>0 = a x\<close> |
|
62271 | 527 |
show "P (a x)" \<proof> |
42921 | 528 |
next |
529 |
case (Suc n) |
|
58618 | 530 |
note hyp = \<open>\<And>x. n = a x \<Longrightarrow> A (a x) \<Longrightarrow> P (a x)\<close> |
531 |
and prem = \<open>A (a x)\<close> |
|
532 |
and defn = \<open>Suc n = a x\<close> |
|
62271 | 533 |
show "P (a x)" \<proof> |
42921 | 534 |
qed |
535 |
end |
|
536 |
||
537 |
||
58618 | 538 |
section \<open>Natural Deduction \label{sec:natural-deduction-synopsis}\<close> |
42920 | 539 |
|
58618 | 540 |
subsection \<open>Rule statements\<close> |
42920 | 541 |
|
58618 | 542 |
text \<open> |
42920 | 543 |
Isabelle/Pure ``theorems'' are always natural deduction rules, |
544 |
which sometimes happen to consist of a conclusion only. |
|
545 |
||
61493 | 546 |
The framework connectives \<open>\<And>\<close> and \<open>\<Longrightarrow>\<close> indicate the |
58618 | 547 |
rule structure declaratively. For example:\<close> |
42920 | 548 |
|
549 |
thm conjI |
|
550 |
thm impI |
|
551 |
thm nat.induct |
|
552 |
||
58618 | 553 |
text \<open> |
42920 | 554 |
The object-logic is embedded into the Pure framework via an implicit |
69597 | 555 |
derivability judgment \<^term>\<open>Trueprop :: bool \<Rightarrow> prop\<close>. |
42920 | 556 |
|
557 |
Thus any HOL formulae appears atomic to the Pure framework, while |
|
558 |
the rule structure outlines the corresponding proof pattern. |
|
559 |
||
560 |
This can be made explicit as follows: |
|
58618 | 561 |
\<close> |
42920 | 562 |
|
563 |
notepad |
|
564 |
begin |
|
565 |
write Trueprop ("Tr") |
|
566 |
||
567 |
thm conjI |
|
568 |
thm impI |
|
569 |
thm nat.induct |
|
570 |
end |
|
571 |
||
58618 | 572 |
text \<open> |
42920 | 573 |
Isar provides first-class notation for rule statements as follows. |
58618 | 574 |
\<close> |
42920 | 575 |
|
576 |
print_statement conjI |
|
577 |
print_statement impI |
|
578 |
print_statement nat.induct |
|
579 |
||
580 |
||
58618 | 581 |
subsubsection \<open>Examples\<close> |
42920 | 582 |
|
58618 | 583 |
text \<open> |
42920 | 584 |
Introductions and eliminations of some standard connectives of |
585 |
the object-logic can be written as rule statements as follows. (The |
|
586 |
proof ``@{command "by"}~@{method blast}'' serves as sanity check.) |
|
58618 | 587 |
\<close> |
42920 | 588 |
|
589 |
lemma "(P \<Longrightarrow> False) \<Longrightarrow> \<not> P" by blast |
|
590 |
lemma "\<not> P \<Longrightarrow> P \<Longrightarrow> Q" by blast |
|
591 |
||
592 |
lemma "P \<Longrightarrow> Q \<Longrightarrow> P \<and> Q" by blast |
|
593 |
lemma "P \<and> Q \<Longrightarrow> (P \<Longrightarrow> Q \<Longrightarrow> R) \<Longrightarrow> R" by blast |
|
594 |
||
595 |
lemma "P \<Longrightarrow> P \<or> Q" by blast |
|
596 |
lemma "Q \<Longrightarrow> P \<or> Q" by blast |
|
597 |
lemma "P \<or> Q \<Longrightarrow> (P \<Longrightarrow> R) \<Longrightarrow> (Q \<Longrightarrow> R) \<Longrightarrow> R" by blast |
|
598 |
||
599 |
lemma "(\<And>x. P x) \<Longrightarrow> (\<forall>x. P x)" by blast |
|
600 |
lemma "(\<forall>x. P x) \<Longrightarrow> P x" by blast |
|
601 |
||
602 |
lemma "P x \<Longrightarrow> (\<exists>x. P x)" by blast |
|
603 |
lemma "(\<exists>x. P x) \<Longrightarrow> (\<And>x. P x \<Longrightarrow> R) \<Longrightarrow> R" by blast |
|
604 |
||
605 |
lemma "x \<in> A \<Longrightarrow> x \<in> B \<Longrightarrow> x \<in> A \<inter> B" by blast |
|
606 |
lemma "x \<in> A \<inter> B \<Longrightarrow> (x \<in> A \<Longrightarrow> x \<in> B \<Longrightarrow> R) \<Longrightarrow> R" by blast |
|
607 |
||
608 |
lemma "x \<in> A \<Longrightarrow> x \<in> A \<union> B" by blast |
|
609 |
lemma "x \<in> B \<Longrightarrow> x \<in> A \<union> B" by blast |
|
610 |
lemma "x \<in> A \<union> B \<Longrightarrow> (x \<in> A \<Longrightarrow> R) \<Longrightarrow> (x \<in> B \<Longrightarrow> R) \<Longrightarrow> R" by blast |
|
611 |
||
612 |
||
58618 | 613 |
subsection \<open>Isar context elements\<close> |
42920 | 614 |
|
58618 | 615 |
text \<open>We derive some results out of the blue, using Isar context |
42920 | 616 |
elements and some explicit blocks. This illustrates their meaning |
58618 | 617 |
wrt.\ Pure connectives, without goal states getting in the way.\<close> |
42920 | 618 |
|
619 |
notepad |
|
620 |
begin |
|
621 |
{ |
|
622 |
fix x |
|
62271 | 623 |
have "B x" \<proof> |
42920 | 624 |
} |
625 |
have "\<And>x. B x" by fact |
|
626 |
||
627 |
next |
|
628 |
||
629 |
{ |
|
630 |
assume A |
|
62271 | 631 |
have B \<proof> |
42920 | 632 |
} |
633 |
have "A \<Longrightarrow> B" by fact |
|
634 |
||
635 |
next |
|
636 |
||
637 |
{ |
|
63039 | 638 |
define x where "x = t" |
62271 | 639 |
have "B x" \<proof> |
42920 | 640 |
} |
641 |
have "B t" by fact |
|
642 |
||
643 |
next |
|
644 |
||
645 |
{ |
|
62271 | 646 |
obtain x :: 'a where "B x" \<proof> |
647 |
have C \<proof> |
|
42920 | 648 |
} |
649 |
have C by fact |
|
650 |
||
651 |
end |
|
652 |
||
653 |
||
58618 | 654 |
subsection \<open>Pure rule composition\<close> |
42920 | 655 |
|
58618 | 656 |
text \<open> |
42920 | 657 |
The Pure framework provides means for: |
658 |
||
61421 | 659 |
\<^item> backward-chaining of rules by @{inference resolution} |
42920 | 660 |
|
61421 | 661 |
\<^item> closing of branches by @{inference assumption} |
42920 | 662 |
|
663 |
||
61493 | 664 |
Both principles involve higher-order unification of \<open>\<lambda>\<close>-terms |
665 |
modulo \<open>\<alpha>\<beta>\<eta>\<close>-equivalence (cf.\ Huet and Miller). |
|
61458 | 666 |
\<close> |
42920 | 667 |
|
668 |
notepad |
|
669 |
begin |
|
670 |
assume a: A and b: B |
|
671 |
thm conjI |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67369
diff
changeset
|
672 |
thm conjI [of A B] \<comment> \<open>instantiation\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67369
diff
changeset
|
673 |
thm conjI [of A B, OF a b] \<comment> \<open>instantiation and composition\<close> |
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67369
diff
changeset
|
674 |
thm conjI [OF a b] \<comment> \<open>composition via unification (trivial)\<close> |
58618 | 675 |
thm conjI [OF \<open>A\<close> \<open>B\<close>] |
42920 | 676 |
|
677 |
thm conjI [OF disjI1] |
|
678 |
end |
|
679 |
||
58618 | 680 |
text \<open>Note: Low-level rule composition is tedious and leads to |
681 |
unreadable~/ unmaintainable expressions in the text.\<close> |
|
42920 | 682 |
|
683 |
||
58618 | 684 |
subsection \<open>Structured backward reasoning\<close> |
42920 | 685 |
|
58618 | 686 |
text \<open>Idea: Canonical proof decomposition via @{command fix}~/ |
42920 | 687 |
@{command assume}~/ @{command show}, where the body produces a |
58618 | 688 |
natural deduction rule to refine some goal.\<close> |
42920 | 689 |
|
690 |
notepad |
|
691 |
begin |
|
692 |
fix A B :: "'a \<Rightarrow> bool" |
|
693 |
||
694 |
have "\<And>x. A x \<Longrightarrow> B x" |
|
695 |
proof - |
|
696 |
fix x |
|
697 |
assume "A x" |
|
62271 | 698 |
show "B x" \<proof> |
42920 | 699 |
qed |
700 |
||
701 |
have "\<And>x. A x \<Longrightarrow> B x" |
|
702 |
proof - |
|
703 |
{ |
|
704 |
fix x |
|
705 |
assume "A x" |
|
62271 | 706 |
show "B x" \<proof> |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67369
diff
changeset
|
707 |
} \<comment> \<open>implicit block structure made explicit\<close> |
58618 | 708 |
note \<open>\<And>x. A x \<Longrightarrow> B x\<close> |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67369
diff
changeset
|
709 |
\<comment> \<open>side exit for the resulting rule\<close> |
42920 | 710 |
qed |
711 |
end |
|
712 |
||
713 |
||
58618 | 714 |
subsection \<open>Structured rule application\<close> |
42920 | 715 |
|
58618 | 716 |
text \<open> |
42920 | 717 |
Idea: Previous facts and new claims are composed with a rule from |
718 |
the context (or background library). |
|
58618 | 719 |
\<close> |
42920 | 720 |
|
721 |
notepad |
|
722 |
begin |
|
62272 | 723 |
assume r\<^sub>1: "A \<Longrightarrow> B \<Longrightarrow> C" \<comment> \<open>simple rule (Horn clause)\<close> |
42920 | 724 |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67369
diff
changeset
|
725 |
have A \<proof> \<comment> \<open>prefix of facts via outer sub-proof\<close> |
42920 | 726 |
then have C |
62272 | 727 |
proof (rule r\<^sub>1) |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67369
diff
changeset
|
728 |
show B \<proof> \<comment> \<open>remaining rule premises via inner sub-proof\<close> |
42920 | 729 |
qed |
730 |
||
731 |
have C |
|
62272 | 732 |
proof (rule r\<^sub>1) |
62271 | 733 |
show A \<proof> |
734 |
show B \<proof> |
|
42920 | 735 |
qed |
736 |
||
62271 | 737 |
have A and B \<proof> |
42920 | 738 |
then have C |
62272 | 739 |
proof (rule r\<^sub>1) |
42920 | 740 |
qed |
741 |
||
62271 | 742 |
have A and B \<proof> |
62272 | 743 |
then have C by (rule r\<^sub>1) |
42920 | 744 |
|
745 |
next |
|
746 |
||
62272 | 747 |
assume r\<^sub>2: "A \<Longrightarrow> (\<And>x. B\<^sub>1 x \<Longrightarrow> B\<^sub>2 x) \<Longrightarrow> C" \<comment> \<open>nested rule\<close> |
42920 | 748 |
|
62271 | 749 |
have A \<proof> |
42920 | 750 |
then have C |
62272 | 751 |
proof (rule r\<^sub>2) |
42920 | 752 |
fix x |
62272 | 753 |
assume "B\<^sub>1 x" |
754 |
show "B\<^sub>2 x" \<proof> |
|
42920 | 755 |
qed |
756 |
||
69597 | 757 |
txt \<open>The compound rule premise \<^prop>\<open>\<And>x. B\<^sub>1 x \<Longrightarrow> B\<^sub>2 x\<close> is better |
42920 | 758 |
addressed via @{command fix}~/ @{command assume}~/ @{command show} |
58618 | 759 |
in the nested proof body.\<close> |
42920 | 760 |
end |
761 |
||
762 |
||
58618 | 763 |
subsection \<open>Example: predicate logic\<close> |
42920 | 764 |
|
58618 | 765 |
text \<open> |
42920 | 766 |
Using the above principles, standard introduction and elimination proofs |
767 |
of predicate logic connectives of HOL work as follows. |
|
58618 | 768 |
\<close> |
42920 | 769 |
|
770 |
notepad |
|
771 |
begin |
|
62271 | 772 |
have "A \<longrightarrow> B" and A \<proof> |
42920 | 773 |
then have B .. |
774 |
||
62271 | 775 |
have A \<proof> |
42920 | 776 |
then have "A \<or> B" .. |
777 |
||
62271 | 778 |
have B \<proof> |
42920 | 779 |
then have "A \<or> B" .. |
780 |
||
62271 | 781 |
have "A \<or> B" \<proof> |
42920 | 782 |
then have C |
783 |
proof |
|
784 |
assume A |
|
62271 | 785 |
then show C \<proof> |
42920 | 786 |
next |
787 |
assume B |
|
62271 | 788 |
then show C \<proof> |
42920 | 789 |
qed |
790 |
||
62271 | 791 |
have A and B \<proof> |
42920 | 792 |
then have "A \<and> B" .. |
793 |
||
62271 | 794 |
have "A \<and> B" \<proof> |
42920 | 795 |
then have A .. |
796 |
||
62271 | 797 |
have "A \<and> B" \<proof> |
42920 | 798 |
then have B .. |
799 |
||
62271 | 800 |
have False \<proof> |
42920 | 801 |
then have A .. |
802 |
||
803 |
have True .. |
|
804 |
||
805 |
have "\<not> A" |
|
806 |
proof |
|
807 |
assume A |
|
62271 | 808 |
then show False \<proof> |
42920 | 809 |
qed |
810 |
||
62271 | 811 |
have "\<not> A" and A \<proof> |
42920 | 812 |
then have B .. |
813 |
||
814 |
have "\<forall>x. P x" |
|
815 |
proof |
|
816 |
fix x |
|
62271 | 817 |
show "P x" \<proof> |
42920 | 818 |
qed |
819 |
||
62271 | 820 |
have "\<forall>x. P x" \<proof> |
42920 | 821 |
then have "P a" .. |
822 |
||
823 |
have "\<exists>x. P x" |
|
824 |
proof |
|
62271 | 825 |
show "P a" \<proof> |
42920 | 826 |
qed |
827 |
||
62271 | 828 |
have "\<exists>x. P x" \<proof> |
42920 | 829 |
then have C |
830 |
proof |
|
831 |
fix a |
|
832 |
assume "P a" |
|
62271 | 833 |
show C \<proof> |
42920 | 834 |
qed |
835 |
||
58618 | 836 |
txt \<open>Less awkward version using @{command obtain}:\<close> |
62271 | 837 |
have "\<exists>x. P x" \<proof> |
42920 | 838 |
then obtain a where "P a" .. |
839 |
end |
|
840 |
||
58618 | 841 |
text \<open>Further variations to illustrate Isar sub-proofs involving |
842 |
@{command show}:\<close> |
|
42920 | 843 |
|
844 |
notepad |
|
845 |
begin |
|
846 |
have "A \<and> B" |
|
61580 | 847 |
proof \<comment> \<open>two strictly isolated subproofs\<close> |
62271 | 848 |
show A \<proof> |
42920 | 849 |
next |
62271 | 850 |
show B \<proof> |
42920 | 851 |
qed |
852 |
||
853 |
have "A \<and> B" |
|
61580 | 854 |
proof \<comment> \<open>one simultaneous sub-proof\<close> |
62271 | 855 |
show A and B \<proof> |
42920 | 856 |
qed |
857 |
||
858 |
have "A \<and> B" |
|
61580 | 859 |
proof \<comment> \<open>two subproofs in the same context\<close> |
62271 | 860 |
show A \<proof> |
861 |
show B \<proof> |
|
42920 | 862 |
qed |
863 |
||
864 |
have "A \<and> B" |
|
61580 | 865 |
proof \<comment> \<open>swapped order\<close> |
62271 | 866 |
show B \<proof> |
867 |
show A \<proof> |
|
42920 | 868 |
qed |
869 |
||
870 |
have "A \<and> B" |
|
61580 | 871 |
proof \<comment> \<open>sequential subproofs\<close> |
62271 | 872 |
show A \<proof> |
873 |
show B using \<open>A\<close> \<proof> |
|
42920 | 874 |
qed |
875 |
end |
|
876 |
||
877 |
||
58618 | 878 |
subsubsection \<open>Example: set-theoretic operators\<close> |
42920 | 879 |
|
61493 | 880 |
text \<open>There is nothing special about logical connectives (\<open>\<and>\<close>, \<open>\<or>\<close>, \<open>\<forall>\<close>, \<open>\<exists>\<close> etc.). Operators from |
45103 | 881 |
set-theory or lattice-theory work analogously. It is only a matter |
42920 | 882 |
of rule declarations in the library; rules can be also specified |
883 |
explicitly. |
|
58618 | 884 |
\<close> |
42920 | 885 |
|
886 |
notepad |
|
887 |
begin |
|
62271 | 888 |
have "x \<in> A" and "x \<in> B" \<proof> |
42920 | 889 |
then have "x \<in> A \<inter> B" .. |
890 |
||
62271 | 891 |
have "x \<in> A" \<proof> |
42920 | 892 |
then have "x \<in> A \<union> B" .. |
893 |
||
62271 | 894 |
have "x \<in> B" \<proof> |
42920 | 895 |
then have "x \<in> A \<union> B" .. |
896 |
||
62271 | 897 |
have "x \<in> A \<union> B" \<proof> |
42920 | 898 |
then have C |
899 |
proof |
|
900 |
assume "x \<in> A" |
|
62271 | 901 |
then show C \<proof> |
42920 | 902 |
next |
903 |
assume "x \<in> B" |
|
62271 | 904 |
then show C \<proof> |
42920 | 905 |
qed |
906 |
||
907 |
next |
|
908 |
have "x \<in> \<Inter>A" |
|
909 |
proof |
|
910 |
fix a |
|
911 |
assume "a \<in> A" |
|
62271 | 912 |
show "x \<in> a" \<proof> |
42920 | 913 |
qed |
914 |
||
62271 | 915 |
have "x \<in> \<Inter>A" \<proof> |
42920 | 916 |
then have "x \<in> a" |
917 |
proof |
|
62271 | 918 |
show "a \<in> A" \<proof> |
42920 | 919 |
qed |
920 |
||
62271 | 921 |
have "a \<in> A" and "x \<in> a" \<proof> |
42920 | 922 |
then have "x \<in> \<Union>A" .. |
923 |
||
62271 | 924 |
have "x \<in> \<Union>A" \<proof> |
42920 | 925 |
then obtain a where "a \<in> A" and "x \<in> a" .. |
926 |
end |
|
927 |
||
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
928 |
|
58618 | 929 |
section \<open>Generalized elimination and cases\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
930 |
|
58618 | 931 |
subsection \<open>General elimination rules\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
932 |
|
58618 | 933 |
text \<open> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
934 |
The general format of elimination rules is illustrated by the |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
935 |
following typical representatives: |
58618 | 936 |
\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
937 |
|
61580 | 938 |
thm exE \<comment> \<open>local parameter\<close> |
939 |
thm conjE \<comment> \<open>local premises\<close> |
|
940 |
thm disjE \<comment> \<open>split into cases\<close> |
|
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
941 |
|
58618 | 942 |
text \<open> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
943 |
Combining these characteristics leads to the following general scheme |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
944 |
for elimination rules with cases: |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
945 |
|
61421 | 946 |
\<^item> prefix of assumptions (or ``major premises'') |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
947 |
|
61421 | 948 |
\<^item> one or more cases that enable to establish the main conclusion |
949 |
in an augmented context |
|
58618 | 950 |
\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
951 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
952 |
notepad |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
953 |
begin |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
954 |
assume r: |
67369 | 955 |
"A\<^sub>1 \<Longrightarrow> A\<^sub>2 \<Longrightarrow> \<comment> \<open>assumptions\<close> |
956 |
(\<And>x y. B\<^sub>1 x y \<Longrightarrow> C\<^sub>1 x y \<Longrightarrow> R) \<Longrightarrow> \<comment> \<open>case 1\<close> |
|
957 |
(\<And>x y. B\<^sub>2 x y \<Longrightarrow> C\<^sub>2 x y \<Longrightarrow> R) \<Longrightarrow> \<comment> \<open>case 2\<close> |
|
958 |
R \<comment> \<open>main conclusion\<close>" |
|
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
959 |
|
62272 | 960 |
have A\<^sub>1 and A\<^sub>2 \<proof> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
961 |
then have R |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
962 |
proof (rule r) |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
963 |
fix x y |
62272 | 964 |
assume "B\<^sub>1 x y" and "C\<^sub>1 x y" |
62271 | 965 |
show ?thesis \<proof> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
966 |
next |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
967 |
fix x y |
62272 | 968 |
assume "B\<^sub>2 x y" and "C\<^sub>2 x y" |
62271 | 969 |
show ?thesis \<proof> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
970 |
qed |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
971 |
end |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
972 |
|
61493 | 973 |
text \<open>Here \<open>?thesis\<close> is used to refer to the unchanged goal |
58618 | 974 |
statement.\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
975 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
976 |
|
58618 | 977 |
subsection \<open>Rules with cases\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
978 |
|
58618 | 979 |
text \<open> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
980 |
Applying an elimination rule to some goal, leaves that unchanged |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
981 |
but allows to augment the context in the sub-proof of each case. |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
982 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
983 |
Isar provides some infrastructure to support this: |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
984 |
|
61421 | 985 |
\<^item> native language elements to state eliminations |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
986 |
|
61421 | 987 |
\<^item> symbolic case names |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
988 |
|
61421 | 989 |
\<^item> method @{method cases} to recover this structure in a |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
990 |
sub-proof |
58618 | 991 |
\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
992 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
993 |
print_statement exE |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
994 |
print_statement conjE |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
995 |
print_statement disjE |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
996 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
997 |
lemma |
62272 | 998 |
assumes A\<^sub>1 and A\<^sub>2 \<comment> \<open>assumptions\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
999 |
obtains |
62272 | 1000 |
(case\<^sub>1) x y where "B\<^sub>1 x y" and "C\<^sub>1 x y" |
1001 |
| (case\<^sub>2) x y where "B\<^sub>2 x y" and "C\<^sub>2 x y" |
|
62271 | 1002 |
\<proof> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1003 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1004 |
|
58618 | 1005 |
subsubsection \<open>Example\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1006 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1007 |
lemma tertium_non_datur: |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1008 |
obtains |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1009 |
(T) A |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1010 |
| (F) "\<not> A" |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1011 |
by blast |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1012 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1013 |
notepad |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1014 |
begin |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1015 |
fix x y :: 'a |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1016 |
have C |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1017 |
proof (cases "x = y" rule: tertium_non_datur) |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1018 |
case T |
62271 | 1019 |
from \<open>x = y\<close> show ?thesis \<proof> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1020 |
next |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1021 |
case F |
62271 | 1022 |
from \<open>x \<noteq> y\<close> show ?thesis \<proof> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1023 |
qed |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1024 |
end |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1025 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1026 |
|
58618 | 1027 |
subsubsection \<open>Example\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1028 |
|
58618 | 1029 |
text \<open> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1030 |
Isabelle/HOL specification mechanisms (datatype, inductive, etc.) |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1031 |
provide suitable derived cases rules. |
58618 | 1032 |
\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1033 |
|
58310 | 1034 |
datatype foo = Foo | Bar foo |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1035 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1036 |
notepad |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1037 |
begin |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1038 |
fix x :: foo |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1039 |
have C |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1040 |
proof (cases x) |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1041 |
case Foo |
62271 | 1042 |
from \<open>x = Foo\<close> show ?thesis \<proof> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1043 |
next |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1044 |
case (Bar a) |
62271 | 1045 |
from \<open>x = Bar a\<close> show ?thesis \<proof> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1046 |
qed |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1047 |
end |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1048 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1049 |
|
62273 | 1050 |
subsection \<open>Elimination statements and case-splitting\<close> |
1051 |
||
1052 |
text \<open> |
|
1053 |
The @{command consider} states rules for generalized elimination and case |
|
1054 |
splitting. This is like a toplevel statement \<^theory_text>\<open>theorem obtains\<close> used within |
|
1055 |
a proof body; or like a multi-branch \<^theory_text>\<open>obtain\<close> without activation of the |
|
1056 |
local context elements yet. |
|
1057 |
||
1058 |
The proof method @{method cases} is able to use such rules with |
|
1059 |
forward-chaining (e.g.\ via \<^theory_text>\<open>then\<close>). This leads to the subsequent pattern |
|
1060 |
for case-splitting in a particular situation within a proof. |
|
1061 |
\<close> |
|
1062 |
||
1063 |
notepad |
|
1064 |
begin |
|
1065 |
consider (a) A | (b) B | (c) C |
|
1066 |
\<proof> \<comment> \<open>typically \<^theory_text>\<open>by auto\<close>, \<^theory_text>\<open>by blast\<close> etc.\<close> |
|
1067 |
then have something |
|
1068 |
proof cases |
|
1069 |
case a |
|
1070 |
then show ?thesis \<proof> |
|
1071 |
next |
|
1072 |
case b |
|
1073 |
then show ?thesis \<proof> |
|
1074 |
next |
|
1075 |
case c |
|
1076 |
then show ?thesis \<proof> |
|
1077 |
qed |
|
1078 |
end |
|
1079 |
||
58618 | 1080 |
subsection \<open>Obtaining local contexts\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1081 |
|
58618 | 1082 |
text \<open>A single ``case'' branch may be inlined into Isar proof text |
69597 | 1083 |
via @{command obtain}. This proves \<^prop>\<open>(\<And>x. B x \<Longrightarrow> thesis) \<Longrightarrow> |
1084 |
thesis\<close> on the spot, and augments the context afterwards.\<close> |
|
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1085 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1086 |
notepad |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1087 |
begin |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1088 |
fix B :: "'a \<Rightarrow> bool" |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1089 |
|
62271 | 1090 |
obtain x where "B x" \<proof> |
58618 | 1091 |
note \<open>B x\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1092 |
|
69597 | 1093 |
txt \<open>Conclusions from this context may not mention \<^term>\<open>x\<close> again!\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1094 |
{ |
62271 | 1095 |
obtain x where "B x" \<proof> |
1096 |
from \<open>B x\<close> have C \<proof> |
|
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1097 |
} |
58618 | 1098 |
note \<open>C\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1099 |
end |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1100 |
|
45103 | 1101 |
end |