Skip to content

Commit f697ccc

Browse files
committed
feat: lemmas about monotonicity and asymptotics of rpow (#6140)
Co-authored-by: Frédéric Dupuis <31101893+dupuisf@users.noreply.github.com>
1 parent a28c6c6 commit f697ccc

File tree

2 files changed

+91
-0
lines changed

2 files changed

+91
-0
lines changed

Mathlib/Analysis/SpecialFunctions/Pow/Asymptotics.lean

Lines changed: 48 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -51,6 +51,54 @@ theorem tendsto_rpow_neg_atTop {y : ℝ} (hy : 0 < y) : Tendsto (fun x : ℝ =>
5151
(tendsto_rpow_atTop hy).inv_tendsto_atTop
5252
#align tendsto_rpow_neg_at_top tendsto_rpow_neg_atTop
5353

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+
54102
/-- The function `x ^ (a / (b * x + c))` tends to `1` at `+∞`, for any real numbers `a`, `b`, and
55103
`c` such that `b` is nonzero. -/
56104
theorem tendsto_rpow_div_mul_add (a b c : ℝ) (hb : 0 ≠ b) :

Mathlib/Analysis/SpecialFunctions/Pow/Real.lean

Lines changed: 43 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -480,6 +480,24 @@ theorem rpow_le_rpow_of_exponent_le (hx : 1 ≤ x) (hyz : y ≤ z) : x ^ y ≤ x
480480
rw [exp_le_exp]; exact mul_le_mul_of_nonneg_left hyz (log_nonneg hx)
481481
#align real.rpow_le_rpow_of_exponent_le Real.rpow_le_rpow_of_exponent_le
482482

483+
theorem rpow_lt_rpow_of_exponent_neg {x y z : ℝ} (hy : 0 < y) (hxy : y < x) (hz : z < 0) :
484+
x ^ z < y ^ z := by
485+
have hx : 0 < x := hy.trans hxy
486+
rw [←neg_neg z, Real.rpow_neg (le_of_lt hx) (-z), Real.rpow_neg (le_of_lt hy) (-z),
487+
inv_lt_inv (rpow_pos_of_pos hx _) (rpow_pos_of_pos hy _)]
488+
exact Real.rpow_lt_rpow (by positivity) hxy <| neg_pos_of_neg hz
489+
490+
theorem rpow_le_rpow_of_exponent_nonpos {x y : ℝ} (hy : 0 < y) (hxy : y ≤ x) (hz : z ≤ 0) :
491+
x ^ z ≤ y ^ z := by
492+
rcases ne_or_eq z 0 with hz_zero | rfl
493+
case inl =>
494+
rcases ne_or_eq x y with hxy' | rfl
495+
case inl =>
496+
exact le_of_lt <| rpow_lt_rpow_of_exponent_neg hy (Ne.lt_of_le (id (Ne.symm hxy')) hxy)
497+
(Ne.lt_of_le hz_zero hz)
498+
case inr => simp
499+
case inr => simp
500+
483501
@[simp]
484502
theorem rpow_le_rpow_left_iff (hx : 1 < x) : x ^ y ≤ x ^ z ↔ y ≤ z := by
485503
have x_pos : 0 < x := lt_trans zero_lt_one hx
@@ -637,6 +655,31 @@ theorem rpow_nat_inv_pow_nat {x : ℝ} (hx : 0 ≤ x) {n : ℕ} (hn : n ≠ 0) :
637655
rw [← rpow_nat_cast, ← rpow_mul hx, inv_mul_cancel hn0, rpow_one]
638656
#align real.rpow_nat_inv_pow_nat Real.rpow_nat_inv_pow_nat
639657

658+
lemma strictMono_rpow_of_base_gt_one {b : ℝ} (hb : 1 < b) :
659+
StrictMono (rpow b) := by
660+
show StrictMono (fun (x:ℝ) => b ^ x)
661+
simp_rw [Real.rpow_def_of_pos (zero_lt_one.trans hb)]
662+
exact exp_strictMono.comp <| StrictMono.const_mul strictMono_id <| Real.log_pos hb
663+
664+
lemma monotone_rpow_of_base_ge_one {b : ℝ} (hb : 1 ≤ b) :
665+
Monotone (rpow b) := by
666+
rcases lt_or_eq_of_le hb with hb | rfl
667+
case inl => exact (strictMono_rpow_of_base_gt_one hb).monotone
668+
case inr => intro _ _ _; simp
669+
670+
lemma strictAnti_rpow_of_base_lt_one {b : ℝ} (hb₀ : 0 < b) (hb₁ : b < 1) :
671+
StrictAnti (rpow b) := by
672+
show StrictAnti (fun (x:ℝ) => b ^ x)
673+
simp_rw [Real.rpow_def_of_pos hb₀]
674+
exact exp_strictMono.comp_strictAnti <| StrictMono.const_mul_of_neg strictMono_id
675+
<| Real.log_neg hb₀ hb₁
676+
677+
lemma antitone_rpow_of_base_le_one {b : ℝ} (hb₀ : 0 < b) (hb₁ : b ≤ 1) :
678+
Antitone (rpow b) := by
679+
rcases lt_or_eq_of_le hb₁ with hb₁ | rfl
680+
case inl => exact (strictAnti_rpow_of_base_lt_one hb₀ hb₁).antitone
681+
case inr => intro _ _ _; simp
682+
640683
end Real
641684

642685
/-!

0 commit comments

Comments
 (0)