Documentation

Mathlib.SetTheory.ZFC.Basic

A model of ZFC #

In this file, we model Zermelo-Fraenkel set theory (+ Choice) using Lean's underlying type theory. We do this in four main steps:

The model #

Other definitions #

Notes #

To avoid confusion between the Lean Set and the ZFC Set, docstrings in this file refer to them respectively as "Set" and "ZFC set".

TODO #

Prove ZFSet.mapDefinableAux computably.

inductive PSet :
Type (u + 1)

The type of pre-sets in universe u. A pre-set is a family of pre-sets indexed by a type in Type u. The ZFC universe is defined as a quotient of this to ensure extensionality.

Instances For
def PSet.Type :
PSetType u

The underlying type of a pre-set

Equations
Instances For
def PSet.Func (x : PSet) :
x.TypePSet

The underlying pre-set family of a pre-set

Equations
@[simp]
theorem PSet.mk_type (α : Type u_1) (A : αPSet) :
(PSet.mk α A).Type = α
@[simp]
theorem PSet.mk_func (α : Type u_1) (A : αPSet) :
(PSet.mk α A).Func = A
@[simp]
theorem PSet.eta (x : PSet) :
PSet.mk x.Type x.Func = x
def PSet.Equiv :
PSetPSetProp

Two pre-sets are extensionally equivalent if every element of the first family is extensionally equivalent to some element of the second family and vice-versa.

Equations
  • (PSet.mk α A).Equiv (PSet.mk α_1 B) = ((∀ (a : α), ∃ (b : α_1), (A a).Equiv (B b)) ∀ (b : α_1), ∃ (a : α), (A a).Equiv (B b))
theorem PSet.equiv_iff {x : PSet} {y : PSet} :
x.Equiv y (∀ (i : x.Type), ∃ (j : y.Type), (x.Func i).Equiv (y.Func j)) ∀ (j : y.Type), ∃ (i : x.Type), (x.Func i).Equiv (y.Func j)
theorem PSet.Equiv.exists_left {x : PSet} {y : PSet} (h : x.Equiv y) (i : x.Type) :
∃ (j : y.Type), (x.Func i).Equiv (y.Func j)
theorem PSet.Equiv.exists_right {x : PSet} {y : PSet} (h : x.Equiv y) (j : y.Type) :
∃ (i : x.Type), (x.Func i).Equiv (y.Func j)
theorem PSet.Equiv.refl (x : PSet) :
x.Equiv x
theorem PSet.Equiv.rfl {x : PSet} :
x.Equiv x
theorem PSet.Equiv.euc {x : PSet} {y : PSet} {z : PSet} :
x.Equiv yz.Equiv yx.Equiv z
theorem PSet.Equiv.symm {x : PSet} {y : PSet} :
x.Equiv yy.Equiv x
theorem PSet.Equiv.comm {x : PSet} {y : PSet} :
x.Equiv y y.Equiv x
theorem PSet.Equiv.trans {x : PSet} {y : PSet} {z : PSet} (h1 : x.Equiv y) (h2 : y.Equiv z) :
x.Equiv z
theorem PSet.equiv_of_isEmpty (x : PSet) (y : PSet) [IsEmpty x.Type] [IsEmpty y.Type] :
x.Equiv y
def PSet.Subset (x : PSet) (y : PSet) :

A pre-set is a subset of another pre-set if every element of the first family is extensionally equivalent to some element of the second family.

Equations
  • x.Subset y = ∀ (a : x.Type), ∃ (b : y.Type), (x.Func a).Equiv (y.Func b)
instance PSet.instIsReflSubset :
IsRefl PSet fun (x1 x2 : PSet) => x1 x2
Equations
instance PSet.instIsTransSubset :
IsTrans PSet fun (x1 x2 : PSet) => x1 x2
Equations
theorem PSet.Equiv.ext (x : PSet) (y : PSet) :
x.Equiv y x y y x
theorem PSet.Subset.congr_left {x : PSet} {y : PSet} {z : PSet} :
x.Equiv y(x z y z)
theorem PSet.Subset.congr_right {x : PSet} {y : PSet} {z : PSet} :
x.Equiv y(z x z y)
def PSet.Mem (y : PSet) (x : PSet) :

x ∈ y as pre-sets if x is extensionally equivalent to a member of the family y.

Equations
  • y.Mem x = ∃ (b : y.Type), x.Equiv (y.Func b)
