@@ -4223,6 +4223,22 @@ Qed.
42234223
42244224End ae_ge0_le_integral.
42254225
4226+ Section integral_bounded.
4227+ Context d {T : measurableType d} {R : realType}.
4228+ Variable mu : {measure set T -> \bar R}.
4229+ Local Open Scope ereal_scope.
4230+
4231+ Lemma integral_le_bound (D : set T) (f : T -> \bar R) (M : \bar R) :
4232+ measurable D -> measurable_fun D f -> 0 <= M ->
4233+ {ae mu, forall x, D x -> `|f x| <= M} ->
4234+ \int[mu]_(x in D) `|f x| <= M * mu D.
4235+ Proof .
4236+ move=> mD mf M0 dfx; rewrite -integral_cst => //.
4237+ by apply: ae_ge0_le_integral => //; exact: measurableT_comp.
4238+ Qed .
4239+
4240+ End integral_bounded.
4241+
42264242Section integral_ae_eq.
42274243Local Open Scope ereal_scope.
42284244Context d (T : measurableType d) (R : realType) (mu : {measure set T -> \bar R}).
@@ -5346,121 +5362,83 @@ Qed.
53465362End sfinite_fubini.
53475363Arguments sfinite_Fubini {d d' X Y R} m1 m2 f.
53485364
5349- Section integral_bounded.
5350-
5351- Lemma integral_le_bound d (R : realType) (T : measurableType d)
5352- (mu : {measure set T -> \bar R}) (D : set T) (f : T -> \bar R) (M : R) :
5353- measurable D -> measurable_fun D f ->
5354- (0 <= M%:E)%E -> ({ae mu, forall x, D x -> `|f x| <= M%:E})%E ->
5355- (\int[mu]_(x in D) `|f x| <= M%:E * mu D)%E.
5356- Proof .
5357- move=> mD intf M0 dfx; rewrite -integral_cst => //.
5358- apply: ae_ge0_le_integral => //.
5359- by apply: measurableT_comp => //; case/integrableP:intf.
5360- Qed .
5361-
5362- End integral_bounded.
5363-
5364- Lemma open_itvoo_subset {R : realType} (A : set R) (x : R) :
5365- open A -> A x -> \forall r \near 0^'+, `]x-r, x+r[ `<=` A.
5366- Proof .
5367- move=> oA Ax; case/(_ _ Ax): oA => _/posnumP[r] /subset_ball_prop_in_itv xrA.
5368- exists r%:num => //= k; rewrite /= distrC subr0 set_itvoo => /ltr_normlW kr k0.
5369- apply: (subset_trans _ xrA); apply: subset_itvW.
5370- by apply: ler_sub => //; apply: ltW.
5371- by apply: ler_add => //; apply: ltW.
5372- Qed .
5373-
5374- Lemma open_itvcc_subset {R : realType} (A : set R) (x : R) :
5375- open A -> A x -> \forall r \near 0^'+, `[x-r, x+r] `<=` A.
5376- Proof .
5377- move=> oA Ax; case/(_ _ Ax): oA => _/posnumP[r].
5378- have -> : (r%:num = 2 * (r%:num/2)) by rewrite mulrC -mulrA mulVf // mulr1.
5379- move/subset_ball_prop_in_itvcc => xrA; exists (r%:num/2) => //= k.
5380- rewrite /= distrC subr0 set_itvcc => /ltr_normlW kr k0.
5381- move=> z /andP [xkz zxk]; apply: xrA => //; rewrite in_itv /=.
5382- apply/andP; split.
5383- by apply: (le_trans _ xkz); apply: ler_sub => //; exact: ltW.
5384- by apply: (le_trans zxk); apply: ler_add => //; exact: ltW.
5385- Qed .
5386-
53875365Section lebesgue_differentiation.
53885366Context (rT : realType).
53895367Let mu := [the measure _ _ of @lebesgue_measure rT].
53905368Let R := [the measurableType _ of measurableTypeR rT].
53915369
5392- Lemma continuous_compact_integrable (f : R -> R^o) (A : set R^o):
5370+ Lemma continuous_compact_integrable (f : R -> R^o) (A : set R^o) :
53935371 compact A -> {within A, continuous f} -> mu.-integrable A (EFin \o f).
53945372Proof .
5395- move=> cptA ctsfA; apply/integrableP; have mA : measurable A.
5396- by apply:closed_measurable; apply: compact_closed => //; exact: Rhausdorff.
5397- split.
5398- by apply: measurableT_comp => //; apply: subspace_continuous_measurable_fun.
5373+ move=> cptA ctsfA; have mA := compact_measurable cptA; apply/integrableP; split.
5374+ by apply: measurableT_comp => //; exact: subspace_continuous_measurable_fun.
53995375have /compact_bounded [M [_ mrt]] := continuous_compact ctsfA cptA.
54005376apply: le_lt_trans.
5401- apply (@integral_le_bound _ _ _ _ _ _ (`|M| + 1)) => //.
5402- by apply: measurableT_comp => //; apply: subspace_continuous_measurable_fun.
5403- apply: aeW => /= z Az; rewrite lee_fin; apply: mrt => //.
5404- apply: (@lt_le_trans _ _ (M + 1)); by rewrite ?ltr_addl // ler_add// ler_norm.
5405- case/compact_bounded: cptA => N [_ N1x]; have AN1: A `<=` `[- (`|N|+1), `|N|+1].
5406- move=> z Az; rewrite set_itvcc /= -ler_norml; apply: N1x => //.
5407- by apply: (@lt_le_trans _ _ (N + 1)); rewrite ?ltr_addl // ler_add// ler_norm.
5377+ apply (@integral_le_bound _ _ _ _ _ _ (`|M| + 1)%:E) => //.
5378+ by apply: measurableT_comp => //; exact: subspace_continuous_measurable_fun.
5379+ by apply: aeW => /= z Az; rewrite lee_fin mrt // ltr_spaddr// ler_norm.
5380+ case/compact_bounded : cptA => N [_ N1x].
5381+ have AN1 : A `<=` `[- (`|N| + 1), `|N| + 1].
5382+ by move=> z Az; rewrite set_itvcc /= -ler_norml N1x// ltr_spaddr// ler_norm.
54085383apply: (@le_lt_trans _ _ (_ * _)%E).
54095384 by rewrite lee_pmul; last by apply: (le_measure _ _ _ AN1); rewrite inE.
5410- by rewrite /= lebesgue_measure_itv hlength_itv /= -fun_if -EFinM ltry.
5385+ by rewrite /= lebesgue_measure_itv hlength_itv /= -fun_if -EFinM ltry.
54115386Qed .
54125387
54135388Lemma lebesgue_differentiation_continuous (f : R -> rT^o) (A : set R) (x : R) :
5414- open A -> mu.-integrable A (EFin \o f) -> {for x, continuous f} -> A x ->
5415- (fun r => 1/(2*r) * \int[mu]_(z in `[x-r,x+r]) f z) @ 0^'+ --> (f x:R^o).
5416- Proof .
5417- have ritv r : 0 < r -> mu (`[x-r,x+r]%classic) = (2*r)%:E.
5389+ open A -> mu.-integrable A (EFin \o f) -> {for x, continuous f} -> A x ->
5390+ (fun r => 1 / (2 * r) * \int[mu]_(z in `[x - r, x + r]) f z) @ 0^'+ -->
5391+ (f x : R^o).
5392+ Proof .
5393+ have ritv r : 0 < r -> mu (`[x - r, x + r]%classic) = (2 * r)%:E.
54185394 move=> /gt0_cp rE; rewrite /= lebesgue_measure_itv hlength_itv /= lte_fin.
54195395 rewrite ler_lt_add // ?rE // -EFinD; congr (_ _).
54205396 by rewrite opprB addrAC [_ - _]addrC addrA subrr add0r mulr2n mulrDl ?mul1r.
5421- move=> oA intf ctsfx Ax; apply: (@cvg_zero rT [pseudoMetricNormedZmodType R of rT^o]).
5422- apply/cvgrPdist_le => eps epos; have := @nbhs_right_gt rT 0; apply: filter_app.
5423- have ? : Filter (nbhs (0:R)^'+) by exact: at_right_proper_filter.
5424- move/cvgrPdist_le/(_ eps epos)/at_right_in_segment : ctsfx; apply: filter_app.
5425- apply: (filter_app _ _ (open_itvcc_subset oA Ax)).
5426- have mA : measurable A by exact: open_measurable.
5397+ move=> oA intf ctsfx Ax.
5398+ apply: (@cvg_zero rT [pseudoMetricNormedZmodType R of rT^o]).
5399+ apply/cvgrPdist_le => eps epos; apply: filter_app (@nbhs_right_gt rT 0).
5400+ have ? : Filter (nbhs (0 : R)^'+) := at_right_proper_filter 0.
5401+ move/cvgrPdist_le/(_ eps epos)/at_right_in_segment : ctsfx; apply: filter_app.
5402+ apply: filter_app (open_itvcc_subset oA Ax).
5403+ have mA : measurable A := open_measurable oA.
54275404near=> r => xrA; rewrite addrfctE opprfctE => feps rp.
5428- have cptxr : compact `[x-r, x + r] by exact: segment_compact.
5429- rewrite distrC subr0; have r20 : 0 <= 1/(2*r) by rewrite ?divr_ge0 // ?mulr_ge0.
5430- have -> : f x = 1/(2*r) * \int[mu]_(z in `[x-r,x+r]) cst (f x) z.
5431- rewrite /Rintegral /= integral_cst // ritv //= [f x * _]mulrC mulrA.
5432- by rewrite div1r mulVr ?mul1r ?unitfE ?mulf_neq0.
5433- have intRf : mu.-integrable `[x-r, x+r] (EFin \o f).
5405+ have cptxr : compact `[x - r, x + r] := @segment_compact _ _ _.
5406+ rewrite distrC subr0.
5407+ have r20 : 0 <= 1 / (2 * r) by rewrite ?divr_ge0 // ?mulr_ge0.
5408+ have -> : f x = 1 / (2 * r) * \int[mu]_(z in `[x - r, x + r]) cst (f x) z.
5409+ rewrite /Rintegral /= integral_cst // ritv //=.
5410+ by rewrite mulrC mul1r -mulrA divff ?mulr1// mulf_neq0.
5411+ have intRf : mu.-integrable `[x - r, x + r] (EFin \o f).
54345412 exact: (@integrableS _ _ _ mu _ _ _ _ _ xrA intf).
54355413rewrite /= -mulrBr -fineB; first last.
5436- - rewrite integral_fune_fin_num // continuous_compact_integrable // .
5437- by move=> ?; exact: cvg_cst.
5438- - by rewrite integral_fune_fin_num // .
5414+ - rewrite integral_fune_fin_num// continuous_compact_integrable// => ? .
5415+ exact: cvg_cst.
5416+ - by rewrite integral_fune_fin_num.
54395417rewrite -integralB_EFin //; first last.
5440- - by apply: continuous_compact_integrable => // ?; exact: cvg_cst.
5418+ by apply: continuous_compact_integrable => // ?; exact: cvg_cst.
54415419under [fun _ => adde _ _ ]eq_fun => ? do rewrite -EFinD.
5442- have int_fx : mu.-integrable `[( x - r)%R, ( x + r)%R ] (fun z => (f z - f x)%:E).
5420+ have int_fx : mu.-integrable `[x - r, x + r] (fun z => (f z - f x)%:E).
54435421 under [fun z => (f z - _)%:E]eq_fun => ? do rewrite EFinB.
5444- apply: integrableB => //; apply: continuous_compact_integrable =>// .
5445- by move=> ?; exact: cvg_cst.
5422+ rewrite integrableB// continuous_compact_integrable// => ? .
5423+ exact: cvg_cst.
54465424rewrite normrM [ `|_/_| ]ger0_norm // -fine_abse //; first last.
54475425 by rewrite integral_fune_fin_num.
5448- suff : (\int[mu]_(z in `[(x- r)%R,(x+ r)%R]) `|(f z - f x)|%:E <=
5426+ suff : (\int[mu]_(z in `[(x - r)%R, (x + r)%R]) `|(f z - f x)|%:E <=
54495427 (2 * r * eps)%:E)%E.
54505428 move=> intfeps; apply: le_trans.
54515429 apply: (ler_pmul r20 _ (le_refl _)); first exact: fine_ge0.
54525430 apply: fine_le; last apply: le_abse_integral => //.
54535431 - by rewrite abse_fin_num; exact: integral_fune_fin_num.
54545432 - by apply: integral_fune_fin_num => //; exact: integrable_abse.
5455- - by case/integrableP:int_fx.
5433+ - by case/integrableP: int_fx.
54565434 rewrite div1r ler_pdivr_mull -[_ * _]/(fine (_%:E)); last exact: mulr_gt0.
54575435 by rewrite fine_le // integral_fune_fin_num // integrable_abse.
54585436apply: le_trans.
5459- apply:(@integral_le_bound _ _ _ _ _ (fun z => (f z - f x)%:E) eps) => //.
5437+ apply: (@integral_le_bound _ _ _ _ _ (fun z => (f z - f x)%:E) eps%:E ) => //.
54605438 - by case/integrableP: int_fx.
54615439 - exact: ltW.
54625440 - by apply: aeW => ? ?; rewrite /= lee_fin distrC; apply: feps.
54635441by rewrite ritv //= -EFinM lee_fin mulrC.
54645442Unshelve. all: by end_near. Qed .
54655443
5466- End lebesgue_differentiation.
5444+ End lebesgue_differentiation.
0 commit comments