Cardinal Numbers #
We define cardinal numbers as a quotient of types under the equivalence relation of equinumerity.
Main definitions #
Cardinal
is the type of cardinal numbers (in a given universe).Cardinal.mk α
or#α
is the cardinality ofα
. The notation#
lives in the localeCardinal
.- Addition
c₁ + c₂
is defined byCardinal.add_def α β : #α + #β = #(α ⊕ β)
. - Multiplication
c₁ * c₂
is defined byCardinal.mul_def : #α * #β = #(α × β)
. - Exponentiation
c₁ ^ c₂
is defined byCardinal.power_def α β : #α ^ #β = #(β → α)
. Cardinal.sum
is the sum of an indexed family of cardinals, i.e. the cardinality of the corresponding sigma type.Cardinal.prod
is the product of an indexed family of cardinals, i.e. the cardinality of the corresponding pi type.Cardinal.aleph0
orℵ₀
is the cardinality ofℕ
. This definition is universe polymorphic:Cardinal.aleph0.{u} : Cardinal.{u}
(contrast withℕ : Type
, which lives in a specific universe). In some cases the universe level has to be given explicitly.
Implementation notes #
- There is a type of cardinal numbers in every universe level:
Cardinal.{u} : Type (u + 1)
is the quotient of types inType u
. The operationCardinal.lift
lifts cardinal numbers to a higher level. - Cardinal arithmetic specifically for infinite cardinals (like
κ * κ = κ
) is in the fileMathlib/SetTheory/Cardinal/Ordinal.lean
. - There is an instance
Pow Cardinal
, but this will only fire if Lean already knows that both the base and the exponent live in the same universe. As a workaround, you can add
to a file. This notation will work even if Lean doesn't know yet that the base and the exponent live in the same universe (but no exponents in other types can be used). (Porting note: This last point might need to be updated.)local infixr:80 " ^' " => @HPow.hPow Cardinal Cardinal Cardinal _
References #
Tags #
cardinal number, cardinal arithmetic, cardinal exponentiation, aleph, Cantor's theorem, König's theorem, Konig's theorem
Definition of cardinals #
The equivalence relation on types given by equivalence (bijective correspondence) of types. Quotienting by this equivalence relation gives the cardinal numbers.
Equations
- Cardinal.isEquivalent = { r := fun (α β : Type ?u.13) => Nonempty (α ≃ β), iseqv := Cardinal.isEquivalent.proof_1 }
Cardinal.{u}
is the type of cardinal numbers in Type u
,
defined as the quotient of Type u
by existence of an equivalence
(a bijection with explicit inverse).
Equations
Instances For
The cardinal number of a type
Equations
Instances For
The cardinal number of a type
Equations
- Cardinal.«term#_» = Lean.ParserDescr.node `Cardinal.«term#_» 1024 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol "#") (Lean.ParserDescr.cat `term 1024))
Instances For
Avoid using Quotient.mk
to construct a Cardinal
directly
The representative of the cardinal of a type is equivalent to the original type.
Equations
Instances For
Alias of Cardinal.mk_congr
.
Lift a function between Type*
s to a function between Cardinal
s.
Equations
- Cardinal.map f hf = Quotient.map f ⋯
Instances For
Lift a binary operation Type* → Type* → Type*
to a binary operation on Cardinal
s.
Equations
- Cardinal.map₂ f hf = Quotient.map₂ f ⋯
Instances For
Lifting cardinals to a higher universe #
The universe lift operation on cardinals. You can specify the universes explicitly with
lift.{u v} : Cardinal.{v} → Cardinal.{max v u}
Equations
- Cardinal.lift.{?u.15, ?u.14} c = Cardinal.map ULift.{?u.15, ?u.14} (fun (x x_1 : Type ?u.14) (e : x ≃ x_1) => Equiv.ulift.trans (e.trans Equiv.ulift.symm)) c
Instances For
lift.{max u v, u}
equals lift.{v, u}
.
Unfortunately, the simp lemma doesn't work.
lift.{max v u, u}
equals lift.{v, u}
.
A cardinal lifted to a lower or equal universe equals itself.
Unfortunately, the simp lemma doesn't work.
A cardinal lifted to the same universe equals itself.
A cardinal lifted to the zero universe equals itself.
A variant of Cardinal.lift_mk_eq
with specialized universes.
Because Lean often can not realize it should use this specialization itself,
we provide this statement separately so you don't have to solve the specialization problem either.
Alias of Cardinal.mk_congr_lift
.
Basic cardinals #
Equations
- Cardinal.instZero = { zero := Cardinal.lift.{?u.3, 0} (Cardinal.mk (Fin 0)) }
Equations
- Cardinal.instInhabited = { default := 0 }
Equations
- Cardinal.instOne = { one := Cardinal.lift.{?u.3, 0} (Cardinal.mk (Fin 1)) }
Equations
- Cardinal.instAdd = { add := Cardinal.map₂ Sum fun (x x_1 x_2 x_3 : Type ?u.7) => Equiv.sumCongr }
Equations
- Cardinal.instNatCast = { natCast := fun (n : ℕ) => Cardinal.lift.{?u.4, 0} (Cardinal.mk (Fin n)) }
Equations
- Cardinal.instMul = { mul := Cardinal.map₂ Prod fun (x x_1 x_2 x_3 : Type ?u.7) => Equiv.prodCongr }
The cardinal exponential. #α ^ #β
is the cardinal of β → α
.
Equations
- Cardinal.instPowCardinal = { pow := Cardinal.map₂ (fun (α β : Type ?u.12) => β → α) fun (x x_1 x_2 x_3 : Type ?u.12) (e₁ : x ≃ x_1) (e₂ : x_2 ≃ x_3) => e₂.arrowCongr e₁ }
The indexed sum of cardinals is the cardinality of the indexed disjoint union, i.e. sigma type.
Equations
- Cardinal.sum f = Cardinal.mk ((i : ι) × Quotient.out (f i))
Instances For
Similar to mk_sigma_congr
with indexing types in different universes. This is not a strict
generalization.
The indexed product of cardinals is the cardinality of the Pi type (dependent product).
Equations
- Cardinal.prod f = Cardinal.mk ((i : ι) → Quotient.out (f i))
Instances For
Similar to mk_pi_congr
with indexing types in different universes. This is not a strict
generalization.
ℵ₀
is the smallest infinite cardinal.
Equations
Instances For
ℵ₀
is the smallest infinite cardinal.
Equations
- Cardinal.termℵ₀ = Lean.ParserDescr.node `Cardinal.termℵ₀ 1024 (Lean.ParserDescr.symbol "ℵ₀")