Constant n
-ary function with value a
.
Equations
- arity.const a (n + 1) = λ (_x : α), arity.const a n
- arity.const a 0 = a
Equations
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.setoid = {r := pSet.equiv, iseqv := pSet.setoid._proof_1}
Equations
Equations
- pSet.has_coe = {coe := pSet.to_set}
The empty pre-set
Equations
Equations
Equations
- pSet.has_singleton = {singleton := λ (s : pSet), {s}}
Equations
- pSet.is_lawful_singleton = {insert_emptyc_eq := pSet.is_lawful_singleton._proof_1}
The n-th von Neumann ordinal
Equations
- pSet.of_nat (n + 1) = (pSet.of_nat n).insert (pSet.of_nat n)
- pSet.of_nat 0 = ∅
The von Neumann ordinal ω
Equations
- pSet.omega = pSet.mk (ulift ℕ) (λ (n : ulift ℕ), pSet.of_nat n.down)
The image of a function
Equations
- pSet.image f (pSet.mk α A) = pSet.mk α (λ (a : α), f (A a))
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
- pSet.arity.equiv a b = ∀ (x y : pSet), x.equiv y → pSet.arity.equiv (a x) (b y)
- pSet.arity.equiv a b = pSet.equiv a b
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
- pSet.resp n = {x // pSet.arity.equiv x x}
Equations
- pSet.resp.inhabited = {default := ⟨arity.const (inhabited.default pSet) n, _⟩}
Equations
- a.equiv b = pSet.arity.equiv a.val b.val
Equations
- pSet.resp.setoid = {r := pSet.resp.equiv n, iseqv := _}
An equivalence-respecting function yields an n-ary Set function.
Equations
- mk : Π (n : ℕ) (f : pSet.resp n), pSet.definable n (pSet.resp.eval n f)
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
Equations
Equations
- pSet.definable.resp (pSet.resp.eval n f) = f
Equations
- classical.all_definable F = pSet.definable.eq_mk ⟨λ (x : pSet), (pSet.definable.resp (F ⟦x⟧)).val, _⟩ _
- classical.all_definable F = let p : ∃ (a : pSet), ⟦a⟧ = F := _ in pSet.definable.eq_mk ⟨classical.some p, _⟩ _
Equations
Equations
insert x y
is the set {x} ∪ y
Equations
- Set.insert = pSet.resp.eval 2 ⟨pSet.insert, Set.insert._proof_1⟩
Equations
Equations
- Set.has_singleton = {singleton := λ (x : Set), {x}}
Equations
- Set.is_lawful_singleton = {insert_emptyc_eq := Set.is_lawful_singleton._proof_1}
omega
is the first infinite von Neumann ordinal
Equations
The powerset operation, the collection of subsets of a set
Equations
- Set.powerset = pSet.resp.eval 1 ⟨pSet.powerset, Set.powerset._proof_1⟩
The union operator, the collection of elements of elements of a set
Equations
- ⋃ = pSet.resp.eval 1 ⟨pSet.Union, Set.Union._proof_1⟩
Equations
Equations
The image of a (definable) set function
Equations
- Set.image f = let r : pSet.resp 1 := pSet.definable.resp f in pSet.resp.eval 1 ⟨pSet.image r.val, _⟩
Equations
- Set.map_definable_aux f = classical.all_definable (λ (y : Set), y.pair (f y))
Equations
Equations
- Class.inhabited = {default := ∅}
Equations
Coerce a set into a class
Equations
- Class.of_Set x = {y : Set | y ∈ x}
Equations
Assert that A
is a set satisfying p
A ∈ B
if A
is a set which is a member of B
Equations
- A.mem B = Class.to_Set B A
Equations
- Class.has_mem = {mem := Class.mem}
Convert a conglomerate (a collection of classes) into a class
Equations
- Class.Cong_to_Class x = {y : Set | ↑y ∈ x}
Convert a class into a conglomerate (a collection of classes)
Equations
- x.Class_to_Cong = {y : Class | y ∈ x}
The power class of a class is the class of all subclasses that are sets
Equations
- x.powerset = Class.Cong_to_Class (𝒫x)
The union of a class is the class of all members of sets in the class
Equations
- x.Union = ⋃₀x.Class_to_Cong
Function value
Equations
- F′A = Class.iota (λ (y : Set), Class.to_Set (λ (x : Set), F (x.pair y)) A)