Documentation

Mathlib.SetTheory.Cardinal.Order

Order on cardinal numbers #

We define the order on cardinal numbers and show its basic properties, including the ordered semiring structure.

Main definitions #

Main instances #

Main Statements #

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
Equations
  • One or more equations did not get rendered due to their size.
theorem Cardinal.le_def (α β : Type u) :
mk α mk β Nonempty (α β)
theorem Cardinal.mk_le_of_injective {α β : Type u} {f : αβ} (hf : Function.Injective f) :
mk α mk β
theorem Cardinal.mk_le_of_surjective {α β : Type u} {f : αβ} (hf : Function.Surjective f) :
mk β mk α
theorem Cardinal.le_mk_iff_exists_set {c : Cardinal.{u}} {α : Type u} :
c mk α ∃ (p : Set α), mk p = c
theorem Cardinal.mk_subtype_le {α : Type u} (p : αProp) :
mk (Subtype p) mk α
theorem Cardinal.mk_set_le {α : Type u} (s : Set α) :
mk s mk α
theorem Cardinal.lift_mk_le' {α : Type u} {β : Type v} :

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}. #

def Cardinal.liftInitialSeg :
InitialSeg (fun (x1 x2 : Cardinal.{u}) => x1 < x2) fun (x1 x2 : Cardinal.{max u v}) => x1 < x2

Cardinal.lift as an InitialSeg.

Equations
  • One or more equations did not get rendered due to their size.
