11/-
22Copyright (c) 2017 Johannes Hölzl. All rights reserved.
33Released under Apache 2.0 license as described in the file LICENSE.
4- Authors: Johannes Hölzl, Mario Carneiro, Yury Kudryashov
4+ Authors: Johannes Hölzl, Mario Carneiro, Yury Kudryashov, Yaël Dillies
55-/
66import algebra.big_operators.intervals
77import algebra.big_operators.order
@@ -16,68 +16,152 @@ import topology.order.basic
1616
1717> THIS FILE IS SYNCHRONIZED WITH MATHLIB4.
1818> Any changes to this file require a corresponding PR to mathlib4.
19+
20+ ## Main declarations
21+
22+ * `bounded_le_nhds_class`: Typeclass stating that neighborhoods are eventually bounded above.
23+ * `bounded_ge_nhds_class`: Typeclass stating that neighborhoods are eventually bounded below.
24+
25+ ## Implementation notes
26+
27+ The same lemmas are true in `ℝ`, `ℝ × ℝ`, `ι → ℝ`, `euclidean_space ι ℝ`. To avoid code
28+ duplication, we provide an ad hoc axiomatisation of the properties we need.
1929-/
2030
2131open filter topological_space
2232open_locale topology classical
2333
2434universes u v
25- variables {α : Type u } {β : Type v }
35+ variables {ι α β R S : Type * } {π : ι → Type * }
2636
27- section liminf_limsup
37+ /-- Ad hoc typeclass stating that neighborhoods are eventually bounded above. -/
38+ class bounded_le_nhds_class (α : Type *) [preorder α] [topological_space α] : Prop :=
39+ (is_bounded_le_nhds (a : α) : (𝓝 a).is_bounded (≤))
40+
41+ /-- Ad hoc typeclass stating that neighborhoods are eventually bounded below. -/
42+ class bounded_ge_nhds_class (α : Type *) [preorder α] [topological_space α] : Prop :=
43+ (is_bounded_ge_nhds (a : α) : (𝓝 a).is_bounded (≥))
44+
45+ section preorder
46+ variables [preorder α] [preorder β] [topological_space α] [topological_space β]
2847
29- section order_closed_topology
30- variables [semilattice_sup α] [topological_space α] [order_topology α]
48+ section bounded_le_nhds_class
49+ variables [bounded_le_nhds_class α] [bounded_le_nhds_class β] {f : filter ι} {u : ι → α} {a : α}
3150
3251lemma is_bounded_le_nhds (a : α) : (𝓝 a).is_bounded (≤) :=
33- (is_top_or_exists_gt a).elim (λ h, ⟨a, eventually_of_forall h⟩) (λ ⟨b, hb⟩, ⟨b, ge_mem_nhds hb⟩)
52+ bounded_le_nhds_class.is_bounded_le_nhds _
3453
35- lemma filter.tendsto.is_bounded_under_le {f : filter β} {u : β → α} {a : α}
36- (h : tendsto u f (𝓝 a)) : f.is_bounded_under (≤) u :=
54+ lemma filter.tendsto.is_bounded_under_le (h : tendsto u f (𝓝 a)) :
55+ f.is_bounded_under (≤) u :=
3756(is_bounded_le_nhds a).mono h
3857
39- lemma filter.tendsto.bdd_above_range_of_cofinite {u : β → α} {a : α}
58+ lemma filter.tendsto.bdd_above_range_of_cofinite [is_directed α (≤)]
4059 (h : tendsto u cofinite (𝓝 a)) : bdd_above (set.range u) :=
4160h.is_bounded_under_le.bdd_above_range_of_cofinite
4261
43- lemma filter.tendsto.bdd_above_range {u : ℕ → α} {a : α}
44- (h : tendsto u at_top (𝓝 a)) : bdd_above (set.range u) :=
62+ lemma filter.tendsto.bdd_above_range [is_directed α (≤)] {u : ℕ → α} (h : tendsto u at_top (𝓝 a)) :
63+ bdd_above (set.range u) :=
4564h.is_bounded_under_le.bdd_above_range
4665
4766lemma is_cobounded_ge_nhds (a : α) : (𝓝 a).is_cobounded (≥) :=
4867(is_bounded_le_nhds a).is_cobounded_flip
4968
50- lemma filter.tendsto.is_cobounded_under_ge {f : filter β} {u : β → α} {a : α}
51- [ne_bot f] (h : tendsto u f (𝓝 a)) : f.is_cobounded_under (≥) u :=
69+ lemma filter.tendsto.is_cobounded_under_ge [ne_bot f] (h : tendsto u f (𝓝 a)) :
70+ f.is_cobounded_under (≥) u :=
5271h.is_bounded_under_le.is_cobounded_flip
5372
54- end order_closed_topology
73+ instance : bounded_ge_nhds_class αᵒᵈ := ⟨@is_bounded_le_nhds α _ _ _⟩
74+
75+ instance : bounded_le_nhds_class (α × β) :=
76+ begin
77+ refine ⟨λ x, _⟩,
78+ obtain ⟨a, ha⟩ := is_bounded_le_nhds x.1 ,
79+ obtain ⟨b, hb⟩ := is_bounded_le_nhds x.2 ,
80+ rw [←@prod.mk.eta _ _ x, nhds_prod_eq],
81+ exact ⟨(a, b), ha.prod_mk hb⟩,
82+ end
83+
84+ instance [finite ι] [Π i, preorder (π i)] [Π i, topological_space (π i)]
85+ [Π i, bounded_le_nhds_class (π i)] : bounded_le_nhds_class (Π i, π i) :=
86+ begin
87+ refine ⟨λ x, _⟩,
88+ rw nhds_pi,
89+ choose f hf using λ i, is_bounded_le_nhds (x i),
90+ exact ⟨f, eventually_pi hf⟩,
91+ end
92+
93+ end bounded_le_nhds_class
5594
56- section order_closed_topology
57- variables [semilattice_inf α] [topological_space α] [order_topology α]
95+ section bounded_ge_nhds_class
96+ variables [bounded_ge_nhds_class α] [bounded_ge_nhds_class β] {f : filter ι} {u : ι → α} {a : α}
5897
59- lemma is_bounded_ge_nhds (a : α) : (𝓝 a).is_bounded (≥) := @is_bounded_le_nhds αᵒᵈ _ _ _ a
98+ lemma is_bounded_ge_nhds (a : α) : (𝓝 a).is_bounded (≥) :=
99+ bounded_ge_nhds_class.is_bounded_ge_nhds _
60100
61- lemma filter.tendsto.is_bounded_under_ge {f : filter β} {u : β → α} {a : α}
62- (h : tendsto u f (𝓝 a)) : f.is_bounded_under (≥) u :=
101+ lemma filter.tendsto.is_bounded_under_ge (h : tendsto u f (𝓝 a)) :
102+ f.is_bounded_under (≥) u :=
63103(is_bounded_ge_nhds a).mono h
64104
65- lemma filter.tendsto.bdd_below_range_of_cofinite {u : β → α} {a : α}
105+ lemma filter.tendsto.bdd_below_range_of_cofinite [is_directed α (≥)]
66106 (h : tendsto u cofinite (𝓝 a)) : bdd_below (set.range u) :=
67107h.is_bounded_under_ge.bdd_below_range_of_cofinite
68108
69- lemma filter.tendsto.bdd_below_range {u : ℕ → α} {a : α}
70- (h : tendsto u at_top (𝓝 a)) : bdd_below (set.range u) :=
109+ lemma filter.tendsto.bdd_below_range [is_directed α (≥)] {u : ℕ → α} (h : tendsto u at_top (𝓝 a)) :
110+ bdd_below (set.range u) :=
71111h.is_bounded_under_ge.bdd_below_range
72112
73113lemma is_cobounded_le_nhds (a : α) : (𝓝 a).is_cobounded (≤) :=
74114(is_bounded_ge_nhds a).is_cobounded_flip
75115
76- lemma filter.tendsto.is_cobounded_under_le {f : filter β} {u : β → α} {a : α}
77- [ne_bot f] (h : tendsto u f (𝓝 a)) : f.is_cobounded_under (≤) u :=
116+ lemma filter.tendsto.is_cobounded_under_le [ne_bot f] (h : tendsto u f (𝓝 a)) :
117+ f.is_cobounded_under (≤) u :=
78118h.is_bounded_under_ge.is_cobounded_flip
79119
80- end order_closed_topology
120+ instance : bounded_le_nhds_class αᵒᵈ := ⟨@is_bounded_ge_nhds α _ _ _⟩
121+
122+ instance : bounded_ge_nhds_class (α × β) :=
123+ begin
124+ refine ⟨λ x, _⟩,
125+ obtain ⟨a, ha⟩ := is_bounded_ge_nhds x.1 ,
126+ obtain ⟨b, hb⟩ := is_bounded_ge_nhds x.2 ,
127+ rw [←@prod.mk.eta _ _ x, nhds_prod_eq],
128+ exact ⟨(a, b), ha.prod_mk hb⟩,
129+ end
130+
131+ instance [finite ι] [Π i, preorder (π i)] [Π i, topological_space (π i)]
132+ [Π i, bounded_ge_nhds_class (π i)] : bounded_ge_nhds_class (Π i, π i) :=
133+ begin
134+ refine ⟨λ x, _⟩,
135+ rw nhds_pi,
136+ choose f hf using λ i, is_bounded_ge_nhds (x i),
137+ exact ⟨f, eventually_pi hf⟩,
138+ end
139+
140+ end bounded_ge_nhds_class
141+
142+ @[priority 100 ] -- See note [lower instance priority]
143+ instance order_top.to_bounded_le_nhds_class [order_top α] : bounded_le_nhds_class α :=
144+ ⟨λ a, is_bounded_le_of_top⟩
145+
146+ @[priority 100 ] -- See note [lower instance priority]
147+ instance order_bot.to_bounded_ge_nhds_class [order_bot α] : bounded_ge_nhds_class α :=
148+ ⟨λ a, is_bounded_ge_of_bot⟩
149+
150+ @[priority 100 ] -- See note [lower instance priority]
151+ instance order_topology.to_bounded_le_nhds_class [is_directed α (≤)] [order_topology α] :
152+ bounded_le_nhds_class α :=
153+ ⟨λ a, (is_top_or_exists_gt a).elim (λ h, ⟨a, eventually_of_forall h⟩) $ Exists.imp $ λ b,
154+ ge_mem_nhds⟩
155+
156+ @[priority 100 ] -- See note [lower instance priority]
157+ instance order_topology.to_bounded_ge_nhds_class [is_directed α (≥)] [order_topology α] :
158+ bounded_ge_nhds_class α :=
159+ ⟨λ a, (is_bot_or_exists_lt a).elim (λ h, ⟨a, eventually_of_forall h⟩) $ Exists.imp $ λ b,
160+ le_mem_nhds⟩
161+
162+ end preorder
163+
164+ section liminf_limsup
81165
82166section conditionally_complete_linear_order
83167variables [conditionally_complete_linear_order α]
@@ -224,7 +308,7 @@ end liminf_limsup
224308
225309section monotone
226310
227- variables {ι R S : Type *} { F : filter ι} [ne_bot F]
311+ variables {F : filter ι} [ne_bot F]
228312 [complete_linear_order R] [topological_space R] [order_topology R]
229313 [complete_linear_order S] [topological_space S] [order_topology S]
230314
@@ -334,7 +418,7 @@ open_locale topology
334418
335419open filter set
336420
337- variables {ι : Type *} {R : Type *} [complete_linear_order R] [topological_space R] [order_topology R]
421+ variables [complete_linear_order R] [topological_space R] [order_topology R]
338422
339423lemma infi_eq_of_forall_le_of_tendsto {x : R} {as : ι → R}
340424 (x_le : ∀ i, x ≤ as i) {F : filter ι} [filter.ne_bot F] (as_lim : filter.tendsto as F (𝓝 x)) :
@@ -349,7 +433,7 @@ lemma supr_eq_of_forall_le_of_tendsto {x : R} {as : ι → R}
349433 (⨆ i, as i) = x :=
350434@infi_eq_of_forall_le_of_tendsto ι (order_dual R) _ _ _ x as le_x F _ as_lim
351435
352- lemma Union_Ici_eq_Ioi_of_lt_of_tendsto {ι : Type *} (x : R) {as : ι → R} (x_lt : ∀ i, x < as i)
436+ lemma Union_Ici_eq_Ioi_of_lt_of_tendsto (x : R) {as : ι → R} (x_lt : ∀ i, x < as i)
353437 {F : filter ι} [filter.ne_bot F] (as_lim : filter.tendsto as F (𝓝 x)) :
354438 (⋃ (i : ι), Ici (as i)) = Ioi x :=
355439begin
@@ -361,10 +445,10 @@ begin
361445 exact Union_Ici_eq_Ioi_infi obs,
362446end
363447
364- lemma Union_Iic_eq_Iio_of_lt_of_tendsto {ι : Type *} (x : R) {as : ι → R} (lt_x : ∀ i, as i < x)
448+ lemma Union_Iic_eq_Iio_of_lt_of_tendsto (x : R) {as : ι → R} (lt_x : ∀ i, as i < x)
365449 {F : filter ι} [filter.ne_bot F] (as_lim : filter.tendsto as F (𝓝 x)) :
366450 (⋃ (i : ι), Iic (as i)) = Iio x :=
367- @Union_Ici_eq_Ioi_of_lt_of_tendsto (order_dual R) _ _ _ ι x as lt_x F _ as_lim
451+ @Union_Ici_eq_Ioi_of_lt_of_tendsto ι Rᵒᵈ _ _ _ _ _ lt_x F _ as_lim
368452
369453end infi_and_supr
370454
0 commit comments