theorem PSet.Mem.mk {α : Type u} (A : αPSet) (a : α) :
A a PSet.mk α A
theorem PSet.func_mem (x : PSet) (i : x.Type) :
x.Func i x
theorem PSet.Mem.ext {x : PSet} {y : PSet} :
(∀ (w : PSet), w x w y)x.Equiv y
theorem PSet.Mem.congr_right {x : PSet} {y : PSet} :
x.Equiv y∀ {w : PSet}, w x w y
theorem PSet.equiv_iff_mem {x : PSet} {y : PSet} :
x.Equiv y ∀ {w : PSet}, w x w y
theorem PSet.Mem.congr_left {x : PSet} {y : PSet} :
x.Equiv y∀ {w : PSet}, x w y w
theorem PSet.mem_of_subset {x : PSet} {y : PSet} {z : PSet} :
x yz xz y
theorem PSet.subset_iff {x : PSet} {y : PSet} :
x y ∀ ⦃z : PSet⦄, z xz y
theorem PSet.mem_wf :
WellFounded fun (x1 x2 : PSet) => x1 x2
instance PSet.instIsAsymmMem :
IsAsymm PSet fun (x1 x2 : PSet) => x1 x2
Equations
instance PSet.instIsIrreflMem :
IsIrrefl PSet fun (x1 x2 : PSet) => x1 x2
Equations
theorem PSet.mem_asymm {x : PSet} {y : PSet} :
x yyx
theorem PSet.mem_irrefl (x : PSet) :
xx
def PSet.toSet (u : PSet) :

Convert a pre-set to a Set of pre-sets.

Equations
@[simp]
theorem PSet.mem_toSet (a : PSet) (u : PSet) :
a u.toSet a u

A nonempty set is one that contains some element.

Equations
  • u.Nonempty = u.toSet.Nonempty
theorem PSet.nonempty_def (u : PSet) :
u.Nonempty ∃ (x : PSet), x u
theorem PSet.nonempty_of_mem {x : PSet} {u : PSet} (h : x u) :
u.Nonempty
@[simp]
theorem PSet.nonempty_toSet_iff {u : PSet} :
u.toSet.Nonempty u.Nonempty
theorem PSet.nonempty_type_iff_nonempty {x : PSet} :
Nonempty x.Type x.Nonempty
theorem PSet.nonempty_of_nonempty_type (x : PSet) [h : Nonempty x.Type] :
x.Nonempty
theorem PSet.Equiv.eq {x : PSet} {y : PSet} :
x.Equiv y x.toSet = y.toSet

Two pre-sets are equivalent iff they have the same members.

The empty pre-set

Equations
@[simp]
theorem PSet.not_mem_empty (x : PSet) :
x
@[simp]
theorem PSet.toSet_empty :
.toSet =
@[simp]
theorem PSet.empty_subset (x : PSet) :
@[simp]
theorem PSet.not_nonempty_empty :
¬.Nonempty
theorem PSet.equiv_empty (x : PSet) [IsEmpty x.Type] :
x.Equiv
def PSet.insert (x : PSet) (y : PSet) :

Insert an element into a pre-set

Equations
Equations
instance PSet.instInhabitedTypeInsert (x : PSet) (y : PSet) :
Inhabited (insert x y).Type
Equations
@[simp]
theorem PSet.mem_insert_iff {x : PSet} {y : PSet} {z : PSet} :
x insert y z x.Equiv y x z
theorem PSet.mem_insert (x : PSet) (y : PSet) :
x insert x y
theorem PSet.mem_insert_of_mem {y : PSet} {z : PSet} (x : PSet) (h : z y) :
z insert x y
@[simp]
theorem PSet.mem_singleton {x : PSet} {y : PSet} :
x {y} x.Equiv y
theorem PSet.mem_pair {x : PSet} {y : PSet} {z : PSet} :
x {y, z} x.Equiv y x.Equiv z

The n-th von Neumann ordinal

Equations

The von Neumann ordinal ω

Equations
def PSet.sep (p : PSetProp) (x : PSet) :

The pre-set separation operation {x ∈ a | p x}

