@@ -601,7 +601,7 @@ theorem summable_of_ratio_test_tendsto_lt_one {α : Type*} [NormedAddCommGroup
601
601
(h : Tendsto (fun n ↦ ‖f (n + 1 )‖ / ‖f n‖) atTop (𝓝 l)) : Summable f := by
602
602
rcases exists_between hl₁ with ⟨r, hr₀, hr₁⟩
603
603
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₁
605
605
rwa [← div_le_iff₀ (norm_pos_iff.mpr h₁)]
606
606
607
607
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
629
629
{f : ℕ → α} {l : ℝ} (hl : 1 < l) (h : Tendsto (fun n ↦ ‖f (n + 1 )‖ / ‖f n‖) atTop (𝓝 l)) :
630
630
¬Summable f := by
631
631
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
633
633
rw [hc, _root_.div_zero] at hn
634
634
linarith
635
635
rcases exists_between hl with ⟨r, hr₀, hr₁⟩
636
636
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₁
638
638
rwa [← le_div_iff₀ (lt_of_le_of_ne (norm_nonneg _) h₁.symm)]
639
639
640
640
section NormedDivisionRing
0 commit comments