author | wenzelm |
Fri, 06 Mar 2015 15:58:56 +0100 | |
changeset 59621 | 291934bac95e |
parent 58882 | 6e2010ab8bd9 |
child 61541 | 846c72206207 |
permissions | -rw-r--r-- |
33026 | 1 |
(* Title: HOL/Isar_Examples/Cantor.thy |
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
2 |
Author: Markus Wenzel, TU Muenchen |
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
3 |
*) |
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
4 |
|
58882 | 5 |
section \<open>Cantor's Theorem\<close> |
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
6 |
|
31758 | 7 |
theory Cantor |
8 |
imports Main |
|
9 |
begin |
|
7955 | 10 |
|
58614 | 11 |
text_raw \<open>\footnote{This is an Isar version of the final example of |
12 |
the Isabelle/HOL manual @{cite "isabelle-HOL"}.}\<close> |
|
7819 | 13 |
|
58614 | 14 |
text \<open>Cantor's Theorem states that every set has more subsets than |
37671 | 15 |
it has elements. It has become a favorite basic example in pure |
12388 | 16 |
higher-order logic since it is so easily expressed: \[\all{f::\alpha |
17 |
\To \alpha \To \idt{bool}} \ex{S::\alpha \To \idt{bool}} |
|
18 |
\all{x::\alpha} f \ap x \not= S\] |
|
19 |
||
20 |
Viewing types as sets, $\alpha \To \idt{bool}$ represents the |
|
21 |
powerset of $\alpha$. This version of the theorem states that for |
|
22 |
every function from $\alpha$ to its powerset, some subset is outside |
|
23 |
its range. The Isabelle/Isar proofs below uses HOL's set theory, |
|
24 |
with the type $\alpha \ap \idt{set}$ and the operator |
|
58614 | 25 |
$\idt{range}::(\alpha \To \beta) \To \beta \ap \idt{set}$.\<close> |
6505 | 26 |
|
55640 | 27 |
theorem "\<exists>S. S \<notin> range (f :: 'a \<Rightarrow> 'a set)" |
10007 | 28 |
proof |
55640 | 29 |
let ?S = "{x. x \<notin> f x}" |
30 |
show "?S \<notin> range f" |
|
10007 | 31 |
proof |
55640 | 32 |
assume "?S \<in> range f" |
12388 | 33 |
then obtain y where "?S = f y" .. |
23373 | 34 |
then show False |
12388 | 35 |
proof (rule equalityCE) |
55640 | 36 |
assume "y \<in> f y" |
37 |
assume "y \<in> ?S" |
|
38 |
then have "y \<notin> f y" .. |
|
58614 | 39 |
with \<open>y : f y\<close> show ?thesis by contradiction |
12388 | 40 |
next |
55640 | 41 |
assume "y \<notin> ?S" |
42 |
assume "y \<notin> f y" |
|
43 |
then have "y \<in> ?S" .. |
|
58614 | 44 |
with \<open>y \<notin> ?S\<close> show ?thesis by contradiction |
10007 | 45 |
qed |
46 |
qed |
|
47 |
qed |
|
6494 | 48 |
|
58614 | 49 |
text \<open>How much creativity is required? As it happens, Isabelle can |
37671 | 50 |
prove this theorem automatically using best-first search. |
51 |
Depth-first search would diverge, but best-first search successfully |
|
52 |
navigates through the large search space. The context of Isabelle's |
|
53 |
classical prover contains rules for the relevant constructs of HOL's |
|
58614 | 54 |
set theory.\<close> |
6505 | 55 |
|
55640 | 56 |
theorem "\<exists>S. S \<notin> range (f :: 'a \<Rightarrow> 'a set)" |
10007 | 57 |
by best |
6494 | 58 |
|
58614 | 59 |
text \<open>While this establishes the same theorem internally, we do not |
37671 | 60 |
get any idea of how the proof actually works. There is currently no |
61 |
way to transform internal system-level representations of Isabelle |
|
12388 | 62 |
proofs back into Isar text. Writing intelligible proof documents |
58614 | 63 |
really is a creative process, after all.\<close> |
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
64 |
|
10007 | 65 |
end |