Skip to content

Commit 13ecdb6

Browse files
committed
deprecate approximation and make its interface accessible
1 parent 937e155 commit 13ecdb6

File tree

6 files changed

+233
-156
lines changed

6 files changed

+233
-156
lines changed

CHANGELOG_UNRELEASED.md

Lines changed: 24 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -12,17 +12,41 @@
1212

1313
- in `mathcomp_extra.v`:
1414
+ lemma `partition_disjoint_bigfcup`
15+
- in `lebesgue_measure.v`:
16+
+ lemma `measurable_indicP`
17+
18+
- in `lebesgue_integral.v`:
19+
+ definition `dyadic_approx` (was `Let A`)
20+
+ definition `integer_approx` (was `Let B`)
21+
+ lemma `measurable_sum`
22+
+ lemma `integrable_indic`
1523

1624
### Changed
25+
26+
- in `lebesgue_integrale.v`
27+
+ change implicits of `measurable_funP`
1728

1829
### Renamed
1930

31+
- in `lebesgue_measure.v`:
32+
+ `measurable_fun_indic` -> `measurable_indic`
33+
+ `emeasurable_fun_sum` -> `emeasurable_sum`
34+
+ `emeasurable_fun_fsum` -> `emeasurable_fsum`
35+
+ `ge0_emeasurable_fun_sum` -> `ge0_emeasurable_sum`
36+
2037
### Generalized
2138

2239
### Deprecated
2340

41+
- in file `lebesgue_integral.v`:
42+
+ lemma `approximation`
43+
2444
### Removed
2545

46+
- in `lebesgue_integral.v`:
47+
+ lemma `measurable_indic` (was uselessly specializing `measurable_fun_indic` (now `measurable_indic`) from `lebesgue_measure.v`)
48+
+ notation `measurable_fun_indic` (deprecation since 0.6.3)
49+
2650
### Infrastructure
2751

2852
### Misc

theories/charge.v

