diff --git a/src/analysis/calculus/deriv.lean b/src/analysis/calculus/deriv.lean index ff8cce5f9b94f..fe86e329005c9 100644 --- a/src/analysis/calculus/deriv.lean +++ b/src/analysis/calculus/deriv.lean @@ -811,6 +811,14 @@ lemma deriv_smul (hc : differentiable_at 𝕜 c x) (hf : differentiable_at 𝕜 deriv (λ y, c y • f y) x = c x • deriv f x + (deriv c x) • f x := (hc.has_deriv_at.smul hf.has_deriv_at).deriv +theorem has_strict_deriv_at.smul_const + (hc : has_strict_deriv_at c c' x) (f : F) : + has_strict_deriv_at (λ y, c y • f) (c' • f) x := +begin + have := hc.smul (has_strict_deriv_at_const x f), + rwa [smul_zero, zero_add] at this, +end + theorem has_deriv_within_at.smul_const (hc : has_deriv_within_at c c' s x) (f : F) : has_deriv_within_at (λ y, c y • f) (c' • f) s x := diff --git a/src/analysis/special_functions/integrals.lean b/src/analysis/special_functions/integrals.lean index 3d58ad889a228..90f35585e559a 100644 --- a/src/analysis/special_functions/integrals.lean +++ b/src/analysis/special_functions/integrals.lean @@ -78,6 +78,15 @@ begin rpow_def_of_pos hx.1, rpow_def_of_neg (by linarith [hx.1] : -x < 0)], } end +lemma interval_integrable_cpow {r : ℂ} (ha : 0 < a) (hb : 0 < b) : + interval_integrable (λ x : ℝ, (x : ℂ) ^ r) volume a b := +begin + refine (complex.continuous_of_real.continuous_on.cpow_const _).interval_integrable, + intros c hc, + left, + exact_mod_cast lt_of_lt_of_le (lt_min ha hb) hc.left, +end + @[simp] lemma interval_integrable_id : interval_integrable (λ x, x) μ a b := continuous_id.interval_integrable a b @@ -230,6 +239,25 @@ begin exact integral_eq_sub_of_has_deriv_at hderiv' (interval_integrable_rpow (or.inr h.2)) }, end +lemma integral_cpow {r : ℂ} (ha : 0 < a) (hb : 0 < b) (hr : r ≠ -1) : + ∫ (x : ℝ) in a..b, (x : ℂ) ^ r = (b ^ (r + 1) - a ^ (r + 1)) / (r + 1) := +begin + suffices : ∀ x ∈ set.interval a b, has_deriv_at (λ x : ℝ, (x : ℂ) ^ (r + 1) / (r + 1)) (x ^ r) x, + { rw sub_div, + exact integral_eq_sub_of_has_deriv_at this (interval_integrable_cpow ha hb) }, + intros x hx, + suffices : has_deriv_at (λ z : ℂ, z ^ (r + 1) / (r + 1)) (x ^ r) x, + { simpa using has_deriv_at.comp x this complex.of_real_clm.has_deriv_at }, + have hx' : 0 < (x : ℂ).re ∨ (x : ℂ).im ≠ 0, + { left, + norm_cast, + calc 0 < min a b : lt_min ha hb ... ≤ x : hx.left, }, + convert ((has_deriv_at_id (x:ℂ)).cpow_const hx').div_const (r + 1), + simp only [id.def, add_sub_cancel, mul_one], rw [mul_comm, mul_div_cancel], + contrapose! hr, rwa add_eq_zero_iff_eq_neg at hr, +end + + lemma integral_zpow {n : ℤ} (h : 0 ≤ n ∨ n ≠ -1 ∧ (0 : ℝ) ∉ [a, b]) : ∫ x in a..b, x ^ n = (b ^ (n + 1) - a ^ (n + 1)) / (n + 1) := begin diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index b83f1332f22ca..79256ad828bee 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -255,6 +255,11 @@ lemma continuous.const_cpow {b : ℂ} (hf : continuous f) (h : b ≠ 0 ∨ ∀ a continuous (λ x, b ^ f x) := continuous_iff_continuous_at.2 $ λ a, (hf.continuous_at.const_cpow $ h.imp id $ λ h, h a) +lemma continuous_on.cpow_const {b : ℂ} (hf : continuous_on f s) + (h : ∀ (a : α), a ∈ s → 0 < (f a).re ∨ (f a).im ≠ 0) : + continuous_on (λ x, (f x) ^ b) s := +hf.cpow continuous_on_const h + end lim namespace real