Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 8e83805

Browse files
committed
chore(analysis/special_functions/pow): +2 lemmas about nnreal.rpow (#4272)
`λ x, x^y` is continuous in more cases than `λ p, p.1^p.2`.
1 parent 3177493 commit 8e83805

File tree

1 file changed

+24
-4
lines changed

1 file changed

+24
-4
lines changed

src/analysis/special_functions/pow.lean

Lines changed: 24 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -7,10 +7,15 @@ import analysis.special_functions.trigonometric
77
import analysis.calculus.extend_deriv
88

99
/-!
10-
# Power function on `ℂ`, `ℝ` and `ℝ⁺`
10+
# Power function on `ℂ`, `ℝ`, `ℝ≥0`, and `ennreal`
1111
12-
We construct the power functions `x ^ y` where `x` and `y` are complex numbers, or `x` and `y` are
13-
real numbers, or `x` is a nonnegative real and `y` is real, and prove their basic properties.
12+
We construct the power functions `x ^ y` where
13+
* `x` and `y` are complex numbers,
14+
* or `x` and `y` are real numbers,
15+
* or `x` is a nonnegative real number and `y` is a real number;
16+
* or `x` is a number from `[0, +∞]` (a.k.a. `ennreal`) and `y` is a real number.
17+
18+
We also prove basic properties of these functions.
1419
-/
1520

1621
noncomputable theory
@@ -880,7 +885,22 @@ open filter
880885
lemma filter.tendsto.nnrpow {α : Type*} {f : filter α} {u : α → ℝ≥0} {v : α → ℝ} {x : ℝ≥0} {y : ℝ}
881886
(hx : tendsto u f (𝓝 x)) (hy : tendsto v f (𝓝 y)) (h : x ≠ 00 < y) :
882887
tendsto (λ a, (u a) ^ (v a)) f (𝓝 (x ^ y)) :=
883-
tendsto.comp (nnreal.continuous_at_rpow h) (tendsto.prod_mk_nhds hx hy)
888+
tendsto.comp (nnreal.continuous_at_rpow h) (hx.prod_mk_nhds hy)
889+
890+
namespace nnreal
891+
892+
lemma continuous_at_rpow_const {x : ℝ≥0} {y : ℝ} (h : x ≠ 00 ≤ y) :
893+
continuous_at (λ z, z^y) x :=
894+
h.elim (λ h, tendsto_id.nnrpow tendsto_const_nhds (or.inl h)) $
895+
λ h, h.eq_or_lt.elim
896+
(λ h, h ▸ by simp only [rpow_zero, continuous_at_const])
897+
(λ h, tendsto_id.nnrpow tendsto_const_nhds (or.inr h))
898+
899+
lemma continuous_rpow_const {y : ℝ} (h : 0 ≤ y) :
900+
continuous (λ x : ℝ≥0, x^y) :=
901+
continuous_iff_continuous_at.2 $ λ x, continuous_at_rpow_const (or.inr h)
902+
903+
end nnreal
884904

885905
namespace ennreal
886906

0 commit comments

Comments
 (0)