src/HOL/Isar_Examples/Cantor.thy
author wenzelm
Tue, 10 Nov 2015 23:41:20 +0100
changeset 61626 c304402cc3df
parent 61541 846c72206207
child 61930 380cbe15cca5
permissions -rw-r--r--
recovered from a9c0572109af;
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
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
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
3edd5f813f01 observe standard theory naming conventions;
wenzelm
parents: 23373
diff changeset
     8
imports Main
3edd5f813f01 observe standard theory naming conventions;
wenzelm
parents: 23373
diff changeset
     9
begin
7955
wenzelm
parents: 7874
diff changeset
    10
58614
7338eb25226c more cartouches;
wenzelm
parents: 55640
diff changeset
    11
text_raw \<open>\footnote{This is an Isar version of the final example of
7338eb25226c more cartouches;
wenzelm
parents: 55640
diff changeset
    12
  the Isabelle/HOL manual @{cite "isabelle-HOL"}.}\<close>
7819
6dd018b6cf3f tuned presentation;
wenzelm
parents: 7800
diff changeset
    13
58614
7338eb25226c more cartouches;
wenzelm
parents: 55640
diff changeset
    14
text \<open>Cantor's Theorem states that every set has more subsets than
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    15
  it has elements.  It has become a favourite basic example in pure
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    16
  higher-order logic since it is so easily expressed:
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    17
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    18
  @{text [display]
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    19
  \<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
    20
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    21
  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
    22
  version of the theorem states that for every function from \<open>\<alpha>\<close> to its
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    23
  powerset, some subset is outside its range. The Isabelle/Isar proofs below
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    24
  uses HOL's set theory, with the type \<open>\<alpha> set\<close> and the operator
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    25
  \<open>range :: (\<alpha> \<Rightarrow> \<beta>) \<Rightarrow> \<beta> set\<close>.\<close>
6505
2863855a6902 elaborated;
wenzelm
parents: 6494
diff changeset
    26
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    27
theorem "\<exists>S. S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9474
diff changeset
    28
proof
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    29
  let ?S = "{x. x \<notin> f x}"
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    30
  show "?S \<notin> range f"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9474
diff changeset
    31
  proof
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    32
    assume "?S \<in> range f"
12388
c845fec1ac94 simplified proof (no longer use swapped rules);
wenzelm
parents: 10007
diff changeset
    33
    then obtain y where "?S = f y" ..
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 16417
diff changeset
    34
    then show False
12388
c845fec1ac94 simplified proof (no longer use swapped rules);
wenzelm
parents: 10007
diff changeset
    35
    proof (rule equalityCE)
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    36
      assume "y \<in> f y"
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    37
      assume "y \<in> ?S"
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    38
      then have "y \<notin> f y" ..
58614
7338eb25226c more cartouches;
wenzelm
parents: 55640
diff changeset
    39
      with \<open>y : f y\<close> show ?thesis by contradiction
12388
c845fec1ac94 simplified proof (no longer use swapped rules);
wenzelm
parents: 10007
diff changeset
    40
    next
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    41
      assume "y \<notin> ?S"
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    42
      assume "y \<notin> f y"
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    43
      then have "y \<in> ?S" ..
58614
7338eb25226c more cartouches;
wenzelm
parents: 55640
diff changeset
    44
      with \<open>y \<notin> ?S\<close> show ?thesis by contradiction
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9474
diff changeset
    45
    qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9474
diff changeset
    46
  qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9474
diff changeset
    47
qed
6494
ab1442d2e4e1 detailed proofs;
wenzelm
parents: 6444
diff changeset
    48
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    49
text \<open>How much creativity is required? As it happens, Isabelle can prove
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    50
  this theorem automatically using best-first search. Depth-first search
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    51
  would diverge, but best-first search successfully navigates through the
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    52
  large search space. The context of Isabelle's classical prover contains
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    53
  rules for the relevant constructs of HOL's set theory.\<close>
6505
2863855a6902 elaborated;
wenzelm
parents: 6494
diff changeset
    54
55640
abc140f21caa tuned proofs;
wenzelm
parents: 37671
diff changeset
    55
theorem "\<exists>S. S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9474
diff changeset
    56
  by best
6494
ab1442d2e4e1 detailed proofs;
wenzelm
parents: 6444
diff changeset
    57
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    58
text \<open>While this establishes the same theorem internally, we do not get any
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    59
  idea of how the proof actually works. There is currently no way to
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    60
  transform internal system-level representations of Isabelle proofs back
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    61
  into Isar text. Writing intelligible proof documents really is a creative
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    62
  process, after all.\<close>
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    63
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9474
diff changeset
    64
end