Equations
  • PSet.sep p x = PSet.mk { a : x.Type // p (x.Func a) } fun (y : { a : x.Type // p (x.Func a) }) => x.Func y
Equations
theorem PSet.mem_sep {p : PSetProp} (H : ∀ (x y : PSet), x.Equiv yp xp y) {x : PSet} {y : PSet} :
y PSet.sep p x y x p y

The pre-set powerset operator

Equations
  • x.powerset = PSet.mk (Set x.Type) fun (p : Set x.Type) => PSet.mk { a : x.Type // p a } fun (y : { a : x.Type // p a }) => x.Func y
@[simp]
theorem PSet.mem_powerset {x : PSet} {y : PSet} :
y x.powerset y x
def PSet.sUnion (a : PSet) :

The pre-set union operator

Equations
  • ⋃₀ a = PSet.mk ((x : a.Type) × (a.Func x).Type) fun (x : (x : a.Type) × (a.Func x).Type) => match x with | x, y => (a.Func x).Func y

The pre-set union operator

Equations
@[simp]
theorem PSet.mem_sUnion {x : PSet} {y : PSet} :
y ⋃₀ x zx, y z
@[simp]
theorem PSet.toSet_sUnion (x : PSet) :
(⋃₀ x).toSet = ⋃₀ (PSet.toSet '' x.toSet)
def PSet.image (f : PSetPSet) (x : PSet) :

The image of a function from pre-sets to pre-sets.

Equations
theorem PSet.mem_image {f : PSetPSet} (H : ∀ (x y : PSet), x.Equiv y(f x).Equiv (f y)) {x : PSet} {y : PSet} :
y PSet.image f x zx, y.Equiv (f z)

Universe lift operation

Equations

Embedding of one universe in another

Equations

Function equivalence is defined so that f ~ g iff ∀ x y, x ~ y → f x ~ g y. This extends to equivalence of n-ary functions.

Equations
def PSet.Resp (n : ) :
Type (max 0 (u + 1))

resp n is the collection of n-ary functions on PSet that respect equivalence, i.e. when the inputs are equivalent the output is as well.

Equations
Instances For
Equations
def PSet.Resp.f {n : } (f : PSet.Resp (n + 1)) (x : PSet) :

The n-ary image of a (n + 1)-ary function respecting equivalence as a function respecting equivalence.

Equations
  • f.f x = f x,
def PSet.Resp.Equiv {n : } (a : PSet.Resp n) (b : PSet.Resp n) :

Function equivalence for functions respecting equivalence. See PSet.Arity.Equiv.

Equations
theorem PSet.Resp.Equiv.refl {n : } (a : PSet.Resp n) :
a.Equiv a
theorem PSet.Resp.Equiv.euc {n : } {a : PSet.Resp n} {b : PSet.Resp n} {c : PSet.Resp n} :
a.Equiv bc.Equiv ba.Equiv c
theorem PSet.Resp.Equiv.symm {n : } {a : PSet.Resp n} {b : PSet.Resp n} :
a.Equiv bb.Equiv a
theorem PSet.Resp.Equiv.trans {n : } {x : PSet.Resp n} {y : PSet.Resp n} {z : PSet.Resp n} (h1 : x.Equiv y) (h2 : y.Equiv z) :
x.Equiv z
instance PSet.Resp.setoid {n : } :
Equations
  • PSet.Resp.setoid = { r := PSet.Resp.Equiv, iseqv := }
def PSet.Resp.evalAux {n : } :
{ f : PSet.Resp nFunction.OfArity ZFSet ZFSet n // ∀ (a b : PSet.Resp n), a.Equiv bf a = f b }

Helper function for PSet.eval.

Equations

An equivalence-respecting function yields an n-ary ZFC set function.

Equations
theorem PSet.Resp.eval_val {n : } {f : PSet.Resp (n + 1)} {x : PSet} :
PSet.Resp.eval (n + 1) f x = PSet.Resp.eval n (f.f x)
class inductive PSet.Definable (n : ) :

A set function is "definable" if it is the image of some n-ary pre-set function. This isn't exactly definability, but is useful as a sufficient condition for functions that have a computable image.

Instances

The evaluation of a function respecting equivalence is definable, by that same function.

Equations

Turns a definable function into a function that respects equivalence.

Equations

All functions are classically definable.

Equations

Turns a pre-set into a ZFC set.

Equations
@[simp]
theorem ZFSet.mk_eq (x : PSet) :
x = ZFSet.mk x
@[simp]
theorem ZFSet.eq {x : PSet} {y : PSet} :
ZFSet.mk x = ZFSet.mk y x.Equiv y
theorem ZFSet.sound {x : PSet} {y : PSet} (h : x.Equiv y) :
theorem ZFSet.exact {x : PSet} {y : PSet} :
ZFSet.mk x = ZFSet.mk yx.Equiv y
@[simp]
theorem ZFSet.eval_mk {n : } {f : PSet.Resp (n + 1)} {x : PSet} :
PSet.Resp.eval (n + 1) f (ZFSet.mk x) = PSet.Resp.eval n (f.f x)
def ZFSet.Mem :
ZFSetZFSetProp

The membership relation for ZFC sets is inherited from the membership relation for pre-sets.

Equations
Equations
@[simp]
theorem ZFSet.mk_mem_iff {x : PSet} {y : PSet} :

Convert a ZFC set into a Set of ZFC sets

Equations
@[simp]
theorem ZFSet.mem_toSet (a : ZFSet) (u : ZFSet) :
a u.toSet a u
instance ZFSet.small_toSet (x : ZFSet) :
Small.{u, u + 1} x.toSet
Equations
  • =

A nonempty set is one that contains some element.

Equations
  • u.Nonempty = u.toSet.Nonempty
theorem ZFSet.nonempty_def (u : ZFSet) :
u.Nonempty ∃ (x : ZFSet), x u
theorem ZFSet.nonempty_of_mem {x : ZFSet} {u : ZFSet} (h : x u) :
u.Nonempty
@[simp]
theorem ZFSet.nonempty_toSet_iff {u : ZFSet} :
u.toSet.Nonempty u.Nonempty
def ZFSet.Subset (x : ZFSet) (y : ZFSet) :

x ⊆ y as ZFC sets means that all members of x are members of y.

Equations
theorem ZFSet.subset_def {x : ZFSet} {y : ZFSet} :
x y ∀ ⦃z : ZFSet⦄, z xz y
instance ZFSet.instIsReflSubset :
IsRefl ZFSet fun (x1 x2 : ZFSet) => x1 x2
Equations
instance ZFSet.instIsTransSubset :
IsTrans ZFSet fun (x1 x2 : ZFSet) => x1 x2
Equations
@[simp]
theorem ZFSet.subset_iff {x : PSet} {y : PSet} :
@[simp]
theorem ZFSet.toSet_subset_iff {x : ZFSet} {y : ZFSet} :
x.toSet y.toSet x y
theorem ZFSet.ext_iff {x : ZFSet} {y : ZFSet} :
x = y ∀ (z : ZFSet), z x z y
theorem ZFSet.ext {x : ZFSet} {y : ZFSet} :
(∀ (z : ZFSet), z x z y)x = y
@[simp]
theorem ZFSet.toSet_inj {x : ZFSet} {y : ZFSet} :
x.toSet = y.toSet x = y

The empty ZFC set

Equations
@[simp]
theorem ZFSet.not_mem_empty (x : ZFSet) :
x
@[simp]
theorem ZFSet.toSet_empty :
.toSet =
@[simp]
theorem ZFSet.empty_subset (x : ZFSet) :
@[simp]
@[simp]
theorem ZFSet.nonempty_mk_iff {x : PSet} :
(ZFSet.mk x).Nonempty x.Nonempty
theorem ZFSet.eq_empty (x : ZFSet) :
x = ∀ (y : ZFSet), yx
theorem ZFSet.eq_empty_or_nonempty (u : ZFSet) :
u = u.Nonempty
Equations
@[simp]
theorem ZFSet.mem_insert_iff {x : ZFSet} {y : ZFSet} {z : ZFSet} :
x insert y z x = y x z
theorem ZFSet.mem_insert (x : ZFSet) (y : ZFSet) :
x insert x y
theorem ZFSet.mem_insert_of_mem {y : ZFSet} {z : ZFSet} (x : ZFSet) (h : z y) :
z insert x y
@[simp]
theorem ZFSet.toSet_insert (x : ZFSet) (y : ZFSet) :
(insert x y).toSet = insert x y.toSet
@[simp]
theorem ZFSet.mem_singleton {x : ZFSet} {y : ZFSet} :
x {y} x = y
@[simp]
theorem ZFSet.toSet_singleton (x : ZFSet) :
{x}.toSet = {x}
theorem ZFSet.insert_nonempty (u : ZFSet) (v : ZFSet) :
(insert u v).Nonempty
theorem ZFSet.singleton_nonempty (u : ZFSet) :
{u}.Nonempty
theorem ZFSet.mem_pair {x : ZFSet} {y : ZFSet} {z : ZFSet} :
x {y, z} x = y x = z

omega is the first infinite von Neumann ordinal

Equations
def ZFSet.sep (p : ZFSetProp) :

{x ∈ a | p x} is the set of elements in a satisfying p

Equations
@[simp]
theorem ZFSet.mem_sep {p : ZFSetProp} {x : ZFSet} {y : ZFSet} :
y ZFSet.sep p x y x p y
@[simp]
theorem ZFSet.toSet_sep (a : ZFSet) (p : ZFSetProp) :
(ZFSet.sep p a).toSet = {x : ZFSet | x a.toSet p x}

The powerset operation, the collection of subsets of a ZFC set

Equations
@[simp]
theorem ZFSet.mem_powerset {x : ZFSet} {y : ZFSet} :
y x.powerset y x
theorem ZFSet.sUnion_lem {α : Type u} {β : Type u} (A : αPSet) (B : βPSet) (αβ : ∀ (a : α), ∃ (b : β), (A a).Equiv (B b)) (a : (⋃₀ PSet.mk α A).Type) :
∃ (b : (⋃₀ PSet.mk β B).Type), ((⋃₀ PSet.mk α A).Func a).Equiv ((⋃₀ PSet.mk β B).Func b)

The union operator, the collection of elements of elements of a ZFC set

Equations

The union operator, the collection of elements of elements of a ZFC set

Equations
noncomputable def ZFSet.sInter (x : ZFSet) :

The intersection operator, the collection of elements in all of the elements of a ZFC set. We special-case ⋂₀ ∅ = ∅.

Equations

The intersection operator, the collection of elements in all of the elements of a ZFC set. We special-case ⋂₀ ∅ = ∅.

Equations
@[simp]
theorem ZFSet.mem_sUnion {x : ZFSet} {y : ZFSet} :
y ⋃₀ x zx, y z
theorem ZFSet.mem_sInter {x : ZFSet} {y : ZFSet} (h : x.Nonempty) :
y ⋂₀ x zx, y z
theorem ZFSet.mem_of_mem_sInter {x : ZFSet} {y : ZFSet} {z : ZFSet} (hy : y ⋂₀ x) (hz : z x) :
y z
theorem ZFSet.mem_sUnion_of_mem {x : ZFSet} {y : ZFSet} {z : ZFSet} (hy : y z) (hz : z x) :
theorem ZFSet.not_mem_sInter_of_not_mem {x : ZFSet} {y : ZFSet} {z : ZFSet} (hy : yz) (hz : z x) :
y⋂₀ x
@[simp]
theorem ZFSet.sUnion_singleton {x : ZFSet} :
⋃₀ {x} = x
@[simp]
theorem ZFSet.sInter_singleton {x : ZFSet} :
⋂₀ {x} = x
@[simp]
theorem ZFSet.toSet_sUnion (x : ZFSet) :
(⋃₀ x).toSet = ⋃₀ (ZFSet.toSet '' x.toSet)
theorem ZFSet.toSet_sInter {x : ZFSet} (h : x.Nonempty) :
(⋂₀ x).toSet = ⋂₀ (ZFSet.toSet '' x.toSet)
@[simp]
theorem ZFSet.singleton_inj {x : ZFSet} {y : ZFSet} :
{x} = {y} x = y
def ZFSet.union (x : ZFSet) (y : ZFSet) :

The binary union operation

Equations
def ZFSet.inter (x : ZFSet) (y : ZFSet) :

The binary intersection operation

Equations
def ZFSet.diff (x : ZFSet) (y : ZFSet) :

The set difference operation

Equations
@[simp]
theorem ZFSet.toSet_union (x : ZFSet) (y : ZFSet) :
(x y).toSet = x.toSet y.toSet
@[simp]
theorem ZFSet.toSet_inter (x : ZFSet) (y : ZFSet) :
(x y).toSet = x.toSet y.toSet
@[simp]
theorem ZFSet.toSet_sdiff (x : ZFSet) (y : ZFSet) :
(x \ y).toSet = x.toSet \ y.toSet
@[simp]
theorem ZFSet.mem_union {x : ZFSet} {y : ZFSet} {z : ZFSet} :
z x y z x z y
@[simp]
theorem ZFSet.mem_inter {x : ZFSet} {y : ZFSet} {z : ZFSet} :
z x y z x z y
@[simp]
theorem ZFSet.mem_diff {x : ZFSet} {y : ZFSet} {z : ZFSet} :
z x \ y z x zy
@[simp]
theorem ZFSet.sUnion_pair {x : ZFSet} {y : ZFSet} :
⋃₀ {x, y} = x y
theorem ZFSet.mem_wf :
WellFounded fun (x1 x2 : ZFSet) => x1 x2
theorem ZFSet.inductionOn {p : ZFSetProp} (x : ZFSet) (h : ∀ (x : ZFSet), (∀ yx, p y)p x) :
p x

Induction on the relation.

instance ZFSet.instIsAsymmMem :
IsAsymm ZFSet fun (x1 x2 : ZFSet) => x1 x2
Equations
instance ZFSet.instIsIrreflMem :
IsIrrefl ZFSet fun (x1 x2 : ZFSet) => x1 x2
Equations
theorem ZFSet.mem_asymm {x : ZFSet} {y : ZFSet} :
x yyx
theorem ZFSet.mem_irrefl (x : ZFSet) :
xx
theorem ZFSet.regularity (x : ZFSet) (h : x ) :
yx, x y =
def ZFSet.image (f : ZFSetZFSet) [PSet.Definable 1 f] :

The image of a (definable) ZFC set function

Equations
theorem ZFSet.image.mk (f : ZFSetZFSet) [H : PSet.Definable 1 f] (x : ZFSet) {y : ZFSet} :
y xf y ZFSet.image f x
@[simp]
theorem ZFSet.mem_image {f : ZFSetZFSet} [H : PSet.Definable 1 f] {x : ZFSet} {y : ZFSet} :
y ZFSet.image f x zx, f z = y
@[simp]
theorem ZFSet.toSet_image (f : ZFSetZFSet) [H : PSet.Definable 1 f] (x : ZFSet) :
(ZFSet.image f x).toSet = f '' x.toSet
noncomputable def ZFSet.range {α : Type u} (f : αZFSet) :

The range of an indexed family of sets. The universes allow for a more general index type without manual use of ULift.

Equations
@[simp]
theorem ZFSet.mem_range {α : Type u} {f : αZFSet} {x : ZFSet} :
@[simp]
theorem ZFSet.toSet_range {α : Type u} (f : αZFSet) :
(ZFSet.range f).toSet = Set.range f
def ZFSet.pair (x : ZFSet) (y : ZFSet) :

Kuratowski ordered pair

Equations
  • x.pair y = {{x}, {x, y}}
@[simp]
theorem ZFSet.toSet_pair (x : ZFSet) (y : ZFSet) :
(x.pair y).toSet = {{x}, {x, y}}
def ZFSet.pairSep (p : ZFSetZFSetProp) (x : ZFSet) (y : ZFSet) :

A subset of pairs {(a, b) ∈ x × y | p a b}

Equations
@[simp]
theorem ZFSet.mem_pairSep {p : ZFSetZFSetProp} {x : ZFSet} {y : ZFSet} {z : ZFSet} :
z ZFSet.pairSep p x y ax, by, z = a.pair b p a b
@[simp]
theorem ZFSet.pair_inj {x : ZFSet} {y : ZFSet} {x' : ZFSet} {y' : ZFSet} :
x.pair y = x'.pair y' x = x' y = y'

The cartesian product, {(a, b) | a ∈ x, b ∈ y}

Equations
@[simp]
theorem ZFSet.mem_prod {x : ZFSet} {y : ZFSet} {z : ZFSet} :
z x.prod y ax, by, z = a.pair b
theorem ZFSet.pair_mem_prod {x : ZFSet} {y : ZFSet} {a : ZFSet} {b : ZFSet} :
a.pair b x.prod y a x b y
def ZFSet.IsFunc (x : ZFSet) (y : ZFSet) (f : ZFSet) :

isFunc x y f is the assertion that f is a subset of x × y which relates to each element of x a unique element of y, so that we can consider f as a ZFC function x → y.

Equations
  • x.IsFunc y f = (f x.prod y zx, ∃! w : ZFSet, z.pair w f)
def ZFSet.funs (x : ZFSet) (y : ZFSet) :

funs x y is y ^ x, the set of all set functions x → y

Equations
  • x.funs y = ZFSet.sep (x.IsFunc y) (x.prod y).powerset
@[simp]
theorem ZFSet.mem_funs {x : ZFSet} {y : ZFSet} {f : ZFSet} :
f x.funs y x.IsFunc y f
noncomputable instance ZFSet.mapDefinableAux (f : ZFSetZFSet) [PSet.Definable 1 f] :
PSet.Definable 1 fun (y : ZFSet) => y.pair (f y)
Equations
noncomputable def ZFSet.map (f : ZFSetZFSet) [PSet.Definable 1 f] :

Graph of a function: map f x is the ZFC function which maps a ∈ x to f a

Equations
@[simp]
theorem ZFSet.mem_map {f : ZFSetZFSet} [PSet.Definable 1 f] {x : ZFSet} {y : ZFSet} :
y ZFSet.map f x zx, z.pair (f z) = y
theorem ZFSet.map_unique {f : ZFSetZFSet} [H : PSet.Definable 1 f] {x : ZFSet} {z : ZFSet} (zx : z x) :
∃! w : ZFSet, z.pair w ZFSet.map f x
@[simp]
theorem ZFSet.map_isFunc {f : ZFSetZFSet} [PSet.Definable 1 f] {x : ZFSet} {y : ZFSet} :
x.IsFunc y (ZFSet.map f x) zx, f z y
@[irreducible]
def ZFSet.Hereditarily (p : ZFSetProp) (x : ZFSet) :

Given a predicate p on ZFC sets. Hereditarily p x means that x has property p and the members of x are all Hereditarily p.

Equations
theorem ZFSet.hereditarily_iff {p : ZFSetProp} {x : ZFSet} :
ZFSet.Hereditarily p x p x yx, ZFSet.Hereditarily p y
theorem ZFSet.Hereditarily.def {p : ZFSetProp} {x : ZFSet} :
ZFSet.Hereditarily p xp x yx, ZFSet.Hereditarily p y

Alias of the forward direction of ZFSet.hereditarily_iff.

theorem ZFSet.Hereditarily.self {p : ZFSetProp} {x : ZFSet} (h : ZFSet.Hereditarily p x) :
p x
theorem ZFSet.Hereditarily.mem {p : ZFSetProp} {x : ZFSet} {y : ZFSet} (h : ZFSet.Hereditarily p x) (hy : y x) :
def Class :
Type (u_1 + 1)

The collection of all classes. We define Class as Set ZFSet, as this allows us to get many instances automatically. However, in practice, we treat it as (the definitionally equal) ZFSet → Prop. This means, the preferred way to state that x : ZFSet belongs to A : Class is to write A x.

Equations
Instances For
Equations
def Class.sep (p : ZFSetProp) (A : Class) :

{x ∈ A | p x} is the class of elements in A satisfying p

Equations
theorem Class.ext_iff {x : Class} {y : Class} :
x = y ∀ (z : ZFSet), x z y z
theorem Class.ext {x : Class} {y : Class} :
(∀ (z : ZFSet), x z y z)x = y

Coerce a ZFC set into a class

Equations

The universal class

Equations
def Class.ToSet (B : Class) (A : Class) :

Assert that A is a ZFC set satisfying B

Equations
def Class.Mem (B : Class) (A : Class) :

A ∈ B if A is a ZFC set which satisfies B

Equations
  • B.Mem A = B.ToSet A
theorem Class.mem_def (A : Class) (B : Class) :
A B ∃ (x : ZFSet), x = A B x
@[simp]
theorem Class.not_mem_empty (x : Class) :
x
@[simp]
@[simp]
theorem Class.mem_univ {A : Class} :
A Class.univ ∃ (x : ZFSet), x = A
@[simp]
theorem Class.eq_univ_iff_forall {A : Class} :
A = Class.univ ∀ (x : ZFSet), A x
theorem Class.eq_univ_of_forall {A : Class} :
(∀ (x : ZFSet), A x)A = Class.univ
theorem Class.mem_wf :
WellFounded fun (x1 x2 : Class) => x1 x2
instance Class.instIsAsymmMem :
IsAsymm Class fun (x1 x2 : Class) => x1 x2
Equations
instance Class.instIsIrreflMem :
IsIrrefl Class fun (x1 x2 : Class) => x1 x2
Equations
theorem Class.mem_asymm {x : Class} {y : Class} :
x yyx
theorem Class.mem_irrefl (x : Class) :
xx

There is no universal set. This is stated as univuniv, meaning that univ (the class of all sets) is proper (does not belong to the class of all sets).

Convert a conglomerate (a collection of classes) into a class

Equations

Convert a class into a conglomerate (a collection of classes)

Equations
@[simp]
theorem Class.classToCong_empty :
.classToCong =

The power class of a class is the class of all subclasses that are ZFC sets

Equations

The union of a class is the class of all members of ZFC sets in the class

Equations

The union of a class is the class of all members of ZFC sets in the class

Equations

The intersection of a class is the class of all members of ZFC sets in the class

Equations

The intersection of a class is the class of all members of ZFC sets in the class

Equations
theorem Class.ofSet.inj {x : ZFSet} {y : ZFSet} (h : x = y) :
x = y
@[simp]
theorem Class.toSet_of_ZFSet (A : Class) (x : ZFSet) :
A.ToSet x A x
@[simp]
theorem Class.coe_mem {x : ZFSet} {A : Class} :
x A A x
@[simp]
theorem Class.coe_apply {x : ZFSet} {y : ZFSet} :
y x x y
@[simp]
theorem Class.coe_subset (x : ZFSet) (y : ZFSet) :
x y x y
@[simp]
theorem Class.coe_sep (p : Class) (x : ZFSet) :
(ZFSet.sep p x) = {y : ZFSet | y x p y}
@[simp]
theorem Class.coe_empty :
=
@[simp]
theorem Class.coe_insert (x : ZFSet) (y : ZFSet) :
(insert x y) = insert x y
@[simp]
theorem Class.coe_union (x : ZFSet) (y : ZFSet) :
(x y) = x y
@[simp]
theorem Class.coe_inter (x : ZFSet) (y : ZFSet) :
(x y) = x y
@[simp]
theorem Class.coe_diff (x : ZFSet) (y : ZFSet) :
(x \ y) = x \ y
@[simp]
theorem Class.coe_powerset (x : ZFSet) :
x.powerset = (↑x).powerset
@[simp]
theorem Class.powerset_apply {A : Class} {x : ZFSet} :
A.powerset x x A
@[simp]
theorem Class.sUnion_apply {x : Class} {y : ZFSet} :
(⋃₀ x) y ∃ (z : ZFSet), x z y z
@[simp]
theorem Class.coe_sUnion (x : ZFSet) :
(⋃₀ x) = ⋃₀ x
@[simp]
theorem Class.mem_sUnion {x : Class} {y : Class} :
y ⋃₀ x zx, y z
theorem Class.sInter_apply {x : Class} {y : ZFSet} :
(⋂₀ x) y ∀ (z : ZFSet), x zy z
@[simp]
theorem Class.coe_sInter {x : ZFSet} (h : x.Nonempty) :
(⋂₀ x) = ⋂₀ x
theorem Class.mem_of_mem_sInter {x : Class} {y : Class} {z : Class} (hy : y ⋂₀ x) (hz : z x) :
y z
theorem Class.mem_sInter {x : Class} {y : Class} (h : Set.Nonempty x) :
y ⋂₀ x zx, y z
theorem Class.eq_univ_of_powerset_subset {A : Class} (hA : A.powerset A) :

An induction principle for sets. If every subset of a class is a member, then the class is universal.

The definite description operator, which is {x} if {y | A y} = {x} and otherwise.

Equations
theorem Class.iota_val (A : Class) (x : ZFSet) (H : ∀ (y : ZFSet), A y y = x) :
A.iota = x
theorem Class.iota_ex (A : Class) :

Unlike the other set constructors, the iota definite descriptor is a set for any set input, but not constructively so, so there is no associated ClassSet function.

def Class.fval (F : Class) (A : Class) :

Function value

Equations
theorem Class.fval_ex (F : Class) (A : Class) :
@[simp]
theorem ZFSet.map_fval {f : ZFSetZFSet} [H : PSet.Definable 1 f] {x : ZFSet} {y : ZFSet} (h : y x) :
(ZFSet.map f x) y = (f y)
noncomputable def ZFSet.choice (x : ZFSet) :

A choice function on the class of nonempty ZFC sets.

Equations
theorem ZFSet.choice_mem_aux (x : ZFSet) (h : x) (y : ZFSet) (yx : y x) :
(Classical.epsilon fun (z : ZFSet) => z y) y
theorem ZFSet.choice_isFunc (x : ZFSet) (h : x) :
x.IsFunc (⋃₀ x) x.choice
theorem ZFSet.choice_mem (x : ZFSet) (h : x) (y : ZFSet) (yx : y x) :
x.choice y y
@[simp]
theorem ZFSet.toSet_equiv_apply_coe (x : ZFSet) :
(ZFSet.toSet_equiv x) = x.toSet
noncomputable def ZFSet.toSet_equiv :

ZFSet.toSet as an equivalence.

Equations
  • One or more equations did not get rendered due to their size.