Skip to content
Closed
Show file tree
Hide file tree
Changes from 35 commits
Commits
Show all changes
39 commits
Select commit Hold shift + click to select a range
89db452
chore(Topology/Algebra/Order/Field): generalize some lemmas to `Semif…
astrainfinita Aug 15, 2025
e912a79
feat: add `ContinuousSMul` instances for `ℚ≥0`
astrainfinita Aug 15, 2025
1fe47cd
Merge branch 'ordered_topological_semifield' into tendsto_inverse_atT…
astrainfinita Aug 15, 2025
8f4823a
chore(Analysis/SpecificLimits/Basic): generalize lemmas
astrainfinita Aug 15, 2025
f70ef66
suggestion
astrainfinita Aug 15, 2025
c22a8d3
fix
astrainfinita Aug 15, 2025
55e9f64
Merge branch 'master' into tendsto_inverse_atTop_nhds_zero_nat
astrainfinita Oct 19, 2025
162b1ab
Merge branch 'master' into rat_continuousSMul
astrainfinita Oct 19, 2025
94eb313
suggestion
astrainfinita Oct 19, 2025
a05c783
chore: remove `MulAction ℚ α → MulAction ℚ≥0 α` instances
astrainfinita Oct 19, 2025
bca0e9b
chore: move `(NN)Rat.cast_smul_eq_(nn)qsmul`
astrainfinita Oct 19, 2025
aa77423
lint
astrainfinita Oct 19, 2025
efbc0b2
more
astrainfinita Oct 19, 2025
308ffc5
shake
astrainfinita Oct 20, 2025
d47fa93
replace existing lemmas
astrainfinita Oct 20, 2025
d0f9960
Revert "suggestion"
astrainfinita Oct 20, 2025
4bcb9de
generalize
astrainfinita Oct 20, 2025
069f8bd
fix
astrainfinita Oct 20, 2025
7fc7755
Update Rat.lean
astrainfinita Oct 20, 2025
9ddc0f4
Merge branch 'cast_smul_eq_qsmul' into tendsto_inverse_atTop_nhds_zer…
astrainfinita Oct 20, 2025
43bfd66
Merge branch 'rat_continuousSMul' into tendsto_inverse_atTop_nhds_zer…
astrainfinita Oct 20, 2025
2b220a7
Update Rat.lean
astrainfinita Oct 20, 2025
0ae2535
more
astrainfinita Oct 20, 2025
4dce365
Merge branch 'nnrat_action' into rat_continuousSMul
astrainfinita Oct 20, 2025
1921802
Merge branch 'rat_continuousSMul' into tendsto_inverse_atTop_nhds_zer…
astrainfinita Oct 20, 2025
45ceab1
fix
astrainfinita Oct 20, 2025
30ecd8f
fix
astrainfinita Oct 20, 2025
28e203b
fix counterexample
astrainfinita Oct 20, 2025
da46c53
shake
astrainfinita Oct 20, 2025
37b229a
Merge branch 'master' into tendsto_inverse_atTop_nhds_zero_nat
astrainfinita Oct 26, 2025
ecf7ff6
more
astrainfinita Oct 27, 2025
2b971f9
shake
astrainfinita Oct 27, 2025
c7ffbce
fix
astrainfinita Oct 27, 2025
43dbbbc
more
astrainfinita Oct 28, 2025
a85bc52
Merge branch 'master' into tendsto_inverse_atTop_nhds_zero_nat
astrainfinita Oct 28, 2025
4c7ddd0
suggestion
astrainfinita Nov 2, 2025
80c6726
Merge branch 'master' into tendsto_inverse_atTop_nhds_zero_nat
astrainfinita Nov 6, 2025
aea227e
fix merge
astrainfinita Nov 6, 2025
e23d505
fix
astrainfinita Nov 6, 2025
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
8 changes: 1 addition & 7 deletions Counterexamples/DiscreteTopologyNonDiscreteUniformity.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,12 +3,7 @@ Copyright (c) 2024 Filippo A. E. Nuccio. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Filippo A. E. Nuccio
-/

import Mathlib.Analysis.SpecificLimits.Basic
import Mathlib.Order.Interval.Set.Basic
import Mathlib.Topology.MetricSpace.Pseudo.Defs
import Mathlib.Topology.MetricSpace.Cauchy
import Mathlib.Topology.UniformSpace.Cauchy

/-!
# Discrete uniformities and discrete topology
Expand Down Expand Up @@ -83,11 +78,10 @@ open Set Function Filter Metric

/- We remove the "usual" instances of (discrete) topological space and of (discrete) uniform space
from `ℕ`. -/
attribute [-instance] instTopologicalSpaceNat instUniformSpaceNat
attribute [-instance] instTopologicalSpaceNat instUniformSpaceNat Nat.instDist

section Metric


noncomputable local instance : PseudoMetricSpace ℕ where
dist := fun n m ↦ |2 ^ (- n : ℤ) - 2 ^ (- m : ℤ)|
dist_self := by simp only [zpow_neg, zpow_natCast, sub_self, abs_zero, implies_true]
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Analysis/CStarAlgebra/Spectrum.lean
Original file line number Diff line number Diff line change
Expand Up @@ -114,7 +114,7 @@ theorem IsSelfAdjoint.spectralRadius_eq_nnnorm {a : A} (ha : IsSelfAdjoint a) :
refine tendsto_nhds_unique ?_ hconst
convert
(spectrum.pow_nnnorm_pow_one_div_tendsto_nhds_spectralRadius (a : A)).comp
(Nat.tendsto_pow_atTop_atTop_of_one_lt one_lt_two) using 1
(tendsto_pow_atTop_atTop_of_one_lt one_lt_two) using 1
refine funext fun n => ?_
rw [Function.comp_apply, ha.nnnorm_pow_two_pow, ENNReal.coe_pow, ← rpow_natCast, ← rpow_mul]
simp
Expand Down
3 changes: 1 addition & 2 deletions Mathlib/Analysis/Normed/Algebra/Exponential.lean
Original file line number Diff line number Diff line change
Expand Up @@ -8,7 +8,6 @@ import Mathlib.Analysis.Analytic.ChangeOrigin
import Mathlib.Analysis.Complex.Basic
import Mathlib.Data.Nat.Choose.Cast
import Mathlib.Analysis.Analytic.OfScalars
import Mathlib.Analysis.SpecificLimits.RCLike

/-!
# Exponential in a Banach algebra
Expand Down Expand Up @@ -390,7 +389,7 @@ theorem expSeries_radius_eq_top : (expSeries 𝕂 𝔸).radius = ∞ := by
inv_div_inv, norm_mul, div_self this, norm_one, one_mul]
apply norm_zero (E := 𝕂) ▸ Filter.Tendsto.norm
apply (Filter.tendsto_add_atTop_iff_nat (f := fun n => (n : 𝕂)⁻¹) 1).mpr
exact RCLike.tendsto_inverse_atTop_nhds_zero_nat 𝕂
exact tendsto_inv_atTop_nhds_zero_nat
· simp [this]

theorem expSeries_radius_pos : 0 < (expSeries 𝕂 𝔸).radius := by
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Analysis/PSeries.lean
Original file line number Diff line number Diff line change
Expand Up @@ -148,7 +148,7 @@ theorem le_tsum_schlomilch (hf : ∀ ⦃m n⦄, 0 < m → m ≤ n → f n ≤ f

theorem le_tsum_condensed (hf : ∀ ⦃m n⦄, 0 < m → m ≤ n → f n ≤ f m) :
∑' k, f k ≤ f 0 + ∑' k : ℕ, 2 ^ k * f (2 ^ k) := by
rw [ENNReal.tsum_eq_iSup_nat' (Nat.tendsto_pow_atTop_atTop_of_one_lt _root_.one_lt_two)]
rw [ENNReal.tsum_eq_iSup_nat' (tendsto_pow_atTop_atTop_of_one_lt _root_.one_lt_two)]
refine iSup_le fun n => (Finset.le_sum_condensed hf n).trans ?_
simp only [nsmul_eq_mul, Nat.cast_pow, Nat.cast_two]
grw [ENNReal.sum_le_tsum]
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Analysis/Real/Hyperreal.lean
Original file line number Diff line number Diff line change
Expand Up @@ -185,7 +185,7 @@ theorem gt_of_tendsto_zero_of_neg {f : ℕ → ℝ} (hf : Tendsto f atTop (𝓝
rw [← neg_neg r, coe_neg]; exact neg_lt_of_tendsto_zero_of_pos hf (neg_pos.mpr hr)

theorem epsilon_lt_pos (x : ℝ) : 0 < x → ε < x :=
lt_of_tendsto_zero_of_pos tendsto_inverse_atTop_nhds_zero_nat
lt_of_tendsto_zero_of_pos tendsto_inv_atTop_nhds_zero_nat

/-- Standard part predicate -/
def IsSt (x : ℝ*) (r : ℝ) :=
Expand Down Expand Up @@ -581,7 +581,7 @@ theorem infinitesimal_of_tendsto_zero {f : ℕ → ℝ} (h : Tendsto f atTop (
isSt_of_tendsto h

theorem infinitesimal_epsilon : Infinitesimal ε :=
infinitesimal_of_tendsto_zero tendsto_inverse_atTop_nhds_zero_nat
infinitesimal_of_tendsto_zero tendsto_inv_atTop_nhds_zero_nat

theorem not_real_of_infinitesimal_ne_zero (x : ℝ*) : Infinitesimal x → x ≠ 0 → ∀ r : ℝ, x ≠ r :=
fun hi hx r hr =>
Expand Down
1 change: 1 addition & 0 deletions Mathlib/Analysis/SpecialFunctions/Complex/LogBounds.lean
Original file line number Diff line number Diff line change
Expand Up @@ -6,6 +6,7 @@ Authors: Michael Stoll
import Mathlib.Analysis.Complex.Convex
import Mathlib.Analysis.SpecialFunctions.Integrals.Basic
import Mathlib.Analysis.Calculus.Deriv.Shift
import Mathlib.Analysis.SpecificLimits.RCLike

/-!
# Estimates for the complex logarithm
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Analysis/SpecialFunctions/OrdinaryHypergeometric.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Edward Watine
-/
import Mathlib.Analysis.Analytic.OfScalars
import Mathlib.Analysis.SpecificLimits.RCLike
import Mathlib.Analysis.RCLike.Basic

/-!
# Ordinary hypergeometric function in a Banach algebra
Expand Down Expand Up @@ -204,6 +204,6 @@ theorem ordinaryHypergeometricSeries_radius_eq_one
(c + k) / (a + k) * ((1 + k) / (b + k)) := by field_simp
simp_rw [this]
apply (mul_one (1 : 𝕂)) ▸ Filter.Tendsto.mul <;>
convert RCLike.tendsto_add_mul_div_add_mul_atTop_nhds _ _ (1 : 𝕂) one_ne_zero <;> simp
convert tendsto_add_mul_div_add_mul_atTop_nhds _ _ (1 : 𝕂) one_ne_zero <;> simp

end RCLike
124 changes: 80 additions & 44 deletions Mathlib/Analysis/SpecificLimits/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -10,6 +10,7 @@ import Mathlib.Order.Iterate
import Mathlib.Topology.Algebra.Algebra
import Mathlib.Topology.Algebra.InfiniteSum.Real
import Mathlib.Topology.Instances.EReal.Lemmas
import Mathlib.Topology.Instances.Rat

/-!
# A collection of specific limit computations
Expand All @@ -27,24 +28,32 @@ open Set Function Filter Finset Metric Topology Nat uniformity NNReal ENNReal

variable {α : Type*} {β : Type*} {ι : Type*}

theorem tendsto_inverse_atTop_nhds_zero_nat : Tendsto (fun n : ℕ ↦ (n : )⁻¹) atTop (𝓝 0) :=
theorem NNRat.tendsto_inv_atTop_nhds_zero_nat : Tendsto (fun n : ℕ ↦ (n : ℚ≥0)⁻¹) atTop (𝓝 0) :=
tendsto_inv_atTop_zero.comp tendsto_natCast_atTop_atTop

theorem tendsto_const_div_atTop_nhds_zero_nat (C : ℝ) :
Tendsto (fun n : ℕ ↦ C / n) atTop (𝓝 0) := by
simpa only [mul_zero] using tendsto_const_nhds.mul tendsto_inverse_atTop_nhds_zero_nat

theorem tendsto_one_div_atTop_nhds_zero_nat : Tendsto (fun n : ℕ ↦ 1 / (n : ℝ)) atTop (𝓝 0) :=
tendsto_const_div_atTop_nhds_zero_nat 1
theorem NNRat.tendsto_algebraMap_inv_atTop_nhds_zero_nat (𝕜 : Type*) [Semiring 𝕜]
[Algebra ℚ≥0 𝕜] [TopologicalSpace 𝕜] [ContinuousSMul ℚ≥0 𝕜] :
Tendsto (algebraMap ℚ≥0 𝕜 ∘ fun n : ℕ ↦ (n : ℚ≥0)⁻¹) atTop (𝓝 0) := by
convert (continuous_algebraMap ℚ≥0 𝕜).continuousAt.tendsto.comp
tendsto_inv_atTop_nhds_zero_nat
rw [map_zero]

theorem NNReal.tendsto_inverse_atTop_nhds_zero_nat :
Tendsto (fun n : ℕ ↦ (n : ℝ≥0)⁻¹) atTop (𝓝 0) := by
rw [← NNReal.tendsto_coe]
exact _root_.tendsto_inverse_atTop_nhds_zero_nat
theorem tendsto_inv_atTop_nhds_zero_nat {𝕜 : Type*} [Semifield 𝕜] [CharZero 𝕜]
[TopologicalSpace 𝕜] [ContinuousSMul ℚ≥0 𝕜] :
Tendsto (fun n : ℕ ↦ (n : 𝕜)⁻¹) atTop (𝓝 0) := by
convert NNRat.tendsto_algebraMap_inv_atTop_nhds_zero_nat 𝕜
simp

theorem NNReal.tendsto_const_div_atTop_nhds_zero_nat (C : ℝ≥0) :
theorem tendsto_const_div_atTop_nhds_zero_nat {𝕜 : Type*} [Semifield 𝕜] [CharZero 𝕜]
[TopologicalSpace 𝕜] [ContinuousSMul ℚ≥0 𝕜] [ContinuousMul 𝕜] (C : 𝕜) :
Tendsto (fun n : ℕ ↦ C / n) atTop (𝓝 0) := by
simpa using tendsto_const_nhds.mul NNReal.tendsto_inverse_atTop_nhds_zero_nat
simpa only [mul_zero, div_eq_mul_inv] using
(tendsto_const_nhds (x := C)).mul tendsto_inv_atTop_nhds_zero_nat

theorem tendsto_one_div_atTop_nhds_zero_nat {𝕜 : Type*} [Semifield 𝕜] [CharZero 𝕜]
[TopologicalSpace 𝕜] [ContinuousSMul ℚ≥0 𝕜] :
Tendsto (fun n : ℕ ↦ 1 / (n : 𝕜)) atTop (𝓝 0) := by
simp [tendsto_inv_atTop_nhds_zero_nat]

theorem EReal.tendsto_const_div_atTop_nhds_zero_nat {C : EReal} (h : C ≠ ⊥) (h' : C ≠ ⊤) :
Tendsto (fun n : ℕ ↦ C / n) atTop (𝓝 0) := by
Expand All @@ -54,30 +63,23 @@ theorem EReal.tendsto_const_div_atTop_nhds_zero_nat {C : EReal} (h : C ≠ ⊥)
rw [this, ← coe_zero, tendsto_coe]
exact _root_.tendsto_const_div_atTop_nhds_zero_nat C.toReal

theorem tendsto_one_div_add_atTop_nhds_zero_nat :
Tendsto (fun n : ℕ ↦ 1 / ((n : ℝ) + 1)) atTop (𝓝 0) :=
suffices Tendsto (fun n : ℕ ↦ 1 / (↑(n + 1) : ℝ)) atTop (𝓝 0) by simpa
(tendsto_add_atTop_iff_nat 1).2 (_root_.tendsto_const_div_atTop_nhds_zero_nat 1)

theorem NNReal.tendsto_algebraMap_inverse_atTop_nhds_zero_nat (𝕜 : Type*) [Semiring 𝕜]
[Algebra ℝ≥0 𝕜] [TopologicalSpace 𝕜] [ContinuousSMul ℝ≥0 𝕜] :
Tendsto (algebraMap ℝ≥0 𝕜 ∘ fun n : ℕ ↦ (n : ℝ≥0)⁻¹) atTop (𝓝 0) := by
convert (continuous_algebraMap ℝ≥0 𝕜).continuousAt.tendsto.comp
tendsto_inverse_atTop_nhds_zero_nat
theorem tendsto_one_div_add_atTop_nhds_zero_nat {𝕜 : Type*} [Semifield 𝕜] [CharZero 𝕜]
[TopologicalSpace 𝕜] [ContinuousSMul ℚ≥0 𝕜] :
Tendsto (fun n : ℕ ↦ 1 / ((n : 𝕜) + 1)) atTop (𝓝 0) :=
suffices Tendsto (fun n : ℕ ↦ 1 / (↑(n + 1) : 𝕜)) atTop (𝓝 0) by simpa
(tendsto_add_atTop_iff_nat 1).2 tendsto_one_div_atTop_nhds_zero_nat

theorem tendsto_algebraMap_inv_atTop_nhds_zero_nat {𝕜 : Type*} (A : Type*)
[Semifield 𝕜] [CharZero 𝕜] [TopologicalSpace 𝕜] [ContinuousSMul ℚ≥0 𝕜]
[Semiring A] [Algebra 𝕜 A] [TopologicalSpace A] [ContinuousSMul 𝕜 A] :
Tendsto (algebraMap 𝕜 A ∘ fun n : ℕ ↦ (n : 𝕜)⁻¹) atTop (𝓝 0) := by
convert (continuous_algebraMap 𝕜 A).continuousAt.tendsto.comp tendsto_inv_atTop_nhds_zero_nat
rw [map_zero]

theorem tendsto_algebraMap_inverse_atTop_nhds_zero_nat (𝕜 : Type*) [Semiring 𝕜] [Algebra ℝ 𝕜]
[TopologicalSpace 𝕜] [ContinuousSMul ℝ 𝕜] :
Tendsto (algebraMap ℝ 𝕜 ∘ fun n : ℕ ↦ (n : ℝ)⁻¹) atTop (𝓝 0) :=
NNReal.tendsto_algebraMap_inverse_atTop_nhds_zero_nat 𝕜

/-- The limit of `n / (n + x)` is 1, for any constant `x` (valid in `ℝ` or any topological division
algebra over `ℝ`, e.g., `ℂ`).

TODO: introduce a typeclass saying that `1 / n` tends to 0 at top, making it possible to get this
statement simultaneously on `ℚ`, `ℝ` and `ℂ`. -/
theorem tendsto_natCast_div_add_atTop {𝕜 : Type*} [DivisionRing 𝕜] [TopologicalSpace 𝕜]
[CharZero 𝕜] [Algebra ℝ 𝕜] [ContinuousSMul ℝ 𝕜] [IsTopologicalDivisionRing 𝕜] (x : 𝕜) :
algebra over `ℚ≥0`, e.g., `ℂ`). -/
theorem tendsto_natCast_div_add_atTop {𝕜 : Type*} [DivisionSemiring 𝕜] [TopologicalSpace 𝕜]
[CharZero 𝕜] [ContinuousSMul ℚ≥0 𝕜] [IsTopologicalSemiring 𝕜] [ContinuousInv₀ 𝕜] (x : 𝕜) :
Tendsto (fun n : ℕ ↦ (n : 𝕜) / (n + x)) atTop (𝓝 1) := by
convert Tendsto.congr' ((eventually_ne_atTop 0).mp (Eventually.of_forall fun n hn ↦ _)) _
· exact fun n : ℕ ↦ 1 / (1 + x / n)
Expand All @@ -88,12 +90,28 @@ theorem tendsto_natCast_div_add_atTop {𝕜 : Type*} [DivisionRing 𝕜] [Topolo
refine tendsto_const_nhds.div (tendsto_const_nhds.add ?_) (by simp)
simp_rw [div_eq_mul_inv]
refine tendsto_const_nhds.mul ?_
have := ((continuous_algebraMap 𝕜).tendsto _).comp tendsto_inverse_atTop_nhds_zero_nat
have := ((continuous_algebraMap ℚ≥0 𝕜).tendsto _).comp tendsto_inv_atTop_nhds_zero_nat
rw [map_zero, Filter.tendsto_atTop'] at this
refine Iff.mpr tendsto_atTop' ?_
intros
simp_all only [comp_apply, map_inv₀, map_natCast]

theorem tendsto_add_mul_div_add_mul_atTop_nhds {𝕜 : Type*} [Semifield 𝕜] [CharZero 𝕜]
[TopologicalSpace 𝕜] [ContinuousSMul ℚ≥0 𝕜] [IsTopologicalSemiring 𝕜] [ContinuousInv₀ 𝕜]
(a b c : 𝕜) {d : 𝕜} (hd : d ≠ 0) :
Tendsto (fun k : ℕ ↦ (a + c * k) / (b + d * k)) atTop (𝓝 (c / d)) := by
apply Filter.Tendsto.congr'
case f₁ => exact fun k ↦ (a * (↑k)⁻¹ + c) / (b * (↑k)⁻¹ + d)
· refine (eventually_ne_atTop 0).mp (Eventually.of_forall ?_)
intro h hx
dsimp
field_simp (discharger := norm_cast)
· apply Filter.Tendsto.div _ _ hd
all_goals
apply zero_add (_ : 𝕜) ▸ Filter.Tendsto.add_const _ _
apply mul_zero (_ : 𝕜) ▸ Filter.Tendsto.const_mul _ _
exact tendsto_inv_atTop_nhds_zero_nat

/-- For any positive `m : ℕ`, `((n % m : ℕ) : ℝ) / (n : ℝ)` tends to `0` as `n` tends to `∞`. -/
theorem tendsto_mod_div_atTop_nhds_zero_nat {m : ℕ} (hm : 0 < m) :
Tendsto (fun n : ℕ => ((n % m : ℕ) : ℝ) / (n : ℝ)) atTop (𝓝 0) := by
Expand Down Expand Up @@ -163,16 +181,14 @@ theorem tendsto_add_one_pow_atTop_atTop_of_pos
not_bddAbove_iff.2 fun _ ↦ Set.exists_range_iff.2 <| add_one_pow_unbounded_of_pos _ h

theorem tendsto_pow_atTop_atTop_of_one_lt
[Ring α] [LinearOrder α] [IsStrictOrderedRing α] [Archimedean α] {r : α}
(h : 1 < r) : Tendsto (fun n : ℕ ↦ r ^ n) atTop atTop :=
sub_add_cancel r 1 ▸ tendsto_add_one_pow_atTop_atTop_of_pos (sub_pos.2 h)

theorem Nat.tendsto_pow_atTop_atTop_of_one_lt {m : ℕ} (h : 1 < m) :
Tendsto (fun n : ℕ ↦ m ^ n) atTop atTop :=
tsub_add_cancel_of_le (le_of_lt h) ▸ tendsto_add_one_pow_atTop_atTop_of_pos (tsub_pos_of_lt h)
[Semiring α] [LinearOrder α] [IsStrictOrderedRing α] [ExistsAddOfLE α] [Archimedean α] {r : α}
(h : 1 < r) : Tendsto (fun n : ℕ ↦ r ^ n) atTop atTop := by
obtain ⟨r, r0, rfl⟩ := exists_pos_add_of_lt' h
rw [add_comm]
exact tendsto_add_one_pow_atTop_atTop_of_pos r0

theorem tendsto_pow_atTop_nhds_zero_of_lt_one {𝕜 : Type*}
[Field 𝕜] [LinearOrder 𝕜] [IsStrictOrderedRing 𝕜] [Archimedean 𝕜]
[Semifield 𝕜] [LinearOrder 𝕜] [IsStrictOrderedRing 𝕜] [ExistsAddOfLE 𝕜] [Archimedean 𝕜]
[TopologicalSpace 𝕜] [OrderTopology 𝕜] {r : 𝕜} (h₁ : 0 ≤ r) (h₂ : r < 1) :
Tendsto (fun n : ℕ ↦ r ^ n) atTop (𝓝 0) :=
h₁.eq_or_lt.elim
Expand Down Expand Up @@ -201,7 +217,7 @@ theorem tendsto_pow_atTop_nhds_zero_of_lt_one {𝕜 : Type*}
· simpa only [← abs_pow] using (tendsto_pow_atTop_nhds_zero_of_lt_one (abs_nonneg r)) h

theorem tendsto_pow_atTop_nhdsWithin_zero_of_lt_one {𝕜 : Type*}
[Field 𝕜] [LinearOrder 𝕜] [IsStrictOrderedRing 𝕜]
[Semifield 𝕜] [LinearOrder 𝕜] [IsStrictOrderedRing 𝕜] [ExistsAddOfLE 𝕜]
[Archimedean 𝕜] [TopologicalSpace 𝕜] [OrderTopology 𝕜] {r : 𝕜} (h₁ : 0 < r) (h₂ : r < 1) :
Tendsto (fun n : ℕ ↦ r ^ n) atTop (𝓝[>] 0) :=
tendsto_inf.2
Expand Down Expand Up @@ -743,3 +759,23 @@ lemma Nat.tendsto_div_const_atTop {n : ℕ} (hn : n ≠ 0) : Tendsto (· / n) at
rw [Tendsto, map_div_atTop_eq_nat n hn.bot_lt]

end

@[deprecated (since := "2025-10-27")]
alias tendsto_inverse_atTop_nhds_zero_nat := tendsto_inv_atTop_nhds_zero_nat

@[deprecated (since := "2025-10-27")]
alias NNReal.tendsto_inverse_atTop_nhds_zero_nat := tendsto_inv_atTop_nhds_zero_nat

@[deprecated (since := "2025-10-27")]
alias NNReal.tendsto_const_div_atTop_nhds_zero_nat := tendsto_const_div_atTop_nhds_zero_nat

@[deprecated (since := "2025-10-27")]
alias NNReal.tendsto_algebraMap_inverse_atTop_nhds_zero_nat :=
tendsto_algebraMap_inv_atTop_nhds_zero_nat

@[deprecated (since := "2025-10-27")]
alias tendsto_algebraMap_inverse_atTop_nhds_zero_nat :=
tendsto_algebraMap_inv_atTop_nhds_zero_nat

@[deprecated (since := "2025-10-27")]
protected alias Nat.tendsto_pow_atTop_atTop_of_one_lt := tendsto_pow_atTop_atTop_of_one_lt
26 changes: 6 additions & 20 deletions Mathlib/Analysis/SpecificLimits/RCLike.lean
Original file line number Diff line number Diff line change
Expand Up @@ -18,11 +18,6 @@ namespace RCLike

variable (𝕜 : Type*) [RCLike 𝕜]

theorem tendsto_inverse_atTop_nhds_zero_nat :
Tendsto (fun n : ℕ => (n : 𝕜)⁻¹) atTop (𝓝 0) := by
convert tendsto_algebraMap_inverse_atTop_nhds_zero_nat 𝕜
simp

theorem tendsto_ofReal_cobounded_cobounded :
Tendsto ofReal (Bornology.cobounded ℝ) (Bornology.cobounded 𝕜) :=
tendsto_norm_atTop_iff_cobounded.mp (mod_cast tendsto_norm_cobounded_atTop)
Expand All @@ -35,19 +30,10 @@ theorem tendsto_ofReal_atBot_cobounded :
Tendsto ofReal atBot (Bornology.cobounded 𝕜) :=
tendsto_norm_atTop_iff_cobounded.mp (mod_cast tendsto_abs_atBot_atTop)

variable {𝕜}

theorem tendsto_add_mul_div_add_mul_atTop_nhds (a b c : 𝕜) {d : 𝕜} (hd : d ≠ 0) :
Tendsto (fun k : ℕ ↦ (a + c * k) / (b + d * k)) atTop (𝓝 (c / d)) := by
apply Filter.Tendsto.congr'
case f₁ => exact fun k ↦ (a * (↑k)⁻¹ + c) / (b * (↑k)⁻¹ + d)
· refine (eventually_ne_atTop 0).mp (Eventually.of_forall ?_)
intro h hx
field_simp (discharger := norm_cast)
· apply Filter.Tendsto.div _ _ hd
all_goals
apply zero_add (_ : 𝕜) ▸ Filter.Tendsto.add_const _ _
apply mul_zero (_ : 𝕜) ▸ Filter.Tendsto.const_mul _ _
exact tendsto_inverse_atTop_nhds_zero_nat 𝕜

end RCLike

@[deprecated (since := "2025-10-27")]
alias RCLike.tendsto_inverse_atTop_nhds_zero_nat := tendsto_inv_atTop_nhds_zero_nat

@[deprecated (since := "2025-10-27")]
alias RCLike.tendsto_add_mul_div_add_mul_atTop_nhds := tendsto_add_mul_div_add_mul_atTop_nhds
Original file line number Diff line number Diff line change
Expand Up @@ -563,7 +563,7 @@ theorem translationNumber_eq_of_tendsto₀ {τ' : ℝ}
(h : Tendsto (fun n : ℕ => f^[n] 0 / n) atTop (𝓝 τ')) : τ f = τ' :=
f.translationNumber_eq_of_tendsto_aux <| by
simpa [Function.comp_def, transnumAuxSeq_def, coe_pow] using
h.comp (Nat.tendsto_pow_atTop_atTop_of_one_lt one_lt_two)
h.comp (tendsto_pow_atTop_atTop_of_one_lt one_lt_two)

theorem translationNumber_eq_of_tendsto₀' {τ' : ℝ}
(h : Tendsto (fun n : ℕ => f^[n + 1] 0 / (n + 1)) atTop (𝓝 τ')) : τ f = τ' :=
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Dynamics/Ergodic/AddCircle.lean
Original file line number Diff line number Diff line change
Expand Up @@ -113,7 +113,7 @@ theorem ergodic_zsmul {n : ℤ} (hn : 1 < |n|) : Ergodic fun y : AddCircle T =>
have hu₁ : ∀ j, (u j +ᵥ s : Set _) =ᵐ[volume] s := fun j => by
rw [vadd_eq_self_of_preimage_zsmul_eq_self hs' (hnu j)]
have hu₂ : Tendsto (fun j => addOrderOf <| u j) atTop atTop := by
simp_rw [hu₀]; exact Nat.tendsto_pow_atTop_atTop_of_one_lt hn
simp_rw [hu₀]; exact tendsto_pow_atTop_atTop_of_one_lt hn
rw [eventuallyConst_set']
exact ae_empty_or_univ_of_forall_vadd_ae_eq_self hs.nullMeasurableSet hu₁ hu₂ }

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/NumberTheory/Padics/Hensel.lean
Original file line number Diff line number Diff line change
Expand Up @@ -306,7 +306,7 @@ private theorem bound' : Tendsto (fun n : ℕ => ‖F.derivative.aeval a‖ * T
exact
tendsto_const_nhds.mul
(Tendsto.comp (tendsto_pow_atTop_nhds_zero_of_lt_one (norm_nonneg _) (T_lt_one hnorm))
(Nat.tendsto_pow_atTop_atTop_of_one_lt (by simp)))
(tendsto_pow_atTop_atTop_of_one_lt (by simp)))

private theorem bound :
∀ {ε}, ε > 0 → ∃ N : ℕ, ∀ {n}, n ≥ N → ‖F.derivative.aeval a‖ * T ^ 2 ^ n < ε := fun hε ↦
Expand Down