| author | wenzelm | 
| Thu, 14 Dec 2017 21:15:04 +0100 | |
| changeset 67204 | 849a838f7e57 | 
| parent 63680 | 6e1e8b5abbfa | 
| child 69597 | ff784d5a5bfb | 
| permissions | -rw-r--r-- | 
| 33026 | 1  | 
(* Title: HOL/Isar_Examples/Cantor.thy  | 
| 61930 | 2  | 
Author: Makarius  | 
| 
6444
 
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  | 
| 63585 | 8  | 
imports Main  | 
| 31758 | 9  | 
begin  | 
| 7955 | 10  | 
|
| 61931 | 11  | 
subsection \<open>Mathematical statement and proof\<close>  | 
12  | 
||
| 61930 | 13  | 
text \<open>  | 
14  | 
Cantor's Theorem states that there is no surjection from  | 
|
15  | 
a set to its powerset. The proof works by diagonalization. E.g.\ see  | 
|
| 63680 | 16  | 
\<^item> \<^url>\<open>http://mathworld.wolfram.com/CantorDiagonalMethod.html\<close>  | 
17  | 
\<^item> \<^url>\<open>https://en.wikipedia.org/wiki/Cantor's_diagonal_argument\<close>  | 
|
| 61930 | 18  | 
\<close>  | 
19  | 
||
| 62523 | 20  | 
theorem Cantor: "\<nexists>f :: 'a \<Rightarrow> 'a set. \<forall>A. \<exists>x. A = f x"  | 
| 61930 | 21  | 
proof  | 
| 61931 | 22  | 
assume "\<exists>f :: 'a \<Rightarrow> 'a set. \<forall>A. \<exists>x. A = f x"  | 
23  | 
then obtain f :: "'a \<Rightarrow> 'a set" where *: "\<forall>A. \<exists>x. A = f x" ..  | 
|
| 61930 | 24  | 
  let ?D = "{x. x \<notin> f x}"
 | 
| 61931 | 25  | 
from * obtain a where "?D = f a" by blast  | 
| 61930 | 26  | 
moreover have "a \<in> ?D \<longleftrightarrow> a \<notin> f a" by blast  | 
27  | 
ultimately show False by blast  | 
|
28  | 
qed  | 
|
29  | 
||
| 7819 | 30  | 
|
| 61931 | 31  | 
subsection \<open>Automated proofs\<close>  | 
32  | 
||
| 61930 | 33  | 
text \<open>  | 
| 61931 | 34  | 
These automated proofs are much shorter, but lack information why and how it  | 
35  | 
works.  | 
|
| 61930 | 36  | 
\<close>  | 
37  | 
||
| 62523 | 38  | 
theorem "\<nexists>f :: 'a \<Rightarrow> 'a set. \<forall>A. \<exists>x. f x = A"  | 
| 61930 | 39  | 
by best  | 
40  | 
||
| 62523 | 41  | 
theorem "\<nexists>f :: 'a \<Rightarrow> 'a set. \<forall>A. \<exists>x. f x = A"  | 
| 61930 | 42  | 
by force  | 
43  | 
||
44  | 
||
| 61931 | 45  | 
subsection \<open>Elementary version in higher-order predicate logic\<close>  | 
46  | 
||
47  | 
text \<open>  | 
|
48  | 
The subsequent formulation bypasses set notation of HOL; it uses elementary  | 
|
49  | 
\<open>\<lambda>\<close>-calculus and predicate logic, with standard introduction and elimination  | 
|
50  | 
rules. This also shows that the proof does not require classical reasoning.  | 
|
51  | 
\<close>  | 
|
52  | 
||
53  | 
lemma iff_contradiction:  | 
|
54  | 
assumes *: "\<not> A \<longleftrightarrow> A"  | 
|
55  | 
shows False  | 
|
56  | 
proof (rule notE)  | 
|
57  | 
show "\<not> A"  | 
|
58  | 
proof  | 
|
59  | 
assume A  | 
|
60  | 
with * have "\<not> A" ..  | 
|
61  | 
from this and \<open>A\<close> show False ..  | 
|
62  | 
qed  | 
|
63  | 
with * show A ..  | 
|
64  | 
qed  | 
|
65  | 
||
| 62523 | 66  | 
theorem Cantor': "\<nexists>f :: 'a \<Rightarrow> 'a \<Rightarrow> bool. \<forall>A. \<exists>x. A = f x"  | 
| 61931 | 67  | 
proof  | 
| 62038 | 68  | 
assume "\<exists>f :: 'a \<Rightarrow> 'a \<Rightarrow> bool. \<forall>A. \<exists>x. A = f x"  | 
| 61931 | 69  | 
then obtain f :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where *: "\<forall>A. \<exists>x. A = f x" ..  | 
70  | 
let ?D = "\<lambda>x. \<not> f x x"  | 
|
71  | 
from * have "\<exists>x. ?D = f x" ..  | 
|
72  | 
then obtain a where "?D = f a" ..  | 
|
| 62038 | 73  | 
then have "?D a \<longleftrightarrow> f a a" by (rule arg_cong)  | 
| 62266 | 74  | 
then have "\<not> f a a \<longleftrightarrow> f a a" .  | 
| 61931 | 75  | 
then show False by (rule iff_contradiction)  | 
76  | 
qed  | 
|
77  | 
||
78  | 
||
| 61930 | 79  | 
subsection \<open>Classic Isabelle/HOL example\<close>  | 
80  | 
||
81  | 
text \<open>  | 
|
| 61931 | 82  | 
The following treatment of Cantor's Theorem follows the classic example from  | 
83  | 
  the early 1990s, e.g.\ see the file @{verbatim "92/HOL/ex/set.ML"} in
 | 
|
84  | 
  Isabelle92 or @{cite \<open>\S18.7\<close> "paulson-isa-book"}. The old tactic scripts
 | 
|
85  | 
synthesize key information of the proof by refinement of schematic goal  | 
|
86  | 
states. In contrast, the Isar proof needs to say explicitly what is proven.  | 
|
| 61930 | 87  | 
|
88  | 
\<^bigskip>  | 
|
89  | 
Cantor's Theorem states that every set has more subsets than it has  | 
|
90  | 
elements. It has become a favourite basic example in pure higher-order logic  | 
|
91  | 
since it is so easily expressed:  | 
|
| 61541 | 92  | 
|
93  | 
  @{text [display]
 | 
|
94  | 
\<open>\<forall>f::\<alpha> \<Rightarrow> \<alpha> \<Rightarrow> bool. \<exists>S::\<alpha> \<Rightarrow> bool. \<forall>x::\<alpha>. f x \<noteq> S\<close>}  | 
|
| 12388 | 95  | 
|
| 61541 | 96  | 
Viewing types as sets, \<open>\<alpha> \<Rightarrow> bool\<close> represents the powerset of \<open>\<alpha>\<close>. This  | 
97  | 
version of the theorem states that for every function from \<open>\<alpha>\<close> to its  | 
|
98  | 
powerset, some subset is outside its range. The Isabelle/Isar proofs below  | 
|
| 61930 | 99  | 
uses HOL's set theory, with the type \<open>\<alpha> set\<close> and the operator \<open>range :: (\<alpha> \<Rightarrow>  | 
100  | 
\<beta>) \<Rightarrow> \<beta> set\<close>.  | 
|
101  | 
\<close>  | 
|
| 6505 | 102  | 
|
| 55640 | 103  | 
theorem "\<exists>S. S \<notin> range (f :: 'a \<Rightarrow> 'a set)"  | 
| 10007 | 104  | 
proof  | 
| 55640 | 105  | 
  let ?S = "{x. x \<notin> f x}"
 | 
