1- (* mathcomp analysis (c) 2022 Inria and AIST. License: CeCILL-C. *)
1+ (* mathcomp analysis (c) 2025 Inria and AIST. License: CeCILL-C. *)
22From HB Require Import structures.
33From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval finmap.
44From mathcomp Require Import rat archimedean ring lra.
55From mathcomp Require Import unstable mathcomp_extra boolp classical_sets.
66From mathcomp Require Import functions cardinality fsbigop interval_inference.
77From mathcomp Require Import reals ereal topology normedtype sequences.
8- From mathcomp Require Import esum measure lebesgue_measure numfun.
9- From mathcomp Require Import lebesgue_integral probability exp kernel charge.
8+ From mathcomp Require Import esum measure lebesgue_measure numfun exp .
9+ From mathcomp Require Import lebesgue_integral trigo probability kernel charge.
1010
1111(**md************************************************************************* *)
1212(* # Semantics of a probabilistic programming language using s-finite kernels *)
@@ -21,7 +21,9 @@ From mathcomp Require Import lebesgue_integral probability exp kernel charge.
2121(* poisson_pdf == Poisson pdf *)
2222(* exponential_pdf == exponential distribution pdf *)
2323(* measurable_sum X Y == the type X + Y, as a measurable type *)
24+ (* ``` *)
2425(* *)
26+ (* ``` *)
2527(* mscore f t := mscale `|f t| \d_tt *)
2628(* kscore f := fun=> mscore f *)
2729(* This is an s-finite kernel. *)
@@ -1466,8 +1468,8 @@ by rewrite /score/= /mscale/= ger0_norm//= f0.
14661468Qed .
14671469
14681470Lemma score_score (f : R -> R) (g : R * unit -> R)
1469- (mf : measurable_fun setT f)
1470- (mg : measurable_fun setT g) :
1471+ (mf : measurable_fun [set: R] f)
1472+ (mg : measurable_fun [set: R * unit] g) :
14711473 letin (score mf) (score mg) =
14721474 score (measurable_funM mf (measurableT_comp mg (pair2_measurable tt))).
14731475Proof .
@@ -2181,3 +2183,134 @@ rewrite letin'E/= /sample; unlock.
21812183rewrite integral_bernoulli ?r0//=.
21822184by rewrite /mscale/= iteE//= iteE//= fail'E mule0 adde0 ger0_norm.
21832185Qed .
2186+
2187+ (* TODO: move to probability.v? *)
2188+ Section gauss.
2189+ Variable R : realType.
2190+ Local Open Scope ring_scope.
2191+
2192+ Definition gauss_pdf := @normal_pdf R 0 1.
2193+
2194+ Lemma normal_pdf_gt0 m s x : 0 < s -> 0 < normal_pdf m s x :> R.
2195+ Proof .
2196+ move=> s0; rewrite /normal_pdf gt_eqF// mulr_gt0 ?expR_gt0// invr_gt0.
2197+ by rewrite sqrtr_gt0 pmulrn_rgt0// mulr_gt0 ?pi_gt0 ?exprn_gt0.
2198+ Qed .
2199+
2200+ Lemma gauss_pdf_gt0 x : 0 < gauss_pdf x.
2201+ Proof . exact: normal_pdf_gt0. Qed .
2202+
2203+ Definition gauss_prob := @normal_prob R 0 1.
2204+
2205+ HB.instance Definition _ := Probability.on gauss_prob.
2206+
2207+ Lemma gauss_prob_dominates : gauss_prob `<< lebesgue_measure.
2208+ Proof . exact: normal_prob_dominates. Qed .
2209+
2210+ Lemma continuous_gauss_pdf x : {for x, continuous gauss_pdf}.
2211+ Proof . exact: continuous_normal_pdf. Qed .
2212+
2213+ End gauss.
2214+
2215+ (* the Lebesgue measure is definable in Staton's language
2216+ [equation (10), Section 4.1, Staton ESOP 2017] *)
2217+ Section gauss_lebesgue.
2218+ Context d (T : measurableType d) (R : realType).
2219+ Notation mu := (@lebesgue_measure R).
2220+
2221+ Let f1 (x : g_sigma_algebraType (R.-ocitv.-measurable)) := (gauss_pdf x)^-1%R.
2222+
2223+ Let f1E (x : R) : f1 x = (Num.sqrt (pi *+ 2) * expR (- (- x ^+ 2 / 2)))%R.
2224+ Proof .
2225+ rewrite /f1 /gauss_pdf /normal_pdf oner_eq0.
2226+ rewrite /normal_peak expr1n mul1r.
2227+ by rewrite /normal_fun subr0 expr1n invfM invrK expRN.
2228+ Qed .
2229+
2230+ Let f1_gt0 (x : R) : (0 < f1 x)%R.
2231+ Proof . by rewrite f1E mulr_gt0 ?expR_gt0// sqrtr_gt0 mulrn_wgt0// pi_gt0. Qed .
2232+
2233+ Lemma measurable_fun_f1 : measurable_fun setT f1.
2234+ Proof .
2235+ apply: continuous_measurable_fun => x.
2236+ apply: (@continuousV _ _ (@gauss_pdf R)).
2237+ by rewrite gt_eqF// gauss_pdf_gt0.
2238+ exact: continuous_gauss_pdf.
2239+ Qed .
2240+
2241+ Lemma integral_mgauss01 : forall U, measurable U ->
2242+ \int[(@gauss_prob R)]_(y in U) (f1 y)%:E =
2243+ \int[mu]_(x0 in U) (gauss_pdf x0 * f1 x0)%:E.
2244+ Proof .
2245+ move=> U mU.
2246+ under [in RHS]eq_integral do rewrite EFinM/= muleC.
2247+ rewrite /=.
2248+ rewrite -(@Radon_Nikodym_SigmaFinite.change_of_variables
2249+ _ _ _ _ (@lebesgue_measure R))//=; last 3 first.
2250+ exact: gauss_prob_dominates.
2251+ by move=> /= x; rewrite lee_fin ltW.
2252+ apply/measurable_EFinP.
2253+ apply: measurable_funTS.
2254+ exact: measurable_fun_f1.
2255+ apply: ae_eq_integral => //=.
2256+ - apply: emeasurable_funM => //.
2257+ apply/measurable_funTS/measurableT_comp => //.
2258+ exact: measurable_fun_f1.
2259+ apply: (measurable_int mu).
2260+ apply: (integrableS _ _ (@subsetT _ _)) => //=.
2261+ apply: Radon_Nikodym_SigmaFinite.f_integrable => /=.
2262+ exact: gauss_prob_dominates.
2263+ - apply: emeasurable_funM => //.
2264+ apply/measurable_funTS/measurableT_comp => //.
2265+ exact: measurable_fun_f1.
2266+ apply/measurable_funTS/measurableT_comp => //.
2267+ exact: measurable_normal_pdf.
2268+ - apply: ae_eqe_mul2l => /=.
2269+ rewrite /Radon_Nikodym_SigmaFinite.f/=.
2270+ case: pselect => [gauss_prob_dom|]; last first.
2271+ by move=> /(_ (@gauss_prob_dominates R)).
2272+ case: cid => //= h [h1 h2 h3] gauss_probE.
2273+ apply: integral_ae_eq => //=.
2274+ + exact: integrableS h3.
2275+ + apply/measurable_funTS/measurableT_comp => //.
2276+ exact: measurable_normal_pdf.
2277+ + move=> E EU mE.
2278+ by rewrite -gauss_probE.
2279+ Qed .
2280+
2281+ Let mf1 : measurable_fun setT f1.
2282+ Proof .
2283+ apply: (measurable_comp (F := [set r : R | r != 0%R])) => //.
2284+ - exact: open_measurable.
2285+ - by move=> /= r [t _ <-]; rewrite gt_eqF// gauss_pdf_gt0.
2286+ - apply: open_continuous_measurable_fun => //.
2287+ by apply/in_setP => x /= x0; exact: inv_continuous.
2288+ - exact: measurable_normal_pdf.
2289+ Qed .
2290+
2291+ Definition staton_lebesgue : R.-sfker T ~> _ :=
2292+ letin (sample_cst (@gauss_prob R : pprobability _ _))
2293+ (letin
2294+ (score (measurableT_comp mf1 macc1of2))
2295+ (ret macc1of3)).
2296+
2297+ Lemma staton_lebesgueE x U : measurable U ->
2298+ staton_lebesgue x U = lebesgue_measure U.
2299+ Proof .
2300+ move=> mU; rewrite [in LHS]/staton_lebesgue/=.
2301+ rewrite [in LHS]letinE /=.
2302+ transitivity (\int[(@gauss_prob R)]_(y in U) (f1 y)%:E).
2303+ rewrite -[in RHS](setTI U) integral_mkcondr/=.
2304+ apply: eq_integral => //= r _.
2305+ rewrite letinE/= ge0_integral_mscale//= ger0_norm//; last first.
2306+ by rewrite invr_ge0// normal_pdf_ge0.
2307+ rewrite integral_dirac// diracT mul1e/= diracE epatch_indic/=.
2308+ by rewrite indicE.
2309+ rewrite integral_mgauss01//.
2310+ transitivity (\int[lebesgue_measure]_(x in U) (\1_U x)%:E).
2311+ apply: eq_integral => /= y yU.
2312+ by rewrite /f1 divrr ?indicE ?yU// unitfE gt_eqF// gauss_pdf_gt0.
2313+ by rewrite integral_indic//= setIid.
2314+ Qed .
2315+
2316+ End gauss_lebesgue.
0 commit comments