Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit ca2a99d

Browse files
committed
feat(set_theory/ordinal_arithmetic): Normal functions evaluated at ω (#11687)
1 parent cbad62c commit ca2a99d

File tree

2 files changed

+31
-0
lines changed

2 files changed

+31
-0
lines changed

src/set_theory/ordinal.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1151,6 +1151,9 @@ wf.conditionally_complete_linear_order_with_bot 0 $ le_antisymm (ordinal.zero_le
11511151
protected theorem not_lt_zero (o : ordinal) : ¬ o < 0 :=
11521152
not_lt_bot
11531153

1154+
theorem eq_zero_or_pos : ∀ a : ordinal, a = 00 < a :=
1155+
eq_bot_or_bot_lt
1156+
11541157
lemma Inf_eq_omin {s : set ordinal} (hs : s.nonempty) :
11551158
Inf s = omin s hs :=
11561159
begin

src/set_theory/ordinal_arithmetic.lean

Lines changed: 28 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1890,6 +1890,9 @@ theorem omega_le {o : ordinal.{u}} : omega ≤ o ↔ ∀ n : ℕ, (n : ordinal)
18901890
let ⟨n, e⟩ := lt_omega.1 h in
18911891
by rw [e, ← succ_le]; exact H (n+1)⟩
18921892

1893+
theorem sup_nat_cast : sup nat.cast = omega :=
1894+
(sup_le.2 $ λ n, (nat_lt_omega n).le).antisymm $ omega_le.2 $ le_sup _
1895+
18931896
theorem nat_lt_limit {o} (h : is_limit o) : ∀ n : ℕ, (n : ordinal) < o
18941897
| 0 := lt_of_le_of_ne (ordinal.zero_le o) h.1.symm
18951898
| (n+1) := h.2 _ (nat_lt_limit n)
@@ -2056,6 +2059,31 @@ le_antisymm
20562059
(λ b hb, le_of_lt (opow_lt_omega h hb)))
20572060
(le_opow_self _ a1)
20582061

2062+
theorem is_normal.apply_omega {f : ordinal.{u} → ordinal.{u}} (hf : is_normal f) :
2063+
sup.{0 u} (f ∘ nat.cast) = f omega :=
2064+
by rw [←sup_nat_cast, is_normal.sup.{0 u u} hf ⟨0⟩]
2065+
2066+
theorem sup_add_nat (o : ordinal.{u}) : sup (λ n : ℕ, o + n) = o + omega :=
2067+
(add_is_normal o).apply_omega
2068+
2069+
theorem sup_mul_nat (o : ordinal) : sup (λ n : ℕ, o * n) = o * omega :=
2070+
begin
2071+
rcases eq_zero_or_pos o with rfl | ho,
2072+
{ rw zero_mul, exact sup_eq_zero_iff.2 (λ n, zero_mul n) },
2073+
{ exact (mul_is_normal ho).apply_omega }
2074+
end
2075+
2076+
theorem sup_opow_nat {o : ordinal.{u}} (ho : 0 < o) :
2077+
sup (λ n : ℕ, o ^ n) = o ^ omega :=
2078+
begin
2079+
rcases lt_or_eq_of_le (one_le_iff_pos.2 ho) with ho₁ | rfl,
2080+
{ exact (opow_is_normal ho₁).apply_omega },
2081+
{ rw one_opow,
2082+
refine le_antisymm (sup_le.2 (λ n, by rw one_opow)) _,
2083+
convert le_sup _ 0,
2084+
rw [nat.cast_zero, opow_zero] }
2085+
end
2086+
20592087
/-! ### Fixed points of normal functions -/
20602088

20612089
/-- The next fixed point function, the least fixed point of the normal function `f` above `a`. -/

0 commit comments

Comments
 (0)