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