Order on cardinal numbers #
We define the order on cardinal numbers and show its basic properties, including the ordered semiring structure.
Main definitions #
- The order
c₁ ≤ c₂
is defined byCardinal.le_def α β : #α ≤ #β ↔ Nonempty (α ↪ β)
. Order.IsSuccLimit c
means thatc
is a (weak) limit cardinal:c ≠ 0 ∧ ∀ x < c, succ x < c
.Cardinal.IsStrongLimit c
means thatc
is a strong limit cardinal:c ≠ 0 ∧ ∀ x < c, 2 ^ x < c
.
Main instances #
- Cardinals form a
CanonicallyOrderedAdd
OrderedCommSemiring
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
(inMathlib.SetTheory.Cardinal.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 #
The current setup interweaves the order structure and the algebraic structure on Cardinal
tightly.
For example, we need to know what a ring is in order to show that 0
is the smallest cardinality.
That is reflected in this file containing both the order and algebra structure.
References #
Tags #
cardinal number, cardinal arithmetic, cardinal exponentiation, aleph, Cantor's theorem, König's theorem, Konig's theorem
Order on cardinals #
We define the order on cardinal numbers by #α ≤ #β
if and only if
there exists an embedding (injective function) from α to β.
Equations
- Cardinal.instLE = { le := fun (q₁ q₂ : Cardinal.{?u.28}) => Quotient.liftOn₂ q₁ q₂ (fun (α β : Type ?u.28) => Nonempty (α ↪ β)) Cardinal.instLE.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.
lift
sends Cardinal.{u}
to an initial segment of Cardinal.{max u v}
. #
Cardinal.lift
as an InitialSeg
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Cardinal.lift
as an OrderEmbedding
.
Instances For
Basic cardinals #
Alias of Cardinal.power_natCast
.
Order properties #
Equations
- One or more equations did not get rendered due to their size.
Equations
Equations
- Cardinal.instWellFoundedRelation = { rel := fun (x1 x2 : Cardinal.{?u.5}) => x1 < x2, wf := Cardinal.lt_wf }
Note that the successor of c
is not the same as c + 1
except in the case of finite c
.
Limit cardinals #
A cardinal is a strong limit if it is not zero and it is closed under powersets.
Note that ℵ₀
is a strong limit by this definition.
- two_power_lt ⦃x : Cardinal.{u_1}⦄ : x < c → 2 ^ x < c
Instances For
Indexed cardinal sum
#
Well-ordering theorem #
An embedding of any type to the set of cardinals in its universe.
Equations
Instances For
Any type can be endowed with a well order, obtained by pulling back the well order over cardinals by some embedding.
Equations
- WellOrderingRel = ⇑embeddingToCardinal ⁻¹'o fun (x1 x2 : Cardinal.{?u.23}) => x1 < x2
Instances For
The well-ordering theorem (or Zermelo's theorem): every type has a well-order
Bounds on suprema #
Indexed cardinal prod
#
König's theorem