Skip to content

Commit 43c9317

Browse files
committed
chore(Topology/Order): rename 2 lemmas (#19261)
... to allow dot notation.
1 parent 1f86e43 commit 43c9317

File tree

15 files changed

+45
-37
lines changed

15 files changed

+45
-37
lines changed

Mathlib/Algebra/Order/Floor/Prime.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -39,7 +39,7 @@ theorem exists_prime_mul_pow_div_factorial_lt_one [LinearOrderedField K] [FloorR
3939
letI := Preorder.topology K
4040
haveI : OrderTopology K := ⟨rfl⟩
4141
((Filter.frequently_atTop.mpr Nat.exists_infinite_primes).and_eventually
42-
(eventually_lt_of_tendsto_lt zero_lt_one
42+
(Filter.Tendsto.eventually_lt_const zero_lt_one
4343
(FloorSemiring.tendsto_mul_pow_div_factorial_sub_atTop a c 1))).forall_exists_of_atTop
4444
(n + 1)
4545

Mathlib/Analysis/Asymptotics/TVS.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -125,7 +125,7 @@ lemma isLittleOTVS_one [ContinuousSMul 𝕜 E] {f : α → E} {l : Filter α} :
125125
rcases NormedField.exists_one_lt_norm 𝕜 with ⟨c, hc⟩
126126
obtain ⟨ε, hε₀, hε⟩ : ∃ ε : ℝ≥0, 0 < ε ∧ (ε * ‖c‖₊ / r : ℝ≥0∞) < 1 := by
127127
apply Eventually.exists_gt
128-
refine eventually_lt_of_tendsto_lt zero_lt_one <| Continuous.tendsto' ?_ _ _ (by simp)
128+
refine Continuous.tendsto' ?_ _ _ (by simp) |>.eventually_lt_const zero_lt_one
129129
fun_prop (disch := intros; first | apply ENNReal.coe_ne_top | positivity)
130130
filter_upwards [hr ε hε₀.ne'] with x hx
131131
refine mem_of_egauge_lt_one hUb (hx.trans_lt ?_)

Mathlib/Analysis/BoundedVariation.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -180,7 +180,7 @@ theorem lowerSemicontinuous_aux {ι : Type*} {F : ι → α → E} {p : Filter
180180
(𝓝 (∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i)))) := by
181181
apply tendsto_finset_sum
182182
exact fun i _ => Tendsto.edist (Ffs (u i.succ) (us i.succ)) (Ffs (u i) (us i))
183-
exact (eventually_gt_of_tendsto_gt hlt this).mono fun i h => h.trans_le (sum_le (F i) n um us)
183+
exact (this.eventually_const_lt hlt).mono fun i h => h.trans_le (sum_le (F i) n um us)
184184

185185
/-- The map `(eVariationOn · s)` is lower semicontinuous for pointwise convergence *on `s`*.
186186
Pointwise convergence on `s` is encoded here as uniform convergence on the family consisting of the

Mathlib/Analysis/Normed/Lp/lpSpace.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -183,7 +183,7 @@ theorem of_exponent_ge {p q : ℝ≥0∞} {f : ∀ i, E i} (hfq : Memℓp f q) (
183183
have hf' := hfq.summable hq
184184
refine .of_norm_bounded_eventually _ hf' (@Set.Finite.subset _ { i | 1 ≤ ‖f i‖ } ?_ _ ?_)
185185
· have H : { x : α | 1 ≤ ‖f x‖ ^ q.toReal }.Finite := by
186-
simpa using eventually_lt_of_tendsto_lt (by norm_num) hf'.tendsto_cofinite_zero
186+
simpa using hf'.tendsto_cofinite_zero.eventually_lt_const (by norm_num)
187187
exact H.subset fun i hi => Real.one_le_rpow hi hq.le
188188
· show ∀ i, ¬|‖f i‖ ^ p.toReal| ≤ ‖f i‖ ^ q.toReal → 1 ≤ ‖f i‖
189189
intro i hi

Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -592,7 +592,7 @@ theorem tendsto_one_plus_div_rpow_exp (t : ℝ) :
592592
have h₁ : (1 : ℝ) / 2 < 1 := by norm_num
593593
have h₂ : Tendsto (fun x : ℝ => 1 + t / x) atTop (𝓝 1) := by
594594
simpa using (tendsto_inv_atTop_zero.const_mul t).const_add 1
595-
refine (eventually_ge_of_tendsto_gt h₁ h₂).mono fun x hx => ?_
595+
refine (h₂.eventually_const_le h₁).mono fun x hx => ?_
596596
have hx' : 0 < 1 + t / x := by linarith
597597
simp [mul_comm x, exp_mul, exp_log hx']
598598

Mathlib/Analysis/SpecificLimits/Normed.lean

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -601,7 +601,7 @@ theorem summable_of_ratio_test_tendsto_lt_one {α : Type*} [NormedAddCommGroup
601601
(h : Tendsto (fun n ↦ ‖f (n + 1)‖ / ‖f n‖) atTop (𝓝 l)) : Summable f := by
602602
rcases exists_between hl₁ with ⟨r, hr₀, hr₁⟩
603603
refine summable_of_ratio_norm_eventually_le hr₁ ?_
604-
filter_upwards [eventually_le_of_tendsto_lt hr₀ h, hf] with _ _ h₁
604+
filter_upwards [h.eventually_le_const hr₀, hf] with _ _ h₁
605605
rwa [← div_le_iff₀ (norm_pos_iff.mpr h₁)]
606606

607607
theorem not_summable_of_ratio_norm_eventually_ge {α : Type*} [SeminormedAddCommGroup α] {f : ℕ → α}
@@ -629,12 +629,12 @@ theorem not_summable_of_ratio_test_tendsto_gt_one {α : Type*} [SeminormedAddCom
629629
{f : ℕ → α} {l : ℝ} (hl : 1 < l) (h : Tendsto (fun n ↦ ‖f (n + 1)‖ / ‖f n‖) atTop (𝓝 l)) :
630630
¬Summable f := by
631631
have key : ∀ᶠ n in atTop, ‖f n‖ ≠ 0 := by
632-
filter_upwards [eventually_ge_of_tendsto_gt hl h] with _ hn hc
632+
filter_upwards [h.eventually_const_le hl] with _ hn hc
633633
rw [hc, _root_.div_zero] at hn
634634
linarith
635635
rcases exists_between hl with ⟨r, hr₀, hr₁⟩
636636
refine not_summable_of_ratio_norm_eventually_ge hr₀ key.frequently ?_
637-
filter_upwards [eventually_ge_of_tendsto_gt hr₁ h, key] with _ _ h₁
637+
filter_upwards [h.eventually_const_le hr₁, key] with _ _ h₁
638638
rwa [← le_div_iff₀ (lt_of_le_of_ne (norm_nonneg _) h₁.symm)]
639639

640640
section NormedDivisionRing

Mathlib/Computability/AkraBazzi/AkraBazzi.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -544,9 +544,9 @@ lemma tendsto_atTop_sumCoeffsExp : Tendsto (fun (p : ℝ) => ∑ i, a i * (b i)
544544
lemma one_mem_range_sumCoeffsExp : 1 ∈ Set.range (fun (p : ℝ) => ∑ i, a i * (b i) ^ p) := by
545545
refine mem_range_of_exists_le_of_exists_ge R.continuous_sumCoeffsExp ?le_one ?ge_one
546546
case le_one =>
547-
exact Eventually.exists <| eventually_le_of_tendsto_lt zero_lt_one R.tendsto_zero_sumCoeffsExp
547+
exact R.tendsto_zero_sumCoeffsExp.eventually_le_const zero_lt_one |>.exists
548548
case ge_one =>
549-
exact Eventually.exists <| R.tendsto_atTop_sumCoeffsExp.eventually_ge_atTop _
549+
exact R.tendsto_atTop_sumCoeffsExp.eventually_ge_atTop _ |>.exists
550550

551551
/-- The function x ↦ ∑ a_i b_i^x is injective. This implies the uniqueness of `p`. -/
552552
lemma injective_sumCoeffsExp : Function.Injective (fun (p : ℝ) => ∑ i, a i * (b i) ^ p) :=

Mathlib/Data/Real/Pi/Irrational.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -283,8 +283,8 @@ private lemma not_irrational_exists_rep {x : ℝ} :
283283
rwa [lt_div_iff₀ (by positivity), zero_mul] at this
284284
have k (n : ℕ) : 0 < (a : ℝ) ^ (2 * n + 1) / n ! := by positivity
285285
have j : ∀ᶠ n : ℕ in atTop, (a : ℝ) ^ (2 * n + 1) / n ! * I n (π / 2) < 1 := by
286-
have := eventually_lt_of_tendsto_lt (show (0 : ℝ) < 1 / 2 by norm_num)
287-
(tendsto_pow_div_factorial_at_top_aux a)
286+
have := (tendsto_pow_div_factorial_at_top_aux a).eventually_lt_const
287+
(show (0 : ℝ) < 1 / 2 by norm_num)
288288
filter_upwards [this] with n hn
289289
rw [lt_div_iff₀ (zero_lt_two : (0 : ℝ) < 2)] at hn
290290
exact hn.trans_le' (mul_le_mul_of_nonneg_left (I_le _) (by positivity))

Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -193,12 +193,12 @@ theorem ae_le_of_forall_setLIntegral_le_of_sigmaFinite₀ [SigmaFinite μ]
193193
have : Tendsto (fun n => g x + u n) atTop (𝓝 (g x + (0 : ℝ≥0))) :=
194194
tendsto_const_nhds.add (ENNReal.tendsto_coe.2 u_lim)
195195
simp only [ENNReal.coe_zero, add_zero] at this
196-
exact eventually_le_of_tendsto_lt hx this
196+
exact this.eventually_le_const hx
197197
have L2 : ∀ᶠ n : ℕ in (atTop : Filter ℕ), g x ≤ (n : ℝ≥0) :=
198-
haveI : Tendsto (fun n : ℕ => ((n : ℝ≥0) : ℝ≥0∞)) atTop (𝓝 ∞) := by
198+
have : Tendsto (fun n : ℕ => ((n : ℝ≥0) : ℝ≥0∞)) atTop (𝓝 ∞) := by
199199
simp only [ENNReal.coe_natCast]
200200
exact ENNReal.tendsto_nat_nhds_top
201-
eventually_ge_of_tendsto_gt (hx.trans_le le_top) this
201+
this.eventually_const_le (hx.trans_le le_top)
202202
apply Set.mem_iUnion.2
203203
exact ((L1.and L2).and (eventually_mem_spanningSets μ x)).exists
204204
refine le_antisymm ?_ bot_le

Mathlib/MeasureTheory/Function/LpSpace.lean

Lines changed: 1 addition & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -673,8 +673,7 @@ theorem exists_eLpNorm_indicator_le (hp : p ≠ ∞) (c : E) {ε : ℝ≥0∞} (
673673
convert (NNReal.continuousAt_rpow_const (Or.inr hp₀')).tendsto.const_mul _
674674
simp [hp₀''.ne']
675675
have hε' : 0 < ε := hε.bot_lt
676-
obtain ⟨δ, hδ, hδε'⟩ :=
677-
NNReal.nhds_zero_basis.eventually_iff.mp (eventually_le_of_tendsto_lt hε' this)
676+
obtain ⟨δ, hδ, hδε'⟩ := NNReal.nhds_zero_basis.eventually_iff.mp (this.eventually_le_const hε')
678677
obtain ⟨η, hη, hηδ⟩ := exists_between hδ
679678
refine ⟨η, hη, ?_⟩
680679
rw [← ENNReal.coe_rpow_of_nonneg _ hp₀', ← ENNReal.coe_mul]

0 commit comments

Comments
 (0)