| author | wenzelm |
| Tue, 10 Nov 2015 23:41:20 +0100 | |
| changeset 61626 | c304402cc3df |
| parent 61541 | 846c72206207 |
| child 61930 | 380cbe15cca5 |
| 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 |
| 61541 | 15 |
it has elements. It has become a favourite basic example in pure |
16 |
higher-order logic since it is so easily expressed: |
|
17 |
||
18 |
@{text [display]
|
|
19 |
\<open>\<forall>f::\<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool. \<exists>S::\<alpha> \<Rightarrow> bool. \<forall>x::\<alpha>. f x \<noteq> S\<close>} |
|
| 12388 | 20 |
|
| 61541 | 21 |
Viewing types as sets, \<open>\<alpha> \<Rightarrow> bool\<close> represents the powerset of \<open>\<alpha>\<close>. This |
22 |
version of the theorem states that for every function from \<open>\<alpha>\<close> to its |
|
23 |
powerset, some subset is outside its range. The Isabelle/Isar proofs below |
|
24 |
uses HOL's set theory, with the type \<open>\<alpha> set\<close> and the operator |
|
25 |
\<open>range :: (\<alpha> \<Rightarrow> \<beta>) \<Rightarrow> \<beta> set\<close>.\<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 |
|
| 61541 | 49 |
text \<open>How much creativity is required? As it happens, Isabelle can prove |
50 |
this theorem automatically using best-first search. Depth-first search |
|
51 |
would diverge, but best-first search successfully navigates through the |
|
52 |
large search space. The context of Isabelle's classical prover contains |
|
53 |
rules for the relevant constructs of HOL's set theory.\<close> |
|
| 6505 | 54 |
|
| 55640 | 55 |
theorem "\<exists>S. S \<notin> range (f :: 'a \<Rightarrow> 'a set)" |
| 10007 | 56 |
by best |
| 6494 | 57 |
|
| 61541 | 58 |
text \<open>While this establishes the same theorem internally, we do not get any |
59 |
idea of how the proof actually works. There is currently no way to |
|
60 |
transform internal system-level representations of Isabelle proofs back |
|
61 |
into Isar text. Writing intelligible proof documents really is a creative |
|
62 |
process, after all.\<close> |
|
|
6444
2ebe9e630cab
Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff
changeset
|
63 |
|
| 10007 | 64 |
end |