Documentation

Mathlib.SetTheory.Cardinal.Defs

Cardinal Numbers #

We define cardinal numbers as a quotient of types under the equivalence relation of equinumerity.

Main definitions #

Implementation notes #

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
def Cardinal :
Type (u + 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
      Instances For
        theorem Cardinal.inductionOn {p : Cardinal.{u_1}Prop} (c : Cardinal.{u_1}) (h : ∀ (α : Type u_1), p (mk α)) :
        p c
        theorem Cardinal.inductionOn₂ {p : Cardinal.{u_1}Cardinal.{u_2}Prop} (c₁ : Cardinal.{u_1}) (c₂ : Cardinal.{u_2}) (h : ∀ (α : Type u_1) (β : Type u_2), p (mk α) (mk β)) :
        p c₁ c₂
        theorem Cardinal.inductionOn₃ {p : Cardinal.{u_1}Cardinal.{u_2}Cardinal.{u_3}Prop} (c₁ : Cardinal.{u_1}) (c₂ : Cardinal.{u_2}) (c₃ : Cardinal.{u_3}) (h : ∀ (α : Type u_1) (β : Type u_2) (γ : Type u_3), p (mk α) (mk β) (mk γ)) :
        p c₁ c₂ c₃
        theorem Cardinal.induction_on_pi {ι : Type u} {p : (ιCardinal.{v})Prop} (f : ιCardinal.{v}) (h : ∀ (f : ιType v), p fun (i : ι) => mk (f i)) :
        p f
        theorem Cardinal.eq {α β : Type u} :
        mk α = mk β Nonempty (α β)
        @[deprecated "No deprecation message was provided." (since := "2024-10-24")]
        theorem Cardinal.mk'_def (α : Type u) :
        α = mk α

        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
          theorem Cardinal.mk_congr {α β : Type u} (e : α β) :
          mk α = mk β
          theorem Equiv.cardinal_eq {α β : Type u} (e : α β) :

          Alias of Cardinal.mk_congr.

          def Cardinal.map (f : Type u → Type v) (hf : (α β : Type u) → α βf α f β) :

          Lift a function between Type*s to a function between Cardinals.

          Equations
          Instances For
            @[simp]
            theorem Cardinal.map_mk (f : Type u → Type v) (hf : (α β : Type u) → α βf α f β) (α : Type u) :
            map f hf (mk α) = mk (f α)
            def Cardinal.map₂ (f : Type u → Type v → Type w) (hf : (α β : Type u) → (γ δ : Type v) → α βγ δf α γ f β δ) :

            Lift a binary operation Type* → Type* → Type* to a binary operation on Cardinals.

            Equations
            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
              Instances For
                @[simp]
                theorem Cardinal.mk_uLift (α : Type u) :

                lift.{max u v, u} equals lift.{v, u}.

                Unfortunately, the simp lemma doesn't work.

                @[deprecated Cardinal.lift_umax (since := "2024-10-24")]

                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.

                @[simp]

                A cardinal lifted to the same universe equals itself.

                @[simp]

                A cardinal lifted to the zero universe equals itself.

                theorem Cardinal.lift_mk_eq' {α : Type u} {β : Type v} :

                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.

                theorem Cardinal.mk_congr_lift {α : Type u} {β : Type v} (e : α β) :

                Basic cardinals #

                @[simp]
                theorem Cardinal.mk_eq_zero (α : Type u) [IsEmpty α] :
                mk α = 0
                theorem Cardinal.mk_eq_zero_iff {α : Type u} :
                mk α = 0 IsEmpty α
                @[simp]
                theorem Cardinal.mk_ne_zero (α : Type u) [Nonempty α] :
                mk α 0
                theorem Cardinal.mk_eq_one (α : Type u) [Subsingleton α] [Nonempty α] :
                mk α = 1
                Equations
                theorem Cardinal.add_def (α β : Type u) :
                mk α + mk β = mk (α β)
                @[simp]
                theorem Cardinal.mk_sum (α : Type u) (β : Type v) :
                mk (α β) = lift.{v, u} (mk α) + lift.{u, v} (mk β)
                @[simp]
                theorem Cardinal.mk_option {α : Type u} :
                mk (Option α) = mk α + 1
                @[simp]
                theorem Cardinal.mk_psum (α : Type u) (β : Type v) :
                mk (α ⊕' β) = lift.{v, u} (mk α) + lift.{u, v} (mk β)
                Equations
                theorem Cardinal.mul_def (α β : Type u) :
                mk α * mk β = mk (α × β)
                @[simp]
                theorem Cardinal.mk_prod (α : Type u) (β : Type v) :
                mk (α × β) = lift.{v, u} (mk α) * lift.{u, v} (mk β)

                The cardinal exponential. #α ^ #β is the cardinal of β → α.

                Equations
                theorem Cardinal.power_def (α β : Type u) :
                mk α ^ mk β = mk (βα)
                theorem Cardinal.mk_arrow (α : Type u) (β : Type v) :
                mk (αβ) = lift.{u, v} (mk β) ^ lift.{v, u} (mk α)
                @[simp]
                theorem Cardinal.power_zero (a : Cardinal.{u_1}) :
                a ^ 0 = 1
                @[simp]
                theorem Cardinal.power_one (a : Cardinal.{u}) :
                a ^ 1 = a
                theorem Cardinal.power_add (a b c : Cardinal.{u_1}) :
                a ^ (b + c) = a ^ b * a ^ c
                @[simp]
                theorem Cardinal.one_power {a : Cardinal.{u_1}} :
                1 ^ a = 1
                @[simp]
                theorem Cardinal.zero_power {a : Cardinal.{u_1}} :
                a 00 ^ a = 0
                theorem Cardinal.mul_power {a b c : Cardinal.{u_1}} :
                (a * b) ^ c = a ^ c * b ^ c

                Indexed cardinal sum #

                The indexed sum of cardinals is the cardinality of the indexed disjoint union, i.e. sigma type.

                Equations
                Instances For
                  @[simp]
                  theorem Cardinal.mk_sigma {ι : Type u_2} (f : ιType u_1) :
                  mk ((i : ι) × f i) = sum fun (i : ι) => mk (f i)
                  theorem Cardinal.mk_sigma_congr_lift {ι : Type v} {ι' : Type v'} {f : ιType w} {g : ι'Type w'} (e : ι ι') (h : ∀ (i : ι), lift.{w', w} (mk (f i)) = lift.{w, w'} (mk (g (e i)))) :
                  lift.{max v' w', max w v} (mk ((i : ι) × f i)) = lift.{max v w, max w' v'} (mk ((i : ι') × g i))
                  theorem Cardinal.mk_sigma_congr {ι ι' : Type u} {f : ιType v} {g : ι'Type v} (e : ι ι') (h : ∀ (i : ι), mk (f i) = mk (g (e i))) :
                  mk ((i : ι) × f i) = mk ((i : ι') × g i)
                  theorem Cardinal.mk_sigma_congr' {ι : Type u} {ι' : Type v} {f : ιType (max w u v)} {g : ι'Type (max w u v)} (e : ι ι') (h : ∀ (i : ι), mk (f i) = mk (g (e i))) :
                  mk ((i : ι) × f i) = mk ((i : ι') × g i)

                  Similar to mk_sigma_congr with indexing types in different universes. This is not a strict generalization.

                  theorem Cardinal.mk_sigma_congrRight {ι : Type u} {f g : ιType v} (h : ∀ (i : ι), mk (f i) = mk (g i)) :
                  mk ((i : ι) × f i) = mk ((i : ι) × g i)
                  theorem Cardinal.mk_psigma_congrRight {ι : Type u} {f g : ιType v} (h : ∀ (i : ι), mk (f i) = mk (g i)) :
                  mk ((i : ι) ×' f i) = mk ((i : ι) ×' g i)
                  theorem Cardinal.mk_psigma_congrRight_prop {ι : Prop} {f g : ιType v} (h : ∀ (i : ι), mk (f i) = mk (g i)) :
                  mk ((i : ι) ×' f i) = mk ((i : ι) ×' g i)
                  theorem Cardinal.mk_sigma_arrow {ι : Type u_3} (α : Type u_1) (f : ιType u_2) :
                  mk (Sigma fα) = mk ((i : ι) → f iα)
                  @[simp]
                  theorem Cardinal.sum_const (ι : Type u) (a : Cardinal.{v}) :
                  (sum fun (x : ι) => a) = lift.{v, u} (mk ι) * lift.{u, v} a
                  theorem Cardinal.sum_const' (ι : Type u) (a : Cardinal.{u}) :
                  (sum fun (x : ι) => a) = mk ι * a
                  @[simp]
                  theorem Cardinal.lift_sum {ι : Type u} (f : ιCardinal.{v}) :
                  lift.{w, max v u} (sum f) = sum fun (i : ι) => lift.{w, v} (f i)
                  theorem Cardinal.sum_nat_eq_add_sum_succ (f : Cardinal.{u}) :
                  sum f = f 0 + sum fun (i : ) => f (i + 1)

                  Indexed cardinal prod #

                  The indexed product of cardinals is the cardinality of the Pi type (dependent product).

                  Equations
                  Instances For
                    @[simp]
                    theorem Cardinal.mk_pi {ι : Type u} (α : ιType v) :
                    mk ((i : ι) → α i) = prod fun (i : ι) => mk (α i)
                    theorem Cardinal.mk_pi_congr_lift {ι : Type v} {ι' : Type v'} {f : ιType w} {g : ι'Type w'} (e : ι ι') (h : ∀ (i : ι), lift.{w', w} (mk (f i)) = lift.{w, w'} (mk (g (e i)))) :
                    lift.{max v' w', max v w} (mk ((i : ι) → f i)) = lift.{max v w, max v' w'} (mk ((i : ι') → g i))
                    theorem Cardinal.mk_pi_congr {ι ι' : Type u} {f : ιType v} {g : ι'Type v} (e : ι ι') (h : ∀ (i : ι), mk (f i) = mk (g (e i))) :
                    mk ((i : ι) → f i) = mk ((i : ι') → g i)
                    theorem Cardinal.mk_pi_congr_prop {ι ι' : Prop} {f : ιType v} {g : ι'Type v} (e : ι ι') (h : ∀ (i : ι), mk (f i) = mk (g )) :
                    mk ((i : ι) → f i) = mk ((i : ι') → g i)
                    theorem Cardinal.mk_pi_congr' {ι : Type u} {ι' : Type v} {f : ιType (max w u v)} {g : ι'Type (max w u v)} (e : ι ι') (h : ∀ (i : ι), mk (f i) = mk (g (e i))) :
                    mk ((i : ι) → f i) = mk ((i : ι') → g i)

                    Similar to mk_pi_congr with indexing types in different universes. This is not a strict generalization.

                    theorem Cardinal.mk_pi_congrRight {ι : Type u} {f g : ιType v} (h : ∀ (i : ι), mk (f i) = mk (g i)) :
                    mk ((i : ι) → f i) = mk ((i : ι) → g i)
                    theorem Cardinal.mk_pi_congrRight_prop {ι : Prop} {f g : ιType v} (h : ∀ (i : ι), mk (f i) = mk (g i)) :
                    mk ((i : ι) → f i) = mk ((i : ι) → g i)
                    @[simp]
                    theorem Cardinal.prod_const (ι : Type u) (a : Cardinal.{v}) :
                    (prod fun (x : ι) => a) = lift.{u, v} a ^ lift.{v, u} (mk ι)
                    theorem Cardinal.prod_const' (ι : Type u) (a : Cardinal.{u}) :
                    (prod fun (x : ι) => a) = a ^ mk ι
                    @[simp]
                    theorem Cardinal.prod_eq_zero {ι : Type u_1} (f : ιCardinal.{u}) :
                    prod f = 0 ∃ (i : ι), f i = 0
                    theorem Cardinal.prod_ne_zero {ι : Type u_1} (f : ιCardinal.{u_2}) :
                    prod f 0 ∀ (i : ι), f i 0
                    theorem Cardinal.power_sum {ι : Type u_1} (a : Cardinal.{max u_1 u_2}) (f : ιCardinal.{max u_1 u_2}) :
                    a ^ sum f = prod fun (i : ι) => a ^ f i
                    @[simp]
                    theorem Cardinal.lift_prod {ι : Type u} (c : ιCardinal.{v}) :
                    lift.{w, max v u} (prod c) = prod fun (i : ι) => lift.{w, v} (c i)

                    The first infinite cardinal aleph0 #

                    ℵ₀ is the smallest infinite cardinal.

                    Equations
                    Instances For

                      ℵ₀ is the smallest infinite cardinal.

                      Equations
                      Instances For

                        Cardinalities of basic sets and types #

                        theorem Cardinal.mk_subtype_of_equiv {α β : Type u} (p : βProp) (e : α β) :
                        mk { a : α // p (e a) } = mk { b : β // p b }