Skip to content

Commit

Permalink
Merge pull request #3244 from mtzguido/cantor
Browse files Browse the repository at this point in the history
Add a proof that the cardinality of Type(u+1) is greater than Type(u)
  • Loading branch information
mtzguido authored Apr 15, 2024
2 parents 76f4550 + f13d403 commit de01633
Show file tree
Hide file tree
Showing 9 changed files with 190 additions and 0 deletions.
1 change: 1 addition & 0 deletions ocaml/fstar-lib/generated/FStar_Cardinality_Cantor.ml

Some generated files are not rendered by default. Learn more about how customized files appear on GitHub.

8 changes: 8 additions & 0 deletions ocaml/fstar-lib/generated/FStar_Cardinality_Universes.ml

Some generated files are not rendered by default. Learn more about how customized files appear on GitHub.

8 changes: 8 additions & 0 deletions ocaml/fstar-lib/generated/FStar_Functions.ml

Some generated files are not rendered by default. Learn more about how customized files appear on GitHub.

24 changes: 24 additions & 0 deletions ulib/FStar.Cardinality.Cantor.fst
Original file line number Diff line number Diff line change
@@ -0,0 +1,24 @@
module FStar.Cardinality.Cantor

open FStar.Functions

let no_surj_powerset (a : Type) (f : a -> powerset a) : Lemma (~(is_surj f)) =
let aux () : Lemma (requires is_surj f) (ensures False) =
(* Cantor's proof: given a supposed surjective f,
we define a set s that cannot be in the image of f. Namely,
the set of x:a such that x is not in f(x). *)
let s : powerset a = fun x -> not (f x x) in
let aux (x : a) : Lemma (requires f x == s) (ensures False) =
// We have f x == s, which means that f x x == not (f x x), contradiction
assert (f x x) // this triggers the SMT appropriately
in
Classical.forall_intro (Classical.move_requires aux)
in
Classical.move_requires aux ()

let no_inj_powerset (a : Type) (f : powerset a -> a) : Lemma (~(is_inj f)) =
let aux () : Lemma (requires is_inj f) (ensures False) =
let g : a -> powerset a = inverse_of_inj f (fun _ -> false) in
no_surj_powerset a g
in
Classical.move_requires aux ()
13 changes: 13 additions & 0 deletions ulib/FStar.Cardinality.Cantor.fsti
Original file line number Diff line number Diff line change
@@ -0,0 +1,13 @@
module FStar.Cardinality.Cantor

(* Cantor's theorem: there is no surjection from a set to its
powerset, and therefore also no injection from the powerset to the
set. *)

open FStar.Functions

val no_surj_powerset (a : Type) (f : a -> powerset a)
: Lemma (~(is_surj f))

val no_inj_powerset (a : Type) (f : powerset a -> a)
: Lemma (~(is_inj f))
38 changes: 38 additions & 0 deletions ulib/FStar.Cardinality.Universes.fst
Original file line number Diff line number Diff line change
@@ -0,0 +1,38 @@
module FStar.Cardinality.Universes

open FStar.Functions
open FStar.Cardinality.Cantor