Instances For
    @[deprecated Cardinal.mem_range_lift_of_le (since := "2024-10-07")]

    Basic cardinals #

    @[simp]
    theorem Cardinal.mk_fintype (α : Type u) [h : Fintype α] :
    mk α = (Fintype.card α)
    theorem Cardinal.power_mul {a b c : Cardinal.{u_1}} :
    a ^ (b * c) = (a ^ b) ^ c
    @[simp]
    theorem Cardinal.power_natCast (a : Cardinal.{u}) (n : ) :
    a ^ n = a ^ n
    @[deprecated Cardinal.power_natCast (since := "2024-10-16")]
    theorem Cardinal.power_cast_right (a : Cardinal.{u}) (n : ) :
    a ^ n = a ^ n

    Alias of Cardinal.power_natCast.

    @[simp]
    @[simp]
    theorem Cardinal.mk_set {α : Type u} :
    mk (Set α) = 2 ^ mk α
    @[simp]
    theorem Cardinal.mk_powerset {α : Type u} (s : Set α) :
    mk ↑(𝒫 s) = 2 ^ mk s

    A variant of Cardinal.mk_set expressed in terms of a Set instead of a Type.

    Order properties #

    Equations
    • One or more equations did not get rendered due to their size.
    theorem Cardinal.power_le_power_left {a b c : Cardinal.{u_1}} :
    a 0b ca ^ b a ^ c
    theorem Cardinal.self_le_power (a : Cardinal.{u_1}) {b : Cardinal.{u_1}} (hb : 1 b) :
    a a ^ b
    theorem Cardinal.cantor (a : Cardinal.{u}) :
    a < 2 ^ a

    Cantor's theorem

    theorem Cardinal.power_le_max_power_one {a b c : Cardinal.{u_1}} (h : b c) :
    a ^ b a ^ c 1
    theorem Cardinal.power_le_power_right {a b c : Cardinal.{u_1}} :
    a ba ^ c b ^ c
    theorem Cardinal.power_pos {a : Cardinal.{u_1}} (b : Cardinal.{u_1}) (ha : 0 < a) :
    0 < a ^ b
    theorem Cardinal.lt_wf :
    WellFounded fun (x1 x2 : Cardinal.{u}) => x1 < x2
    @[simp]

    Note that the successor of c is not the same as c + 1 except in the case of finite c.

    Equations

    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.

    Instances For

      Indexed cardinal sum #

      theorem Cardinal.le_sum {ι : Type u_1} (f : ιCardinal.{max u_1 u_2}) (i : ι) :
      f i sum f
      theorem Cardinal.iSup_le_sum {ι : Type u_1} (f : ιCardinal.{max u_1 u_2}) :
      @[simp]
      theorem Cardinal.sum_add_distrib {ι : Type u_1} (f g : ιCardinal.{u_2}) :
      sum (f + g) = sum f + sum g
      @[simp]
      theorem Cardinal.sum_add_distrib' {ι : Type u_1} (f g : ιCardinal.{u_2}) :
      (sum fun (i : ι) => f i + g i) = sum f + sum g
      theorem Cardinal.sum_le_sum {ι : Type u_1} (f g : ιCardinal.{u_2}) (H : ∀ (i : ι), f i g i) :
      sum f sum g
      theorem Cardinal.mk_le_mk_mul_of_mk_preimage_le {α β : Type u} {c : Cardinal.{u}} (f : αβ) (hf : ∀ (b : β), mk ↑(f ⁻¹' {b}) c) :
      mk α mk β * c
      theorem Cardinal.lift_mk_le_lift_mk_mul_of_lift_mk_preimage_le {α : Type u} {β : Type v} {c : Cardinal.{max u v}} (f : αβ) (hf : ∀ (b : β), lift.{v, u} (mk ↑(f ⁻¹' {b})) c) :

      Well-ordering theorem #

      An embedding of any type to the set of cardinals in its universe.

      Equations
      Instances For
        def WellOrderingRel {α : Type u} :
        ααProp

        Any type can be endowed with a well order, obtained by pulling back the well order over cardinals by some embedding.

        Equations
        Instances For
          instance IsWellOrder.subtype_nonempty {α : Type u} :
          Nonempty { r : ααProp // IsWellOrder α r }
          theorem exists_wellOrder (α : Type u) :
          ∃ (x : LinearOrder α), WellFoundedLT α

          The well-ordering theorem (or Zermelo's theorem): every type has a well-order

          Bounds on suprema #

          theorem Cardinal.exists_eq_of_iSup_eq_of_not_isSuccPrelimit {ι : Type u} (f : ιCardinal.{v}) (ω : Cardinal.{v}) ( : ¬Order.IsSuccPrelimit ω) (h : ⨆ (i : ι), f i = ω) :
          ∃ (i : ι), f i = ω
          theorem Cardinal.exists_eq_of_iSup_eq_of_not_isSuccLimit {ι : Type u} [ : Nonempty ι] (f : ιCardinal.{v}) (hf : BddAbove (Set.range f)) {c : Cardinal.{v}} (hc : ¬Order.IsSuccLimit c) (h : ⨆ (i : ι), f i = c) :
          ∃ (i : ι), f i = c

          Indexed cardinal prod #

          theorem Cardinal.sum_lt_prod {ι : Type u_1} (f g : ιCardinal.{u_2}) (H : ∀ (i : ι), f i < g i) :
          sum f < prod g

          König's theorem

          theorem Cardinal.prod_le_prod {ι : Type u_1} (f g : ιCardinal.{u_2}) (H : ∀ (i : ι), f i g i) :

          The first infinite cardinal aleph0 #

          Properties about the cast from #

          theorem Cardinal.mk_fin (n : ) :
          mk (Fin n) = n
          @[simp]
          theorem Cardinal.lift_natCast (n : ) :
          lift.{u, v} n = n
          @[simp]
          theorem Cardinal.lift_eq_nat_iff {a : Cardinal.{u}} {n : } :
          lift.{v, u} a = n a = n
          @[simp]
          theorem Cardinal.nat_eq_lift_iff {n : } {a : Cardinal.{u}} :
          n = lift.{v, u} a n = a
          @[simp]
          theorem Cardinal.lift_le_nat_iff {a : Cardinal.{u}} {n : } :
          lift.{v, u} a n a n
          @[simp]
          theorem Cardinal.nat_le_lift_iff {n : } {a : Cardinal.{u}} :
          n lift.{v, u} a n a
          @[simp]
          theorem Cardinal.lift_lt_nat_iff {a : Cardinal.{u}} {n : } :
          lift.{v, u} a < n a < n
          @[simp]
          theorem Cardinal.nat_lt_lift_iff {n : } {a : Cardinal.{u}} :
          n < lift.{v, u} a n < a
          theorem Cardinal.mk_coe_finset {α : Type u} {s : Finset α} :
          mk { x : α // x s } = s.card
          theorem Cardinal.card_le_of_finset {α : Type u_1} (s : Finset α) :
          s.card mk α