author | wenzelm |
Fri, 01 Jan 2016 14:44:52 +0100 | |
changeset 62026 | ea3b1b0413b4 |
parent 61656 | cfabbc083977 |
child 62271 | 4cfe65cfd369 |
permissions | -rw-r--r-- |
61656 | 1 |
(*:maxLineLen=78:*) |
2 |
||
42917 | 3 |
theory Synopsis |
4 |
imports Base Main |
|
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> |
42918 | 29 |
def x \<equiv> "t::'a" |
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> |
42917 | 56 |
have b: B sorry |
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 |
{ |
|
159 |
have a: A sorry |
|
160 |
have b: B sorry |
|
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 |
{ |
|
177 |
have a: A sorry |
|
178 |
next |
|
179 |
have b: B |
|
180 |
proof - |
|
181 |
show B sorry |
|
182 |
next |
|
183 |
have c: C sorry |
|
184 |
next |
|
185 |
have d: D sorry |
|
186 |
qed |
|
187 |
} |
|
188 |
||
58618 | 189 |
txt \<open>Alternative version with explicit parentheses everywhere:\<close> |
42917 | 190 |
|
191 |
{ |
|
192 |
{ |
|
193 |
have a: A sorry |
|
194 |
} |
|
195 |
{ |
|
196 |
have b: B |
|
197 |
proof - |
|
198 |
{ |
|
199 |
show B sorry |
|
200 |
} |
|
201 |
{ |
|
202 |
have c: C sorry |
|
203 |
} |
|
204 |
{ |
|
205 |
have d: D sorry |
|
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> |
42919 | 217 |
For example, see @{file "~~/src/HOL/Isar_Examples/Group.thy"}. |
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 |
|
238 |
show ?thesis sorry |
|
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> |
42919 | 263 |
have "a = b" sorry |
264 |
also |
|
265 |
have "b = c" sorry |
|
266 |
also |
|
267 |
have "c = d" sorry |
|
268 |
finally |
|
269 |
have "a = d" . |
|
270 |
||
61493 | 271 |
txt \<open>Variant using the \<open>\<dots>\<close> abbreviation:\<close> |
42919 | 272 |
have "a = b" sorry |
273 |
also |
|
274 |
have "\<dots> = c" sorry |
|
275 |
also |
|
276 |
have "\<dots> = d" sorry |
|
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 - |
|
283 |
have "a = b" sorry |
|
284 |
also |
|
285 |
have "\<dots> = c" sorry |
|
286 |
also |
|
287 |
have "\<dots> = d" sorry |
|
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 |
||
295 |
have "a < b" sorry |
|
296 |
also |
|
45814 | 297 |
have "b \<le> c" sorry |
42919 | 298 |
also |
299 |
have "c = d" sorry |
|
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 |
||
58618 | 316 |
subsection \<open>Degenerate calculations and bigstep reasoning\<close> |
42919 | 317 |
|
61493 | 318 |
text \<open>The Idea is to append \<open>this\<close> to \<open>calculation\<close>, |
58618 | 319 |
without rule composition.\<close> |
42919 | 320 |
|
321 |
notepad |
|
322 |
begin |
|
58618 | 323 |
txt \<open>A vacuous proof:\<close> |
42919 | 324 |
have A sorry |
325 |
moreover |
|
326 |
have B sorry |
|
327 |
moreover |
|
328 |
have C sorry |
|
329 |
ultimately |
|
330 |
have A and B and C . |
|
331 |
next |
|
58618 | 332 |
txt \<open>Slightly more content (trivial bigstep reasoning):\<close> |
42919 | 333 |
have A sorry |
334 |
moreover |
|
335 |
have B sorry |
|
336 |
moreover |
|
337 |
have C sorry |
|
338 |
ultimately |
|
339 |
have "A \<and> B \<and> C" by blast |
|
340 |
next |
|
58618 | 341 |
txt \<open>More ambitious bigstep reasoning involving structured results:\<close> |
42919 | 342 |
have "A \<or> B \<or> C" sorry |
343 |
moreover |
|
344 |
{ assume A have R sorry } |
|
345 |
moreover |
|
346 |
{ assume B have R sorry } |
|
347 |
moreover |
|
348 |
{ assume C have R sorry } |
|
349 |
ultimately |
|
61580 | 350 |
have R by blast \<comment> \<open>``big-bang integration'' of proof blocks (occasionally fragile)\<close> |
42919 | 351 |
end |
352 |
||
42920 | 353 |
|
58618 | 354 |
section \<open>Induction\<close> |
42921 | 355 |
|
58618 | 356 |
subsection \<open>Induction as Natural Deduction\<close> |
42921 | 357 |
|
58618 | 358 |
text \<open>In principle, induction is just a special case of Natural |
42921 | 359 |
Deduction (see also \secref{sec:natural-deduction-synopsis}). For |
58618 | 360 |
example:\<close> |
42921 | 361 |
|
362 |
thm nat.induct |
|
363 |
print_statement nat.induct |
|
364 |
||
365 |
notepad |
|
366 |
begin |
|
367 |
fix n :: nat |
|
368 |
have "P n" |
|
61580 | 369 |
proof (rule nat.induct) \<comment> \<open>fragile rule application!\<close> |
42921 | 370 |
show "P 0" sorry |
371 |
next |
|
372 |
fix n :: nat |
|
373 |
assume "P n" |
|
374 |
show "P (Suc n)" sorry |
|
375 |
qed |
|
376 |
end |
|
377 |
||
58618 | 378 |
text \<open> |
42921 | 379 |
In practice, much more proof infrastructure is required. |
380 |
||
381 |
The proof method @{method induct} provides: |
|
382 |
||
61421 | 383 |
\<^item> implicit rule selection and robust instantiation |
42921 | 384 |
|
61421 | 385 |
\<^item> context elements via symbolic case names |
42921 | 386 |
|
61421 | 387 |
\<^item> support for rule-structured induction statements, with local |
388 |
parameters, premises, etc. |
|
58618 | 389 |
\<close> |
42921 | 390 |
|
391 |
notepad |
|
392 |
begin |
|
393 |
fix n :: nat |
|
394 |
have "P n" |
|
395 |
proof (induct n) |
|
396 |
case 0 |
|
397 |
show ?case sorry |
|
398 |
next |
|
399 |
case (Suc n) |
|
400 |
from Suc.hyps show ?case sorry |
|
401 |
qed |
|
402 |
end |
|
403 |
||
404 |
||
58618 | 405 |
subsubsection \<open>Example\<close> |
42921 | 406 |
|
58618 | 407 |
text \<open> |
42921 | 408 |
The subsequent example combines the following proof patterns: |
409 |
||
61421 | 410 |
\<^item> outermost induction (over the datatype structure of natural |
42921 | 411 |
numbers), to decompose the proof problem in top-down manner |
412 |
||
61421 | 413 |
\<^item> calculational reasoning (\secref{sec:calculations-synopsis}) |
42921 | 414 |
to compose the result in each case |
415 |
||
61421 | 416 |
\<^item> solving local claims within the calculation by simplification |
58618 | 417 |
\<close> |
42921 | 418 |
|
419 |
lemma |
|
420 |
fixes n :: nat |
|
421 |
shows "(\<Sum>i=0..n. i) = n * (n + 1) div 2" |
|
422 |
proof (induct n) |
|
423 |
case 0 |
|
424 |
have "(\<Sum>i=0..0. i) = (0::nat)" by simp |
|
425 |
also have "\<dots> = 0 * (0 + 1) div 2" by simp |
|
426 |
finally show ?case . |
|
427 |
next |
|
428 |
case (Suc n) |
|
429 |
have "(\<Sum>i=0..Suc n. i) = (\<Sum>i=0..n. i) + (n + 1)" by simp |
|
430 |
also have "\<dots> = n * (n + 1) div 2 + (n + 1)" by (simp add: Suc.hyps) |
|
431 |
also have "\<dots> = (n * (n + 1) + 2 * (n + 1)) div 2" by simp |
|
432 |
also have "\<dots> = (Suc n * (Suc n + 1)) div 2" by simp |
|
433 |
finally show ?case . |
|
434 |
qed |
|
435 |
||
58618 | 436 |
text \<open>This demonstrates how induction proofs can be done without |
437 |
having to consider the raw Natural Deduction structure.\<close> |
|
42921 | 438 |
|
439 |
||
58618 | 440 |
subsection \<open>Induction with local parameters and premises\<close> |
42921 | 441 |
|
58618 | 442 |
text \<open>Idea: Pure rule statements are passed through the induction |
42921 | 443 |
rule. This achieves convenient proof patterns, thanks to some |
444 |
internal trickery in the @{method induct} method. |
|
445 |
||
61493 | 446 |
Important: Using compact HOL formulae with \<open>\<forall>/\<longrightarrow>\<close> is a |
42921 | 447 |
well-known anti-pattern! It would produce useless formal noise. |
58618 | 448 |
\<close> |
42921 | 449 |
|
450 |
notepad |
|
451 |
begin |
|
452 |
fix n :: nat |
|
453 |
fix P :: "nat \<Rightarrow> bool" |
|
454 |
fix Q :: "'a \<Rightarrow> nat \<Rightarrow> bool" |
|
455 |
||
456 |
have "P n" |
|
457 |
proof (induct n) |
|
458 |
case 0 |
|
459 |
show "P 0" sorry |
|
460 |
next |
|
461 |
case (Suc n) |
|
58618 | 462 |
from \<open>P n\<close> show "P (Suc n)" sorry |
42921 | 463 |
qed |
464 |
||
465 |
have "A n \<Longrightarrow> P n" |
|
466 |
proof (induct n) |
|
467 |
case 0 |
|
58618 | 468 |
from \<open>A 0\<close> show "P 0" sorry |
42921 | 469 |
next |
470 |
case (Suc n) |
|
58618 | 471 |
from \<open>A n \<Longrightarrow> P n\<close> |
472 |
and \<open>A (Suc n)\<close> show "P (Suc n)" sorry |
|
42921 | 473 |
qed |
474 |
||
475 |
have "\<And>x. Q x n" |
|
476 |
proof (induct n) |
|
477 |
case 0 |
|
478 |
show "Q x 0" sorry |
|
479 |
next |
|
480 |
case (Suc n) |
|
58618 | 481 |
from \<open>\<And>x. Q x n\<close> show "Q x (Suc n)" sorry |
482 |
txt \<open>Local quantification admits arbitrary instances:\<close> |
|
483 |
note \<open>Q a n\<close> and \<open>Q b n\<close> |
|
42921 | 484 |
qed |
485 |
end |
|
486 |
||
487 |
||
58618 | 488 |
subsection \<open>Implicit induction context\<close> |
42921 | 489 |
|
58618 | 490 |
text \<open>The @{method induct} method can isolate local parameters and |
42921 | 491 |
premises directly from the given statement. This is convenient in |
492 |
practical applications, but requires some understanding of what is |
|
58618 | 493 |
going on internally (as explained above).\<close> |
42921 | 494 |
|
495 |
notepad |
|
496 |
begin |
|
497 |
fix n :: nat |
|
498 |
fix Q :: "'a \<Rightarrow> nat \<Rightarrow> bool" |
|
499 |
||
500 |
fix x :: 'a |
|
501 |
assume "A x n" |
|
502 |
then have "Q x n" |
|
503 |
proof (induct n arbitrary: x) |
|
504 |
case 0 |
|
58618 | 505 |
from \<open>A x 0\<close> show "Q x 0" sorry |
42921 | 506 |
next |
507 |
case (Suc n) |
|
61580 | 508 |
from \<open>\<And>x. A x n \<Longrightarrow> Q x n\<close> \<comment> \<open>arbitrary instances can be produced here\<close> |
58618 | 509 |
and \<open>A x (Suc n)\<close> show "Q x (Suc n)" sorry |
42921 | 510 |
qed |
511 |
end |
|
512 |
||
513 |
||
58618 | 514 |
subsection \<open>Advanced induction with term definitions\<close> |
42921 | 515 |
|
58618 | 516 |
text \<open>Induction over subexpressions of a certain shape are delicate |
42921 | 517 |
to formalize. The Isar @{method induct} method provides |
518 |
infrastructure for this. |
|
519 |
||
520 |
Idea: sub-expressions of the problem are turned into a defined |
|
521 |
induction variable; often accompanied with fixing of auxiliary |
|
58618 | 522 |
parameters in the original expression.\<close> |
42921 | 523 |
|
524 |
notepad |
|
525 |
begin |
|
526 |
fix a :: "'a \<Rightarrow> nat" |
|
527 |
fix A :: "nat \<Rightarrow> bool" |
|
528 |
||
529 |
assume "A (a x)" |
|
530 |
then have "P (a x)" |
|
531 |
proof (induct "a x" arbitrary: x) |
|
532 |
case 0 |
|
58618 | 533 |
note prem = \<open>A (a x)\<close> |
534 |
and defn = \<open>0 = a x\<close> |
|
42921 | 535 |
show "P (a x)" sorry |
536 |
next |
|
537 |
case (Suc n) |
|
58618 | 538 |
note hyp = \<open>\<And>x. n = a x \<Longrightarrow> A (a x) \<Longrightarrow> P (a x)\<close> |
539 |
and prem = \<open>A (a x)\<close> |
|
540 |
and defn = \<open>Suc n = a x\<close> |
|
42921 | 541 |
show "P (a x)" sorry |
542 |
qed |
|
543 |
end |
|
544 |
||
545 |
||
58618 | 546 |
section \<open>Natural Deduction \label{sec:natural-deduction-synopsis}\<close> |
42920 | 547 |
|
58618 | 548 |
subsection \<open>Rule statements\<close> |
42920 | 549 |
|
58618 | 550 |
text \<open> |
42920 | 551 |
Isabelle/Pure ``theorems'' are always natural deduction rules, |
552 |
which sometimes happen to consist of a conclusion only. |
|
553 |
||
61493 | 554 |
The framework connectives \<open>\<And>\<close> and \<open>\<Longrightarrow>\<close> indicate the |
58618 | 555 |
rule structure declaratively. For example:\<close> |
42920 | 556 |
|
557 |
thm conjI |
|
558 |
thm impI |
|
559 |
thm nat.induct |
|
560 |
||
58618 | 561 |
text \<open> |
42920 | 562 |
The object-logic is embedded into the Pure framework via an implicit |
563 |
derivability judgment @{term "Trueprop :: bool \<Rightarrow> prop"}. |
|
564 |
||
565 |
Thus any HOL formulae appears atomic to the Pure framework, while |
|
566 |
the rule structure outlines the corresponding proof pattern. |
|
567 |
||
568 |
This can be made explicit as follows: |
|
58618 | 569 |
\<close> |
42920 | 570 |
|
571 |
notepad |
|
572 |
begin |
|
573 |
write Trueprop ("Tr") |
|
574 |
||
575 |
thm conjI |
|
576 |
thm impI |
|
577 |
thm nat.induct |
|
578 |
end |
|
579 |
||
58618 | 580 |
text \<open> |
42920 | 581 |
Isar provides first-class notation for rule statements as follows. |
58618 | 582 |
\<close> |
42920 | 583 |
|
584 |
print_statement conjI |
|
585 |
print_statement impI |
|
586 |
print_statement nat.induct |
|
587 |
||
588 |
||
58618 | 589 |
subsubsection \<open>Examples\<close> |
42920 | 590 |
|
58618 | 591 |
text \<open> |
42920 | 592 |
Introductions and eliminations of some standard connectives of |
593 |
the object-logic can be written as rule statements as follows. (The |
|
594 |
proof ``@{command "by"}~@{method blast}'' serves as sanity check.) |
|
58618 | 595 |
\<close> |
42920 | 596 |
|
597 |
lemma "(P \<Longrightarrow> False) \<Longrightarrow> \<not> P" by blast |
|
598 |
lemma "\<not> P \<Longrightarrow> P \<Longrightarrow> Q" by blast |
|
599 |
||
600 |
lemma "P \<Longrightarrow> Q \<Longrightarrow> P \<and> Q" by blast |
|
601 |
lemma "P \<and> Q \<Longrightarrow> (P \<Longrightarrow> Q \<Longrightarrow> R) \<Longrightarrow> R" by blast |
|
602 |
||
603 |
lemma "P \<Longrightarrow> P \<or> Q" by blast |
|
604 |
lemma "Q \<Longrightarrow> P \<or> Q" by blast |
|
605 |
lemma "P \<or> Q \<Longrightarrow> (P \<Longrightarrow> R) \<Longrightarrow> (Q \<Longrightarrow> R) \<Longrightarrow> R" by blast |
|
606 |
||
607 |
lemma "(\<And>x. P x) \<Longrightarrow> (\<forall>x. P x)" by blast |
|
608 |
lemma "(\<forall>x. P x) \<Longrightarrow> P x" by blast |
|
609 |
||
610 |
lemma "P x \<Longrightarrow> (\<exists>x. P x)" by blast |
|
611 |
lemma "(\<exists>x. P x) \<Longrightarrow> (\<And>x. P x \<Longrightarrow> R) \<Longrightarrow> R" by blast |
|
612 |
||
613 |
lemma "x \<in> A \<Longrightarrow> x \<in> B \<Longrightarrow> x \<in> A \<inter> B" by blast |
|
614 |
lemma "x \<in> A \<inter> B \<Longrightarrow> (x \<in> A \<Longrightarrow> x \<in> B \<Longrightarrow> R) \<Longrightarrow> R" by blast |
|
615 |
||
616 |
lemma "x \<in> A \<Longrightarrow> x \<in> A \<union> B" by blast |
|
617 |
lemma "x \<in> B \<Longrightarrow> x \<in> A \<union> B" by blast |
|
618 |
lemma "x \<in> A \<union> B \<Longrightarrow> (x \<in> A \<Longrightarrow> R) \<Longrightarrow> (x \<in> B \<Longrightarrow> R) \<Longrightarrow> R" by blast |
|
619 |
||
620 |
||
58618 | 621 |
subsection \<open>Isar context elements\<close> |
42920 | 622 |
|
58618 | 623 |
text \<open>We derive some results out of the blue, using Isar context |
42920 | 624 |
elements and some explicit blocks. This illustrates their meaning |
58618 | 625 |
wrt.\ Pure connectives, without goal states getting in the way.\<close> |
42920 | 626 |
|
627 |
notepad |
|
628 |
begin |
|
629 |
{ |
|
630 |
fix x |
|
631 |
have "B x" sorry |
|
632 |
} |
|
633 |
have "\<And>x. B x" by fact |
|
634 |
||
635 |
next |
|
636 |
||
637 |
{ |
|
638 |
assume A |
|
639 |
have B sorry |
|
640 |
} |
|
641 |
have "A \<Longrightarrow> B" by fact |
|
642 |
||
643 |
next |
|
644 |
||
645 |
{ |
|
646 |
def x \<equiv> t |
|
647 |
have "B x" sorry |
|
648 |
} |
|
649 |
have "B t" by fact |
|
650 |
||
651 |
next |
|
652 |
||
653 |
{ |
|
654 |
obtain x :: 'a where "B x" sorry |
|
655 |
have C sorry |
|
656 |
} |
|
657 |
have C by fact |
|
658 |
||
659 |
end |
|
660 |
||
661 |
||
58618 | 662 |
subsection \<open>Pure rule composition\<close> |
42920 | 663 |
|
58618 | 664 |
text \<open> |
42920 | 665 |
The Pure framework provides means for: |
666 |
||
61421 | 667 |
\<^item> backward-chaining of rules by @{inference resolution} |
42920 | 668 |
|
61421 | 669 |
\<^item> closing of branches by @{inference assumption} |
42920 | 670 |
|
671 |
||
61493 | 672 |
Both principles involve higher-order unification of \<open>\<lambda>\<close>-terms |
673 |
modulo \<open>\<alpha>\<beta>\<eta>\<close>-equivalence (cf.\ Huet and Miller). |
|
61458 | 674 |
\<close> |
42920 | 675 |
|
676 |
notepad |
|
677 |
begin |
|
678 |
assume a: A and b: B |
|
679 |
thm conjI |
|
61580 | 680 |
thm conjI [of A B] \<comment> "instantiation" |
681 |
thm conjI [of A B, OF a b] \<comment> "instantiation and composition" |
|
682 |
thm conjI [OF a b] \<comment> "composition via unification (trivial)" |
|
58618 | 683 |
thm conjI [OF \<open>A\<close> \<open>B\<close>] |
42920 | 684 |
|
685 |
thm conjI [OF disjI1] |
|
686 |
end |
|
687 |
||
58618 | 688 |
text \<open>Note: Low-level rule composition is tedious and leads to |
689 |
unreadable~/ unmaintainable expressions in the text.\<close> |
|
42920 | 690 |
|
691 |
||
58618 | 692 |
subsection \<open>Structured backward reasoning\<close> |
42920 | 693 |
|
58618 | 694 |
text \<open>Idea: Canonical proof decomposition via @{command fix}~/ |
42920 | 695 |
@{command assume}~/ @{command show}, where the body produces a |
58618 | 696 |
natural deduction rule to refine some goal.\<close> |
42920 | 697 |
|
698 |
notepad |
|
699 |
begin |
|
700 |
fix A B :: "'a \<Rightarrow> bool" |
|
701 |
||
702 |
have "\<And>x. A x \<Longrightarrow> B x" |
|
703 |
proof - |
|
704 |
fix x |
|
705 |
assume "A x" |
|
706 |
show "B x" sorry |
|
707 |
qed |
|
708 |
||
709 |
have "\<And>x. A x \<Longrightarrow> B x" |
|
710 |
proof - |
|
711 |
{ |
|
712 |
fix x |
|
713 |
assume "A x" |
|
714 |
show "B x" sorry |
|
61580 | 715 |
} \<comment> "implicit block structure made explicit" |
58618 | 716 |
note \<open>\<And>x. A x \<Longrightarrow> B x\<close> |
61580 | 717 |
\<comment> "side exit for the resulting rule" |
42920 | 718 |
qed |
719 |
end |
|
720 |
||
721 |
||
58618 | 722 |
subsection \<open>Structured rule application\<close> |
42920 | 723 |
|
58618 | 724 |
text \<open> |
42920 | 725 |
Idea: Previous facts and new claims are composed with a rule from |
726 |
the context (or background library). |
|
58618 | 727 |
\<close> |
42920 | 728 |
|
729 |
notepad |
|
730 |
begin |
|
61580 | 731 |
assume r1: "A \<Longrightarrow> B \<Longrightarrow> C" \<comment> \<open>simple rule (Horn clause)\<close> |
42920 | 732 |
|
61580 | 733 |
have A sorry \<comment> "prefix of facts via outer sub-proof" |
42920 | 734 |
then have C |
735 |
proof (rule r1) |
|
61580 | 736 |
show B sorry \<comment> "remaining rule premises via inner sub-proof" |
42920 | 737 |
qed |
738 |
||
739 |
have C |
|
740 |
proof (rule r1) |
|
741 |
show A sorry |
|
742 |
show B sorry |
|
743 |
qed |
|
744 |
||
745 |
have A and B sorry |
|
746 |
then have C |
|
747 |
proof (rule r1) |
|
748 |
qed |
|
749 |
||
750 |
have A and B sorry |
|
751 |
then have C by (rule r1) |
|
752 |
||
753 |
next |
|
754 |
||
61580 | 755 |
assume r2: "A \<Longrightarrow> (\<And>x. B1 x \<Longrightarrow> B2 x) \<Longrightarrow> C" \<comment> \<open>nested rule\<close> |
42920 | 756 |
|
757 |
have A sorry |
|
758 |
then have C |
|
759 |
proof (rule r2) |
|
760 |
fix x |
|
761 |
assume "B1 x" |
|
762 |
show "B2 x" sorry |
|
763 |
qed |
|
764 |
||
58618 | 765 |
txt \<open>The compound rule premise @{prop "\<And>x. B1 x \<Longrightarrow> B2 x"} is better |
42920 | 766 |
addressed via @{command fix}~/ @{command assume}~/ @{command show} |
58618 | 767 |
in the nested proof body.\<close> |
42920 | 768 |
end |
769 |
||
770 |
||
58618 | 771 |
subsection \<open>Example: predicate logic\<close> |
42920 | 772 |
|
58618 | 773 |
text \<open> |
42920 | 774 |
Using the above principles, standard introduction and elimination proofs |
775 |
of predicate logic connectives of HOL work as follows. |
|
58618 | 776 |
\<close> |
42920 | 777 |
|
778 |
notepad |
|
779 |
begin |
|
780 |
have "A \<longrightarrow> B" and A sorry |
|
781 |
then have B .. |
|
782 |
||
783 |
have A sorry |
|
784 |
then have "A \<or> B" .. |
|
785 |
||
786 |
have B sorry |
|
787 |
then have "A \<or> B" .. |
|
788 |
||
789 |
have "A \<or> B" sorry |
|
790 |
then have C |
|
791 |
proof |
|
792 |
assume A |
|
793 |
then show C sorry |
|
794 |
next |
|
795 |
assume B |
|
796 |
then show C sorry |
|
797 |
qed |
|
798 |
||
799 |
have A and B sorry |
|
800 |
then have "A \<and> B" .. |
|
801 |
||
802 |
have "A \<and> B" sorry |
|
803 |
then have A .. |
|
804 |
||
805 |
have "A \<and> B" sorry |
|
806 |
then have B .. |
|
807 |
||
808 |
have False sorry |
|
809 |
then have A .. |
|
810 |
||
811 |
have True .. |
|
812 |
||
813 |
have "\<not> A" |
|
814 |
proof |
|
815 |
assume A |
|
816 |
then show False sorry |
|
817 |
qed |
|
818 |
||
819 |
have "\<not> A" and A sorry |
|
820 |
then have B .. |
|
821 |
||
822 |
have "\<forall>x. P x" |
|
823 |
proof |
|
824 |
fix x |
|
825 |
show "P x" sorry |
|
826 |
qed |
|
827 |
||
828 |
have "\<forall>x. P x" sorry |
|
829 |
then have "P a" .. |
|
830 |
||
831 |
have "\<exists>x. P x" |
|
832 |
proof |
|
833 |
show "P a" sorry |
|
834 |
qed |
|
835 |
||
836 |
have "\<exists>x. P x" sorry |
|
837 |
then have C |
|
838 |
proof |
|
839 |
fix a |
|
840 |
assume "P a" |
|
841 |
show C sorry |
|
842 |
qed |
|
843 |
||
58618 | 844 |
txt \<open>Less awkward version using @{command obtain}:\<close> |
42920 | 845 |
have "\<exists>x. P x" sorry |
846 |
then obtain a where "P a" .. |
|
847 |
end |
|
848 |
||
58618 | 849 |
text \<open>Further variations to illustrate Isar sub-proofs involving |
850 |
@{command show}:\<close> |
|
42920 | 851 |
|
852 |
notepad |
|
853 |
begin |
|
854 |
have "A \<and> B" |
|
61580 | 855 |
proof \<comment> \<open>two strictly isolated subproofs\<close> |
42920 | 856 |
show A sorry |
857 |
next |
|
858 |
show B sorry |
|
859 |
qed |
|
860 |
||
861 |
have "A \<and> B" |
|
61580 | 862 |
proof \<comment> \<open>one simultaneous sub-proof\<close> |
42920 | 863 |
show A and B sorry |
864 |
qed |
|
865 |
||
866 |
have "A \<and> B" |
|
61580 | 867 |
proof \<comment> \<open>two subproofs in the same context\<close> |
42920 | 868 |
show A sorry |
869 |
show B sorry |
|
870 |
qed |
|
871 |
||
872 |
have "A \<and> B" |
|
61580 | 873 |
proof \<comment> \<open>swapped order\<close> |
42920 | 874 |
show B sorry |
875 |
show A sorry |
|
876 |
qed |
|
877 |
||
878 |
have "A \<and> B" |
|
61580 | 879 |
proof \<comment> \<open>sequential subproofs\<close> |
42920 | 880 |
show A sorry |
58618 | 881 |
show B using \<open>A\<close> sorry |
42920 | 882 |
qed |
883 |
end |
|
884 |
||
885 |
||
58618 | 886 |
subsubsection \<open>Example: set-theoretic operators\<close> |
42920 | 887 |
|
61493 | 888 |
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 | 889 |
set-theory or lattice-theory work analogously. It is only a matter |
42920 | 890 |
of rule declarations in the library; rules can be also specified |
891 |
explicitly. |
|
58618 | 892 |
\<close> |
42920 | 893 |
|
894 |
notepad |
|
895 |
begin |
|
896 |
have "x \<in> A" and "x \<in> B" sorry |
|
897 |
then have "x \<in> A \<inter> B" .. |
|
898 |
||
899 |
have "x \<in> A" sorry |
|
900 |
then have "x \<in> A \<union> B" .. |
|
901 |
||
902 |
have "x \<in> B" sorry |
|
903 |
then have "x \<in> A \<union> B" .. |
|
904 |
||
905 |
have "x \<in> A \<union> B" sorry |
|
906 |
then have C |
|
907 |
proof |
|
908 |
assume "x \<in> A" |
|
909 |
then show C sorry |
|
910 |
next |
|
911 |
assume "x \<in> B" |
|
912 |
then show C sorry |
|
913 |
qed |
|
914 |
||
915 |
next |
|
916 |
have "x \<in> \<Inter>A" |
|
917 |
proof |
|
918 |
fix a |
|
919 |
assume "a \<in> A" |
|
920 |
show "x \<in> a" sorry |
|
921 |
qed |
|
922 |
||
923 |
have "x \<in> \<Inter>A" sorry |
|
924 |
then have "x \<in> a" |
|
925 |
proof |
|
926 |
show "a \<in> A" sorry |
|
927 |
qed |
|
928 |
||
929 |
have "a \<in> A" and "x \<in> a" sorry |
|
930 |
then have "x \<in> \<Union>A" .. |
|
931 |
||
932 |
have "x \<in> \<Union>A" sorry |
|
933 |
then obtain a where "a \<in> A" and "x \<in> a" .. |
|
934 |
end |
|
935 |
||
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
936 |
|
58618 | 937 |
section \<open>Generalized elimination and cases\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
938 |
|
58618 | 939 |
subsection \<open>General elimination rules\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
940 |
|
58618 | 941 |
text \<open> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
942 |
The general format of elimination rules is illustrated by the |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
943 |
following typical representatives: |
58618 | 944 |
\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
945 |
|
61580 | 946 |
thm exE \<comment> \<open>local parameter\<close> |
947 |
thm conjE \<comment> \<open>local premises\<close> |
|
948 |
thm disjE \<comment> \<open>split into cases\<close> |
|
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
949 |
|
58618 | 950 |
text \<open> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
951 |
Combining these characteristics leads to the following general scheme |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
952 |
for elimination rules with cases: |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
953 |
|
61421 | 954 |
\<^item> prefix of assumptions (or ``major premises'') |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
955 |
|
61421 | 956 |
\<^item> one or more cases that enable to establish the main conclusion |
957 |
in an augmented context |
|
58618 | 958 |
\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
959 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
960 |
notepad |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
961 |
begin |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
962 |
assume r: |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
963 |
"A1 \<Longrightarrow> A2 \<Longrightarrow> (* assumptions *) |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
964 |
(\<And>x y. B1 x y \<Longrightarrow> C1 x y \<Longrightarrow> R) \<Longrightarrow> (* case 1 *) |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
965 |
(\<And>x y. B2 x y \<Longrightarrow> C2 x y \<Longrightarrow> R) \<Longrightarrow> (* case 2 *) |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
966 |
R (* main conclusion *)" |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
967 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
968 |
have A1 and A2 sorry |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
969 |
then have R |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
970 |
proof (rule r) |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
971 |
fix x y |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
972 |
assume "B1 x y" and "C1 x y" |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
973 |
show ?thesis sorry |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
974 |
next |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
975 |
fix x y |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
976 |
assume "B2 x y" and "C2 x y" |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
977 |
show ?thesis sorry |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
978 |
qed |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
979 |
end |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
980 |
|
61493 | 981 |
text \<open>Here \<open>?thesis\<close> is used to refer to the unchanged goal |
58618 | 982 |
statement.\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
983 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
984 |
|
58618 | 985 |
subsection \<open>Rules with cases\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
986 |
|
58618 | 987 |
text \<open> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
988 |
Applying an elimination rule to some goal, leaves that unchanged |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
989 |
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
|
990 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
991 |
Isar provides some infrastructure to support this: |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
992 |
|
61421 | 993 |
\<^item> native language elements to state eliminations |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
994 |
|
61421 | 995 |
\<^item> symbolic case names |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
996 |
|
61421 | 997 |
\<^item> method @{method cases} to recover this structure in a |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
998 |
sub-proof |
58618 | 999 |
\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1000 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1001 |
print_statement exE |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1002 |
print_statement conjE |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1003 |
print_statement disjE |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1004 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1005 |
lemma |
61580 | 1006 |
assumes A1 and A2 \<comment> \<open>assumptions\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1007 |
obtains |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1008 |
(case1) x y where "B1 x y" and "C1 x y" |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1009 |
| (case2) x y where "B2 x y" and "C2 x y" |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1010 |
sorry |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1011 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1012 |
|
58618 | 1013 |
subsubsection \<open>Example\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1014 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1015 |
lemma tertium_non_datur: |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1016 |
obtains |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1017 |
(T) A |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1018 |
| (F) "\<not> A" |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1019 |
by blast |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1020 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1021 |
notepad |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1022 |
begin |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1023 |
fix x y :: 'a |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1024 |
have C |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1025 |
proof (cases "x = y" rule: tertium_non_datur) |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1026 |
case T |
58618 | 1027 |
from \<open>x = y\<close> show ?thesis sorry |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1028 |
next |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1029 |
case F |
58618 | 1030 |
from \<open>x \<noteq> y\<close> show ?thesis sorry |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1031 |
qed |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1032 |
end |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1033 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1034 |
|
58618 | 1035 |
subsubsection \<open>Example\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1036 |
|
58618 | 1037 |
text \<open> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1038 |
Isabelle/HOL specification mechanisms (datatype, inductive, etc.) |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1039 |
provide suitable derived cases rules. |
58618 | 1040 |
\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1041 |
|
58310 | 1042 |
datatype foo = Foo | Bar foo |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1043 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1044 |
notepad |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1045 |
begin |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1046 |
fix x :: foo |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1047 |
have C |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1048 |
proof (cases x) |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1049 |
case Foo |
58618 | 1050 |
from \<open>x = Foo\<close> show ?thesis sorry |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1051 |
next |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1052 |
case (Bar a) |
58618 | 1053 |
from \<open>x = Bar a\<close> show ?thesis sorry |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1054 |
qed |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1055 |
end |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1056 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1057 |
|
58618 | 1058 |
subsection \<open>Obtaining local contexts\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1059 |
|
58618 | 1060 |
text \<open>A single ``case'' branch may be inlined into Isar proof text |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1061 |
via @{command obtain}. This proves @{prop "(\<And>x. B x \<Longrightarrow> thesis) \<Longrightarrow> |
58618 | 1062 |
thesis"} on the spot, and augments the context afterwards.\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1063 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1064 |
notepad |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1065 |
begin |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1066 |
fix B :: "'a \<Rightarrow> bool" |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1067 |
|
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1068 |
obtain x where "B x" sorry |
58618 | 1069 |
note \<open>B x\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1070 |
|
58618 | 1071 |
txt \<open>Conclusions from this context may not mention @{term x} again!\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1072 |
{ |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1073 |
obtain x where "B x" sorry |
58618 | 1074 |
from \<open>B x\<close> have C sorry |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1075 |
} |
58618 | 1076 |
note \<open>C\<close> |
42922
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1077 |
end |
91e229959d4c
some material on "Generalized elimination and cases";
wenzelm
parents:
42921
diff
changeset
|
1078 |
|
45103 | 1079 |
end |