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

Commit 2f4f8ad

Browse files
vihdzpfpvandoorn
andcommitted
feat(set_theory/principal): Principal ordinals are unbounded (#11755)
Amazingly, this theorem requires no conditions on the operation. Co-authored-by: Floris van Doorn <fpvdoorn@gmail.com>
1 parent 50ee3d5 commit 2f4f8ad

File tree

2 files changed

+40
-1
lines changed

2 files changed

+40
-1
lines changed

src/set_theory/ordinal_arithmetic.lean

Lines changed: 4 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -2107,7 +2107,10 @@ le_sup _ n
21072107
theorem le_nfp_self (f a) : a ≤ nfp f a :=
21082108
iterate_le_nfp f a 0
21092109

2110-
theorem is_normal.lt_nfp {f} (H : is_normal f) {a b} : f b < nfp f a ↔ b < nfp f a :=
2110+
theorem lt_nfp {f : ordinal → ordinal} {a b} : a < nfp f b ↔ ∃ n, a < (f^[n]) b :=
2111+
lt_sup
2112+
2113+
protected theorem is_normal.lt_nfp {f} (H : is_normal f) {a b} : f b < nfp f a ↔ b < nfp f a :=
21112114
lt_sup.trans $ iff.trans
21122115
(by exact
21132116
⟨λ ⟨n, h⟩, ⟨n, lt_of_le_of_lt (H.le_self _) h⟩,

src/set_theory/principal.lean

Lines changed: 36 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -19,6 +19,8 @@ We define principal or indecomposable ordinals, and we prove the standard proper
1919

2020
universe u
2121

22+
noncomputable theory
23+
2224
namespace ordinal
2325

2426
/-! ### Principal ordinals -/
@@ -67,4 +69,38 @@ theorem nfp_le_of_principal {op : ordinal → ordinal → ordinal}
6769
{a o : ordinal} (hao : a < o) (ho : principal op o) : nfp (op a) a ≤ o :=
6870
nfp_le.2 $ λ n, le_of_lt (ho.iterate_lt hao n)
6971

72+
/-! ### Principal ordinals are unbounded -/
73+
74+
/-- The least strict upper bound of `op` applied to all pairs of ordinals less than `o`. This is
75+
essentially a two-argument version of `ordinal.blsub`. -/
76+
def blsub₂ (op : ordinal → ordinal → ordinal) (o : ordinal) : ordinal :=
77+
lsub (λ x : o.out.α × o.out.α, op (typein o.out.r x.1) (typein o.out.r x.2))
78+
79+
theorem lt_blsub₂ (op : ordinal → ordinal → ordinal) {o : ordinal} {a b : ordinal} (ha : a < o)
80+
(hb : b < o) : op a b < blsub₂ op o :=
81+
begin
82+
convert lt_lsub _ (prod.mk (enum o.out.r a (by rwa type_out)) (enum o.out.r b (by rwa type_out))),
83+
simp only [typein_enum]
84+
end
85+
86+
theorem principal_nfp_blsub₂ (op : ordinal → ordinal → ordinal) (o : ordinal) :
87+
principal op (nfp (blsub₂.{u u} op) o) :=
88+
begin
89+
intros a b ha hb,
90+
rw lt_nfp at *,
91+
cases ha with m hm,
92+
cases hb with n hn,
93+
cases le_total ((blsub₂.{u u} op)^[m] o) ((blsub₂.{u u} op)^[n] o) with h h,
94+
{ use n + 1,
95+
rw function.iterate_succ',
96+
exact lt_blsub₂ op (hm.trans_le h) hn },
97+
{ use m + 1,
98+
rw function.iterate_succ',
99+
exact lt_blsub₂ op hm (hn.trans_le h) },
100+
end
101+
102+
theorem unbounded_principal (op : ordinal → ordinal → ordinal) :
103+
set.unbounded (<) {o | principal op o} :=
104+
λ o, ⟨_, principal_nfp_blsub₂ op o, (le_nfp_self _ o).not_lt⟩
105+
70106
end ordinal

0 commit comments

Comments
 (0)