@@ -6,6 +6,7 @@ Authors: Sébastien Gouëzel, Johannes Hölzl, Rémy Degenne
6
6
import Mathlib.Algebra.BigOperators.Group.Finset
7
7
import Mathlib.Algebra.Order.Group.Defs
8
8
import Mathlib.Algebra.Order.Group.Unbundled.Abs
9
+ import Mathlib.Algebra.Order.GroupWithZero.Unbundled
9
10
import Mathlib.Order.Filter.Cofinite
10
11
import Mathlib.Order.Hom.CompleteLattice
11
12
@@ -111,6 +112,22 @@ theorem IsBoundedUnder.comp {l : Filter γ} {q : β → β → Prop} {u : γ →
111
112
section Preorder
112
113
variable [Preorder α] {f : Filter β} {u : β → α} {s : Set β}
113
114
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
+
114
131
lemma isBoundedUnder_iff_eventually_bddAbove :
115
132
f.IsBoundedUnder (· ≤ ·) u ↔ ∃ s, BddAbove (u '' s) ∧ ∀ᶠ x in f, x ∈ s := by
116
133
constructor
@@ -281,6 +298,57 @@ theorem isCobounded_principal (s : Set α) :
281
298
theorem IsCobounded.mono (h : f ≤ g) : f.IsCobounded r → g.IsCobounded r
282
299
| ⟨b, hb⟩ => ⟨b, fun a ha => hb a (h ha)⟩
283
300
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
+
284
352
end Relation
285
353
286
354
section add_and_sum
@@ -325,19 +393,69 @@ lemma isBoundedUnder_le_add [Add R] [AddLeftMono R] [AddRightMono R]
325
393
326
394
lemma isBoundedUnder_le_sum {κ : Type *} [AddCommMonoid R] [AddLeftMono R] [AddRightMono R]
327
395
{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
330
398
331
399
lemma isBoundedUnder_ge_sum {κ : Type *} [AddCommMonoid R] [AddLeftMono R] [AddRightMono R]
332
400
{u : κ → α → R} (s : Finset κ) :
333
401
(∀ 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
338
428
339
429
end add_and_sum
340
430
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
+
341
459
section Nonempty
342
460
variable [Preorder α] [Nonempty α] {f : Filter β} {u : β → α}
343
461
@@ -1638,50 +1756,6 @@ section frequently_bounded
1638
1756
1639
1757
variable {R S : Type *} {F : Filter R} [LinearOrder R] [LinearOrder S]
1640
1758
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
-
1685
1759
lemma Monotone.frequently_ge_map_of_frequently_ge {f : R → S} (f_incr : Monotone f)
1686
1760
{l : R} (freq_ge : ∃ᶠ x in F, l ≤ x) :
1687
1761
∃ᶠ 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
1733
1807
1734
1808
end frequently_bounded
1735
1809
1736
- set_option linter.style.longFile 1800
1810
+ set_option linter.style.longFile 1900
0 commit comments