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 : #α * #β = #(α × β)
. - The order
c₁ ≤ c₂
is defined byCardinal.le_def α β : #α ≤ #β ↔ Nonempty (α ↪ β)
. - Exponentiation
c₁ ^ c₂
is defined byCardinal.power_def α β : #α ^ #β = #(β → α)
. Cardinal.isLimit c
means thatc
is a (weak) limit cardinal:c ≠ 0 ∧ ∀ x < c, succ x < c
.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.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.powerlt a b
ora ^< b
is defined as the supremum ofa ^ c
forc < b
.
Main instances #
- Cardinals form a
CanonicallyOrderedCommSemiring
with the aforementioned sum and product. - Cardinals form a
SuccOrder
. UseOrder.succ c
for the smallest cardinal greater thanc
. - The less than relation on cardinals forms a well-order.
- Cardinals form a
ConditionallyCompleteLinearOrderBot
. Bounded sets for cardinals in universeu
are precisely the sets indexed by some type in universeu
, seeCardinal.bddAbove_iff_small
. One can usesSup
for the cardinal supremum, andsInf
for the minimum of a set of cardinals.
Main Statements #
- Cantor's theorem:
Cardinal.cantor c : c < 2 ^ c
. - König's theorem:
Cardinal.sum_lt_prod
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 fileSetTheory/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
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) => 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
- Cardinal.«term#_» = Lean.ParserDescr.node `Cardinal.term#_ 1024 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol "#") (Lean.ParserDescr.cat `term 1024))
Instances For
Equations
The representative of the cardinal of a type is equivalent to the original type.
Equations
- Cardinal.outMkEquiv = Nonempty.some ⋯
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
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, v} c = Cardinal.map ULift.{u, v} (fun (x x_1 : Type v) (e : x ≃ x_1) => Equiv.ulift.trans (e.trans Equiv.ulift.symm)) c
Instances For
lift.{max u v, u}
equals lift.{v, u}
.
lift.{max v u, u}
equals lift.{v, u}
.
A cardinal lifted to a lower or equal universe equals itself.
A cardinal lifted to the same universe equals itself.
A cardinal lifted to the zero universe equals itself.
We define the order on cardinal numbers by #α ≤ #β
if and only if
there exists an embedding (injective function) from α to β.
Equations
- Cardinal.instLECardinal = { le := fun (q₁ q₂ : Cardinal.{u}) => Quotient.liftOn₂ q₁ q₂ (fun (α β : Type u) => Nonempty (α ↪ β)) Cardinal.instLECardinal.proof_1 }
Equations
- One or more equations did not get rendered due to their size.
A variant of Cardinal.lift_mk_le
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.
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.
Cardinal.lift
as an OrderEmbedding
.
Instances For
Equations
- Cardinal.instZeroCardinal = { zero := Cardinal.lift.{u, 0} (Cardinal.mk (Fin 0)) }
Equations
- Cardinal.instInhabitedCardinal = { default := 0 }
Equations
- Cardinal.instOneCardinal = { one := Cardinal.lift.{u, 0} (Cardinal.mk (Fin 1)) }
Equations
Alias of the reverse direction of Cardinal.mk_le_one_iff_set_subsingleton
.
Equations
- Cardinal.instAddCardinal = { add := Cardinal.map₂ Sum fun (x x_1 x_2 x_3 : Type u) => Equiv.sumCongr }
Equations
- Cardinal.instNatCastCardinal = { natCast := fun (n : ℕ) => Cardinal.lift.{u, 0} (Cardinal.mk (Fin n)) }
Equations
- Cardinal.instMulCardinal = { mul := Cardinal.map₂ Prod fun (x x_1 x_2 x_3 : Type u) => Equiv.prodCongr }
The cardinal exponential. #α ^ #β
is the cardinal of β → α
.
Equations
- Cardinal.instPowCardinal = { pow := Cardinal.map₂ (fun (α β : Type u) => β → α) fun (x x_1 x_2 x_3 : Type u) (e₁ : x ≃ x_1) (e₂ : x_2 ≃ x_3) => Equiv.arrowCongr e₂ e₁ }
Porting note (#11229): Deprecated section. Remove.
Porting note (#11229): Deprecated section. Remove.
A variant of Cardinal.mk_set
expressed in terms of a Set
instead of a Type
.
Equations
Equations
Equations
- One or more equations did not get rendered due to their size.
Equations
- One or more equations did not get rendered due to their size.
Equations
- One or more equations did not get rendered due to their size.
Equations
- One or more equations did not get rendered due to their size.
Equations
Equations
- Cardinal.instDistribLatticeCardinal = inferInstance
Equations
- Cardinal.instWellFoundedRelationCardinal = { rel := fun (x x_1 : Cardinal.{u}) => x < x_1, wf := Cardinal.lt_wf }
Equations
- Cardinal.wo = ⋯
Note that the successor of c
is not the same as c + 1
except in the case of finite c
.
Equations
- One or more equations did not get rendered due to their size.
A cardinal is a limit if it is not zero or a successor cardinal. Note that ℵ₀
is a limit
cardinal by this definition, but 0
isn't.
Use IsSuccLimit
if you want to include the c = 0
case.
Equations
- Cardinal.IsLimit c = (c ≠ 0 ∧ Order.IsSuccLimit c)
Instances For
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
The range of an indexed cardinal function, whose outputs live in a higher universe than the inputs, is always bounded above.
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
A set of cardinals is bounded above iff it's small, i.e. it corresponds to a usual ZFC set.
A variant of ciSup_of_empty
but with 0
on the RHS for convenience
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
The lift of a supremum is the supremum of the lifts.
The lift of a supremum is the supremum of the lifts.
To prove that the lift of a supremum is bounded by some cardinal t
,
it suffices to show that the lift of each cardinal is bounded by t
.
To prove an inequality between the lifts to a common universe of two different supremums, it suffices to show that the lift of each cardinal from the smaller supremum if bounded by the lift of some cardinal from the larger supremum.
A variant of lift_iSup_le_lift_iSup
with universes specialized via w = v
and w' = v'
.
This is sometimes necessary to avoid universe unification issues.
ℵ₀
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 "ℵ₀")
Instances For
Properties about the cast from ℕ
#
Alias of the reverse direction of Cardinal.lt_aleph0_iff_set_finite
.
Alias of the reverse direction of Cardinal.le_aleph0_iff_set_countable
.
Equations
See also Cardinal.nsmul_lt_aleph0_iff_of_ne_zero
if you already have n ≠ 0
.
See also Cardinal.nsmul_lt_aleph0_iff
for a hypothesis-free version.
See also Cardinal.aleph0_le_mul_iff
.
See also Cardinal.aleph0_le_mul_iff'
.
König's theorem
The cardinality of a union is at most the sum of the cardinalities of the two sets.
The function a ^< b
, defined as the supremum of a ^ c
for c < b
.
Instances For
The function a ^< b
, defined as the supremum of a ^ c
for c < b
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The cardinality of a nontrivial module over a ring is at least the cardinality of the ring if there are no zero divisors (for instance if the ring is a field)