Lines changed: 10 additions & 12 deletions
Original file line numberDiff line numberDiff line change
@@ -1837,7 +1837,7 @@ have int_f_nuT : \int[mu]_x f x = nu setT.
18371837
by apply: eq_eseriesr => i _; rewrite int_f_E// setTI.
18381838
rewrite -UET measure_bigcup//.
18391839
by apply: eq_eseriesl => // x; rewrite in_setT.
1840-
have mf : measurable_fun setT f by exact: ge0_emeasurable_fun_sum.
1840+
have mf : measurable_fun setT f by exact: ge0_emeasurable_sum.
18411841
have fi : mu.-integrable setT f.
18421842
apply/integrableP; split => //.
18431843
under eq_integral do (rewrite gee0_abs; last exact: nneseries_ge0).
@@ -1921,29 +1921,29 @@ Lemma change_of_variables f E : (forall x, 0 <= f x) ->
19211921
\int[mu]_(x in E) (f x * ('d nu '/d mu) x) = \int[nu]_(x in E) f x.
19221922
Proof.
19231923
move=> f0 mE mf; set g := 'd nu '/d mu.
1924-
have [h [ndh hf]] := approximation mE mf (fun x _ => f0 x).
1924+
pose h := nnsfun_approx mE mf.
19251925
have -> : \int[nu]_(x in E) f x =
19261926
lim (\int[nu]_(x in E) (EFin \o h n) x @[n --> \oo]).
19271927
have fE x : E x -> f x = lim ((EFin \o h n) x @[n --> \oo]).
1928-
by move=> Ex; apply/esym/cvg_lim => //; exact: hf.
1928+
by move=> Ex; apply/esym/cvg_lim => //; exact: cvg_nnsfun_approx.
19291929
under eq_integral => x /[!inE] /fE -> //.
19301930
apply: monotone_convergence => //.
19311931
- move=> n; apply/measurable_EFinP.
19321932
by apply: (measurable_funS measurableT) => //; exact/measurable_funP.
19331933
- by move=> n x Ex //=; rewrite lee_fin.
1934-
- by move=> x Ex a b /ndh /=; rewrite lee_fin => /lefP.
1934+
- by move=> x Ex a b ab; rewrite lee_fin; exact/lefP/nd_nnsfun_approx.
19351935
have -> : \int[mu]_(x in E) (f \* g) x =
19361936
lim (\int[mu]_(x in E) ((EFin \o h n) \* g) x @[n --> \oo]).
19371937
have fg x :E x -> f x * g x = lim (((EFin \o h n) \* g) x @[n --> \oo]).
19381938
by move=> Ex; apply/esym/cvg_lim => //; apply: cvgeMr;
1939-
[exact: f_fin_num|exact: hf].
1939+
[exact: f_fin_num|exact: cvg_nnsfun_approx].
19401940
under eq_integral => x /[!inE] /fg -> //.
19411941
apply: monotone_convergence => [//| | |].
19421942
- move=> n; apply/emeasurable_funM; apply/measurable_funTS.
19431943
exact/measurable_EFinP.
19441944
exact: measurable_int (f_integrable _).
19451945
- by move=> n x Ex //=; rewrite mule_ge0 ?lee_fin//=; exact: f_ge0.
1946-
- by move=> x Ex a b /ndh /= /lefP hahb; rewrite lee_wpmul2r ?lee_fin// f_ge0.
1946+
- by move=> x Ex a b ab/=; rewrite lee_wpmul2r ?lee_fin ?f_ge0//; exact/lefP/nd_nnsfun_approx.
19471947
suff suf n : \int[mu]_(x in E) ((EFin \o h n) x * g x) =
19481948
\int[nu]_(x in E) (EFin \o h n) x.
19491949
by under eq_fun do rewrite suf.
@@ -2005,18 +2005,16 @@ Local Notation "'d nu '/d mu" := (f nu mu).
20052005
Lemma chain_rule E : nu `<< mu -> mu `<< la -> measurable E ->
20062006
ae_eq la E ('d nu '/d la) ('d nu '/d mu \* 'd mu '/d la).
20072007
Proof.
2008-
move=> numu mula mE; have nula := measure_dominates_trans numu mula.
2008+
move=> numu mula mE.
20092009
have mf : measurable_fun E ('d nu '/d mu).
20102010
exact/measurable_funTS/(measurable_int _ (f_integrable _)).
2011-
have [h [ndh hf]] := approximation mE mf (fun x _ => f_ge0 numu x).
20122011
apply: integral_ae_eq => //.
2013-
- apply: (integrableS measurableT) => //.
2014-
apply: f_integrable.
2015-
exact: (measure_dominates_trans numu mula).
2012+
- apply: (integrableS measurableT) => //; apply: f_integrable.
2013+
exact: measure_dominates_trans numu mula.
20162014
- apply: emeasurable_funM => //.
20172015
exact/measurable_funTS/(measurable_int _ (f_integrable _)).
20182016
- move=> A AE mA; rewrite change_of_variables//.
2019-
+ by rewrite -!f_integral.
2017+
+ by rewrite -!f_integral//; exact: measure_dominates_trans numu mula.
20202018
+ exact: f_ge0.
20212019
+ exact: measurable_funS mf.
20222020
Qed.

theories/kernel.v

Lines changed: 31 additions & 23 deletions
Original file line numberDiff line numberDiff line change
@@ -131,8 +131,7 @@ Definition kseries : X -> {measure set Y -> \bar R} :=
131131
Lemma measurable_fun_kseries (U : set Y) :
132132
measurable U -> measurable_fun [set: X] (kseries ^~ U).
133133
Proof.
134-
move=> mU.
135-
by apply: ge0_emeasurable_fun_sum => // n _; exact/measurable_kernel.
134+
by move=> mU; apply: ge0_emeasurable_sum => // n _; exact/measurable_kernel.
136135
Qed.
137136

138137
HB.instance Definition _ :=
@@ -546,7 +545,7 @@ rewrite [X in measurable_fun _ X](_ : _ = (fun x => \sum_(r \in range (k_ n))
546545
apply/measurable_EFinP/measurableT_comp => [//|].
547546
exact/measurableT_comp.
548547
- by move=> m y _; rewrite nnfun_muleindic_ge0.
549-
apply: emeasurable_fun_fsum => // r.
548+
apply: emeasurable_fsum => // r.
550549
rewrite [X in measurable_fun _ X](_ : _ = fun x => r%:E *
551550
\int[l x]_y (\1_(k_ n @^-1` [set r]) (x, y))%:E); last first.
552551
apply/funext => x; under eq_integral do rewrite EFinM.
@@ -571,24 +570,30 @@ Lemma measurable_fun_integral_finite_kernel (l : R.-fker X ~> Y)
571570
(k0 : forall z, 0 <= k z) (mk : measurable_fun [set: X * Y] k) :
572571
measurable_fun [set: X] (fun x => \int[l x]_y k (x, y)).
573572
Proof.
574-
have [k_ [ndk_ k_k]] := approximation measurableT mk (fun x _ => k0 x).
575-
apply: (measurable_fun_xsection_integral ndk_ (k_k ^~ Logic.I)) => n r.
576-
have [l_ hl_] := measure_uub l.
577-
by apply: measurable_fun_xsection_finite_kernel => // /[!inE].
573+
pose k_ := nnsfun_approx measurableT mk.
574+
apply: (@measurable_fun_xsection_integral _ k_).
575+
- by move=> a b ab; exact/nd_nnsfun_approx.
576+
- by move=> z; exact/cvg_nnsfun_approx.
577+
- move => n r.
578+
have [l_ hl_] := measure_uub l.
579+
by apply: measurable_fun_xsection_finite_kernel => // /[!inE].
578580
Qed.
579581

580582
Lemma measurable_fun_integral_sfinite_kernel (l : R.-sfker X ~> Y)
581583
(k0 : forall t, 0 <= k t) (mk : measurable_fun [set: X * Y] k) :
582584
measurable_fun [set: X] (fun x => \int[l x]_y k (x, y)).
583585
Proof.
584-
have [k_ [ndk_ k_k]] := approximation measurableT mk (fun xy _ => k0 xy).
585-
apply: (measurable_fun_xsection_integral ndk_ (k_k ^~ Logic.I)) => n r.
586-
have [l_ hl_] := sfinite_kernel l.
587-
rewrite (_ : (fun x => _) = (fun x =>
588-
mseries (l_ ^~ x) 0 (xsection (k_ n @^-1` [set r]) x))); last first.
589-
by apply/funext => x; rewrite hl_//; exact/measurable_xsection.
590-
apply: ge0_emeasurable_fun_sum => // m _.
591-
by apply: measurable_fun_xsection_finite_kernel => // /[!inE].
586+
pose k_ := nnsfun_approx measurableT mk.
587+
apply: (@measurable_fun_xsection_integral _ k_).
588+
- by move=> a b ab; exact/nd_nnsfun_approx.
589+
- by move=> z; exact/cvg_nnsfun_approx.
590+
- move => n r.
591+
have [l_ hl_] := sfinite_kernel l.
592+
rewrite (_ : (fun x => _) = (fun x =>
593+
mseries (l_ ^~ x) 0 (xsection (k_ n @^-1` [set r]) x))); last first.
594+
by apply/funext => x; rewrite hl_//; exact/measurable_xsection.
595+
apply: ge0_emeasurable_sum => // m _.
596+
by apply: measurable_fun_xsection_finite_kernel => // /[!inE].
592597
Qed.
593598

594599
End measurable_fun_integral_finite_sfinite.
@@ -1007,8 +1012,11 @@ Lemma measurable_fun_integral_kernel
10071012
(k : Y -> \bar R) (k0 : forall z, 0 <= k z) (mk : measurable_fun [set: Y] k) :
10081013
measurable_fun [set: X] (fun x => \int[l x]_y k y).
10091014
Proof.
1010-
have [k_ [ndk_ k_k]] := approximation measurableT mk (fun x _ => k0 x).
1011-
by apply: (measurable_fun_preimage_integral ndk_ k_k) => n r; exact/ml.
1015+
pose k_ := nnsfun_approx measurableT mk.
1016+
apply: (@measurable_fun_preimage_integral _ _ _ _ _ _ k_) => //.
1017+
- by move=> a b ab; exact/nd_nnsfun_approx.
1018+
- by move=> z _; exact/cvg_nnsfun_approx.
1019+
- by move=> n r; exact/ml.
10121020
Qed.
10131021

10141022
End measurable_fun_integral_kernel.
@@ -1077,13 +1085,13 @@ Lemma integral_kcomp x f : (forall z, 0 <= f z) -> measurable_fun [set: Z] f ->
10771085
\int[kcomp l k x]_z f z = \int[l x]_y (\int[k (x, y)]_z f z).
10781086
Proof.
10791087
move=> f0 mf.
1080-
have [f_ [ndf_ f_f]] := approximation measurableT mf (fun z _ => f0 z).
1088+
pose f_ := nnsfun_approx measurableT mf.
10811089
transitivity (\int[kcomp l k x]_z (lim ((f_ n z)%:E @[n --> \oo]))).
1082-
by apply/eq_integral => z _; apply/esym/cvg_lim => //=; exact: f_f.
1090+
by apply/eq_integral => z _; apply/esym/cvg_lim => //=; exact: cvg_nnsfun_approx.
10831091
rewrite monotone_convergence//; last 3 first.
10841092
by move=> n; exact/measurable_EFinP.
10851093
by move=> n z _; rewrite lee_fin.
1086-
by move=> z _ a b /ndf_ /lefP ab; rewrite lee_fin.
1094+
by move=> z _ a b ab; rewrite lee_fin; exact/lefP/nd_nnsfun_approx.
10871095
rewrite (_ : (fun _ => _) =
10881096
(fun n => \int[l x]_y (\int[k (x, y)]_z (f_ n z)%:E)))//; last first.
10891097
by apply/funext => n; rewrite integral_kcomp_nnsfun.
@@ -1099,12 +1107,12 @@ transitivity (\int[l x]_y lim ((\int[k (x, y)]_z (f_ n z)%:E) @[n --> \oo])).
10991107
+ exact/measurable_EFinP.
11001108
+ by move=> z _; rewrite lee_fin.
11011109
+ exact/measurable_EFinP.
1102-
+ by move: ab => /ndf_ /lefP ab z _; rewrite lee_fin.
1110+
+ by move=> z _; rewrite lee_fin; exact/lefP/nd_nnsfun_approx.
11031111
apply: eq_integral => y _; rewrite -monotone_convergence//; last 3 first.
11041112
- by move=> n; exact/measurable_EFinP.
11051113
- by move=> n z _; rewrite lee_fin.
1106-
- by move=> z _ a b /ndf_ /lefP; rewrite lee_fin.
1107-
by apply: eq_integral => z _; apply/cvg_lim => //; exact: f_f.
1114+
- by move=> z _ a b ab; rewrite lee_fin; exact/lefP/nd_nnsfun_approx.
1115+
by apply: eq_integral => z _; apply/cvg_lim => //; exact: cvg_nnsfun_approx.
11081116
Qed.
11091117

11101118
End integral_kcomp.

0 commit comments

Comments
 (0)