Skip to content

Commit 7a8a81c

Browse files
committed
feat: liminf and limsup add/mul lemmas in Real (#18365)
Lemmas such as `limsup u f + liminf v f ≤ limsup (u + v) f` (and its multiplicative avatars) in (basically) reals. These lemmas had already been done in EReal/ENNReal. The strategy is the same; using reals brings some complexity with the gestion of `IsBoundedUnder`/`IsCoboundedUnder` properties. Summary of the modifications: * In [Mathlib/Algebra/Order/Group/DenselyOrdered.lean](master...D-Thomine/limsup_add_mul#diff-87f5a269787b67b9ec09b3e5d2d10ab4a5cf94fb0ea67087ca8efe323e26d787): proof of `le_mul_of_forall_lt` and `mul_le_of_forall_lt`, as well as their additive versions, which are key for the final results on (R, +). * In [Mathlib/Algebra/Order/Field/Basic.lean](master...D-Thomine/limsup_add_mul#diff-38676ab53a1bf62d5eb4d1a87a4c7607f056a6c13968c1683148ddb178863343): proof of `le_mul_of_forall_lt'` and `mul_le_of_forall_lt_of_nonneg`, which are key for the final results on (R+, *). * In [Mathlib/Order/LiminfLimsup.lean](master...D-Thomine/limsup_add_mul#diff-15958880e4301dca2ca609bc984fb6192aa21a3013677c7ddb10b4d885927026): four lemmas from `IsCobounded.frequently_le` to `IsCobounded.of_frequently_le` are moved to the beginning of the file, so they can be used within. Four new elementary lemmas about `IsBoundedUnder` which I found quite practical. Four new lemmas about boundedness/coboundedness of sums and products. * In [Mathlib/Topology/Algebra/Order/LiminfLimsup.lean](master...D-Thomine/limsup_add_mul#diff-47f3840d4d3e20e25976271e7e13b4542be1cd0c5fd147e72bb1a0736440624b): eight new lemmas about addition/multiplication and limsup/liminfs.
1 parent 3571e2c commit 7a8a81c

File tree

4 files changed

+338
-55
lines changed

4 files changed

+338
-55
lines changed

Mathlib/Algebra/Order/Field/Basic.lean

Lines changed: 41 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -747,6 +747,47 @@ theorem le_of_forall_sub_le (h : ∀ ε > 0, b - ε ≤ a) : b ≤ a := by
747747
simpa only [@and_comm ((0 : α) < _), lt_sub_iff_add_lt, gt_iff_lt] using
748748
exists_add_lt_and_pos_of_lt h
749749

750+
private lemma exists_lt_mul_left_of_nonneg {a b c : α} (ha : 0 ≤ a) (hc : 0 ≤ c) (h : c < a * b) :
751+
∃ a' ∈ Set.Ico 0 a, c < a' * b := by
752+
rcases eq_or_lt_of_le ha with rfl | ha
753+
· rw [zero_mul] at h; exact (not_le_of_lt h hc).rec
754+
rcases lt_trichotomy b 0 with hb | rfl | hb
755+
· exact (not_le_of_lt (h.trans (mul_neg_of_pos_of_neg ha hb)) hc).rec
756+
· rw [mul_zero] at h; exact (not_le_of_lt h hc).rec
757+
· obtain ⟨a', ha', a_a'⟩ := exists_between ((div_lt_iff₀ hb).2 h)
758+
exact ⟨a', ⟨(div_nonneg hc hb.le).trans ha'.le, a_a'⟩, (div_lt_iff₀ hb).1 ha'⟩
759+
760+
private lemma exists_lt_mul_right_of_nonneg {a b c : α} (ha : 0 ≤ a) (hc : 0 ≤ c) (h : c < a * b) :
761+
∃ b' ∈ Set.Ico 0 b, c < a * b' := by
762+
have hb : 0 < b := pos_of_mul_pos_right (hc.trans_lt h) ha
763+
simp_rw [mul_comm a] at h ⊢
764+
exact exists_lt_mul_left_of_nonneg hb.le hc h
765+
766+
private lemma exists_mul_left_lt₀ {a b c : α} (hc : a * b < c) : ∃ a' > a, a' * b < c := by
767+
rcases le_or_lt b 0 with hb | hb
768+
· obtain ⟨a', ha'⟩ := exists_gt a
769+
exact ⟨a', ha', hc.trans_le' (antitone_mul_right hb ha'.le)⟩
770+
· obtain ⟨a', ha', hc'⟩ := exists_between ((lt_div_iff₀ hb).2 hc)
771+
exact ⟨a', ha', (lt_div_iff₀ hb).1 hc'⟩
772+
773+
private lemma exists_mul_right_lt₀ {a b c : α} (hc : a * b < c) : ∃ b' > b, a * b' < c := by
774+
simp_rw [mul_comm a] at hc ⊢; exact exists_mul_left_lt₀ hc
775+
776+
lemma le_mul_of_forall_lt₀ {a b c : α} (h : ∀ a' > a, ∀ b' > b, c ≤ a' * b') : c ≤ a * b := by
777+
refine le_of_forall_le_of_dense fun d hd ↦ ?_
778+
obtain ⟨a', ha', hd⟩ := exists_mul_left_lt₀ hd
779+
obtain ⟨b', hb', hd⟩ := exists_mul_right_lt₀ hd
780+
exact (h a' ha' b' hb').trans hd.le
781+
782+
lemma mul_le_of_forall_lt_of_nonneg {a b c : α} (ha : 0 ≤ a) (hc : 0 ≤ c)
783+
(h : ∀ a' ≥ 0, a' < a → ∀ b' ≥ 0, b' < b → a' * b' ≤ c) : a * b ≤ c := by
784+
refine le_of_forall_ge_of_dense fun d d_ab ↦ ?_
785+
rcases lt_or_le d 0 with hd | hd
786+
· exact hd.le.trans hc
787+
obtain ⟨a', ha', d_ab⟩ := exists_lt_mul_left_of_nonneg ha hd d_ab
788+
obtain ⟨b', hb', d_ab⟩ := exists_lt_mul_right_of_nonneg ha'.1 hd d_ab
789+
exact d_ab.le.trans (h a' ha'.1 ha'.2 b' hb'.1 hb'.2)
790+
750791
theorem mul_self_inj_of_nonneg (a0 : 0 ≤ a) (b0 : 0 ≤ b) : a * a = b * b ↔ a = b :=
751792
mul_self_eq_mul_self_iff.trans <|
752793
or_iff_left_of_imp fun h => by

Mathlib/Algebra/Order/Group/DenselyOrdered.lean

Lines changed: 50 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -11,7 +11,6 @@ import Mathlib.Algebra.Order.Monoid.Unbundled.ExistsOfLE
1111
# Lemmas about densely linearly ordered groups.
1212
-/
1313

14-
1514
variable {α : Type*}
1615

1716
section DenselyOrdered
@@ -38,3 +37,53 @@ theorem le_iff_forall_lt_one_mul_le : a ≤ b ↔ ∀ ε < 1, a * ε ≤ b :=
3837
le_iff_forall_one_lt_le_mul (α := αᵒᵈ)
3938

4039
end DenselyOrdered
40+
41+
section DenselyOrdered
42+
43+
@[to_additive]
44+
private lemma exists_lt_mul_left [Group α] [LT α] [DenselyOrdered α]
45+
[CovariantClass α α (Function.swap (· * ·)) (· < ·)] {a b c : α} (hc : c < a * b) :
46+
∃ a' < a, c < a' * b := by
47+
obtain ⟨a', hc', ha'⟩ := exists_between (div_lt_iff_lt_mul.2 hc)
48+
exact ⟨a', ha', div_lt_iff_lt_mul.1 hc'⟩
49+
50+
@[to_additive]
51+
private lemma exists_lt_mul_right [CommGroup α] [LT α] [DenselyOrdered α]
52+
[CovariantClass α α (· * ·) (· < ·)] {a b c : α} (hc : c < a * b) :
53+
∃ b' < b, c < a * b' := by
54+
obtain ⟨a', hc', ha'⟩ := exists_between (div_lt_iff_lt_mul'.2 hc)
55+
exact ⟨a', ha', div_lt_iff_lt_mul'.1 hc'⟩
56+
57+
@[to_additive]
58+
private lemma exists_mul_left_lt [Group α] [LT α] [DenselyOrdered α]
59+
[CovariantClass α α (Function.swap (· * ·)) (· < ·)] {a b c : α} (hc : a * b < c) :
60+
∃ a' > a, a' * b < c := by
61+
obtain ⟨a', ha', hc'⟩ := exists_between (lt_div_iff_mul_lt.2 hc)
62+
exact ⟨a', ha', lt_div_iff_mul_lt.1 hc'⟩
63+
64+
@[to_additive]
65+
private lemma exists_mul_right_lt [CommGroup α] [LT α] [DenselyOrdered α]
66+
[CovariantClass α α (· * ·) (· < ·)] {a b c : α} (hc : a * b < c) :
67+
∃ b' > b, a * b' < c := by
68+
obtain ⟨a', ha', hc'⟩ := exists_between (lt_div_iff_mul_lt'.2 hc)
69+
exact ⟨a', ha', lt_div_iff_mul_lt'.1 hc'⟩
70+
71+
@[to_additive]
72+
lemma le_mul_of_forall_lt [CommGroup α] [LinearOrder α] [CovariantClass α α (· * ·) (· ≤ ·)]
73+
[DenselyOrdered α] {a b c : α} (h : ∀ a' > a, ∀ b' > b, c ≤ a' * b') :
74+
c ≤ a * b := by
75+
refine le_of_forall_le_of_dense fun d hd ↦ ?_
76+
obtain ⟨a', ha', hd⟩ := exists_mul_left_lt hd
77+
obtain ⟨b', hb', hd⟩ := exists_mul_right_lt hd
78+
exact (h a' ha' b' hb').trans hd.le
79+
80+
@[to_additive]
81+
lemma mul_le_of_forall_lt [CommGroup α] [LinearOrder α] [CovariantClass α α (· * ·) (· ≤ ·)]
82+
[DenselyOrdered α] {a b c : α} (h : ∀ a' < a, ∀ b' < b, a' * b' ≤ c) :
83+
a * b ≤ c := by
84+
refine le_of_forall_ge_of_dense fun d hd ↦ ?_
85+
obtain ⟨a', ha', hd⟩ := exists_lt_mul_left hd
86+
obtain ⟨b', hb', hd⟩ := exists_lt_mul_right hd
87+
exact hd.le.trans (h a' ha' b' hb')
88+
89+
end DenselyOrdered