106  | 
show "?S \<notin> range f"  | 
|
| 10007 | 107  | 
proof  | 
| 55640 | 108  | 
assume "?S \<in> range f"  | 
| 12388 | 109  | 
then obtain y where "?S = f y" ..  | 
| 23373 | 110  | 
then show False  | 
| 12388 | 111  | 
proof (rule equalityCE)  | 
| 55640 | 112  | 
assume "y \<in> f y"  | 
113  | 
assume "y \<in> ?S"  | 
|
114  | 
then have "y \<notin> f y" ..  | 
|
| 61930 | 115  | 
with \<open>y \<in> f y\<close> show ?thesis by contradiction  | 
| 12388 | 116  | 
next  | 
| 55640 | 117  | 
assume "y \<notin> ?S"  | 
118  | 
assume "y \<notin> f y"  | 
|
119  | 
then have "y \<in> ?S" ..  | 
|
| 58614 | 120  | 
with \<open>y \<notin> ?S\<close> show ?thesis by contradiction  | 
| 10007 | 121  | 
qed  | 
122  | 
qed  | 
|
123  | 
qed  | 
|
| 6494 | 124  | 
|
| 61930 | 125  | 
text \<open>  | 
126  | 
How much creativity is required? As it happens, Isabelle can prove this  | 
|
127  | 
theorem automatically using best-first search. Depth-first search would  | 
|
128  | 
diverge, but best-first search successfully navigates through the large  | 
|
129  | 
search space. The context of Isabelle's classical prover contains rules for  | 
|
130  | 
the relevant constructs of HOL's set theory.  | 
|
131  | 
\<close>  | 
|
| 6505 | 132  | 
|
| 55640 | 133  | 
theorem "\<exists>S. S \<notin> range (f :: 'a \<Rightarrow> 'a set)"  | 
| 10007 | 134  | 
by best  | 
| 6494 | 135  | 
|
| 10007 | 136  | 
end  |