@@ -51,6 +51,54 @@ theorem tendsto_rpow_neg_atTop {y : ℝ} (hy : 0 < y) : Tendsto (fun x : ℝ =>
51
51
(tendsto_rpow_atTop hy).inv_tendsto_atTop
52
52
#align tendsto_rpow_neg_at_top tendsto_rpow_neg_atTop
53
53
54
+ open Asymptotics in
55
+ lemma tendsto_rpow_atTop_of_base_lt_one (b : ℝ) (hb₀ : -1 < b) (hb₁ : b < 1 ) :
56
+ Tendsto (rpow b) atTop (𝓝 (0 :ℝ)) := by
57
+ show Tendsto (fun z => b^z) atTop (𝓝 0 )
58
+ rcases lt_trichotomy b 0 with hb|rfl|hb
59
+ case inl => -- b < 0
60
+ simp_rw [Real.rpow_def_of_nonpos hb.le, hb.ne, ite_false]
61
+ rw [←isLittleO_const_iff (c := (1 :ℝ)) one_ne_zero, (one_mul (1 : ℝ)).symm]
62
+ refine IsLittleO.mul_isBigO ?exp ?cos
63
+ case exp =>
64
+ rw [isLittleO_const_iff one_ne_zero]
65
+ refine tendsto_exp_atBot.comp <| (tendsto_const_mul_atBot_of_neg ?_).mpr tendsto_id
66
+ rw [←log_neg_eq_log, log_neg_iff (by linarith)]
67
+ linarith
68
+ case cos =>
69
+ rw [isBigO_iff]
70
+ exact ⟨1 , eventually_of_forall fun x => by simp [Real.abs_cos_le_one]⟩
71
+ case inr.inl => -- b = 0
72
+ refine Tendsto.mono_right ?_ (Iff.mpr pure_le_nhds_iff rfl)
73
+ rw [tendsto_pure]
74
+ filter_upwards [eventually_ne_atTop 0 ] with _ hx
75
+ simp [hx]
76
+ case inr.inr => -- b > 0
77
+ simp_rw [Real.rpow_def_of_pos hb]
78
+ refine tendsto_exp_atBot.comp <| (tendsto_const_mul_atBot_of_neg ?_).mpr tendsto_id
79
+ exact (log_neg_iff hb).mpr hb₁
80
+
81
+ lemma tendsto_rpow_atTop_of_base_gt_one (b : ℝ) (hb : 1 < b) :
82
+ Tendsto (rpow b) atBot (𝓝 (0 :ℝ)) := by
83
+ show Tendsto (fun z => b^z) atBot (nhds 0 )
84
+ simp_rw [Real.rpow_def_of_pos (by positivity : 0 < b)]
85
+ refine tendsto_exp_atBot.comp <| (tendsto_const_mul_atBot_of_pos ?_).mpr tendsto_id
86
+ exact (log_pos_iff (by positivity)).mpr <| by aesop
87
+
88
+ lemma tendsto_rpow_atBot_of_base_lt_one (b : ℝ) (hb₀ : 0 < b) (hb₁ : b < 1 ) :
89
+ Tendsto (rpow b) atBot atTop := by
90
+ show Tendsto (fun z => b^z) atBot atTop
91
+ simp_rw [Real.rpow_def_of_pos (by positivity : 0 < b)]
92
+ refine tendsto_exp_atTop.comp <| (tendsto_const_mul_atTop_iff_neg <| tendsto_id (α := ℝ)).mpr ?_
93
+ exact (log_neg_iff hb₀).mpr hb₁
94
+
95
+ lemma tendsto_rpow_atBot_of_base_gt_one (b : ℝ) (hb : 1 < b) : Tendsto (rpow b) atBot (𝓝 0 ) := by
96
+ show Tendsto (fun z => b^z) atBot (𝓝 0 )
97
+ simp_rw [Real.rpow_def_of_pos (by positivity : 0 < b)]
98
+ refine tendsto_exp_atBot.comp <| (tendsto_const_mul_atBot_iff_pos <| tendsto_id (α := ℝ)).mpr ?_
99
+ exact (log_pos_iff (by positivity)).mpr <| by aesop
100
+
101
+
54
102
/-- The function `x ^ (a / (b * x + c))` tends to `1` at `+∞`, for any real numbers `a`, `b`, and
55
103
`c` such that `b` is nonzero. -/
56
104
theorem tendsto_rpow_div_mul_add (a b c : ℝ) (hb : 0 ≠ b) :
0 commit comments