Skip to content

Commit d8399fe

Browse files
committed
feat: port SetTheory.Ordinal.Topology (#2342)
1 parent 79902bf commit d8399fe

File tree

2 files changed

+237
-0
lines changed

2 files changed

+237
-0
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -970,6 +970,7 @@ import Mathlib.SetTheory.Cardinal.SchroederBernstein
970970
import Mathlib.SetTheory.Lists
971971
import Mathlib.SetTheory.Ordinal.Arithmetic
972972
import Mathlib.SetTheory.Ordinal.Basic
973+
import Mathlib.SetTheory.Ordinal.Topology
973974
import Mathlib.Tactic.Abel
974975
import Mathlib.Tactic.Alias
975976
import Mathlib.Tactic.ApplyFun
Lines changed: 236 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,236 @@
1+
/-
2+
Copyright (c) 2022 Violeta Hernández Palacios. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Violeta Hernández Palacios
5+
6+
! This file was ported from Lean 3 source module set_theory.ordinal.topology
7+
! leanprover-community/mathlib commit 740acc0e6f9adf4423f92a485d0456fc271482da
8+
! Please do not edit these lines, except to modify the commit id
9+
! if you have ported upstream changes.
10+
-/
11+
import Mathlib.SetTheory.Ordinal.Arithmetic
12+
import Mathlib.Topology.Order.Basic
13+
14+
/-!
15+
### Topology of ordinals
16+
17+
We prove some miscellaneous results involving the order topology of ordinals.
18+
19+
### Main results
20+
21+
* `Ordinal.isClosed_iff_sup` / `Ordinal.isClosed_iff_bsup`: A set of ordinals is closed iff it's
22+
closed under suprema.
23+
* `Ordinal.isNormal_iff_strictMono_and_continuous`: A characterization of normal ordinal
24+
functions.
25+
* `Ordinal.enumOrd_isNormal_iff_isClosed`: The function enumerating the ordinals of a set is
26+
normal iff the set is closed.
27+
-/
28+
29+
30+
noncomputable section
31+
32+
universe u v
33+
34+
open Cardinal Order Topology
35+
36+
namespace Ordinal
37+
38+
variable {s : Set Ordinal.{u}} {a : Ordinal.{u}}
39+
40+
instance : TopologicalSpace Ordinal.{u} := Preorder.topology Ordinal.{u}
41+
instance : OrderTopology Ordinal.{u} := ⟨rfl⟩
42+
43+
theorem isOpen_singleton_iff : IsOpen ({a} : Set Ordinal) ↔ ¬IsLimit a := by
44+
refine' ⟨fun h ⟨h₀, hsucc⟩ => _, fun ha => _⟩
45+
· obtain ⟨b, c, hbc, hbc'⟩ :=
46+
(mem_nhds_iff_exists_Ioo_subset' ⟨0, Ordinal.pos_iff_ne_zero.2 h₀⟩ ⟨_, lt_succ a⟩).1
47+
(h.mem_nhds rfl)
48+
have hba := hsucc b hbc.1
49+
exact hba.ne (hbc' ⟨lt_succ b, hba.trans hbc.2⟩)
50+
· rcases zero_or_succ_or_limit a with (rfl | ⟨b, rfl⟩ | ha')
51+
· rw [← bot_eq_zero, ← Set.Iic_bot, ← Iio_succ]
52+
exact isOpen_Iio
53+
· rw [← Set.Icc_self, Icc_succ_left, ← Ioo_succ_right]
54+
exact isOpen_Ioo
55+
· exact (ha ha').elim
56+
#align ordinal.is_open_singleton_iff Ordinal.isOpen_singleton_iff
57+
58+
-- porting note: todo: generalize to a `SuccOrder`
59+
theorem nhds_right' (a : Ordinal) : 𝓝[>] a = ⊥ := (covby_succ a).nhdsWithin_Ioi
60+
61+
-- todo: generalize to a `SuccOrder`
62+
theorem nhds_left'_eq_nhds_ne (a : Ordinal) : 𝓝[<] a = 𝓝[≠] a := by
63+
rw [← nhds_left'_sup_nhds_right', nhds_right', sup_bot_eq]
64+
65+
-- todo: generalize to a `SuccOrder`
66+
theorem nhds_left_eq_nhds (a : Ordinal) : 𝓝[≤] a = 𝓝 a := by
67+
rw [← nhds_left_sup_nhds_right', nhds_right', sup_bot_eq]
68+
69+
-- todo: generalize to a `SuccOrder`
70+
theorem nhdsBasis_Ioc (h : a ≠ 0) : (𝓝 a).HasBasis (· < a) (Set.Ioc · a) :=
71+
nhds_left_eq_nhds a ▸ nhdsWithin_Iic_basis' ⟨0, h.bot_lt⟩
72+
73+
-- todo: generalize to a `SuccOrder`
74+
theorem nhds_eq_pure : 𝓝 a = pure a ↔ ¬IsLimit a :=
75+
(isOpen_singleton_iff_nhds_eq_pure _).symm.trans isOpen_singleton_iff
76+
77+
-- todo: generalize `Ordinal.IsLimit` and this lemma to a `SuccOrder`
78+
theorem isOpen_iff : IsOpen s ↔ ∀ o ∈ s, IsLimit o → ∃ a < o, Set.Ioo a o ⊆ s := by
79+
refine isOpen_iff_mem_nhds.trans <| forall₂_congr fun o ho => ?_
80+
by_cases ho' : IsLimit o
81+
· simp only [(nhdsBasis_Ioc ho'.1).mem_iff, ho', true_implies]
82+
refine exists_congr fun a => and_congr_right fun ha => ?_
83+
simp only [← Set.Ioo_insert_right ha, Set.insert_subset, ho, true_and]
84+
· simp [nhds_eq_pure.2 ho', ho, ho']
85+
#align ordinal.is_open_iff Ordinal.isOpen_iff
86+
87+
open List Set in
88+
theorem mem_closure_tfae (a : Ordinal.{u}) (s : Set Ordinal) :
89+
TFAE [a ∈ closure s,
90+
a ∈ closure (s ∩ Iic a),
91+
(s ∩ Iic a).Nonempty ∧ supₛ (s ∩ Iic a) = a,
92+
∃ t, t ⊆ s ∧ t.Nonempty ∧ BddAbove t ∧ supₛ t = a,
93+
∃ (o : Ordinal.{u}), o ≠ 0 ∧ ∃ (f : ∀ x < o, Ordinal),
94+
(∀ x hx, f x hx ∈ s) ∧ bsup.{u, u} o f = a,
95+
∃ (ι : Type u), Nonempty ι ∧ ∃ f : ι → Ordinal, (∀ i, f i ∈ s) ∧ sup.{u, u} f = a] := by
96+
apply_rules [tfae_of_cycle, Chain.cons, Chain.nil]
97+
· simp only [mem_closure_iff_nhdsWithin_neBot, inter_comm s, nhdsWithin_inter', nhds_left_eq_nhds]
98+
exact id
99+
· intro h
100+
cases' (s ∩ Iic a).eq_empty_or_nonempty with he hne
101+
· simp [he] at h
102+
· refine ⟨hne, (isLUB_of_mem_closure ?_ h).csupₛ_eq hne⟩
103+
exact fun x hx => hx.2
104+
· exact fun h => ⟨_, inter_subset_left _ _, h.1, bddAbove_Iic.mono (inter_subset_right _ _), h.2
105+
· rintro ⟨t, hts, hne, hbdd, rfl⟩
106+
have hlub : IsLUB t (supₛ t) := isLUB_csupₛ hne hbdd
107+
let ⟨y, hyt⟩ := hne
108+
classical
109+
refine ⟨succ (supₛ t), succ_ne_zero _, fun x _ => if x ∈ t then x else y, fun x _ => ?_, ?_⟩
110+
· simp only
111+
split_ifs with h <;> exact hts ‹_›
112+
· refine le_antisymm (bsup_le fun x _ => ?_) (csupₛ_le hne fun x hx => ?_)
113+
· split_ifs <;> exact hlub.1 ‹_›
114+
· refine (if_pos hx).symm.trans_le (le_bsup _ _ <| (hlub.1 hx).trans_lt (lt_succ _))
115+
· rintro ⟨o, h₀, f, hfs, rfl⟩
116+
exact ⟨_, out_nonempty_iff_ne_zero.2 h₀, familyOfBFamily o f, fun _ => hfs _ _, rfl⟩
117+
· rintro ⟨ι, hne, f, hfs, rfl⟩
118+
rw [sup, supᵢ]
119+
exact closure_mono (range_subset_iff.2 hfs) <| csupₛ_mem_closure (range_nonempty f)
120+
(bddAbove_range.{u, u} f)
121+
122+
theorem mem_closure_iff_sup :
123+
a ∈ closure s ↔
124+
∃ (ι : Type u) (_ : Nonempty ι) (f : ι → Ordinal), (∀ i, f i ∈ s) ∧ sup.{u, u} f = a :=
125+
((mem_closure_tfae a s).out 0 5).trans <| by simp only [exists_prop]
126+
#align ordinal.mem_closure_iff_sup Ordinal.mem_closure_iff_sup
127+
128+
theorem mem_closed_iff_sup (hs : IsClosed s) :
129+
a ∈ s ↔ ∃ (ι : Type u) (_hι : Nonempty ι) (f : ι → Ordinal),
130+
(∀ i, f i ∈ s) ∧ sup.{u, u} f = a :=
131+
by rw [← mem_closure_iff_sup, hs.closure_eq]
132+
#align ordinal.mem_closed_iff_sup Ordinal.mem_closed_iff_sup
133+
134+
theorem mem_closure_iff_bsup :
135+
a ∈ closure s ↔
136+
∃ (o : Ordinal) (_ho : o ≠ 0) (f : ∀ a < o, Ordinal),
137+
(∀ i hi, f i hi ∈ s) ∧ bsup.{u, u} o f = a :=
138+
((mem_closure_tfae a s).out 0 4).trans <| by simp only [exists_prop]
139+
#align ordinal.mem_closure_iff_bsup Ordinal.mem_closure_iff_bsup
140+
141+
theorem mem_closed_iff_bsup (hs : IsClosed s) :
142+
a ∈ s ↔
143+
∃ (o : Ordinal) (_ho : o ≠ 0) (f : ∀ a < o, Ordinal),
144+
(∀ i hi, f i hi ∈ s) ∧ bsup.{u, u} o f = a :=
145+
by rw [← mem_closure_iff_bsup, hs.closure_eq]
146+
#align ordinal.mem_closed_iff_bsup Ordinal.mem_closed_iff_bsup
147+
148+
theorem isClosed_iff_sup :
149+
IsClosed s ↔
150+
∀ {ι : Type u}, Nonempty ι → ∀ f : ι → Ordinal, (∀ i, f i ∈ s) → sup.{u, u} f ∈ s := by
151+
use fun hs ι hι f hf => (mem_closed_iff_sup hs).2 ⟨ι, hι, f, hf, rfl⟩
152+
rw [← closure_subset_iff_isClosed]
153+
intro h x hx
154+
rcases mem_closure_iff_sup.1 hx with ⟨ι, hι, f, hf, rfl⟩
155+
exact h hι f hf
156+
#align ordinal.is_closed_iff_sup Ordinal.isClosed_iff_sup
157+
158+
theorem isClosed_iff_bsup :
159+
IsClosed s ↔
160+
∀ {o : Ordinal}, o ≠ 0 → ∀ f : ∀ a < o, Ordinal,
161+
(∀ i hi, f i hi ∈ s) → bsup.{u, u} o f ∈ s := by
162+
rw [isClosed_iff_sup]
163+
refine' ⟨fun H o ho f hf => H (out_nonempty_iff_ne_zero.2 ho) _ _, fun H ι hι f hf => _⟩
164+
· exact fun i => hf _ _
165+
· rw [← bsup_eq_sup]
166+
apply H (type_ne_zero_iff_nonempty.2 hι)
167+
exact fun i hi => hf _
168+
#align ordinal.is_closed_iff_bsup Ordinal.isClosed_iff_bsup
169+
170+
theorem isLimit_of_mem_frontier (ha : a ∈ frontier s) : IsLimit a := by
171+
simp only [frontier_eq_closure_inter_closure, Set.mem_inter_iff, mem_closure_iff] at ha
172+
by_contra h
173+
rw [← isOpen_singleton_iff] at h
174+
rcases ha.1 _ h rfl with ⟨b, hb, hb'⟩
175+
rcases ha.2 _ h rfl with ⟨c, hc, hc'⟩
176+
rw [Set.mem_singleton_iff] at *
177+
subst hb; subst hc
178+
exact hc' hb'
179+
#align ordinal.is_limit_of_mem_frontier Ordinal.isLimit_of_mem_frontier
180+
181+
theorem isNormal_iff_strictMono_and_continuous (f : Ordinal.{u} → Ordinal.{u}) :
182+
IsNormal f ↔ StrictMono f ∧ Continuous f := by
183+
refine' ⟨fun h => ⟨h.strictMono, _⟩, _⟩
184+
· rw [continuous_def]
185+
intro s hs
186+
rw [isOpen_iff] at *
187+
intro o ho ho'
188+
rcases hs _ ho (h.isLimit ho') with ⟨a, ha, has⟩
189+
rw [← IsNormal.bsup_eq.{u, u} h ho', lt_bsup] at ha
190+
rcases ha with ⟨b, hb, hab⟩
191+
exact
192+
⟨b, hb, fun c hc =>
193+
Set.mem_preimage.2 (has ⟨hab.trans (h.strictMono hc.1), h.strictMono hc.2⟩)⟩
194+
· rw [isNormal_iff_strictMono_limit]
195+
rintro ⟨h, h'⟩
196+
refine' ⟨h, fun o ho a h => _⟩
197+
suffices : o ∈ f ⁻¹' Set.Iic a
198+
exact Set.mem_preimage.1 this
199+
rw [mem_closed_iff_sup (IsClosed.preimage h' (@isClosed_Iic _ _ _ _ a))]
200+
exact
201+
⟨_, out_nonempty_iff_ne_zero.2 ho.1, typein (· < ·), fun i => h _ (typein_lt_self i),
202+
sup_typein_limit ho.2
203+
#align ordinal.is_normal_iff_strict_mono_and_continuous Ordinal.isNormal_iff_strictMono_and_continuous
204+
205+
theorem enumOrd_isNormal_iff_isClosed (hs : s.Unbounded (· < ·)) :
206+
IsNormal (enumOrd s) ↔ IsClosed s := by
207+
have Hs := enumOrd_strictMono hs
208+
refine'
209+
fun h => isClosed_iff_sup.2 fun {ι} hι f hf => _, fun h =>
210+
(isNormal_iff_strictMono_limit _).2 ⟨Hs, fun a ha o H => _⟩⟩
211+
· let g : ι → Ordinal.{u} := fun i => (enumOrdOrderIso hs).symm ⟨_, hf i⟩
212+
suffices enumOrd s (sup.{u, u} g) = sup.{u, u} f
213+
by
214+
rw [← this]
215+
exact enumOrd_mem hs _
216+
rw [@IsNormal.sup.{u, u, u} _ h ι g hι]
217+
congr
218+
ext x
219+
change ((enumOrdOrderIso hs) _).val = f x
220+
rw [OrderIso.apply_symm_apply]
221+
· rw [isClosed_iff_bsup] at h
222+
suffices : enumOrd s a ≤ bsup.{u, u} a fun b (_ : b < a) => enumOrd s b
223+
exact this.trans (bsup_le H)
224+
cases' enumOrd_surjective hs _
225+
(h ha.1 (fun b _ => enumOrd s b) fun b _ => enumOrd_mem hs b) with
226+
b hb
227+
rw [← hb]
228+
apply Hs.monotone
229+
by_contra' hba
230+
apply (Hs (lt_succ b)).not_le
231+
rw [hb]
232+
exact le_bsup.{u, u} _ _ (ha.2 _ hba)
233+
#align ordinal.enum_ord_is_normal_iff_is_closed Ordinal.enumOrd_isNormal_iff_isClosed
234+
235+
end Ordinal
236+

0 commit comments

Comments
 (0)