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

Commit d04e034

Browse files
feat(measure_theory/interval_integral): FTC-2 (#4945)
The second fundamental theorem of calculus and supporting lemmas Co-authored-by: Yury G. Kudryashov <urkud@urkud.name>
1 parent 46c35cc commit d04e034

File tree

2 files changed

+145
-3
lines changed

2 files changed

+145
-3
lines changed

src/analysis/calculus/mean_value.lean

Lines changed: 51 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -409,6 +409,57 @@ theorem norm_image_sub_le_of_norm_deriv_le_segment_01 {C : ℝ}
409409
by simpa only [sub_zero, mul_one]
410410
using norm_image_sub_le_of_norm_deriv_le_segment hf bound 1 (right_mem_Icc.2 zero_le_one)
411411

412+
theorem constant_of_has_deriv_right_zero (hcont : continuous_on f (Icc a b))
413+
(hderiv : ∀ x ∈ Ico a b, has_deriv_within_at f 0 (Ici x) x) :
414+
∀ x ∈ Icc a b, f x = f a :=
415+
by simpa only [zero_mul, norm_le_zero_iff, sub_eq_zero] using
416+
λ x hx, norm_image_sub_le_of_norm_deriv_right_le_segment
417+
hcont hderiv (λ y hy, by rw norm_le_zero_iff) x hx
418+
419+
theorem constant_of_deriv_within_zero (hdiff : differentiable_on ℝ f (Icc a b))
420+
(hderiv : ∀ x ∈ Ico a b, deriv_within f (Icc a b) x = 0) :
421+
∀ x ∈ Icc a b, f x = f a :=
422+
begin
423+
have H : ∀ x ∈ Ico a b, ∥deriv_within f (Icc a b) x∥ ≤ 0 :=
424+
by simpa only [norm_le_zero_iff] using λ x hx, hderiv x hx,
425+
simpa only [zero_mul, norm_le_zero_iff, sub_eq_zero] using
426+
λ x hx, norm_image_sub_le_of_norm_deriv_le_segment hdiff H x hx,
427+
end
428+
429+
variables {f' g : ℝ → E}
430+
431+
/-- If two continuous functions on `[a, b]` have the same right derivative and are equal at `a`,
432+
then they are equal everywhere on `[a, b]`. -/
433+
theorem eq_of_has_deriv_right_eq
434+
(derivf : ∀ x ∈ Ico a b, has_deriv_within_at f (f' x) (Ici x) x)
435+
(derivg : ∀ x ∈ Ico a b, has_deriv_within_at g (f' x) (Ici x) x)
436+
(fcont : continuous_on f (Icc a b)) (gcont : continuous_on g (Icc a b))
437+
(hi : f a = g a) :
438+
∀ y ∈ Icc a b, f y = g y :=
439+
begin
440+
simp only [← @sub_eq_zero _ _ (f _)] at hi ⊢,
441+
exact hi ▸ constant_of_has_deriv_right_zero (fcont.sub gcont)
442+
(λ y hy, by simpa only [sub_self] using (derivf y hy).sub (derivg y hy)),
443+
end
444+
445+
/-- If two differentiable functions on `[a, b]` have the same derivative within `[a, b]` everywhere
446+
on `[a, b)` and are equal at `a`, then they are equal everywhere on `[a, b]`. -/
447+
theorem eq_of_deriv_within_eq (fdiff : differentiable_on ℝ f (Icc a b))
448+
(gdiff : differentiable_on ℝ g (Icc a b))
449+
(hderiv : eq_on (deriv_within f (Icc a b)) (deriv_within g (Icc a b)) (Ico a b))
450+
(hi : f a = g a) :
451+
∀ y ∈ Icc a b, f y = g y :=
452+
begin
453+
have A : ∀ y ∈ Ico a b, has_deriv_within_at f (deriv_within f (Icc a b) y) (Ici y) y :=
454+
λ y hy, (fdiff y (mem_Icc_of_Ico hy)).has_deriv_within_at.nhds_within
455+
(Icc_mem_nhds_within_Ici hy),
456+
have B : ∀ y ∈ Ico a b, has_deriv_within_at g (deriv_within g (Icc a b) y) (Ici y) y :=
457+
λ y hy, (gdiff y (mem_Icc_of_Ico hy)).has_deriv_within_at.nhds_within
458+
(Icc_mem_nhds_within_Ici hy),
459+
exact eq_of_has_deriv_right_eq A (λ y hy, (hderiv hy).symm ▸ B y hy) fdiff.continuous_on
460+
gdiff.continuous_on hi
461+
end
462+
412463
end
413464

414465
/-! ### Vector-valued functions `f : E → F` -/

src/measure_theory/interval_integral.lean

Lines changed: 94 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -5,7 +5,8 @@ Author: Yury G. Kudryashov
55
-/
66
import measure_theory.set_integral
77
import measure_theory.lebesgue_measure
8-
import analysis.calculus.deriv
8+
import analysis.calculus.fderiv_measurable
9+
import analysis.calculus.mean_value
910

1011
/-!
1112
# Integral over an interval
@@ -203,10 +204,14 @@ begin
203204
all_goals
204205
{ refine measure_theory.integrable_on.mono_set _ Ioc_subset_Icc_self,
205206
refine continuous_on.integrable_on_compact compact_Icc hum (hu.mono _) },
206-
{ exact Icc_subset_interval },
207-
{ exact Icc_subset_interval' }
207+
exacts [Icc_subset_interval, Icc_subset_interval']
208208
end
209209

210+
lemma continuous_on.interval_integrable_of_Icc {u : ℝ → E} {a b : ℝ} (h : a ≤ b)
211+
(hu : continuous_on u (Icc a b)) (hum : measurable u) :
212+
interval_integrable u μ a b :=
213+
continuous_on.interval_integrable ((interval_of_le h).symm ▸ hu) hum
214+
210215
/-- A continuous function on `ℝ` is `interval_integrable` with respect to any locally finite measure
211216
`ν` on ℝ. -/
212217
lemma continuous.interval_integrable [borel_space E] {u : ℝ → E} (hu : continuous u) (a b : ℝ) :
@@ -1214,4 +1219,90 @@ lemma deriv_within_integral_left (hf : interval_integrable f volume a b)
12141219
deriv_within (λ u, ∫ x in u..b, f x) s a = -f a :=
12151220
(integral_has_deriv_within_at_left hf ha).deriv_within hs
12161221

1222+
1223+
/-~ ### Theorems pertaining to FTC-2. -/
1224+
1225+
variables {f' : ℝ → E}
1226+
1227+
/-- The integral of a continuous function is differentiable on a real set `s`. -/
1228+
theorem differentiable_on_integral_of_continuous {s : set ℝ}
1229+
(hintg : ∀ x ∈ s, interval_integrable f volume a x) (hcont : continuous f) :
1230+
differentiable_on ℝ (λ u, ∫ x in a..u, f x) s :=
1231+
λ y hy, (integral_has_deriv_at_right (hintg y hy)
1232+
hcont.continuous_at).differentiable_at.differentiable_within_at
1233+
1234+
/-- The integral of a continuous function is continuous on a real set `s`. This is true even
1235+
without the assumption of continuity, but a proof of that fact does not yet exist in mathlib. -/
1236+
theorem continuous_on_integral_of_continuous {s : set ℝ}
1237+
(hintg : ∀ x ∈ s, interval_integrable f volume a x) (hcont : continuous f) :
1238+
continuous_on (λ u, ∫ x in a..u, f x) s :=
1239+
(differentiable_on_integral_of_continuous hintg hcont).continuous_on
1240+
1241+
/-- Fundamental theorem of calculus-2: If `f : ℝ → E` is continuous on `[a, b]` and has a right
1242+
derivative at `f' x` for all `x` in `[a, b)`, and `f'` is continuous on `[a, b]` and measurable,
1243+
then `∫ y in a..b, f' y` equals `f b - f a`. -/
1244+
theorem integral_eq_sub_of_has_deriv_right_of_le (hab : a ≤ b) (hcont : continuous_on f (Icc a b))
1245+
(hderiv : ∀ x ∈ Ico a b, has_deriv_within_at f (f' x) (Ici x) x)
1246+
(hcont' : continuous_on f' (Icc a b)) (hmeas' : measurable f') :
1247+
∫ y in a..b, f' y = f b - f a :=
1248+
begin
1249+
refine eq_sub_of_add_eq (eq_of_has_deriv_right_eq (λ y hy, _) hderiv
1250+
(λ y hy, _) hcont (by simp) _ (right_mem_Icc.2 hab)),
1251+
{ refine (integral_has_deriv_within_at_right _ _).add_const _,
1252+
{ refine (hcont'.mono _).interval_integrable hmeas',
1253+
simp [hy.1, Icc_subset_Icc_right hy.2.le] },
1254+
{ exact (hcont' _ (mem_Icc_of_Ico hy)).mono_of_mem (Icc_mem_nhds_within_Ioi hy) } },
1255+
{ -- TODO: prove that integral of any integrable function is continuous, and use here
1256+
letI : tendsto_Ixx_class Ioc (𝓟 (Icc a b)) (𝓟 (Ioc a b)) :=
1257+
tendsto_Ixx_class_principal.2 (λ x hx y hy, Ioc_subset_Ioc hx.1 hy.2),
1258+
haveI : is_measurably_generated (𝓝[Ioc a b] y) :=
1259+
is_measurable_Ioc.nhds_within_is_measurably_generated y,
1260+
letI : FTC_filter y (𝓝[Icc a b] y) (𝓝[Ioc a b] y) := ⟨pure_le_nhds_within hy, inf_le_left⟩,
1261+
refine (integral_has_deriv_within_at_right _ _).continuous_within_at.add
1262+
continuous_within_at_const,
1263+
{ exact (hcont'.mono $ Icc_subset_Icc_right hy.2).interval_integrable_of_Icc hy.1 hmeas' },
1264+
{ exact (hcont' y hy).mono Ioc_subset_Icc_self } }
1265+
end
1266+
1267+
/-- Fundamental theorem of calculus-2: If `f : ℝ → E` is continuous on `[a, b]` (where `a ≤ b`) and
1268+
has a right derivative at `f' x` for all `x` in `[a, b)`, and `f'` is continuous on `[a, b]` and
1269+
measurable, then `∫ y in a..b, f' y` equals `f b - f a`. -/
1270+
theorem integral_eq_sub_of_has_deriv_right (hcont : continuous_on f (interval a b))
1271+
(hderiv : ∀ x ∈ Ico (min a b) (max a b), has_deriv_within_at f (f' x) (Ici x) x)
1272+
(hcont' : continuous_on f' (interval a b)) (hmeas' : measurable f') :
1273+
∫ y in a..b, f' y = f b - f a :=
1274+
begin
1275+
cases le_total a b with hab hab,
1276+
{ simp only [interval_of_le, min_eq_left, max_eq_right, hab] at hcont hcont' hderiv,
1277+
exact integral_eq_sub_of_has_deriv_right_of_le hab hcont hderiv hcont' hmeas' },
1278+
{ simp only [interval_of_ge, min_eq_right, max_eq_left, hab] at hcont hcont' hderiv,
1279+
rw [integral_symm, integral_eq_sub_of_has_deriv_right_of_le hab hcont hderiv hcont' hmeas',
1280+
neg_sub] }
1281+
end
1282+
1283+
/-- Fundamental theorem of calculus-2: If `f : ℝ → E` is continuous on `[a, b]` and has a derivative
1284+
at `f' x` for all `x` in `[a, b)`, and `f'` is continuous on `[a, b]` and measurable, then
1285+
`∫ y in a..b, f' y` equals `f b - f a`. -/
1286+
theorem integral_eq_sub_of_has_deriv_at' (hcont : continuous_on f (interval a b))
1287+
(hderiv : ∀ x ∈ Ico (min a b) (max a b), has_deriv_at f (f' x) x)
1288+
(hcont' : continuous_on f' (interval a b)) (hmeas' : measurable f') :
1289+
∫ y in a..b, f' y = f b - f a :=
1290+
integral_eq_sub_of_has_deriv_right hcont (λ x hx, (hderiv x hx).has_deriv_within_at) hcont' hmeas'
1291+
1292+
/-- Fundamental theorem of calculus-2: If `f : ℝ → E` has a derivative at `f' x` for all `x` in
1293+
`[a, b)` and `f'` is continuous on `[a, b]` and measurable, then `∫ y in a..b, f' y` equals
1294+
`f b - f a`. -/
1295+
theorem integral_eq_sub_of_has_deriv_at (hderiv : ∀ x ∈ interval a b, has_deriv_at f (f' x) x)
1296+
(hcont' : continuous_on f' (interval a b)) (hmeas' : measurable f') :
1297+
∫ y in a..b, f' y = f b - f a :=
1298+
integral_eq_sub_of_has_deriv_at' (λ x hx, (hderiv x hx).continuous_at.continuous_within_at)
1299+
(λ x hx, hderiv _ (mem_Icc_of_Ico hx)) hcont' hmeas'
1300+
1301+
/-- Fundamental theorem of calculus-2: If `f : ℝ → E` is differentiable at every `x` in `[a, b]` and
1302+
its derivative is continuous on `[a, b]`, then `∫ y in a..b, deriv f y` equals `f b - f a`. -/
1303+
theorem integral_deriv_eq_sub (hderiv : ∀ x ∈ interval a b, differentiable_at ℝ f x)
1304+
(hcont' : continuous_on (deriv f) (interval a b)) :
1305+
∫ y in a..b, deriv f y = f b - f a :=
1306+
integral_eq_sub_of_has_deriv_at (λ x hx, (hderiv x hx).has_deriv_at) hcont' (measurable_deriv f)
1307+
12171308
end interval_integral

0 commit comments

Comments
 (0)