@@ -131,8 +131,7 @@ Definition kseries : X -> {measure set Y -> \bar R} :=
131131Lemma measurable_fun_kseries (U : set Y) :
132132 measurable U -> measurable_fun [set: X] (kseries ^~ U).
133133Proof .
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.
136135Qed .
137136
138137HB.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.
550549rewrite [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)).
573572Proof .
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].
578580Qed .
579581
580582Lemma 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)).
583585Proof .
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].
592597Qed .
593598
594599End 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).
10091014Proof .
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.
10121020Qed .
10131021
10141022End 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).
10781086Proof .
10791087move=> f0 mf.
1080- have [ f_ [ndf_ f_f]] := approximation measurableT mf (fun z _ => f0 z) .
1088+ pose f_ := nnsfun_approx measurableT mf.
10811089transitivity (\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 .
10831091rewrite 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 .
10871095rewrite (_ : (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 .
11031111apply: 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 .
11081116Qed .
11091117
11101118End integral_kcomp.
0 commit comments