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

Commit f32fda7

Browse files
committed
feat(set_theory/ordinal_arithmetic): More lsub and blsub lemmas (#11848)
We prove variants of `sup_typein`, which serve as analogs for `blsub_id`. We also prove `sup_eq_lsub_or_sup_succ_eq_lsub`, which combines `sup_le_lsub` and `lsub_le_sup_succ`.
1 parent b7360f9 commit f32fda7

File tree

1 file changed

+37
-5
lines changed

1 file changed

+37
-5
lines changed

src/set_theory/ordinal_arithmetic.lean

Lines changed: 37 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -1067,22 +1067,25 @@ end
10671067

10681068
theorem lt_bsup_of_limit {o : ordinal} {f : Π a < o, ordinal}
10691069
(hf : ∀ {a a'} (ha : a < o) (ha' : a' < o), a < a' → f a ha < f a' ha')
1070-
(ho : o.is_limit) (i h) : f i h < bsup o f :=
1071-
lt_of_lt_of_le (hf _ _ $ lt_succ_self i) (le_bsup f i.succ $ ho.2 _ h)
1070+
(ho : ∀ a < o, succ a < o) (i h) : f i h < bsup o f :=
1071+
(hf _ _ $ lt_succ_self i).trans_le (le_bsup f i.succ $ ho _ h)
10721072

10731073
theorem bsup_zero {o : ordinal} (ho : o = 0) (f : Π a < o, ordinal) : bsup o f = 0 :=
10741074
bsup_eq_zero_iff.2 (λ i hi, by { subst ho, exact (ordinal.not_lt_zero i hi).elim })
10751075

10761076
theorem bsup_const {o : ordinal} (ho : o ≠ 0) (a : ordinal) : bsup o (λ _ _, a) = a :=
10771077
le_antisymm (bsup_le.2 (λ _ _, le_rfl)) (le_bsup _ 0 (ordinal.pos_iff_ne_zero.2 ho))
10781078

1079-
theorem bsup_id {o} (ho : is_limit o) : bsup.{u u} o (λ x _, x) = o :=
1079+
theorem bsup_id_limit {o} (ho : ∀ a < o, succ a < o) : bsup.{u u} o (λ x _, x) = o :=
10801080
le_antisymm (bsup_le.2 (λ i hi, hi.le))
10811081
(not_lt.1 (λ h, (lt_bsup_of_limit.{u u} (λ _ _ _ _, id) ho _ h).false))
10821082

1083+
theorem bsup_id_succ (o) : bsup.{u u} (succ o) (λ x _, x) = o :=
1084+
le_antisymm (bsup_le.2 $ (λ o, lt_succ.1)) (le_bsup _ o (lt_succ_self o))
1085+
10831086
theorem is_normal.bsup_eq {f} (H : is_normal f) {o : ordinal} (h : is_limit o) :
10841087
bsup.{u} o (λ x _, f x) = f o :=
1085-
by { rw [←is_normal.bsup.{u u} H (λ x _, x) h.1, bsup_id h] }
1088+
by { rw [←is_normal.bsup.{u u} H (λ x _, x) h.1, bsup_id_limit h.2] }
10861089

10871090
/-- The least strict upper bound of a family of ordinals. -/
10881091
def lsub {ι} (f : ι → ordinal) : ordinal :=
@@ -1100,6 +1103,14 @@ sup_le.2 $ λ i, le_of_lt (lt_lsub f i)
11001103
theorem lsub_le_sup_succ {ι} (f : ι → ordinal) : lsub f ≤ succ (sup f) :=
11011104
lsub_le.2 $ λ i, lt_succ.2 (le_sup f i)
11021105

1106+
theorem sup_eq_lsub_or_sup_succ_eq_lsub {ι} (f : ι → ordinal) :
1107+
sup f = lsub f ∨ (sup f).succ = lsub f :=
1108+
begin
1109+
cases eq_or_lt_of_le (sup_le_lsub f),
1110+
{ exact or.inl h },
1111+
{ exact or.inr ((succ_le.2 h).antisymm (lsub_le_sup_succ f)) }
1112+
end
1113+
11031114
theorem sup_succ_le_lsub {ι} (f : ι → ordinal) : (sup f).succ ≤ lsub f ↔ ∃ i, f i = sup f :=
11041115
begin
11051116
refine ⟨λ h, _, _⟩,
@@ -1189,6 +1200,10 @@ bsup_le.2 (λ i h, le_of_lt (lt_blsub f i h))
11891200
theorem blsub_le_bsup_succ {o} (f : Π a < o, ordinal) : blsub o f ≤ (bsup o f).succ :=
11901201
blsub_le.2 (λ i h, lt_succ.2 (le_bsup f i h))
11911202

1203+
theorem bsup_eq_blsub_or_succ_bsup_eq_blsub {o} (f : Π a < o, ordinal) :
1204+
bsup o f = blsub o f ∨ succ (bsup o f) = blsub o f :=
1205+
by { rw [bsup_eq_sup, blsub_eq_lsub], exact sup_eq_lsub_or_sup_succ_eq_lsub _ }
1206+
11921207
theorem bsup_succ_le_blsub {o} (f : Π a < o, ordinal) :
11931208
(bsup o f).succ ≤ blsub o f ↔ ∃ i hi, f i hi = bsup o f :=
11941209
begin
@@ -1231,7 +1246,7 @@ by rw [blsub_le, lsub_le]; exact
12311246
theorem blsub_const {o : ordinal} (ho : o ≠ 0) (a : ordinal) : blsub.{u v} o (λ _ _, a) = a + 1 :=
12321247
bsup_const.{u v} ho a.succ
12331248

1234-
theorem blsub_id {o} : blsub.{u u} o (λ x _, x) = o :=
1249+
theorem blsub_id (o) : blsub.{u u} o (λ x _, x) = o :=
12351250
begin
12361251
apply le_antisymm,
12371252
{ rw blsub_le,
@@ -1240,6 +1255,23 @@ begin
12401255
exact (lt_blsub.{u u} (λ x _, x) _ h).false
12411256
end
12421257

1258+
theorem lsub_typein (o : ordinal) : lsub.{u u} (typein o.out.r) = o :=
1259+
by { have := blsub_id o, rwa blsub_eq_lsub at this }
1260+
1261+
theorem sup_typein_limit {o : ordinal} (ho : ∀ a, a < o → succ a < o) :
1262+
sup.{u u} (typein o.out.r) = o :=
1263+
by rw (sup_eq_lsub_iff_succ.{u u} (typein o.out.r)).2; rwa lsub_typein o
1264+
1265+
theorem sup_typein_succ (o : ordinal) : sup.{u u} (typein o.succ.out.r) = o :=
1266+
begin
1267+
cases sup_eq_lsub_or_sup_succ_eq_lsub.{u u} (typein o.succ.out.r) with h h,
1268+
{ rw sup_eq_lsub_iff_succ at h,
1269+
simp only [lsub_typein] at h,
1270+
exact (h o (lt_succ_self o)).false.elim },
1271+
rw [←succ_inj, h],
1272+
exact lsub_typein _
1273+
end
1274+
12431275
end ordinal
12441276

12451277
/-! ### Results about injectivity and surjectivity -/

0 commit comments

Comments
 (0)