Skip to content

Commit a577026

Browse files
committed
feat: some simp lemmas to compute more cardinals (#7660)
1 parent 84f5ad0 commit a577026

File tree

3 files changed

+47
-4
lines changed

3 files changed

+47
-4
lines changed

Mathlib/SetTheory/Cardinal/Basic.lean

Lines changed: 20 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1656,6 +1656,16 @@ theorem nat_mul_aleph0 {n : ℕ} (hn : n ≠ 0) : ↑n * ℵ₀ = ℵ₀ :=
16561656
theorem aleph0_mul_nat {n : ℕ} (hn : n ≠ 0) : ℵ₀ * n = ℵ₀ := by rw [mul_comm, nat_mul_aleph0 hn]
16571657
#align cardinal.aleph_0_mul_nat Cardinal.aleph0_mul_nat
16581658

1659+
-- See note [no_index around OfNat.ofNat]
1660+
@[simp]
1661+
theorem ofNat_mul_aleph0 {n : ℕ} [Nat.AtLeastTwo n] : no_index (OfNat.ofNat n) * ℵ₀ = ℵ₀ :=
1662+
nat_mul_aleph0 (OfNat.ofNat_ne_zero n)
1663+
1664+
-- See note [no_index around OfNat.ofNat]
1665+
@[simp]
1666+
theorem aleph0_mul_ofNat {n : ℕ} [Nat.AtLeastTwo n] : ℵ₀ * no_index (OfNat.ofNat n) = ℵ₀ :=
1667+
aleph0_mul_nat (OfNat.ofNat_ne_zero n)
1668+
16591669
@[simp]
16601670
theorem add_le_aleph0 {c₁ c₂ : Cardinal} : c₁ + c₂ ≤ ℵ₀ ↔ c₁ ≤ ℵ₀ ∧ c₂ ≤ ℵ₀ :=
16611671
fun h => ⟨le_self_add.trans h, le_add_self.trans h⟩, fun h =>
@@ -1671,6 +1681,16 @@ theorem aleph0_add_nat (n : ℕ) : ℵ₀ + n = ℵ₀ :=
16711681
theorem nat_add_aleph0 (n : ℕ) : ↑n + ℵ₀ = ℵ₀ := by rw [add_comm, aleph0_add_nat]
16721682
#align cardinal.nat_add_aleph_0 Cardinal.nat_add_aleph0
16731683

1684+
-- See note [no_index around OfNat.ofNat]
1685+
@[simp]
1686+
theorem ofNat_add_aleph0 {n : ℕ} [Nat.AtLeastTwo n] : no_index (OfNat.ofNat n) + ℵ₀ = ℵ₀ :=
1687+
nat_add_aleph0 n
1688+
1689+
-- See note [no_index around OfNat.ofNat]
1690+
@[simp]
1691+
theorem aleph0_add_ofNat {n : ℕ} [Nat.AtLeastTwo n] : ℵ₀ + no_index (OfNat.ofNat n) = ℵ₀ :=
1692+
aleph0_add_nat n
1693+
16741694
/-- This function sends finite cardinals to the corresponding natural, and infinite cardinals
16751695
to 0. -/
16761696
def toNat : ZeroHom Cardinal ℕ where

Mathlib/SetTheory/Cardinal/Continuum.lean

Lines changed: 23 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -27,7 +27,7 @@ open Cardinal
2727

2828
/-- Cardinality of continuum. -/
2929
def continuum : Cardinal.{u} :=
30-
2 ^ aleph0.{u}
30+
2 ^ ℵ₀
3131
#align cardinal.continuum Cardinal.continuum
3232

3333
-- mathport name: Cardinal.continuum
@@ -131,19 +131,29 @@ theorem continuum_add_aleph0 : 𝔠 + ℵ₀ = 𝔠 :=
131131

132132
@[simp]
133133
theorem continuum_add_self : 𝔠 + 𝔠 = 𝔠 :=
134-
add_eq_right aleph0_le_continuum le_rfl
134+
add_eq_self aleph0_le_continuum
135135
#align cardinal.continuum_add_self Cardinal.continuum_add_self
136136

137137
@[simp]
138138
theorem nat_add_continuum (n : ℕ) : ↑n + 𝔠 = 𝔠 :=
139-
add_eq_right aleph0_le_continuum (nat_lt_continuum n).le
139+
nat_add_eq n aleph0_le_continuum
140140
#align cardinal.nat_add_continuum Cardinal.nat_add_continuum
141141

142142
@[simp]
143143
theorem continuum_add_nat (n : ℕ) : 𝔠 + n = 𝔠 :=
144144
(add_comm _ _).trans (nat_add_continuum n)
145145
#align cardinal.continuum_add_nat Cardinal.continuum_add_nat
146146

147+
-- See note [no_index around OfNat.ofNat]
148+
@[simp]
149+
theorem ofNat_add_continuum {n : ℕ} [Nat.AtLeastTwo n] : no_index (OfNat.ofNat n) + 𝔠 = 𝔠 :=
150+
nat_add_continuum n
151+
152+
-- See note [no_index around OfNat.ofNat]
153+
@[simp]
154+
theorem continuum_add_ofNat {n : ℕ} [Nat.AtLeastTwo n] : 𝔠 + no_index (OfNat.ofNat n) = 𝔠 :=
155+
continuum_add_nat n
156+
147157
/-!
148158
### Multiplication
149159
-/
@@ -174,6 +184,16 @@ theorem continuum_mul_nat {n : ℕ} (hn : n ≠ 0) : 𝔠 * n = 𝔠 :=
174184
(mul_comm _ _).trans (nat_mul_continuum hn)
175185
#align cardinal.continuum_mul_nat Cardinal.continuum_mul_nat
176186

187+
-- See note [no_index around OfNat.ofNat]
188+
@[simp]
189+
theorem ofNat_mul_continuum {n : ℕ} [Nat.AtLeastTwo n] : no_index (OfNat.ofNat n) * 𝔠 = 𝔠 :=
190+
nat_mul_continuum (OfNat.ofNat_ne_zero n)
191+
192+
-- See note [no_index around OfNat.ofNat]
193+
@[simp]
194+
theorem continuum_mul_ofNat {n : ℕ} [Nat.AtLeastTwo n] : 𝔠 * no_index (OfNat.ofNat n) = 𝔠 :=
195+
continuum_mul_nat (OfNat.ofNat_ne_zero n)
196+
177197
/-!
178198
### Power
179199
-/

Mathlib/SetTheory/Cardinal/Ordinal.lean

Lines changed: 4 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -810,8 +810,11 @@ theorem add_nat_eq {a : Cardinal} (n : ℕ) (ha : ℵ₀ ≤ a) : a + n = a :=
810810
add_eq_left ha ((nat_lt_aleph0 _).le.trans ha)
811811
#align cardinal.add_nat_eq Cardinal.add_nat_eq
812812

813+
theorem nat_add_eq {a : Cardinal} (n : ℕ) (ha : ℵ₀ ≤ a) : n + a = a := by
814+
rw [add_comm, add_nat_eq n ha]
815+
813816
theorem add_one_eq {a : Cardinal} (ha : ℵ₀ ≤ a) : a + 1 = a :=
814-
add_eq_left ha (one_le_aleph0.trans ha)
817+
add_one_of_aleph0_le ha
815818
#align cardinal.add_one_eq Cardinal.add_one_eq
816819

817820
--Porting note: removed `simp`, `simp` can prove it

0 commit comments

Comments
 (0)