Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - feat: improper integration by parts #10874

Closed
wants to merge 9 commits into from
Closed
6 changes: 6 additions & 0 deletions Mathlib/MeasureTheory/Integral/FundThmCalculus.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1340,6 +1340,8 @@ section Parts

variable [NormedRing A] [NormedAlgebra ℝ A] [CompleteSpace A]

/-- For improper integrals, see `MeasureTheory.integral_deriv_mul_eq_sub`,
`MeasureTheory.integral_Ioi_deriv_mul_eq_sub`, and `MeasureTheory.integral_Iic_deriv_mul_eq_sub`. -/
theorem integral_deriv_mul_eq_sub {u v u' v' : ℝ → A} (hu : ∀ x ∈ uIcc a b, HasDerivAt u (u' x) x)
(hv : ∀ x ∈ uIcc a b, HasDerivAt v (v' x) x) (hu' : IntervalIntegrable u' volume a b)
(hv' : IntervalIntegrable v' volume a b) :
Expand All @@ -1349,6 +1351,10 @@ theorem integral_deriv_mul_eq_sub {u v u' v' : ℝ → A} (hu : ∀ x ∈ uIcc a
(hv'.continuousOn_mul (HasDerivAt.continuousOn hu))
#align interval_integral.integral_deriv_mul_eq_sub intervalIntegral.integral_deriv_mul_eq_sub

/-- **Integration by parts**. For improper integrals, see
`MeasureTheory.integral_mul_deriv_eq_deriv_mul`,
`MeasureTheory.integral_Ioi_mul_deriv_eq_deriv_mul`,
and `MeasureTheory.integral_Iic_mul_deriv_eq_deriv_mul`. -/
theorem integral_mul_deriv_eq_deriv_mul {u v u' v' : ℝ → A}
(hu : ∀ x ∈ uIcc a b, HasDerivAt u (u' x) x) (hv : ∀ x ∈ uIcc a b, HasDerivAt v (v' x) x)
(hu' : IntervalIntegrable u' volume a b) (hv' : IntervalIntegrable v' volume a b) :
Expand Down
92 changes: 92 additions & 0 deletions Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1054,4 +1054,96 @@ theorem integrableOn_Ioi_comp_mul_right_iff (f : ℝ → E) (c : ℝ) {a : ℝ}

end IoiIntegrability

/-!
## Integration by parts
-/

section IntegrationByParts

variable {A : Type*} [NormedRing A] [NormedAlgebra ℝ A] [CompleteSpace A]
{a b : ℝ} {a' b' : A} {u : ℝ → A} {v : ℝ → A} {u' : ℝ → A} {v' : ℝ → A}

/-- For finite intervals, see: `intervalIntegral.integral_deriv_mul_eq_sub`. -/
theorem integral_deriv_mul_eq_sub
(hu : ∀ x, HasDerivAt u (u' x) x) (hv : ∀ x, HasDerivAt v (v' x) x)
(huv : Integrable (u' * v + u * v'))
(h_bot : Tendsto (u * v) atBot (𝓝 a')) (h_top : Tendsto (u * v) atTop (𝓝 b')) :
∫ (x : ℝ), u' x * v x + u x * v' x = b' - a' :=
integral_of_hasDerivAt_of_tendsto (fun x ↦ (hu x).mul (hv x)) huv h_bot h_top

/-- **Integration by parts on (-∞, ∞).**
For finite intervals, see: `intervalIntegral.integral_mul_deriv_eq_deriv_mul`. -/
theorem integral_mul_deriv_eq_deriv_mul
(hu : ∀ x, HasDerivAt u (u' x) x) (hv : ∀ x, HasDerivAt v (v' x) x)
(huv' : Integrable (u * v')) (hu'v : Integrable (u' * v))
(h_bot : Tendsto (u * v) atBot (𝓝 a')) (h_top : Tendsto (u * v) atTop (𝓝 b')) :
∫ (x : ℝ), u x * v' x = b' - a' - ∫ (x : ℝ), u' x * v x := by
rw [Pi.mul_def] at huv' hu'v
rw [eq_sub_iff_add_eq, ← integral_add huv' hu'v]
simpa only [add_comm] using integral_deriv_mul_eq_sub hu hv (hu'v.add huv') h_bot h_top

/-- For finite intervals, see: `intervalIntegral.integral_deriv_mul_eq_sub`. -/
theorem integral_Ioi_deriv_mul_eq_sub
(hu : ∀ x ∈ Ioi a, HasDerivAt u (u' x) x) (hv : ∀ x ∈ Ioi a, HasDerivAt v (v' x) x)
(huv : IntegrableOn (u' * v + u * v') (Ioi a))
(h_zero : Tendsto (u * v) (𝓝[>] a) (𝓝 a')) (h_infty : Tendsto (u * v) atTop (𝓝 b')) :
∫ (x : ℝ) in Ioi a, u' x * v x + u x * v' x = b' - a' := by
rw [← Ici_diff_left] at h_zero
let f := Function.update (u * v) a a'
Copy link
Member

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I would argue that this should instead be handled directly by integral_Ioi_of_hasDerivAt_of_tendsto, which should assume existence of a limit rather than continuity. We could then also avoid changing the function by applying FTC on smaller intervals not touching the left bound. If you want to tinker with it that would be great, otherwise can you leave a TODO/open an issue about it?

Copy link
Collaborator Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I added the TODO comment and started tinkering here https://github.com/leanprover-community/mathlib4/pull/11226/files

have hderiv : ∀ x ∈ Ioi a, HasDerivAt f (u' x * v x + u x * v' x) x := by
intro x hx
apply ((hu x hx).mul (hv x hx)).congr_of_eventuallyEq
filter_upwards [Ioi_mem_nhds hx] with x (hx : a < x)
exact Function.update_noteq (ne_of_gt hx) a' (u * v)
have htendsto : Tendsto f atTop (𝓝 b') := by
apply h_infty.congr'
filter_upwards [Ioi_mem_atTop a] with x (hx : a < x)
exact (Function.update_noteq (ne_of_gt hx) a' (u * v)).symm
llllvvuu marked this conversation as resolved.
Show resolved Hide resolved
simpa using integral_Ioi_of_hasDerivAt_of_tendsto
(continuousWithinAt_update_same.mpr h_zero) hderiv huv htendsto

/-- **Integration by parts on (a, ∞).**
For finite intervals, see: `intervalIntegral.integral_mul_deriv_eq_deriv_mul`. -/
theorem integral_Ioi_mul_deriv_eq_deriv_mul
(hu : ∀ x ∈ Ioi a, HasDerivAt u (u' x) x) (hv : ∀ x ∈ Ioi a, HasDerivAt v (v' x) x)
(huv' : IntegrableOn (u * v') (Ioi a)) (hu'v : IntegrableOn (u' * v) (Ioi a))
(h_zero : Tendsto (u * v) (𝓝[>] a) (𝓝 a')) (h_infty : Tendsto (u * v) atTop (𝓝 b')) :
∫ (x : ℝ) in Ioi a, u x * v' x = b' - a' - ∫ (x : ℝ) in Ioi a, u' x * v x := by
rw [Pi.mul_def] at huv' hu'v
rw [eq_sub_iff_add_eq, ← integral_add huv' hu'v]
simpa only [add_comm] using integral_Ioi_deriv_mul_eq_sub hu hv (hu'v.add huv') h_zero h_infty

/-- For finite intervals, see: `intervalIntegral.integral_deriv_mul_eq_sub`. -/
theorem integral_Iic_deriv_mul_eq_sub
(hu : ∀ x ∈ Iio a, HasDerivAt u (u' x) x) (hv : ∀ x ∈ Iio a, HasDerivAt v (v' x) x)
(huv : IntegrableOn (u' * v + u * v') (Iic a))
(h_zero : Tendsto (u * v) (𝓝[<] a) (𝓝 a')) (h_infty : Tendsto (u * v) atBot (𝓝 b')) :
∫ (x : ℝ) in Iic a, u' x * v x + u x * v' x = a' - b' := by
rw [← Iic_diff_right] at h_zero
let f := Function.update (u * v) a a'
have hderiv : ∀ x ∈ Iio a, HasDerivAt f (u' x * v x + u x * v' x) x := by
intro x hx
apply ((hu x hx).mul (hv x hx)).congr_of_eventuallyEq
filter_upwards [Iio_mem_nhds hx] with x (hx : x < a)
exact Function.update_noteq (ne_of_lt hx) a' (u * v)
have htendsto : Tendsto f atBot (𝓝 b') := by
apply h_infty.congr'
filter_upwards [Iio_mem_atBot a] with x (hx : x < a)
exact (Function.update_noteq (ne_of_lt hx) a' (u * v)).symm
simpa using integral_Iic_of_hasDerivAt_of_tendsto
(continuousWithinAt_update_same.mpr h_zero) hderiv huv htendsto

/-- **Integration by parts on (∞, a].**
For finite intervals, see: `intervalIntegral.integral_mul_deriv_eq_deriv_mul`. -/
theorem integral_Iic_mul_deriv_eq_deriv_mul
(hu : ∀ x ∈ Iio a, HasDerivAt u (u' x) x) (hv : ∀ x ∈ Iio a, HasDerivAt v (v' x) x)
(huv' : IntegrableOn (u * v') (Iic a)) (hu'v : IntegrableOn (u' * v) (Iic a))
(h_zero : Tendsto (u * v) (𝓝[<] a) (𝓝 a')) (h_infty : Tendsto (u * v) atBot (𝓝 b')) :
∫ (x : ℝ) in Iic a, u x * v' x = a' - b' - ∫ (x : ℝ) in Iic a, u' x * v x := by
rw [Pi.mul_def] at huv' hu'v
rw [eq_sub_iff_add_eq, ← integral_add huv' hu'v]
simpa only [add_comm] using integral_Iic_deriv_mul_eq_sub hu hv (hu'v.add huv') h_zero h_infty

end IntegrationByParts

end MeasureTheory
Loading