Mathlib/Order/LiminfLimsup.lean

Lines changed: 125 additions & 51 deletions
Original file line numberDiff line numberDiff line change
@@ -6,6 +6,7 @@ Authors: Sébastien Gouëzel, Johannes Hölzl, Rémy Degenne
66
import Mathlib.Algebra.BigOperators.Group.Finset
77
import Mathlib.Algebra.Order.Group.Defs
88
import Mathlib.Algebra.Order.Group.Unbundled.Abs
9+
import Mathlib.Algebra.Order.GroupWithZero.Unbundled
910
import Mathlib.Order.Filter.Cofinite
1011
import Mathlib.Order.Hom.CompleteLattice
1112

@@ -111,6 +112,22 @@ theorem IsBoundedUnder.comp {l : Filter γ} {q : β → β → Prop} {u : γ →
111112
section Preorder
112113
variable [Preorder α] {f : Filter β} {u : β → α} {s : Set β}
113114

115+
lemma IsBoundedUnder.eventually_le (h : IsBoundedUnder (· ≤ ·) f u) :
116+
∃ a, ∀ᶠ x in f, u x ≤ a := by
117+
obtain ⟨a, ha⟩ := h
118+
use a
119+
exact eventually_map.1 ha
120+
121+
lemma IsBoundedUnder.eventually_ge (h : IsBoundedUnder (· ≥ ·) f u) :
122+
∃ a, ∀ᶠ x in f, a ≤ u x :=
123+
IsBoundedUnder.eventually_le (α := αᵒᵈ) h
124+
125+
lemma isBoundedUnder_of_eventually_le {a : α} (h : ∀ᶠ x in f, u x ≤ a) :
126+
IsBoundedUnder (· ≤ ·) f u := ⟨a, h⟩
127+
128+
lemma isBoundedUnder_of_eventually_ge {a : α} (h : ∀ᶠ x in f, a ≤ u x) :
129+
IsBoundedUnder (· ≥ ·) f u := ⟨a, h⟩
130+
114131
lemma isBoundedUnder_iff_eventually_bddAbove :
115132
f.IsBoundedUnder (· ≤ ·) u ↔ ∃ s, BddAbove (u '' s) ∧ ∀ᶠ x in f, x ∈ s := by
116133
constructor
@@ -281,6 +298,57 @@ theorem isCobounded_principal (s : Set α) :
281298
theorem IsCobounded.mono (h : f ≤ g) : f.IsCobounded r → g.IsCobounded r
282299
| ⟨b, hb⟩ => ⟨b, fun a ha => hb a (h ha)⟩
283300

301+
/-- For nontrivial filters in linear orders, coboundedness for `≤` implies frequent boundedness
302+
from below. -/
303+
lemma IsCobounded.frequently_ge [LinearOrder α] [NeBot f] (cobdd : IsCobounded (· ≤ ·) f) :
304+
∃ l, ∃ᶠ x in f, l ≤ x := by
305+
obtain ⟨t, ht⟩ := cobdd
306+
rcases isBot_or_exists_lt t with tbot | ⟨t', ht'⟩
307+
· exact ⟨t, .of_forall fun r ↦ tbot r⟩
308+
refine ⟨t', fun ev ↦ ?_⟩
309+
specialize ht t' (by filter_upwards [ev] with _ h using (not_le.mp h).le)
310+
exact not_lt_of_le ht ht'
311+
312+
/-- For nontrivial filters in linear orders, coboundedness for `≥` implies frequent boundedness
313+
from above. -/
314+
lemma IsCobounded.frequently_le [LinearOrder α] [NeBot f] (cobdd : IsCobounded (· ≥ ·) f) :
315+
∃ u, ∃ᶠ x in f, x ≤ u :=
316+
cobdd.frequently_ge (α := αᵒᵈ)
317+
318+
/-- In linear orders, frequent boundedness from below implies coboundedness for `≤`. -/
319+
lemma IsCobounded.of_frequently_ge [LinearOrder α] {l : α} (freq_ge : ∃ᶠ x in f, l ≤ x) :
320+
IsCobounded (· ≤ ·) f := by
321+
rcases isBot_or_exists_lt l with lbot | ⟨l', hl'⟩
322+
· exact ⟨l, fun x _ ↦ lbot x⟩
323+
refine ⟨l', fun u hu ↦ ?_⟩
324+
obtain ⟨w, l_le_w, w_le_u⟩ := (freq_ge.and_eventually hu).exists
325+
exact hl'.le.trans (l_le_w.trans w_le_u)
326+
327+
/-- In linear orders, frequent boundedness from above implies coboundedness for `≥`. -/
328+
lemma IsCobounded.of_frequently_le [LinearOrder α] {u : α} (freq_le : ∃ᶠ r in f, r ≤ u) :
329+
IsCobounded (· ≥ ·) f :=
330+
IsCobounded.of_frequently_ge (α := αᵒᵈ) freq_le
331+
332+
lemma IsCoboundedUnder.frequently_ge [LinearOrder α] {f : Filter ι} [NeBot f] {u : ι → α}
333+
(h : IsCoboundedUnder (· ≤ ·) f u) :
334+
∃ a, ∃ᶠ x in f, a ≤ u x :=
335+
IsCobounded.frequently_ge h
336+
337+
lemma IsCoboundedUnder.frequently_le [LinearOrder α] {f : Filter ι} [NeBot f] {u : ι → α}
338+
(h : IsCoboundedUnder (· ≥ ·) f u) :
339+
∃ a, ∃ᶠ x in f, u x ≤ a :=
340+
IsCobounded.frequently_le h
341+
342+
lemma IsCoboundedUnder.of_frequently_ge [LinearOrder α] {f : Filter ι} {u : ι → α}
343+
{a : α} (freq_ge : ∃ᶠ x in f, a ≤ u x) :
344+
IsCoboundedUnder (· ≤ ·) f u :=
345+
IsCobounded.of_frequently_ge freq_ge
346+
347+
lemma IsCoboundedUnder.of_frequently_le [LinearOrder α] {f : Filter ι} {u : ι → α}
348+
{a : α} (freq_le : ∃ᶠ x in f, u x ≤ a) :
349+
IsCoboundedUnder (· ≥ ·) f u :=
350+
IsCobounded.of_frequently_le freq_le
351+
284352
end Relation
285353

286354
section add_and_sum
@@ -325,19 +393,69 @@ lemma isBoundedUnder_le_add [Add R] [AddLeftMono R] [AddRightMono R]
325393

326394
lemma isBoundedUnder_le_sum {κ : Type*} [AddCommMonoid R] [AddLeftMono R] [AddRightMono R]
327395
{u : κ → α → R} (s : Finset κ) :
328-
(∀ k ∈ s, f.IsBoundedUnder (· ≤ ·) (u k)) → f.IsBoundedUnder (· ≤ ·) (∑ k ∈ s, u k) := by
329-
apply isBoundedUnder_sum (fun _ _ ↦ isBoundedUnder_le_add) le_rfl
396+
(∀ k ∈ s, f.IsBoundedUnder (· ≤ ·) (u k)) → f.IsBoundedUnder (· ≤ ·) (∑ k ∈ s, u k) :=
397+
fun h ↦ isBoundedUnder_sum (fun _ _ ↦ isBoundedUnder_le_add) le_rfl s h
330398

331399
lemma isBoundedUnder_ge_sum {κ : Type*} [AddCommMonoid R] [AddLeftMono R] [AddRightMono R]
332400
{u : κ → α → R} (s : Finset κ) :
333401
(∀ k ∈ s, f.IsBoundedUnder (· ≥ ·) (u k)) →
334-
f.IsBoundedUnder (· ≥ ·) (∑ k ∈ s, u k) := by
335-
haveI aux : CovariantClass R R (fun a b ↦ a + b) (· ≥ ·) :=
336-
{ elim := fun x _ _ hy ↦ add_le_add_left hy x }
337-
apply isBoundedUnder_sum (fun _ _ ↦ isBoundedUnder_ge_add) le_rfl
402+
f.IsBoundedUnder (· ≥ ·) (∑ k ∈ s, u k) :=
403+
fun h ↦ isBoundedUnder_sum (fun _ _ ↦ isBoundedUnder_ge_add) le_rfl s h
404+
405+
end add_and_sum
406+
407+
section add_and_sum
408+
409+
variable {α : Type*} {R : Type*} [LinearOrder R] [Add R] {f : Filter α} [f.NeBot]
410+
[CovariantClass R R (fun a b ↦ a + b) (· ≤ ·)] [CovariantClass R R (fun a b ↦ b + a) (· ≤ ·)]
411+
{u v : α → R}
412+
413+
lemma isCoboundedUnder_ge_add (hu : f.IsBoundedUnder (· ≤ ·) u)
414+
(hv : f.IsCoboundedUnder (· ≥ ·) v) :
415+
f.IsCoboundedUnder (· ≥ ·) (u + v) := by
416+
obtain ⟨U, hU⟩ := hu.eventually_le
417+
obtain ⟨V, hV⟩ := hv.frequently_le
418+
apply IsCoboundedUnder.of_frequently_le (a := U + V)
419+
exact (hV.and_eventually hU).mono fun x hx ↦ add_le_add hx.2 hx.1
420+
421+
lemma isCoboundedUnder_le_add (hu : f.IsBoundedUnder (· ≥ ·) u)
422+
(hv : f.IsCoboundedUnder (· ≤ ·) v) :
423+
f.IsCoboundedUnder (· ≤ ·) (u + v) := by
424+
obtain ⟨U, hU⟩ := hu.eventually_ge
425+
obtain ⟨V, hV⟩ := hv.frequently_ge
426+
apply IsCoboundedUnder.of_frequently_ge (a := U + V)
427+
exact (hV.and_eventually hU).mono fun x hx ↦ add_le_add hx.2 hx.1
338428

339429
end add_and_sum
340430

431+
section mul
432+
433+
lemma isBoundedUnder_le_mul_of_nonneg [Mul α] [Zero α] [Preorder α] [PosMulMono α]
434+
[MulPosMono α] {f : Filter ι} {u v : ι → α} (h₁ : 0 ≤ᶠ[f] u)
435+
(h₂ : IsBoundedUnder (fun x1 x2 ↦ x1 ≤ x2) f u)
436+
(h₃ : 0 ≤ᶠ[f] v)
437+
(h₄ : IsBoundedUnder (fun x1 x2 ↦ x1 ≤ x2) f v) :
438+
IsBoundedUnder (fun x1 x2 ↦ x1 ≤ x2) f (u * v) := by
439+
obtain ⟨U, hU⟩ := h₂.eventually_le
440+
obtain ⟨V, hV⟩ := h₄.eventually_le
441+
refine isBoundedUnder_of_eventually_le (a := U * V) ?_
442+
filter_upwards [hU, hV, h₁, h₃] with x x_U x_V u_0 v_0
443+
exact mul_le_mul x_U x_V v_0 (u_0.trans x_U)
444+
445+
lemma isCoboundedUnder_ge_mul_of_nonneg [Mul α] [Zero α] [LinearOrder α] [PosMulMono α]
446+
[MulPosMono α] {f : Filter ι} [f.NeBot] {u v : ι → α} (h₁ : 0 ≤ᶠ[f] u)
447+
(h₂ : IsBoundedUnder (fun x1 x2 ↦ x1 ≤ x2) f u)
448+
(h₃ : 0 ≤ᶠ[f] v)
449+
(h₄ : IsCoboundedUnder (fun x1 x2 ↦ x1 ≥ x2) f v) :
450+
IsCoboundedUnder (fun x1 x2 ↦ x1 ≥ x2) f (u * v) := by
451+
obtain ⟨U, hU⟩ := h₂.eventually_le
452+
obtain ⟨V, hV⟩ := h₄.frequently_le
453+
exact IsCoboundedUnder.of_frequently_le (a := U * V)
454+
<| (hV.and_eventually (hU.and (h₁.and h₃))).mono fun x ⟨x_V, x_U, u_0, v_0⟩ ↦
455+
mul_le_mul x_U x_V v_0 (u_0.trans x_U)
456+
457+
end mul
458+
341459
section Nonempty
342460
variable [Preorder α] [Nonempty α] {f : Filter β} {u : β → α}
343461

@@ -1638,50 +1756,6 @@ section frequently_bounded
16381756

16391757
variable {R S : Type*} {F : Filter R} [LinearOrder R] [LinearOrder S]
16401758

1641-
namespace Filter
1642-
1643-
/-- For nontrivial filters in linear orders, coboundedness for `≤` implies frequent boundedness
1644-
from below. -/
1645-
lemma IsCobounded.frequently_ge [NeBot F] (cobdd : IsCobounded (· ≤ ·) F) :
1646-
∃ l, ∃ᶠ x in F, l ≤ x := by
1647-
obtain ⟨t, ht⟩ := cobdd
1648-
by_cases tbot : IsBot t
1649-
· refine ⟨t, Frequently.of_forall fun r ↦ tbot r⟩
1650-
obtain ⟨t', ht'⟩ : ∃ t', t' < t := by
1651-
by_contra!
1652-
exact tbot this
1653-
refine ⟨t', ?_⟩
1654-
intro ev
1655-
specialize ht t' (by filter_upwards [ev] with _ h using (not_le.mp h).le)
1656-
apply lt_irrefl t' <| lt_of_lt_of_le ht' ht
1657-
1658-
/-- For nontrivial filters in linear orders, coboundedness for `≥` implies frequent boundedness
1659-
from above. -/
1660-
lemma IsCobounded.frequently_le [NeBot F] (cobdd : IsCobounded (· ≥ ·) F) :
1661-
∃ u, ∃ᶠ x in F, x ≤ u :=
1662-
cobdd.frequently_ge (R := Rᵒᵈ)
1663-
1664-
/-- In linear orders, frequent boundedness from below implies coboundedness for `≤`. -/
1665-
lemma IsCobounded.of_frequently_ge {l : R} (freq_ge : ∃ᶠ x in F, l ≤ x) :
1666-
IsCobounded (· ≤ ·) F := by
1667-
by_cases lbot : IsBot l
1668-
· refine ⟨l, fun x _ ↦ lbot x⟩
1669-
obtain ⟨l', hl'⟩ : ∃ l', l' < l := by
1670-
by_contra!
1671-
exact lbot this
1672-
refine ⟨l', ?_⟩
1673-
intro u hu
1674-
have key : ∃ᶠ x in F, l ≤ x ∧ x ≤ u := Frequently.and_eventually freq_ge hu
1675-
obtain ⟨w, l_le_w, w_le_u⟩ := key.exists
1676-
exact hl'.le.trans <| l_le_w.trans w_le_u
1677-
1678-
/-- In linear orders, frequent boundedness from above implies coboundedness for `≥`. -/
1679-
lemma IsCobounded.of_frequently_le {u : R} (freq_le : ∃ᶠ r in F, r ≤ u) :
1680-
IsCobounded (· ≥ ·) F :=
1681-
IsCobounded.of_frequently_ge (R := Rᵒᵈ) freq_le
1682-
1683-
end Filter
1684-
16851759
lemma Monotone.frequently_ge_map_of_frequently_ge {f : R → S} (f_incr : Monotone f)
16861760
{l : R} (freq_ge : ∃ᶠ x in F, l ≤ x) :
16871761
∃ᶠ x' in F.map f, f l ≤ x' := by
@@ -1733,4 +1807,4 @@ lemma Antitone.isCoboundedUnder_ge_of_isCobounded {f : R → S} (f_decr : Antito
17331807

17341808
end frequently_bounded
17351809

1736-
set_option linter.style.longFile 1800
1810+
set_option linter.style.longFile 1900

0 commit comments

Comments
 (0)