(* This type is an injection of all powersets of Type u (i.e. Type u -> bool
functions) into Type (u+1) *)
noeq
type type_powerset : (Type u#a -> bool) -> Type u#(a+1) =
| Mk : f:(Type u#a -> bool) -> type_powerset f

let aux_inj_type_powerset (f1 f2 : powerset (Type u#a))
: Lemma (requires type_powerset f1 == type_powerset f2)
(ensures f1 == f2)
=
assert (type_powerset f1 == type_powerset f2);
let xx1 : type_powerset f1 = Mk f1 in
let xx2 : type_powerset f2 = coerce_eq () xx1 in
assert (xx1 === xx2);
assert (f1 == f2)

let inj_type_powerset () : Lemma (is_inj type_powerset) =
Classical.forall_intro_2 (fun f1 -> Classical.move_requires (aux_inj_type_powerset f1))

(* The general structure of this proof is:
1- We know there is an injection of powerset(Type(u)) into Type(u+1) (see type_powerset above)
2- We know there is NO injection from powerset(Type(u)) into Type(u) (see no_inj_powerset)
3- Therefore, there cannot be an injection from Type(u+1) into Type(u), otherwise we would
compose it with the first injection and obtain the second impossible injection.
*)
let no_inj_universes (f : Type u#(a+1) -> Type u#a) : Lemma (~(is_inj f)) =
let aux () : Lemma (requires is_inj f) (ensures False) =
let comp : powerset (Type u#a) -> Type u#a = fun x -> f (type_powerset x) in
inj_type_powerset ();
inj_comp type_powerset f;
no_inj_powerset _ comp
in
Classical.move_requires aux ()
8 changes: 8 additions & 0 deletions ulib/FStar.Cardinality.Universes.fsti
Original file line number Diff line number Diff line change
@@ -0,0 +1,8 @@
module FStar.Cardinality.Universes

open FStar.Functions
open FStar.Cardinality.Cantor

(* Prove that there can be no injection from Type (u+1) into Type u *)
val no_inj_universes (f : Type u#(a+1) -> Type u#a)
: Lemma (~(is_inj f))
41 changes: 41 additions & 0 deletions ulib/FStar.Functions.fst
Original file line number Diff line number Diff line change
@@ -0,0 +1,41 @@
module FStar.Functions

let inj_comp (#a #b #c : _) (f : a -> b) (g : b -> c)
: Lemma (requires is_inj f /\ is_inj g)
(ensures is_inj (fun x -> g (f x)))
= ()

let surj_comp (#a #b #c : _) (f : a -> b) (g : b -> c)
: Lemma (requires is_surj f /\ is_surj g)
(ensures is_surj (fun x -> g (f x)))
= ()

let lem_surj (#a #b : _) (f : a -> b) (y : b)
: Lemma (requires is_surj f) (ensures in_image f y)
= ()

let inverse_of_bij #a #b f =
(* Construct the inverse from indefinite description + choice. *)
let g0 (y:b) : GTot (x:a{f x == y}) =
FStar.IndefiniteDescription.indefinite_description_ghost a
(fun (x:a) -> f x == y)
in
(* Prove it's surjective *)
let aux (x:a) : Lemma (exists (y:b). g0 y == x) =
assert (g0 (f x) == x)
in
Classical.forall_intro aux;
Ghost.Pull.pull g0

let inverse_of_inj #a #b f def =
(* f is a bijection into its image, obtain its inverse *)
let f' : a -> image_of f = fun x -> f x in
let g_partial = inverse_of_bij #a #(image_of f) f' in
(* extend the inverse to the full domain b *)
let g : b -> GTot a =
fun (y:b) ->
if FStar.StrongExcludedMiddle.strong_excluded_middle (in_image f y)
then g_partial y
else def
in
Ghost.Pull.pull g
49 changes: 49 additions & 0 deletions ulib/FStar.Functions.fsti
Original file line number Diff line number Diff line change
@@ -0,0 +1,49 @@
module FStar.Functions

(* This module contains basic definitions and lemmas
about functions and sets. *)

let is_inj (#a #b : _) (f : a -> b) : prop =
forall (x1 x2 : a). f x1 == f x2 ==> x1 == x2

let is_surj (#a #b : _) (f : a -> b) : prop =
forall (y:b). exists (x:a). f x == y

let is_bij (#a #b : _) (f : a -> b) : prop =
is_inj f /\ is_surj f

let in_image (#a #b : _) (f : a -> b) (y : b) : prop =
exists (x:a). f x == y

let image_of (#a #b : _) (f : a -> b) : Type =
y:b{in_image f y}

(* g inverses f *)
let is_inverse_of (#a #b : _) (g : b -> a) (f : a -> b) =
forall (x:a). g (f x) == x

let powerset (a:Type u#aa) : Type u#aa = a -> bool

val inj_comp (#a #b #c : _) (f : a -> b) (g : b -> c)
: Lemma (requires is_inj f /\ is_inj g)
(ensures is_inj (fun x -> g (f x)))

val surj_comp (#a #b #c : _) (f : a -> b) (g : b -> c)
: Lemma (requires is_surj f /\ is_surj g)
(ensures is_surj (fun x -> g (f x)))

val lem_surj (#a #b : _) (f : a -> b) (y : b)
: Lemma (requires is_surj f) (ensures in_image f y)

(* An bijection has a perfect inverse. *)
val inverse_of_bij (#a #b : _) (f : a -> b)
: Ghost (b -> a)
(requires is_bij f)
(ensures fun g -> is_bij g /\ g `is_inverse_of` f /\ f `is_inverse_of` g)

(* An injective function has an inverse (as long as the domain is non-empty),
and this inverse is surjective. *)
val inverse_of_inj (#a #b : _) (f : a -> b{is_inj f}) (def : a)
: Ghost (b -> a)
(requires is_inj f)
(ensures fun g -> is_surj g /\ g `is_inverse_of` f)

0 comments on commit de01633

Please sign in to comment.