Skip to content

Commit

Permalink
chore(SetTheory/Ordinal/Arithmetic): omega0_isLimit → `isLimit_omeg…
Browse files Browse the repository at this point in the history
…a0` (#17744)

I missed this rename in #17673.
  • Loading branch information
vihdzp committed Oct 17, 2024
1 parent b1267a6 commit 08804ac
Show file tree
Hide file tree
Showing 6 changed files with 35 additions and 29 deletions.
6 changes: 3 additions & 3 deletions Mathlib/SetTheory/Cardinal/Aleph.lean
Original file line number Diff line number Diff line change
Expand Up @@ -249,7 +249,7 @@ theorem aleph'_limit {o : Ordinal} (ho : o.IsLimit) : aleph' o = ⨆ a : Iio o,
@[simp]
theorem aleph'_omega0 : aleph' ω = ℵ₀ :=
eq_of_forall_ge_iff fun c => by
simp only [aleph'_le_of_limit omega0_isLimit, lt_omega0, exists_imp, aleph0_le]
simp only [aleph'_le_of_limit isLimit_omega0, lt_omega0, exists_imp, aleph0_le]
exact forall_swap.trans (forall_congr' fun n => by simp only [forall_eq, aleph'_nat])

@[deprecated (since := "2024-09-30")]
Expand Down Expand Up @@ -333,12 +333,12 @@ instance nonempty_toType_aleph (o : Ordinal) : Nonempty (ℵ_ o).ord.toType := b
exact fun h => (ord_injective h).not_gt (aleph_pos o)

theorem ord_aleph_isLimit (o : Ordinal) : (ℵ_ o).ord.IsLimit :=
ord_isLimit <| aleph0_le_aleph _
isLimit_ord <| aleph0_le_aleph _

instance (o : Ordinal) : NoMaxOrder (ℵ_ o).ord.toType :=
toType_noMax_of_succ_lt (ord_aleph_isLimit o).2

theorem exists_aleph {c : Cardinal} : ℵ₀ ≤ c ↔ ∃ o, c = aleph o :=
theorem exists_aleph {c : Cardinal} : ℵ₀ ≤ c ↔ ∃ o, c = ℵ_ o :=
fun h =>
⟨aleph'.symm c - ω, by
rw [aleph_eq_aleph', Ordinal.add_sub_cancel_of_le, aleph'.apply_symm_apply]
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/SetTheory/Cardinal/Arithmetic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -82,7 +82,7 @@ theorem mul_eq_self {c : Cardinal} (h : ℵ₀ ≤ c) : c * c = c := by
· exact (mul_lt_aleph0 qo qo).trans_le ol
· suffices (succ (typein LT.lt (g p))).card < ⟦α⟧ from (IH _ this qo).trans_lt this
rw [← lt_ord]
apply (ord_isLimit ol).2
apply (isLimit_ord ol).2
rw [mk'_def, e]
apply typein_lt_type

Expand Down
12 changes: 6 additions & 6 deletions Mathlib/SetTheory/Cardinal/Cofinality.lean
Original file line number Diff line number Diff line change
Expand Up @@ -681,7 +681,7 @@ theorem aleph_cof {o : Ordinal} (ho : o.IsLimit) : (aleph o).ord.cof = o.cof :=

@[simp]
theorem cof_omega0 : cof ω = ℵ₀ :=
(aleph0_le_cof.2 omega0_isLimit).antisymm' <| by
(aleph0_le_cof.2 isLimit_omega0).antisymm' <| by
rw [← card_omega0]
apply cof_le_card

Expand Down Expand Up @@ -862,7 +862,7 @@ theorem mk_bounded_subset {α : Type*} (h : ∀ x < #α, (2^x) < #α) {r : α
· refine @mk_le_of_injective α _ (fun x => Subtype.mk {x} ?_) ?_
· apply bounded_singleton
rw [← hr]
apply ord_isLimit ha
apply isLimit_ord ha
· intro a b hab
simpa [singleton_eq_singleton_iff] using hab

Expand All @@ -881,7 +881,7 @@ theorem mk_subset_mk_lt_cof {α : Type*} (h : ∀ x < #α, (2^x) < #α) :
exact lt_cof_type hs
· refine @mk_le_of_injective α _ (fun x => Subtype.mk {x} ?_) ?_
· rw [mk_singleton]
exact one_lt_aleph0.trans_le (aleph0_le_cof.2 (ord_isLimit h'.aleph0_le))
exact one_lt_aleph0.trans_le (aleph0_le_cof.2 (isLimit_ord h'.aleph0_le))
· intro a b hab
simpa [singleton_eq_singleton_iff] using hab

Expand Down Expand Up @@ -917,7 +917,7 @@ theorem isRegular_succ {c : Cardinal.{u}} (h : ℵ₀ ≤ c) : IsRegular (succ c
(by
cases' Quotient.exists_rep (@succ Cardinal _ _ c) with α αe; simp only [mk'_def] at αe
rcases ord_eq α with ⟨r, wo, re⟩
have := ord_isLimit (h.trans (le_succ _))
have := isLimit_ord (h.trans (le_succ _))
rw [← αe, re] at this ⊢
rcases cof_eq' r this with ⟨S, H, Se⟩
rw [← Se]
Expand Down Expand Up @@ -1123,7 +1123,7 @@ theorem derivFamily_lt_ord_lift {ι} {f : ι → Ordinal → Ordinal} {c} (hc :
rw [derivFamily_succ]
exact
nfpFamily_lt_ord_lift hω (by rwa [hc.cof_eq]) hf
((ord_isLimit hc.1).2 _ (hb ((lt_succ b).trans hb')))
((isLimit_ord hc.1).2 _ (hb ((lt_succ b).trans hb')))
| H₃ b hb H =>
intro hb'
rw [derivFamily_limit f hb]
Expand Down Expand Up @@ -1170,7 +1170,7 @@ theorem univ_inaccessible : IsInaccessible univ.{u, v} :=
theorem lt_power_cof {c : Cardinal.{u}} : ℵ₀ ≤ c → c < (c^cof c.ord) :=
Quotient.inductionOn c fun α h => by
rcases ord_eq α with ⟨r, wo, re⟩
have := ord_isLimit h
have := isLimit_ord h
rw [mk'_def, re] at this ⊢
rcases cof_eq' r this with ⟨S, H, Se⟩
have := sum_lt_prod (fun a : S => #{ x // r x a }) (fun _ => #α) fun i => ?_
Expand Down
20 changes: 13 additions & 7 deletions Mathlib/SetTheory/Ordinal/Arithmetic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -2263,13 +2263,16 @@ theorem one_lt_omega0 : 1 < ω := by simpa only [Nat.cast_one] using nat_lt_omeg
@[deprecated (since := "2024-09-30")]
alias one_lt_omega := one_lt_omega0

theorem omega0_isLimit : IsLimit ω :=
theorem isLimit_omega0 : IsLimit ω :=
⟨omega0_ne_zero, fun o h => by
let ⟨n, e⟩ := lt_omega0.1 h
rw [e]; exact nat_lt_omega0 (n + 1)⟩

@[deprecated (since := "2024-10-14")]
alias omega0_isLimit := isLimit_omega0

@[deprecated (since := "2024-09-30")]
alias omega_isLimit := omega0_isLimit
alias omega_isLimit := isLimit_omega0

theorem omega0_le {o : Ordinal} : ω ≤ o ↔ ∀ n : ℕ, ↑n ≤ o :=
fun h n => (nat_lt_omega0 _).le.trans h, fun H =>
Expand Down Expand Up @@ -2306,14 +2309,14 @@ theorem isLimit_iff_omega0_dvd {a : Ordinal} : IsLimit a ↔ a ≠ 0 ∧ ω ∣
refine ⟨fun l => ⟨l.1, ⟨a / ω, le_antisymm ?_ (mul_div_le _ _)⟩⟩, fun h => ?_⟩
· refine (limit_le l).2 fun x hx => le_of_lt ?_
rw [← div_lt omega0_ne_zero, ← succ_le_iff, le_div omega0_ne_zero, mul_succ,
add_le_of_limit omega0_isLimit]
add_le_of_limit isLimit_omega0]
intro b hb
rcases lt_omega0.1 hb with ⟨n, rfl⟩
exact
(add_le_add_right (mul_div_le _ _) _).trans
(lt_sub.1 <| nat_lt_limit (isLimit_sub l hx) _).le
· rcases h with ⟨a0, b, rfl⟩
refine isLimit_mul_left omega0_isLimit (Ordinal.pos_iff_ne_zero.2 <| mt ?_ a0)
refine isLimit_mul_left isLimit_omega0 (Ordinal.pos_iff_ne_zero.2 <| mt ?_ a0)
intro e
simp only [e, mul_zero]

Expand Down Expand Up @@ -2394,7 +2397,7 @@ namespace Cardinal

open Ordinal

theorem ord_isLimit {c} (co : ℵ₀ ≤ c) : (ord c).IsLimit := by
theorem isLimit_ord {c} (co : ℵ₀ ≤ c) : (ord c).IsLimit := by
refine ⟨fun h => aleph0_ne_zero ?_, fun a => lt_imp_lt_of_le_imp_le fun h => ?_⟩
· rw [← Ordinal.le_zero, ord_le] at h
simpa only [card_zero, nonpos_iff_eq_zero] using co.trans h
Expand All @@ -2403,10 +2406,13 @@ theorem ord_isLimit {c} (co : ℵ₀ ≤ c) : (ord c).IsLimit := by
rw [← ord_le, ← le_succ_of_isLimit, ord_le]
· exact co.trans h
· rw [ord_aleph0]
exact Ordinal.omega0_isLimit
exact Ordinal.isLimit_omega0

@[deprecated (since := "2024-10-14")]
alias ord_isLimit := isLimit_ord

theorem noMaxOrder {c} (h : ℵ₀ ≤ c) : NoMaxOrder c.ord.toType :=
toType_noMax_of_succ_lt (ord_isLimit h).2
toType_noMax_of_succ_lt (isLimit_ord h).2

end Cardinal

Expand Down
14 changes: 7 additions & 7 deletions Mathlib/SetTheory/Ordinal/Notation.lean
Original file line number Diff line number Diff line change
Expand Up @@ -577,7 +577,7 @@ theorem repr_mul : ∀ (o₁ o₂) [NF o₁] [NF o₂], repr (o₁ * o₂) = rep
rw [← mul_assoc]
congr 2
have := mt repr_inj.1 e0
rw [add_mul_limit ao (isLimit_opow_left omega0_isLimit this), mul_assoc,
rw [add_mul_limit ao (isLimit_opow_left isLimit_omega0 this), mul_assoc,
mul_omega0_dvd (natCast_pos.2 n₁.pos) (nat_lt_omega0 _)]
simpa using opow_dvd_opow ω (one_le_iff_ne_zero.2 this)

Expand Down Expand Up @@ -790,7 +790,7 @@ theorem repr_opow_aux₁ {e a} [Ne : NF e] [Na : NF a] {a' : Ordinal} (e0 : repr
have := omega0_le_oadd e n a
rw [repr] at this
refine le_antisymm ?_ (opow_le_opow_left _ this)
apply (opow_le_of_limit ((opow_pos _ omega0_pos).trans_le this).ne' omega0_isLimit).2
apply (opow_le_of_limit ((opow_pos _ omega0_pos).trans_le this).ne' isLimit_omega0).2
intro b l
have := (No.below_of_lt (lt_succ _)).repr_lt
rw [repr] at this
Expand All @@ -801,8 +801,8 @@ theorem repr_opow_aux₁ {e a} [Ne : NF e] [Na : NF a] {a' : Ordinal} (e0 : repr
· apply (mul_le_mul_left' (le_succ b) _).trans
rw [← add_one_eq_succ, add_mul_succ _ (one_add_of_omega0_le h), add_one_eq_succ, succ_le_iff,
Ordinal.mul_lt_mul_iff_left (Ordinal.pos_iff_ne_zero.2 e0)]
exact omega0_isLimit.2 _ l
· apply (principal_mul_omega0 (omega0_isLimit.2 _ h) l).le.trans
exact isLimit_omega0.2 _ l
· apply (principal_mul_omega0 (isLimit_omega0.2 _ h) l).le.trans
simpa using mul_le_mul_right' (one_le_iff_ne_zero.2 e0) ω

section
Expand Down Expand Up @@ -966,7 +966,7 @@ private theorem exists_lt_add {α} [hα : Nonempty α] {o : Ordinal} {f : α →

private theorem exists_lt_mul_omega0' {o : Ordinal} ⦃a⦄ (h : a < o * ω) :
∃ i : ℕ, a < o * ↑i + o := by
obtain ⟨i, hi, h'⟩ := (lt_mul_of_limit omega0_isLimit).1 h
obtain ⟨i, hi, h'⟩ := (lt_mul_of_limit isLimit_omega0).1 h
obtain ⟨i, rfl⟩ := lt_omega0.1 hi
exact ⟨i, h'.trans_le (le_add_right _ _)⟩

Expand Down Expand Up @@ -1025,13 +1025,13 @@ theorem fundamentalSequence_has_prop (o) : FundamentalSequenceProp o (fundamenta
· exact ⟨rfl, inferInstance⟩
· have := opow_pos (repr a') omega0_pos
refine
⟨isLimit_mul this omega0_isLimit, fun i =>
⟨isLimit_mul this isLimit_omega0, fun i =>
⟨this, ?_, fun H => @NF.oadd_zero _ _ (iha.2 H.fst)⟩, exists_lt_mul_omega0'⟩
rw [← mul_succ, ← natCast_succ, Ordinal.mul_lt_mul_iff_left this]
apply nat_lt_omega0
· have := opow_pos (repr a') omega0_pos
refine
⟨isLimit_add _ (isLimit_mul this omega0_isLimit), fun i => ⟨this, ?_, ?_⟩,
⟨isLimit_add _ (isLimit_mul this isLimit_omega0), fun i => ⟨this, ?_, ?_⟩,
exists_lt_add exists_lt_mul_omega0'⟩
· rw [← mul_succ, ← natCast_succ, Ordinal.mul_lt_mul_iff_left this]
apply nat_lt_omega0
Expand Down
10 changes: 5 additions & 5 deletions Mathlib/SetTheory/Ordinal/Principal.lean
Original file line number Diff line number Diff line change
Expand Up @@ -194,7 +194,7 @@ theorem add_omega0_opow {a b : Ordinal} (h : a < ω ^ b) : a + ω ^ b = ω ^ b :
· rw [opow_zero, ← succ_zero, lt_succ_iff, Ordinal.le_zero] at h
rw [h, zero_add]
· rw [opow_succ] at h
rcases (lt_mul_of_limit omega0_isLimit).1 h with ⟨x, xo, ax⟩
rcases (lt_mul_of_limit isLimit_omega0).1 h with ⟨x, xo, ax⟩
apply (add_le_add_right ax.le _).trans
rw [opow_succ, ← mul_add, add_omega0 xo]
· rcases (lt_opow_of_limit omega0_ne_zero l).1 h with ⟨x, xb, ax⟩
Expand Down Expand Up @@ -228,7 +228,7 @@ theorem principal_add_iff_zero_or_omega0_opow {o : Ordinal} :
fun ⟨b, e⟩ => e.symm ▸ fun a => add_omega0_opow⟩
have := H _ h
have := lt_opow_succ_log_self one_lt_omega0 o
rw [opow_succ, lt_mul_of_limit omega0_isLimit] at this
rw [opow_succ, lt_mul_of_limit isLimit_omega0] at this
rcases this with ⟨a, ao, h'⟩
rcases lt_omega0.1 ao with ⟨n, rfl⟩
clear ao
Expand Down Expand Up @@ -347,7 +347,7 @@ theorem mul_lt_omega0_opow {a b c : Ordinal} (c0 : 0 < c) (ha : a < ω ^ c) (hb
· exact (lt_irrefl _).elim c0
· rw [opow_succ] at ha
obtain ⟨n, hn, an⟩ :=
((isNormal_mul_right <| opow_pos _ omega0_pos).limit_lt omega0_isLimit).1 ha
((isNormal_mul_right <| opow_pos _ omega0_pos).limit_lt isLimit_omega0).1 ha
apply (mul_le_mul_right' (le_of_lt an) _).trans_lt
rw [opow_succ, mul_assoc, mul_lt_mul_iff_left (opow_pos _ omega0_pos)]
exact principal_mul_omega0 hn hb
Expand All @@ -366,7 +366,7 @@ theorem mul_omega0_opow_opow {a b : Ordinal} (a0 : 0 < a) (h : a < ω ^ ω ^ b)
exact mul_omega0 a0 h
· apply le_antisymm
· obtain ⟨x, xb, ax⟩ :=
(lt_opow_of_limit omega0_ne_zero (isLimit_opow_left omega0_isLimit b0)).1 h
(lt_opow_of_limit omega0_ne_zero (isLimit_opow_left isLimit_omega0 b0)).1 h
apply (mul_le_mul_right' (le_of_lt ax) _).trans
rw [← opow_add, add_omega0_opow xb]
· conv_lhs => rw [← one_mul (ω ^ _)]
Expand Down Expand Up @@ -442,7 +442,7 @@ theorem principal_opow_omega0 : Principal (· ^ ·) ω := fun a b ha hb =>
alias principal_opow_omega := principal_opow_omega0

theorem opow_omega0 {a : Ordinal} (a1 : 1 < a) (h : a < ω) : a ^ ω = ω :=
((opow_le_of_limit (one_le_iff_ne_zero.1 <| le_of_lt a1) omega0_isLimit).2 fun _ hb =>
((opow_le_of_limit (one_le_iff_ne_zero.1 <| le_of_lt a1) isLimit_omega0).2 fun _ hb =>
(principal_opow_omega0 h hb).le).antisymm
(right_le_opow _ a1)

Expand Down

0 comments on commit 08804ac

Please sign in to comment.