src/HOL/Isar_Examples/Cantor.thy
author wenzelm
Tue, 07 Oct 2014 20:59:46 +0200
changeset 58614 7338eb25226c
parent 55640 abc140f21caa
child 58882 6e2010ab8bd9
permissions -rw-r--r--
more cartouches; more antiquotations;

(*  Title:      HOL/Isar_Examples/Cantor.thy
    Author:     Markus Wenzel, TU Muenchen
*)

header \<open>Cantor's Theorem\<close>

theory Cantor
imports Main
begin

text_raw \<open>\footnote{This is an Isar version of the final example of
  the Isabelle/HOL manual @{cite "isabelle-HOL"}.}\<close>

text \<open>Cantor's Theorem states that every set has more subsets than
  it has elements.  It has become a favorite basic example in pure
  higher-order logic since it is so easily expressed: \[\all{f::\alpha
  \To \alpha \To \idt{bool}} \ex{S::\alpha \To \idt{bool}}
  \all{x::\alpha} f \ap x \not= S\]

  Viewing types as sets, $\alpha \To \idt{bool}$ represents the
  powerset of $\alpha$.  This version of the theorem states that for
  every function from $\alpha$ to its powerset, some subset is outside
  its range.  The Isabelle/Isar proofs below uses HOL's set theory,
  with the type $\alpha \ap \idt{set}$ and the operator
  $\idt{range}::(\alpha \To \beta) \To \beta \ap \idt{set}$.\<close>

theorem "\<exists>S. S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
proof
  let ?S = "{x. x \<notin> f x}"
  show "?S \<notin> range f"
  proof
    assume "?S \<in> range f"
    then obtain y where "?S = f y" ..
    then show False
    proof (rule equalityCE)
      assume "y \<in> f y"
      assume "y \<in> ?S"
      then have "y \<notin> f y" ..
      with \<open>y : f y\<close> show ?thesis by contradiction
    next
      assume "y \<notin> ?S"
      assume "y \<notin> f y"
      then have "y \<in> ?S" ..
      with \<open>y \<notin> ?S\<close> show ?thesis by contradiction
    qed
  qed
qed

text \<open>How much creativity is required?  As it happens, Isabelle can
  prove this theorem automatically using best-first search.
  Depth-first search would diverge, but best-first search successfully
  navigates through the large search space.  The context of Isabelle's
  classical prover contains rules for the relevant constructs of HOL's
  set theory.\<close>

theorem "\<exists>S. S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
  by best

text \<open>While this establishes the same theorem internally, we do not
  get any idea of how the proof actually works.  There is currently no
  way to transform internal system-level representations of Isabelle
  proofs back into Isar text.  Writing intelligible proof documents
  really is a creative process, after all.\<close>

end